Module TBrandContext


Section TBrandContext.
  Require Import String.
  Require Import List.
  Require Import ListSet.
  Require Import Sumbool.
  Require Import Arith.
  Require Import Bool.
  Require Import Morphisms.
  Require Import Basics.
  Require Import EquivDec.
  Require Import Eqdep_dec.

  Require Import Utils.
  Require Import ForeignType.
  Require Import RType.
  Require Import BrandRelation.

Implicitly, everything derives from Any
  Definition Any {ftype:foreign_type} {br:brand_relation} := Brand nil.

  Section Brand_Context.
    Context {ftype:foreign_type}.
    Context {br:brand_relation}.
Represents a mapping from brands to types
  Class brand_context :=
     mkBrand_context {
         brand_context_types : list (string*rtype);
         brand_context_types_sorted : is_list_sorted ODT_lt_dec (domain brand_context_types) = true
      }.

  Lemma brand_context_nodup {m:brand_context} : NoDup (domain brand_context_types).
Proof.
    eapply is_list_sorted_NoDup.
    - eapply StringOrder.lt_strorder.
    - apply brand_context_types_sorted.
  Qed.

  Lemma brand_context_ext m pf1 pf1':
    mkBrand_context m pf1 = mkBrand_context m pf1'.
Proof.
    f_equal.
    - apply UIP_dec. apply bool_dec.
  Defined.
  
  Lemma brand_context_fequal {mm₂}:
    @brand_context_types m₁ = @brand_context_types m₂ -> m₁ = m₂.
Proof.
      destruct m₁; destruct m₂; simpl; intros; subst.
      apply brand_context_ext.
  Qed.

  Definition brand_context_Rec (m:brand_context) := Rec Open _ (brand_context_types_sorted).

  Lemma brand_context_Rec_inj (m1 m2:brand_context) :
    brand_context_Rec m1 = brand_context_Rec m2 -> m1 = m2.
Proof.
    inversion 1; rtype_equalizer; subst.
    apply brand_context_fequal; trivial.
  Qed.

  End Brand_Context.
  End TBrandContext.