Documentation

Mathlib.CategoryTheory.Enriched.Basic

Enriched categories #

We set up the basic theory of V-enriched categories, for V an arbitrary monoidal category.

We do not assume here that V is a concrete category, so there does not need to be an "honest" underlying category!

Use X ⟶[V] Y to obtain the V object of morphisms from X to Y.

This file contains the definitions of V-enriched categories and V-functors.

We don't yet define the V-object of natural transformations between a pair of V-functors (this requires limits in V), but we do provide a presheaf isomorphic to the Yoneda embedding of this object.

We verify that when V = Type v, all these notion reduce to the usual ones.

References #

class CategoryTheory.EnrichedCategory (V : Type v) [Category.{w, v} V] [MonoidalCategory V] (C : Type u₁) :
Type (max (max u₁ v) w)

A V-category is a category enriched in a monoidal category V.

Note that we do not assume that V is a concrete category, so there may not be an "honest" underlying category at all!

Instances

    X ⟶[V] Y is the V object of morphisms from X to Y.

    Equations
      Instances For

        The 𝟙_ V-shaped generalized element giving the identity in a V-enriched category.

        Equations
          Instances For

            The composition V-morphism for a V-enriched category.

            Equations
              Instances For

                A type synonym for C, which should come equipped with a V-enriched category structure. In a moment we will equip this with the W-enriched category structure obtained by applying the functor F : LaxMonoidalFunctor V W to each hom object.

                Equations
                  Instances For

                    Construct an honest category from a Type v-enriched category.

                    Equations
                      Instances For

                        Construct a Type v-enriched category from an honest category.

                        Equations
                          Instances For

                            We verify that an enriched category in Type u is just the same thing as an honest category.

                            Equations
                              Instances For

                                A type synonym for C, which should come equipped with a V-enriched category structure. In a moment we will equip this with the (honest) category structure so that X ⟶ Y is (𝟙_ W) ⟶ (X ⟶[W] Y).

                                We obtain this category by transporting the enrichment in V along the lax monoidal functor coyonedaTensorUnit, then using the equivalence of Type-enriched categories with honest categories.

                                This is sometimes called the "underlying" category of an enriched category, although some care is needed as the functor coyonedaTensorUnit, which always exists, does not necessarily coincide with "the forgetful functor" from V to Type, if such exists. When V is any of Type, Top, AddCommGroup, or Module R, coyonedaTensorUnit is just the usual forgetful functor, however. For V = Algebra R, the usual forgetful functor is coyoneda of R[X], not of R. (Perhaps we should have a typeclass for this situation: ConcreteMonoidal?)

                                Equations
                                  Instances For

                                    Typecheck an object of C as an object of ForgetEnrichment W C.

                                    Equations
                                      Instances For

                                        Typecheck an object of ForgetEnrichment W C as an object of C.

                                        Equations
                                          Instances For
                                            @[simp]
                                            theorem CategoryTheory.ForgetEnrichment.to_of {C : Type u₁} (W : Type v) [Category.{w, v} W] [MonoidalCategory W] [EnrichedCategory W C] (X : C) :
                                            «to» W (of W X) = X

                                            Typecheck a (𝟙_ W)-shaped W-morphism as a morphism in ForgetEnrichment W C.

                                            Equations
                                              Instances For

                                                Typecheck a morphism in ForgetEnrichment W C as a (𝟙_ W)-shaped W-morphism.

                                                Equations
                                                  Instances For
                                                    @[simp]
                                                    @[simp]

                                                    The identity in the "underlying" category of an enriched category.

                                                    structure CategoryTheory.EnrichedFunctor (V : Type v) [Category.{w, v} V] [MonoidalCategory V] (C : Type u₁) [EnrichedCategory V C] (D : Type u₂) [EnrichedCategory V D] :
                                                    Type (max (max u₁ u₂) w)

                                                    A V-functor F between V-enriched categories has a V-morphism from X ⟶[V] Y to F.obj X ⟶[V] F.obj Y, satisfying the usual axioms.

                                                    Instances For
                                                      @[simp]
                                                      theorem CategoryTheory.EnrichedFunctor.map_id_assoc {V : Type v} [Category.{w, v} V] [MonoidalCategory V] {C : Type u₁} [EnrichedCategory V C] {D : Type u₂} [EnrichedCategory V D] (self : EnrichedFunctor V C D) (X : C) {Z : V} (h : EnrichedCategory.Hom (self.obj X) (self.obj X) Z) :
                                                      @[simp]
                                                      theorem CategoryTheory.EnrichedFunctor.map_comp_assoc {V : Type v} [Category.{w, v} V] [MonoidalCategory V] {C : Type u₁} [EnrichedCategory V C] {D : Type u₂} [EnrichedCategory V D] (self : EnrichedFunctor V C D) (X Y Z : C) {Z✝ : V} (h : EnrichedCategory.Hom (self.obj X) (self.obj Z) Z✝) :
                                                      CategoryStruct.comp (eComp V X Y Z) (CategoryStruct.comp (self.map X Z) h) = CategoryStruct.comp (MonoidalCategoryStruct.tensorHom (self.map X Y) (self.map Y Z)) (CategoryStruct.comp (eComp V (self.obj X) (self.obj Y) (self.obj Z)) h)

                                                      The identity enriched functor.

                                                      Equations
                                                        Instances For
                                                          @[simp]
                                                          theorem CategoryTheory.EnrichedFunctor.id_obj (V : Type v) [Category.{w, v} V] [MonoidalCategory V] (C : Type u₁) [EnrichedCategory V C] (X : C) :
                                                          (id V C).obj X = X
                                                          @[simp]
                                                          theorem CategoryTheory.EnrichedFunctor.id_map (V : Type v) [Category.{w, v} V] [MonoidalCategory V] (C : Type u₁) [EnrichedCategory V C] (x✝ x✝¹ : C) :
                                                          (id V C).map x✝ x✝¹ = CategoryStruct.id (EnrichedCategory.Hom x✝ x✝¹)

                                                          Composition of enriched functors.

                                                          Equations
                                                            Instances For
                                                              @[simp]
                                                              theorem CategoryTheory.EnrichedFunctor.comp_map (V : Type v) [Category.{w, v} V] [MonoidalCategory V] {C : Type u₁} {D : Type u₂} {E : Type u₃} [EnrichedCategory V C] [EnrichedCategory V D] [EnrichedCategory V E] (F : EnrichedFunctor V C D) (G : EnrichedFunctor V D E) (x✝ x✝¹ : C) :
                                                              (comp V F G).map x✝ x✝¹ = CategoryStruct.comp (F.map x✝ x✝¹) (G.map (F.obj x✝) (F.obj x✝¹))
                                                              @[simp]
                                                              theorem CategoryTheory.EnrichedFunctor.comp_obj (V : Type v) [Category.{w, v} V] [MonoidalCategory V] {C : Type u₁} {D : Type u₂} {E : Type u₃} [EnrichedCategory V C] [EnrichedCategory V D] [EnrichedCategory V E] (F : EnrichedFunctor V C D) (G : EnrichedFunctor V D E) (X : C) :
                                                              (comp V F G).obj X = G.obj (F.obj X)
                                                              theorem CategoryTheory.EnrichedFunctor.ext (V : Type v) [Category.{w, v} V] [MonoidalCategory V] {C : Type u₁} {D : Type u₂} [EnrichedCategory V C] [EnrichedCategory V D] {F G : EnrichedFunctor V C D} (h_obj : ∀ (X : C), F.obj X = G.obj X) (h_map : ∀ (X Y : C), CategoryStruct.comp (F.map X Y) (eqToHom ) = G.map X Y) :
                                                              F = G

                                                              An enriched functor induces an honest functor of the underlying categories, by mapping the (𝟙_ W)-shaped morphisms.

                                                              Equations
                                                                Instances For

                                                                  We now turn to natural transformations between V-functors.

                                                                  The mostly commonly encountered definition of an enriched natural transformation is a collection of morphisms

                                                                  (𝟙_ W) ⟶ (F.obj X ⟶[V] G.obj X)
                                                                  

                                                                  satisfying an appropriate analogue of the naturality square. (c.f. https://ncatlab.org/nlab/show/enriched+natural+transformation)

                                                                  This is the same thing as a natural transformation F.forget ⟶ G.forget.

                                                                  We formalize this as EnrichedNatTrans F G, which is a Type.

                                                                  However, there's also something much nicer: with appropriate additional hypotheses, there is a V-object EnrichedNatTransObj F G which contains more information, and from which one can recover EnrichedNatTrans F G ≃ (𝟙_ V) ⟶ EnrichedNatTransObj F G.

                                                                  Using these as the hom-objects, we can build a V-enriched category with objects the V-functors.

                                                                  For EnrichedNatTransObj to exist, it suffices to have V braided and complete.

                                                                  Before assuming V is complete, we assume it is braided and define a presheaf enrichedNatTransYoneda F G which is isomorphic to the Yoneda embedding of EnrichedNatTransObj F G whether or not that object actually exists.

                                                                  This presheaf has components (enrichedNatTransYoneda F G).obj A what we call the A-graded enriched natural transformations, which are collections of morphisms

                                                                  A ⟶ (F.obj X ⟶[V] G.obj X)
                                                                  

                                                                  satisfying a similar analogue of the naturality square, this time incorporating a half-braiding on A.

                                                                  (We actually define EnrichedNatTrans F G as the special case A := 𝟙_ V with the trivial half-braiding, and when defining enrichedNatTransYoneda F G we use the half-braidings coming from the ambient braiding on V.)

                                                                  structure CategoryTheory.GradedNatTrans {V : Type v} [Category.{w, v} V] [MonoidalCategory V] {C : Type u₁} [EnrichedCategory V C] {D : Type u₂} [EnrichedCategory V D] (A : Center V) (F G : EnrichedFunctor V C D) :
                                                                  Type (max u₁ w)

                                                                  The type of A-graded natural transformations between V-functors F and G. This is the type of morphisms in V from A to the V-object of natural transformations.

                                                                  Instances For
                                                                    theorem CategoryTheory.GradedNatTrans.ext_iff {V : Type v} {inst✝ : Category.{w, v} V} {inst✝¹ : MonoidalCategory V} {C : Type u₁} {inst✝² : EnrichedCategory V C} {D : Type u₂} {inst✝³ : EnrichedCategory V D} {A : Center V} {F G : EnrichedFunctor V C D} {x y : GradedNatTrans A F G} :
                                                                    x = y x.app = y.app
                                                                    theorem CategoryTheory.GradedNatTrans.ext {V : Type v} {inst✝ : Category.{w, v} V} {inst✝¹ : MonoidalCategory V} {C : Type u₁} {inst✝² : EnrichedCategory V C} {D : Type u₂} {inst✝³ : EnrichedCategory V D} {A : Center V} {F G : EnrichedFunctor V C D} {x y : GradedNatTrans A F G} (app : x.app = y.app) :
                                                                    x = y
                                                                    @[reducible, inline]
                                                                    abbrev CategoryTheory.EnrichedNatTrans {V : Type v} [Category.{w, v} V] [MonoidalCategory V] {C : Type u₁} [EnrichedCategory V C] {D : Type u₂} [EnrichedCategory V D] (F G : EnrichedFunctor V C D) :
                                                                    Type (max u₁ w)

                                                                    A natural transformation between two enriched functors is a 𝟙_ V-graded natural transformation.

                                                                    Equations
                                                                      Instances For

                                                                        The identity natural transformation on an enriched functor.

                                                                        Equations
                                                                          Instances For
                                                                            @[simp]
                                                                            theorem CategoryTheory.EnrichedNatTrans.id_app {V : Type v} [Category.{w, v} V] [MonoidalCategory V] {C : Type u₁} [EnrichedCategory V C] {D : Type u₂} [EnrichedCategory V D] (F : EnrichedFunctor V C D) (X : C) :
                                                                            (id F).app X = eId V (F.obj X)

                                                                            The naturality condition of an enriched natural transformation from F to G as an equality of morphisms Hom X Y ⟶ Hom (F.obj X) (G.obj Y) for X, Y : C.

                                                                            The composition of enriched natural transformations.

                                                                            Equations
                                                                              Instances For

                                                                                A presheaf isomorphic to the Yoneda embedding of the V-object of natural transformations from F to G.

                                                                                Equations
                                                                                  Instances For
                                                                                    @[simp]
                                                                                    theorem CategoryTheory.enrichedNatTransYoneda_map_app {V : Type v} [Category.{w, v} V] [MonoidalCategory V] {C : Type u₁} [EnrichedCategory V C] {D : Type u₂} [EnrichedCategory V D] [BraidedCategory V] (F G : EnrichedFunctor V C D) {X✝ Y✝ : Vᵒᵖ} (f : X✝ Y✝) (σ : GradedNatTrans ((Center.ofBraided V).obj (Opposite.unop X✝)) F G) (X : C) :

                                                                                    We verify that an enriched functor between Type v enriched categories is just the same thing as an honest functor.

                                                                                    Equations
                                                                                      Instances For
                                                                                        @[simp]
                                                                                        theorem CategoryTheory.enrichedFunctorTypeEquivFunctor_symm_apply_map {C : Type u₁} [𝒞 : EnrichedCategory (Type v) C] {D : Type u₂} [𝒟 : EnrichedCategory (Type v) D] (F : Functor C D) (x✝ x✝¹ : C) (f : EnrichedCategory.Hom x✝ x✝¹) :
                                                                                        @[simp]
                                                                                        theorem CategoryTheory.enrichedFunctorTypeEquivFunctor_apply_map {C : Type u₁} [𝒞 : EnrichedCategory (Type v) C] {D : Type u₂} [𝒟 : EnrichedCategory (Type v) D] (F : EnrichedFunctor (Type v) C D) {X✝ Y✝ : C} (f : X✝ Y✝) :

                                                                                        We verify that the presheaf representing natural transformations between Type v-enriched functors is actually represented by the usual type of natural transformations!

                                                                                        Equations
                                                                                          Instances For