X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FReificationCategory.v;h=0227f6099e73eb43f94a3ddf6b5cd0ce33e33fb1;hp=b4d444abb8bc25471064c1e415c01d099b6f8678;hb=d2526b193694dd7a5c7ab9d80d6b6656a7459bb9;hpb=f9297f3fba59884fe98eb4f9257a1fb7552bfd0a diff --git a/src/ReificationCategory.v b/src/ReificationCategory.v index b4d444a..0227f60 100644 --- a/src/ReificationCategory.v +++ b/src/ReificationCategory.v @@ -26,25 +26,44 @@ Require Import Reification. 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 ReificationOrIdentity : SMMEs -> SMMEs -> Type := | roi_id : forall smme:SMMEs, ReificationOrIdentity smme smme | roi_reif : forall s1 s2:SMMEs, Reification s1 s2 (mon_i s2) -> ReificationOrIdentity s1 s2. +Definition reificationOrIdentityFobj s1 s2 (f:ReificationOrIdentity s1 s2) : s1 -> s2 := + match f with + | roi_id s => (fun x => x) + | roi_reif s1 s2 f => reification_rstar_obj f + end. + Definition reificationOrIdentityFunc - : forall s1 s2, ReificationOrIdentity s1 s2 -> { fobj : _ & Functor s1 s2 fobj }. + : forall s1 s2 (f:ReificationOrIdentity s1 s2), Functor (enr_v s1) (enr_v s2) (reificationOrIdentityFobj s1 s2 f). intros. - destruct X. - exists (fun x => x). + destruct f. apply functor_id. - exists (reification_rstar_obj r). - apply r. + unfold reificationOrIdentityFobj. + apply (reification_rstar_f r). Defined. Definition compose_reifications (s0 s1 s2:SMMEs) : Reification s0 s1 (mon_i s1) -> Reification s1 s2 (mon_i s2) -> Reification s0 s2 (mon_i s2). intros. refine - {| reification_rstar := MonoidalFunctorsCompose _ _ _ _ _ (reification_rstar X) (reification_rstar X0) + {| reification_rstar_f := reification_rstar_f X >>>> reification_rstar_f X0 + ; reification_rstar := MonoidalFunctorsCompose _ _ _ _ _ (reification_rstar X) (reification_rstar X0) ; reification_r := fun K => (reification_r X K) >>>> (reification_r X0 (mon_i s1)) |}. intro K. @@ -54,7 +73,28 @@ Definition compose_reifications (s0 s1 s2:SMMEs) : clear q. set (reification_commutes X K) as comm1. set (reification_commutes X0 (mon_i s1)) as comm2. - admit. + set (RepresentableFunctor s0 K) as a in *. + set (reification_rstar_f X) as a' in *. + set (reification_rstar_f X0) as x in *. + set (reification_r X K) as b in *. + set (reification_r X0 (mon_i s1)) as c in *. + set (RepresentableFunctor s2 (mon_i s2)) as c' in *. + set (RepresentableFunctor s1 (mon_i s1)) as b' in *. + apply (ni_comp(F2:=b >>>> (b' >>>> x))). + apply (@ni_respects _ _ _ _ _ _ _ _ _ _ b _ b _ (c >>>> c') _ (b' >>>> x)). + apply ni_id. + apply comm2. + eapply ni_comp. + eapply ni_inv. + apply (ni_associativity b b' x). + eapply ni_inv. + eapply ni_comp. + eapply ni_inv. + apply (ni_associativity a a' x). + apply (@ni_respects _ _ _ _ _ _ _ _ _ _ (a >>>> a') _ (b >>>> b') _ x _ x). + apply ni_inv. + apply comm1. + apply ni_id. Defined. Definition reificationOrIdentityComp @@ -70,7 +110,7 @@ Definition reificationOrIdentityComp Definition MorphismsOfCategoryOfReifications : @SmallFunctors SMMEs. refine {| small_func := ReificationOrIdentity ; small_func_id := fun s => roi_id s - ; small_func_func := fun smme1 smme2 f => projT2 (reificationOrIdentityFunc _ _ f) + ; small_func_func := fun smme1 smme2 f => reificationOrIdentityFunc _ _ f ; small_func_comp := reificationOrIdentityComp |}; intros; simpl. apply if_id.