X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FProgrammingLanguageArrow.v;h=f6a3940a864d22edcb0e4fd5306da07c7baaf987;hp=6bf0f24678306312daf12ba7140727177a41b809;hb=2d963cf6994fa510fe67d5bf3852ffcc8090496c;hpb=ec8ee5cde986e5b38bcae38cda9e63eba94f1d9f diff --git a/src/ProgrammingLanguageArrow.v b/src/ProgrammingLanguageArrow.v index 6bf0f24..f6a3940 100644 --- a/src/ProgrammingLanguageArrow.v +++ b/src/ProgrammingLanguageArrow.v @@ -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. @@ -31,38 +34,41 @@ 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... *) - Definition ArrowsAreGeneralizedArrows (Host:ProgrammingLanguageSMME) - {mf}{mn}{cc}{kehom}{CC} - (arrow:ArrowInProgrammingLanguage Host _ _ CC mf mn cc kehom) : GeneralizedArrowInLanguage. + (* a host language: *) + Context `(Host:ProgrammingLanguage). - Definition TwoLevelLanguage := Reification Guest Host (me_i Host). + (* ... for which Types(L) is cartesian: *) + Context (center_of_TypesL:MonoidalCat (TypesL_PreMonoidal Host)). - Context (GuestHost:TwoLevelLanguage). + (* along with a full subcategory of Z(Types(L)) *) + Context {P}(VK:FullSubcategory (Center (TypesL_PreMonoidal Host)) P). - Definition FlatObject (x:TypesL _ _ Host) := - forall y1 y2, not ((reification_r_obj GuestHost y1 y2)=x). + Context (Pobj_unit : P []). + Context (Pobj_closed : forall a b, P a → P b → P (bin_obj(BinoidalCat:=Center_is_PreMonoidal (TypesL_PreMonoidal Host)) a b)). + Definition VKM := + PreMonoidalFullSubcategory_PreMonoidal (Center_is_PreMonoidal (TypesL_PreMonoidal Host)) VK Pobj_unit Pobj_closed. - Definition FlatSubCategory := FullSubcategory (TypesL _ _ Host) FlatObject. + (* a premonoidal category enriched in aforementioned full subcategory *) + Context (Kehom:(Center (TypesL_PreMonoidal Host)) -> (Center (TypesL_PreMonoidal Host)) -> @ob _ _ VK). + Context (KE :@ECategory (@ob _ _ VK) (@hom _ _ VK) VK _ VKM (pmon_I VKM) VKM (Center (TypesL_PreMonoidal Host)) Kehom). + Context {kbo :(Center (TypesL_PreMonoidal Host)) -> (Center (TypesL_PreMonoidal Host)) -> (Center (TypesL_PreMonoidal Host))}. + Context (kbc :@BinoidalCat (Center (TypesL_PreMonoidal Host)) _ (Underlying KE) kbo). - Section Flattening. + Definition one' := @one _ _ _ (car_terminal(CartesianCat:=CC)) + Context (K :@PreMonoidalCat _ _ KE kbo kbc ). + Context (CC:CartesianCat (Center (TypesL_PreMonoidal Host))). - 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 (TypesL_PreMonoidal Host)) _ _ _ _ K. - End Flattening. + Definition ArrowsAreGeneralizedArrows (arrow:ArrowInProgrammingLanguage) : GeneralizedArrow K_enrichment Host. + refine {| ga_functor_monoidal := inclusion_functor_monoidal VK |}. + Defined. End GArrowInLanguage.