update to account for coq-categories changes
[coq-hetmet.git] / src / ProgrammingLanguage.v
index de7c7f0..7831fad 100644 (file)
@@ -18,6 +18,8 @@ Require Import Enrichment_ch2_8.
 Require Import Subcategories_ch7_1.
 Require Import NaturalTransformations_ch7_4.
 Require Import NaturalIsomorphisms_ch7_5.
+Require Import BinoidalCategories.
+Require Import PreMonoidalCategories.
 Require Import MonoidalCategories_ch7_8.
 Require Import Coherence_ch7_8.
 Require Import Enrichment_ch2_8.
@@ -27,12 +29,6 @@ Require Import FunctorCategories_ch7_7.
 Require Import NaturalDeduction.
 Require Import NaturalDeductionCategory.
 
-Require Import FreydCategories.
-
-Require Import Reification.
-Require Import GeneralizedArrow.
-Require Import GeneralizedArrowFromReification.
-
 Section Programming_Language.
 
   Context {T    : Type}.               (* types of the language *)
@@ -77,13 +73,14 @@ Section Programming_Language.
       apply pl_subst.
       Defined.
 
+    Existing Instance pl_eqv.
     Definition TypesL : ECategory JudgmentsL (Tree ??T) (fun x y => [x|=y]).
       refine
       {| eid   := identityProof
        ; ecomp := cutProof
       |}; intros.
-      apply MonoidalCat_all_central.
-      apply MonoidalCat_all_central.
+      apply (mon_commutative(MonoidalCat:=JudgmentsL)).
+      apply (mon_commutative(MonoidalCat:=JudgmentsL)).
       unfold identityProof; unfold cutProof; simpl.
       apply nd_cut_left_identity.
       unfold identityProof; unfold cutProof; simpl.
@@ -95,7 +92,7 @@ Section Programming_Language.
 
     Definition Types_first c : EFunctor TypesL TypesL (fun x => x,,c ).
       refine {| efunc := fun x y => (@se_expand_right _ _ _ _ _ _ _ (@pl_sequent_join PL) c x y) |}.
-      intros; apply MonoidalCat_all_central.
+      intros; apply (mon_commutative(MonoidalCat:=JudgmentsL)).
       intros. unfold ehom. unfold hom. unfold identityProof. unfold eid. simpl. unfold identityProof.
       apply se_reflexive_right.
       intros. unfold ehom. unfold comp. simpl. unfold cutProof.
@@ -109,7 +106,7 @@ Section Programming_Language.
     Definition Types_second c : EFunctor TypesL TypesL (fun x => c,,x).
       eapply Build_EFunctor.
       instantiate (1:=(fun x y => ((@se_expand_left _ _ _ _ _ _ _ (@pl_sequent_join PL) c x y)))).
-      intros; apply MonoidalCat_all_central.
+      intros; apply (mon_commutative(MonoidalCat:=JudgmentsL)).
       intros. unfold ehom. unfold hom. unfold identityProof. unfold eid. simpl. unfold identityProof.
       apply se_reflexive_left.
       intros. unfold ehom. unfold comp. simpl. unfold cutProof.
@@ -127,23 +124,52 @@ Section Programming_Language.
          |}.
       Defined.
 
-    Definition Types_assoc a b : Types_second a >>>> Types_first b <~~~> Types_first b >>>> Types_second a.
+    Instance Types_assoc_iso a b c : Isomorphic(C:=TypesL) ((a,,b),,c) (a,,(b,,c)) :=
+      { iso_forward  := nd_seq_reflexive _ ;; tsr_ant_cossa _ a b c
+      ; iso_backward := nd_seq_reflexive _ ;; tsr_ant_assoc _ a b c
+      }.
+      admit.
+      admit.
+      Defined.
+
+    Instance Types_assoc a b : Types_second a >>>> Types_first b <~~~> Types_first b >>>> Types_second a :=
+      { ni_iso := fun c => Types_assoc_iso a c b }.
       admit.
       Defined.
 
-    Definition Types_cancelr   : Types_first [] <~~~> functor_id _.
+    Instance Types_cancelr_iso a : Isomorphic(C:=TypesL) (a,,[]) a :=
+      { iso_forward  := nd_seq_reflexive _ ;; tsr_ant_rlecnac _ a
+      ; iso_backward := nd_seq_reflexive _ ;; tsr_ant_cancelr _ a
+      }.
+      admit.
       admit.
       Defined.
 
-    Definition Types_cancell   : Types_second [] <~~~> functor_id _.
+    Instance Types_cancelr   : Types_first [] <~~~> functor_id _ :=
+      { ni_iso := Types_cancelr_iso }.
       admit.
       Defined.
 
-    Definition Types_assoc_ll a b : Types_second (a,,b) <~~~> Types_second b >>>> Types_second a.
+    Instance Types_cancell_iso a : Isomorphic(C:=TypesL) ([],,a) a :=
+      { iso_forward  := nd_seq_reflexive _ ;; tsr_ant_llecnac _ a
+      ; iso_backward := nd_seq_reflexive _ ;; tsr_ant_cancell _ a
+      }.
+      admit.
       admit.
       Defined.
 
