X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FReificationCategory.v;h=e57df4da353e70a7a775dce485a32d43d81e6a83;hp=0227f6099e73eb43f94a3ddf6b5cd0ce33e33fb1;hb=ec8ee5cde986e5b38bcae38cda9e63eba94f1d9f;hpb=d2526b193694dd7a5c7ab9d80d6b6656a7459bb9 diff --git a/src/ReificationCategory.v b/src/ReificationCategory.v index 0227f60..e57df4d 100644 --- a/src/ReificationCategory.v +++ b/src/ReificationCategory.v @@ -26,10 +26,13 @@ Require Import Reification. Require Import WeakFunctorCategory. Require Import SmallSMMEs. -(* Technically reifications form merely a *semicategory* (no identity +(* + * 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 + * portion is identical to the original semicategory (closing under + * composition after putting in the identity maps never produces any + * additional maps) * * Also, technically this category has ALL enrichments (not just the * surjective monic monoidal ones), though there maps OUT OF only the @@ -67,19 +70,19 @@ Definition compose_reifications (s0 s1 s2:SMMEs) : ; reification_r := fun K => (reification_r X K) >>>> (reification_r X0 (mon_i s1)) |}. intro K. - set (ni_associativity (reification_r X K) (reification_r X0 (mon_i s1)) (RepresentableFunctor s2 (mon_i s2))) as q. + set (ni_associativity (reification_r X K) (reification_r X0 (mon_i s1)) (HomFunctor s2 (mon_i s2))) as q. eapply ni_comp. apply q. clear q. set (reification_commutes X K) as comm1. set (reification_commutes X0 (mon_i s1)) as comm2. - set (RepresentableFunctor s0 K) as a in *. + set (HomFunctor 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 *. + set (HomFunctor s2 (mon_i s2)) as c' in *. + set (HomFunctor 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.