X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FProgrammingLanguageArrow.v;h=b613455609918a52475b21b18e56a1c09e1c17c8;hp=a4f40e3b3287847268912f5cbd73f30196557fc1;hb=64d416692bda1d36c33b5efa245d46dcf546ad4a;hpb=562e94b529f34fb3854be7914a49190c5243c55a diff --git a/src/ProgrammingLanguageArrow.v b/src/ProgrammingLanguageArrow.v index a4f40e3..b613455 100644 --- a/src/ProgrammingLanguageArrow.v +++ b/src/ProgrammingLanguageArrow.v @@ -33,91 +33,6 @@ Require Import FreydCategories. 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... *) @@ -133,10 +48,9 @@ Section ArrowInLanguage. (* 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). -Check (Underlying KE). + 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 {kbo:center_of_TypesL -> center_of_TypesL -> center_of_TypesL}. @@ -153,7 +67,6 @@ Check (Underlying KE). Defined. Context (K_surjective:SurjectiveEnrichment K_enrichment). *) -Check (@FreydCategory). Definition ArrowInProgrammingLanguage := @FreydCategory _ _ _ _ _ _ center_of_TypesL _ _ _ _ K. @@ -163,4 +76,3 @@ Check (@FreydCategory). Defined. End GArrowInLanguage. -*) \ No newline at end of file