X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FGeneralizedArrowCategory.v;h=745484b91ca9639873fd681e18d9ea63cea2a1ef;hp=a9bedbb15ba3decb525f3826a1dc0fa359f09f68;hb=68f41d71d573b422b04ed3f4a3eb3ab41de09a79;hpb=f9297f3fba59884fe98eb4f9257a1fb7552bfd0a diff --git a/src/GeneralizedArrowCategory.v b/src/GeneralizedArrowCategory.v index a9bedbb..745484b 100644 --- a/src/GeneralizedArrowCategory.v +++ b/src/GeneralizedArrowCategory.v @@ -18,40 +18,52 @@ 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 MonoidalCategories_ch7_8. Require Import Coherence_ch7_8. Require Import Enrichment_ch2_8. +Require Import Enrichments. Require Import RepresentableStructure_ch7_2. Require Import GeneralizedArrow. Require Import WeakFunctorCategory. -Require Import SmallSMMEs. +(* + * Technically reifications form merely a *semicategory* (no identity + * maps), but one can always freely adjoin identity maps (and nothing + * else) to a semicategory to get a category whose non-identity-map + * portion is identical to the original semicategory + * + * Also, technically this category has ALL enrichments (not just the + * surjective monic monoidal ones), though there maps OUT OF only the + * surjective enrichments and INTO only the monic monoidal + * enrichments. It's a big pain to do this in Coq, but sort of might + * matter in real life: a language with really severe substructural + * restrictions might fail to be monoidally enriched, meaning we can't + * use it as a host language. But that's for the next paper... + *) Inductive GeneralizedArrowOrIdentity : SMMEs -> SMMEs -> Type := | gaoi_id : forall smme:SMMEs, GeneralizedArrowOrIdentity smme smme | gaoi_ga : forall s1 s2:SMMEs, GeneralizedArrow s1 s2 -> GeneralizedArrowOrIdentity s1 s2. -Definition generalizedArrowOrIdentityFunc - : forall s1 s2, GeneralizedArrowOrIdentity s1 s2 -> { fobj : _ & Functor s1 s2 fobj }. - intros. - destruct X. - exists (fun x => x). - apply functor_id. - eapply existT. - apply (g >>>> RepresentableFunctor s2 (mon_i s2)). - Defined. +Definition generalizedArrowOrIdentityFobj (s1 s2:SMMEs) (f:GeneralizedArrowOrIdentity s1 s2) : s1 -> s2 := + match f in GeneralizedArrowOrIdentity S1 S2 return S1 -> S2 with + | gaoi_id s => fun x => x + | gaoi_ga s1 s2 f => fun a => ehom(ECategory:=s2) (enr_c_i (smme_e s2)) (ga_functor_obj f a) + end. -Definition compose_generalizedArrows (s0 s1 s2:SMMEs) : - GeneralizedArrow s0 s1 -> GeneralizedArrow s1 s2 -> GeneralizedArrow s0 s2. - intro g01. - intro g12. - refine - {| ga_functor := g01 >>>> RepresentableFunctor s1 (mon_i s1) >>>> g12 |}. - apply MonoidalFunctorsCompose. - apply MonoidalFunctorsCompose. - apply (ga_functor_monoidal g01). - apply (me_mf s1). - apply (ga_functor_monoidal g12). - Defined. +Definition generalizedArrowOrIdentityFunc s1 s2 (f:GeneralizedArrowOrIdentity s1 s2) + : Functor s1 s2 (generalizedArrowOrIdentityFobj _ _ f) := + match f with + | gaoi_id s => functor_id _ + | gaoi_ga s1 s2 f => ga_functor f >>>> HomFunctor s2 (enr_c_i s2) + end. + +Instance compose_generalizedArrows (s0 s1 s2:SMMEs) + (g01:GeneralizedArrow s0 s1)(g12:GeneralizedArrow s1 s2) : GeneralizedArrow s0 s2 := + { ga_functor_monoidal := g01 >>⊗>> smme_mon s1 >>⊗>> g12 }. + apply ga_host_lang_pure. + Defined. Definition generalizedArrowOrIdentityComp : forall s1 s2 s3, GeneralizedArrowOrIdentity s1 s2 -> GeneralizedArrowOrIdentity s2 s3 -> GeneralizedArrowOrIdentity s1 s3. @@ -66,7 +78,7 @@ Definition generalizedArrowOrIdentityComp Definition MorphismsOfCategoryOfGeneralizedArrows : @SmallFunctors SMMEs. refine {| small_func := GeneralizedArrowOrIdentity ; small_func_id := fun s => gaoi_id s - ; small_func_func := fun smme1 smme2 f => projT2 (generalizedArrowOrIdentityFunc _ _ f) + ; small_func_func := fun smme1 smme2 f => generalizedArrowOrIdentityFunc _ _ f ; small_func_comp := generalizedArrowOrIdentityComp |}; intros; simpl. apply if_id. @@ -78,7 +90,6 @@ Definition MorphismsOfCategoryOfGeneralizedArrows : @SmallFunctors SMMEs. apply if_right_identity. unfold mf_F. idtac. - unfold mf_f. apply if_associativity. Defined.