-    Definition Types_assoc_rr a b : Types_first (a,,b) <~~~> Types_first a >>>> Types_first b.
+    Instance Types_cancell   : Types_second [] <~~~> functor_id _ :=
+      { ni_iso := Types_cancell_iso }.
+      admit.
+      Defined.
+
+    Instance Types_assoc_ll a b : Types_second (a,,b) <~~~> Types_second b >>>> Types_second a :=
+      { ni_iso := fun c => Types_assoc_iso a b c }.
+      admit.
+      Defined.
+
+    Instance Types_assoc_rr a b : Types_first (a,,b) <~~~> Types_first a >>>> Types_first b :=
+      { ni_iso := fun c => iso_inv _ _ (Types_assoc_iso c a b) }.
       admit.
       Defined.
 
@@ -156,10 +182,14 @@ Section Programming_Language.
         }.
         admit. (* pentagon law *)
         admit. (* triangle law *)
-        admit. (* assoc_rr/assoc coherence *)
-        admit. (* assoc_ll/assoc coherence *)
+        intros; simpl; reflexivity.
+        intros; simpl; reflexivity.
+        admit.  (* assoc   central *)
+        admit.  (* cancelr central *)
+        admit.  (* cancell central *)
         Defined.
 
+    (*
     Definition TypesEnrichedInJudgments : Enrichment.
       refine {| enr_c := TypesL |}.
       Defined.
@@ -168,14 +198,19 @@ Section Programming_Language.
     {
     }.
 
+    Lemma CartesianEnrMonoidal (e:Enrichment) `(C:CartesianCat(Ob:= _)(Hom:= _)(C:=Underlying (enr_c e))) : MonoidalEnrichment e.
+      admit.
+      Defined.
+
     (* need to prove that if we have cartesian tuples we have cartesian contexts *)
     Definition LanguagesWithProductsAreSMME : HasProductTypes -> SurjectiveMonicMonoidalEnrichment TypesEnrichedInJudgments.
       admit.
       Defined.
-
+      *)
   End LanguageCategory.
-End Programming_Language.
 
+End Programming_Language.
+(*
 Structure ProgrammingLanguageSMME :=
 { plsmme_t       : Type
 ; plsmme_judg    : Type
@@ -186,56 +221,5 @@ Structure ProgrammingLanguageSMME :=
 }.
 Coercion plsmme_pl   : ProgrammingLanguageSMME >-> ProgrammingLanguage.
 Coercion plsmme_smme : ProgrammingLanguageSMME >-> SurjectiveMonicMonoidalEnrichment.
-
-Section ArrowInLanguage.
-  Context  (Host:ProgrammingLanguageSMME).
-  Context `(CC:CartesianCat (me_mon Host)).
-  Context `(K:@ECategory _ _ _ _ _ _ (@car_mn _ _ _ _ _ _ _ CC) C Kehom).
-  Context `(pmc:PreMonoidalCat K bobj mobj (@one _ _ _ (cartesian_terminal C))).
-    (* FIXME *)
-    (*
-    Definition ArrowInProgrammingLanguage := 
-      @FreydCategory _ _ _ _ _ _ (@car_mn _ _ _ _ _ _ _ CC) _ _ _ _ pmc.
-      *)
-End ArrowInLanguage.
-
-Section GArrowInLanguage.
-  Context (Guest:ProgrammingLanguageSMME).
-  Context (Host :ProgrammingLanguageSMME).
-  Definition GeneralizedArrowInLanguage := GeneralizedArrow Guest Host.
-
-  (* FIXME
-  Definition ArrowsAreGeneralizedArrows : ArrowInProgrammingLanguage -> GeneralizedArrowInLanguage.
-  *)
-  Definition TwoLevelLanguage := Reification Guest Host (me_i Host).
-
-  Context (GuestHost:TwoLevelLanguage).
-
-  Definition FlatObject (x:TypesL _ _ Host) :=
-    forall y1 y2, not ((reification_r_obj GuestHost y1 y2)=x).
-
-  Definition FlatSubCategory := FullSubcategory (TypesL _ _ Host) FlatObject.
-
-  Section Flattening.
-
-    Context  (F:Retraction (TypesL _ _ Host) FlatSubCategory).
-    Definition FlatteningOfReification := garrow_from_reification Guest Host GuestHost >>>> F.
-    Lemma FlatteningIsNotDestructive : 
-      FlatteningOfReification >>>> retraction_retraction F >>>> RepresentableFunctor _ (me_i Host) ~~~~ GuestHost.
-      admit.
-      Qed.
-
-  End Flattening.
-
-End GArrowInLanguage.
-
-Inductive NLevelLanguage : nat -> ProgrammingLanguageSMME -> Type :=
-| NLevelLanguage_zero : forall lang,    NLevelLanguage O lang
-| NLevelLanguage_succ : forall (L1 L2:ProgrammingLanguageSMME) n,
-                          TwoLevelLanguage L1 L2 -> NLevelLanguage n L1 -> NLevelLanguage (S n) L2.
-
-Definition OmegaLevelLanguage : Type :=
-  { f : nat -> ProgrammingLanguageSMME
-  & forall n, TwoLevelLanguage (f n) (f (S n)) }.
-  
+*) 
 Implicit Arguments ND [ Judgment ].