X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FProgrammingLanguageArrow.v;h=b613455609918a52475b21b18e56a1c09e1c17c8;hp=6bf0f24678306312daf12ba7140727177a41b809;hb=64d416692bda1d36c33b5efa245d46dcf546ad4a;hpb=ec8ee5cde986e5b38bcae38cda9e63eba94f1d9f diff --git a/src/ProgrammingLanguageArrow.v b/src/ProgrammingLanguageArrow.v index 6bf0f24..b613455 100644 --- a/src/ProgrammingLanguageArrow.v +++ b/src/ProgrammingLanguageArrow.v @@ -31,38 +31,48 @@ Require Import ProgrammingLanguage. Require Import ProgrammingLanguageGeneralizedArrow. Require Import FreydCategories. +Require Import GeneralizedArrow. + 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:ProgrammingLanguage). + + (* ... for which Types(L) is cartesian: *) + Context {MF}(center_of_TypesL:MonoidalCat (TypesL _ _ Host) (fun x => (fst_obj _ _ x),,(snd_obj _ _ x)) MF []). + + (* along with a monoidal subcategory of Z(Types(L)) *) + Context (VK:MonoidalSubCat center_of_TypesL). - Definition ArrowsAreGeneralizedArrows (Host:ProgrammingLanguageSMME) - {mf}{mn}{cc}{kehom}{CC} - (arrow:ArrowInProgrammingLanguage Host _ _ CC mf mn cc kehom) : GeneralizedArrowInLanguage. + (* a premonoidal category enriched in aforementioned full subcategory *) + Context (Kehom:center_of_TypesL -> center_of_TypesL -> @ob _ _ VK). - Definition TwoLevelLanguage := Reification Guest Host (me_i Host). + 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 (GuestHost:TwoLevelLanguage). + Context {kbo:center_of_TypesL -> center_of_TypesL -> center_of_TypesL}. - Definition FlatObject (x:TypesL _ _ Host) := - forall y1 y2, not ((reification_r_obj GuestHost y1 y2)=x). + Context (kbc:@BinoidalCat center_of_TypesL _ (Underlying KE) kbo). - Definition FlatSubCategory := FullSubcategory (TypesL _ _ Host) FlatObject. + Check (@PreMonoidalCat) + Definition one' := @one _ _ _ (car_terminal(CartesianCat:=CC)) + Context (K:@PreMonoidalCat _ _ KE kbo kbc ). + Context (CC:CartesianCat center_of_TypesL). - 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.