clean up hints for NaturalDeduction, split ProgrammingLanguage into multiple files
[coq-hetmet.git] / src / ProgrammingLanguageArrow.v
index a4f40e3..5386d3e 100644 (file)
@@ -18,6 +18,9 @@ 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 PreMonoidalCenter.
 Require Import MonoidalCategories_ch7_8.
 Require Import Coherence_ch7_8.
 Require Import Enrichment_ch2_8.
@@ -27,140 +30,46 @@ Require Import FunctorCategories_ch7_7.
 Require Import NaturalDeduction.
 Require Import NaturalDeductionCategory.
 
-Require Import ProgrammingLanguage.
-Require Import ProgrammingLanguageGeneralizedArrow.
+Require Import ProgrammingLanguageCategory.
 Require Import FreydCategories.
-
+Require Import Enrichments.
 Require Import GeneralizedArrow.
 
-Section ExtendFunctor.
-
-  Context `(F:Functor).
-  Context  (P:c1 -> Prop).
-
-  Definition domain_subcat := FullSubcategory c1 P.
-
-  Definition functor_restricts_to_full_subcat_on_domain_fobj (a:domain_subcat) : c2 :=
-    F (projT1 a).
-
-  Definition functor_restricts_to_full_subcat_on_domain_fmor (a b:domain_subcat)(f:a~~{domain_subcat}~~>b) :
-    (functor_restricts_to_full_subcat_on_domain_fobj a)~~{c2}~~>(functor_restricts_to_full_subcat_on_domain_fobj b) :=
-    F \ (projT1 f).
-
-  Lemma functor_restricts_to_full_subcat_on_domain : Functor domain_subcat c2 functor_restricts_to_full_subcat_on_domain_fobj.
-    refine {| fmor := functor_restricts_to_full_subcat_on_domain_fmor |};
-      unfold functor_restricts_to_full_subcat_on_domain_fmor; simpl; intros.
-    setoid_rewrite H; reflexivity.
-    setoid_rewrite fmor_preserves_id; reflexivity.
-    setoid_rewrite <- fmor_preserves_comp; reflexivity.
-    Defined.
-
-End ExtendFunctor.
-
-Section MonoidalSubCat.
-
-  (* a monoidal subcategory is a full subcategory, closed under tensor and containing the unit object *)
-  Class MonoidalSubCat {Ob}{Hom}{C:Category Ob Hom}{MFobj}{MF}{MI}(MC:MonoidalCat C MFobj MF MI) :=
-  { msc_P                   :  MC -> Prop
-  ; msc_closed_under_tensor :  forall o1 o2, msc_P o1 -> msc_P o2 -> msc_P (MC (pair_obj o1 o2))
-  ; msc_contains_unit       :  msc_P (mon_i MC)
-  ; msc_subcat              := FullSubcategory MC msc_P
-  }.
-  Local Coercion msc_subcat : MonoidalSubCat >-> SubCategory.
-
-  Context `(MSC:MonoidalSubCat).
-
-  (* any full subcategory of a monoidal category, , is itself monoidal *)
-  Definition mf_restricts_to_full_subcat_on_domain_fobj (a:MSC ×× MSC) : MSC.
-    destruct a.
-    destruct o.
-    destruct o0.
-    set (MC (pair_obj x x0)) as m'.
-    exists m'.
-    apply msc_closed_under_tensor; auto.
-    Defined.
-
-  Definition mf_restricts_to_full_subcat_on_domain_fmor
-    {a}{b}
-    (f:a~~{MSC ×× MSC}~~>b)
-    :
-    (mf_restricts_to_full_subcat_on_domain_fobj a)~~{MSC}~~>(mf_restricts_to_full_subcat_on_domain_fobj b).
-    destruct a as [[a1 a1pf] [a2 a2pf]].
-    destruct b as [[b1 b1pf] [b2 b2pf]].
-    destruct f as [[f1 f1pf] [f2 f2pf]].
-    simpl in *.
-    exists (MC \ (pair_mor (pair_obj a1 a2) (pair_obj b1 b2) f1 f2)); auto.
-    Defined.
-
-  Lemma mf_restricts_to_full_subcat_on_domain : Functor (MSC ×× MSC) MSC
-    mf_restricts_to_full_subcat_on_domain_fobj.
-    refine {| fmor := fun a b f => mf_restricts_to_full_subcat_on_domain_fmor f |};
-      unfold functor_restricts_to_full_subcat_on_domain_fmor; simpl; intros.
-    admit.
-    admit.
-    admit.
-    Defined.
-
-  Definition subcat_i : MSC.
-    exists (mon_i MC).
-    apply msc_contains_unit.
-    Defined.
-
-  Lemma full_subcat_is_monoidal : MonoidalCat MSC _ mf_restricts_to_full_subcat_on_domain subcat_i.
-    admit.
-    Defined.
-
-  Lemma inclusion_functor_monoidal : MonoidalFunctor full_subcat_is_monoidal MC (InclusionFunctor _ MSC).
-    admit.
-    Defined.
-
-End MonoidalSubCat.
-Coercion full_subcat_is_monoidal : MonoidalSubCat >-> MonoidalCat.
-
-(*
 Section ArrowInLanguage.
 
   (* an Arrow In A Programming Language consists of... *)
 
   (* a host language: *)
-  Context `(Host:ProgrammingLanguage).
-
-  (* ... for which Types(L) is cartesian: *)
-  Context {MF}(center_of_TypesL:MonoidalCat (TypesL _ _ Host) (fun x => (fst_obj _ _ x),,(snd_obj _ _ x)) MF []).
+  Context `(Host         : ProgrammingLanguage).
 
-  (* along with a monoidal subcategory of Z(Types(L)) *)
-  Context (VK:MonoidalSubCat center_of_TypesL).
+  Context  (Host_Monoidal  : MonoidalCat  (TypesL_PreMonoidal Host)).
+  Context  (Host_Cartesian : CartesianCat Host_Monoidal).
 
-  (* a premonoidal category enriched in aforementioned full subcategory *)
-  Context (Kehom:center_of_TypesL -> center_of_TypesL -> @ob _ _ VK).
-Check (@ECategory).
-  Context (KE:@ECategory (@ob _ _ VK) (@hom _ _ VK) VK _ VK (mon_i (full_subcat_is_monoidal VK)) (full_subcat_is_monoidal VK) center_of_TypesL Kehom).
+  Context
+    {P}
+    (Pobj_unit   : P [])
+    (Pobj_closed : forall a b, P a → P b → P (bin_obj(BinoidalCat:=Center_is_PreMonoidal (TypesL_PreMonoidal Host)) a b)).
 
-Check (Underlying KE).
+  Context (VK : FullSubcategory Host_Cartesian P).
 
-  Context {kbo:center_of_TypesL -> center_of_TypesL -> center_of_TypesL}.
+  Context  ehom KE (K_bin:@EBinoidalCat _ _ VK _ _ _
+          (PreMonoidalFullSubcategory_PreMonoidal Host_Cartesian VK Pobj_unit Pobj_closed)
+          (TypesL Host) ehom KE (bin_obj(BinoidalCat:=Host_Monoidal))).
 
-  Context (kbc:@BinoidalCat center_of_TypesL _ (Underlying KE) kbo).
+  Context (K_premonoidal:PreMonoidalCat K_bin (one(TerminalObject:=Host_Cartesian))).
 
-  Check (@PreMonoidalCat)
-  Definition one' := @one _ _ _ (car_terminal(CartesianCat:=CC))
-  Context (K:@PreMonoidalCat _ _ KE kbo kbc ).
-  Context (CC:CartesianCat center_of_TypesL).
+  Definition ArrowInProgrammingLanguage :=
+    @FreydCategory _ _ _ _ _ _ _ _ Host_Cartesian _ _ K_bin K_premonoidal.
 
-  (*
   Definition K_enrichment : Enrichment.
-    refine {| enr_c := KE |}.
-    Defined.
-  Context (K_surjective:SurjectiveEnrichment K_enrichment).
-    *)
-Check (@FreydCategory).
-
-  Definition ArrowInProgrammingLanguage :=
-    @FreydCategory _ _ _ _ _ _ center_of_TypesL _ _ _ _ K.
+    refine
+      {| enr_c_pm  := K_premonoidal
+       ; enr_v_mon := MonoidalFullSubcategory_Monoidal Host_Cartesian _ _ VK
+      |}.
+      Defined.
 
-  Definition ArrowsAreGeneralizedArrows (arrow:ArrowInProgrammingLanguage) : GeneralizedArrow K_enrichment Host.
-    refine {| ga_functor_monoidal := inclusion_functor_monoidal VK |}.
-    Defined.
+  Instance ArrowsAreGeneralizedArrows : GeneralizedArrow K_enrichment (TypesEnrichedInJudgments Host) :=
+    { ga_functor_monoidal :=
+        PreMonoidalFullSubcategoryInclusionFunctor_PreMonoidal Host_Cartesian VK Pobj_unit Pobj_closed Host_Cartesian }.
 
-End GArrowInLanguage.
-*)
\ No newline at end of file
+End ArrowInLanguage.