X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FReification.v;h=932d78ef0890c816d8016435dda06ab796bc916e;hp=a56ded8e745fd5636585deb80d3302fca990c83d;hb=64d416692bda1d36c33b5efa245d46dcf546ad4a;hpb=76f4613eaa5989e29bfd59d716c216ee5386c5f7 diff --git a/src/Reification.v b/src/Reification.v index a56ded8..932d78e 100644 --- a/src/Reification.v +++ b/src/Reification.v @@ -1,8 +1,8 @@ (*********************************************************************************************************************************) (* Reification: *) (* *) -(* A reification is a functor R from one enrichING category A to another enrichING category B which forms a commuting square *) -(* with every pair of hom-functors Hom(X,-):a->A and Hom(Y,-):b->B up to natural isomorphism. *) +(* A reification is a functor R from one enrichING category A to another enrichING category B which, for every X, forms *) +(* a commuting square with Hom(X,-):a->A and Hom(I,-):b->B (up to natural isomorphism). *) (* *) (*********************************************************************************************************************************) @@ -18,25 +18,28 @@ 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 Enrichments. Require Import Enrichment_ch2_8. Require Import RepresentableStructure_ch7_2. -Opaque RepresentableFunctor. +Opaque HomFunctor. Opaque functor_comp. Structure Reification (K:Enrichment) (C:Enrichment) (CI:C) := { reification_r_obj : K -> K -> C ; reification_rstar_obj : enr_v K -> enr_v C ; reification_r : forall k:K, Functor K C (reification_r_obj k) ; reification_rstar_f : Functor (enr_v K) (enr_v C) reification_rstar_obj -; reification_rstar : MonoidalFunctor (enr_v_mon K) (enr_v_mon C) reification_rstar_f -; reification_commutes : forall k, reification_r k >>>> RepresentableFunctor C CI <~~~> RepresentableFunctor K k >>>> reification_rstar_f +; reification_rstar : PreMonoidalFunctor (enr_v_mon K) (enr_v_mon C) reification_rstar_f +; reification_commutes : ∀ k, reification_r k >>>> HomFunctor C CI <~~~> HomFunctor K k >>>> reification_rstar_f }. -Transparent RepresentableFunctor. +Transparent HomFunctor. Transparent functor_comp. -Coercion reification_rstar : Reification >-> MonoidalFunctor. +Coercion reification_rstar : Reification >-> PreMonoidalFunctor. Implicit Arguments Reification [ ]. Implicit Arguments reification_r_obj [ K C CI ]. Implicit Arguments reification_r [ K C CI ].