From: Adam Megacz Date: Sat, 2 Apr 2011 20:16:22 +0000 (-0700) Subject: re-arrange ProgrammingLanguage X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=commitdiff_plain;h=e68b13536be2d1def208bde68dbbcdc4c1097d16 re-arrange ProgrammingLanguage --- diff --git a/src/ProgrammingLanguageArrow.v b/src/ProgrammingLanguageArrow.v index 6bf0f24..f4a0d8e 100644 --- a/src/ProgrammingLanguageArrow.v +++ b/src/ProgrammingLanguageArrow.v @@ -31,38 +31,126 @@ Require Import ProgrammingLanguage. Require Import ProgrammingLanguageGeneralizedArrow. 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. - Section ArrowInLanguage. - Context {MF}{mn:MonoidalCat TypesL (fun x => (fst_obj _ _ x),,(snd_obj _ _ x)) MF []} (CC:CartesianCat mn). - Context {Kehom}(K:@ECategory _ _ TypesL _ mn [] mn TypesL Kehom). - Context {bc:BinoidalCat (Underlying K) (@T_Branch _)}. - Context (pmc:@PreMonoidalCat _ _ _ _ bc (@one _ _ _ (car_terminal(CartesianCat:=CC)))). - Definition ArrowInProgrammingLanguage := @FreydCategory _ _ _ _ _ _ mn _ _ _ _ pmc. - End ArrowInLanguage. + (* an Arrow In A Programming Language consists of... *) + + (* a host language: *) + Context (Host:ProgrammingLanguageSMME). + + (* ... for which Types(L) is cartesian: *) + Context {MF}(center_of_TypesL:MonoidalCat (TypesL _ _ Host) (fun x => (fst_obj _ _ x),,(snd_obj _ _ x)) MF []). - Definition ArrowsAreGeneralizedArrows (Host:ProgrammingLanguageSMME) - {mf}{mn}{cc}{kehom}{CC} - (arrow:ArrowInProgrammingLanguage Host _ _ CC mf mn cc kehom) : GeneralizedArrowInLanguage. + (* along with a monoidal subcategory of Z(Types(L)) *) + Context (VK:MonoidalSubCat center_of_TypesL). - Definition TwoLevelLanguage := Reification Guest Host (me_i Host). + (* a premonoidal category enriched in aforementioned full subcategory *) + Context {Kehom:center_of_TypesL -> center_of_TypesL -> VK}. - Context (GuestHost:TwoLevelLanguage). + Context {KE:@ECategory _ _ VK _ _ _ VK center_of_TypesL Kehom}. - Definition FlatObject (x:TypesL _ _ Host) := - forall y1 y2, not ((reification_r_obj GuestHost y1 y2)=x). + Context (CC:CartesianCat center_of_TypesL). - Definition FlatSubCategory := FullSubcategory (TypesL _ _ Host) FlatObject. + Context {kbo}{kbc}(K:@PreMonoidalCat _ _ KE kbo kbc (@one _ _ _ (car_terminal(CartesianCat:=CC)))). - Section Flattening. + (* + Definition K_enrichment : Enrichment. + refine {| enr_c := KE |}. + Defined. + Context (K_surjective:SurjectiveEnrichment K_enrichment). + *) - Context (F:Retraction (TypesL _ _ Host) FlatSubCategory). - Definition FlatteningOfReification := garrow_from_reification Guest Host GuestHost >>>> F. - Lemma FlatteningIsNotDestructive : - FlatteningOfReification >>>> retraction_retraction F >>>> HomFunctor _ (me_i Host) ~~~~ GuestHost. - admit. - Qed. + Definition ArrowInProgrammingLanguage := + @FreydCategory _ _ _ _ _ _ center_of_TypesL _ _ _ _ K. - End Flattening. + Definition ArrowsAreGeneralizedArrows (arrow:ArrowInProgrammingLanguage) : GeneralizedArrow K_enrichment Host. + refine {| ga_functor_monoidal := inclusion_functor_monoidal VK |}. + Defined. End GArrowInLanguage. diff --git a/src/ProgrammingLanguageGeneralizedArrow.v b/src/ProgrammingLanguageGeneralizedArrow.v index 620ca9f..bac2397 100644 --- a/src/ProgrammingLanguageGeneralizedArrow.v +++ b/src/ProgrammingLanguageGeneralizedArrow.v @@ -29,34 +29,20 @@ Require Import NaturalDeductionCategory. Require Import FreydCategories. +Require Import Reification. +Require Import GeneralizedArrow. Require Import GeneralizedArrowFromReification. -Section GArrowInLanguage. +Require Import ProgrammingLanguage. - Definition GeneralizedArrowInLanguage (Guest:ProgrammingLanguageSMME) (Host :ProgrammingLanguageSMME) - := GeneralizedArrow Guest Host. - - Definition ArrowsAreGeneralizedArrows (Host:ProgrammingLanguageSMME) - {mf}{mn}{cc}{kehom}{CC} - (arrow:ArrowInProgrammingLanguage Host _ _ CC mf mn cc kehom) : GeneralizedArrowInLanguage. - - Definition TwoLevelLanguage := Reification Guest Host (me_i Host). +Require Import ReificationsAndGeneralizedArrows. +Require Import ReificationFromGeneralizedArrow. - Context (GuestHost:TwoLevelLanguage). +Section ProgrammingLanguageGeneralizedArrow. - Definition FlatObject (x:TypesL _ _ Host) := - forall y1 y2, not ((reification_r_obj GuestHost y1 y2)=x). + Context (Guest:ProgrammingLanguageSMME) (Host :ProgrammingLanguageSMME). - 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 >>>> HomFunctor _ (me_i Host) ~~~~ GuestHost. - admit. - Qed. + Definition GeneralizedArrowInLanguage + := GeneralizedArrow Guest Host. - End Flattening. +End ProgrammingLanguageGeneralizedArrow. -End GArrowInLanguage. diff --git a/src/ProgrammingLanguageReification.v b/src/ProgrammingLanguageReification.v index 589f748..ab8bb90 100644 --- a/src/ProgrammingLanguageReification.v +++ b/src/ProgrammingLanguageReification.v @@ -29,6 +29,9 @@ Require Import NaturalDeduction. Require Import NaturalDeductionCategory. Require Import ProgrammingLanguage. +Definition TwoLevelLanguage (Guest Host:ProgrammingLanguageSMME) + := Reification Guest Host (me_i Host). + Inductive NLevelLanguage : nat -> ProgrammingLanguageSMME -> Type := | NLevelLanguage_zero : forall lang, NLevelLanguage O lang | NLevelLanguage_succ : forall (L1 L2:ProgrammingLanguageSMME) n, @@ -37,3 +40,4 @@ Inductive NLevelLanguage : nat -> ProgrammingLanguageSMME -> Type := Definition OmegaLevelLanguage : Type := { f : nat -> ProgrammingLanguageSMME & forall n, TwoLevelLanguage (f n) (f (S n)) }. +