Documentation

Mathlib.CategoryTheory.Bicategory.CatEnriched

The strict bicategory associated to a Cat-enriched category #

If C is a type with a EnrichedCategory Cat C structure, then it has hom-categories, whose objects define 1-dimensional arrows on C and whose morphisms define 2-dimensional arrows between these. The enriched category axioms equip this data with the structure of a strict bicategory.

We define a type alias CatEnriched C for a type C with a EnrichedCategory Cat C structure. We provide this with an instance of a strict bicategory structure constructing Bicategory.Strict (CatEnriched C).

If C is a type with a EnrichedOrdinaryCategory Cat C structure, then it has an Enrichred Cat C structure, so the previous construction would again produce a strict bicategory. However, in this setting C is also given a Category C structure, together with an equivalence between this category and the underlying category of the Enriched Cat C, and in examples the given category structure is the preferred one.

Thus, we define a type alias CatEnrichedOrdinary C for a type C with an EnrichedOrdinaryCategory Cat C structure. We provide this with an instance of a strict bicategory structure extending the category structure provided by the given instance Category C constructing Bicategory.Strict (CatEnrichedOrdinary C).

A type synonym for C, which should come equipped with a Cat-enriched category structure. This converts it to a strict bicategory where Category (X ⟶ Y) is (X ⟶[Cat] Y).

Equations
Instances For

    Any enriched category has an underlying category structure defined by ForgetEnrichment. This is equivalent but not definitionally equal the category structure constructed here, which is more canonically associated to the data of an EnrichedCategory Cat structure.

    Equations
    • One or more equations did not get rendered due to their size.
    theorem CategoryTheory.CatEnriched.comp_eq {C : Type u_1} [EnrichedCategory Cat C] {X Y Z : CatEnriched C} (f : X Y) (g : Y Z) :
    def CategoryTheory.CatEnriched.hComp {C : Type u_1} [EnrichedCategory Cat C] {a b c : CatEnriched C} {f f' : a b} {g g' : b c} (η : f f') (θ : g g') :

    The horizonal composition on 2-morphisms is defined using the action on arrows of the composition bifunctor from the enriched category structure.

    Equations
    Instances For
      @[simp]
      theorem CategoryTheory.CatEnriched.eqToHom_hComp_eqToHom {C : Type u_1} [EnrichedCategory Cat C] {a b c : CatEnriched C} {f f' : a b} (α : f = f') {g g' : b c} (β : g = g') :
      hComp (eqToHom α) (eqToHom β) = eqToHom
      @[simp]
      theorem CategoryTheory.CatEnriched.hComp_comp {C : Type u_1} [EnrichedCategory Cat C] {a b c : CatEnriched C} {f₁ f₂ f₃ : a b} {g₁ g₂ g₃ : b c} (η : f₁ f₂) (η' : f₂ f₃) (θ : g₁ g₂) (θ' : g₂ g₃) :

      The interchange law for horizontal and vertical composition of 2-cells in a bicategory.

      The action on objects of the EnrichedCategory Cat coherences proves the category axioms.

      Equations

      The category instance on CatEnriched C promotes it to a Cat enriched ordinary category.

      Equations
      • One or more equations did not get rendered due to their size.
      theorem CategoryTheory.CatEnriched.hComp_assoc_heq {C : Type u_1} [EnrichedCategory Cat C] {a b c d : CatEnriched C} {f f' : a b} {g g' : b c} {h h' : c d} (η : f f') (θ : g g') (κ : h h') :
      hComp (hComp η θ) κ hComp η (hComp θ κ)
      theorem CategoryTheory.CatEnriched.hComp_assoc {C : Type u_1} [EnrichedCategory Cat C] {a b c d : CatEnriched C} {f f' : a b} {g g' : b c} {h h' : c d} (η : f f') (θ : g g') (κ : h h') :
      Equations
      • One or more equations did not get rendered due to their size.

      As the associator and left and right unitors are defined as eqToIso of category axioms, the bicategory structure on CatEnriched C is strict.

      A type synonym for C, which should come equipped with a Cat-enriched category structure. This converts it to a strict bicategory where Category (X ⟶ Y) is (X ⟶[Cat] Y).

      Equations
      Instances For

        The forgetful map from the type alias associated to EnrichedOrdinaryCategory Cat C and the type alias associated to EnrichedCategory Cat C is the identity on underlying types.

        Equations
        Instances For

          The hom-types in a Cat-enriched ordinary category are equivalent to the types underlying the hom-categories.

          Equations
          Instances For

            The 2-cells between a parallel pair of 1-cells f g in CatEnrichedOrdinary C are defined to be the morphisms in the hom-categories provided by the EnrichedCategory Cat C structure between the corresponding objects.

            Instances For

              A 2-cell in CatEnrichedOrdinary C has a corresponding "base" 2-cell in CatEnriched C.

              Equations
              Instances For

                A 2-cell in CatEnriched C can be "made" into a 2-cell in CatEnrichedOrdinary C.

                Equations
                Instances For
                  Equations
                  • One or more equations did not get rendered due to their size.
                  theorem CategoryTheory.CatEnrichedOrdinary.Hom.ext {C : Type u} [Category.{v, u} C] [EnrichedOrdinaryCategory Cat C] {X Y : CatEnrichedOrdinary C} {f g : X Y} (α β : f g) (H : base α = base β) :
                  α = β

                  A Cat-enriched ordinary category comes with hom-categories X ⟶[Cat] Y whose underlying type of objects is equivalent to the type X ⟶ Y defined by the category structure on C. The following definition transfers the category structure to the latter type of objects.

                  Equations

                  The horizonal composition on 2-morphisms is defined using the action on arrows of the composition bifunctor from the enriched category structure.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem CategoryTheory.CatEnrichedOrdinary.eqToHom_hComp_eqToHom {C : Type u} [Category.{v, u} C] [EnrichedOrdinaryCategory Cat C] {a b c : CatEnrichedOrdinary C} {f f' : a b} (α : f = f') {g g' : b c} (β : g = g') :
                    hComp (eqToHom α) (eqToHom β) = eqToHom
                    @[simp]
                    theorem CategoryTheory.CatEnrichedOrdinary.hComp_comp {C : Type u} [Category.{v, u} C] [EnrichedOrdinaryCategory Cat C] {a b c : CatEnrichedOrdinary C} {f₁ f₂ f₃ : a b} {g₁ g₂ g₃ : b c} (η : f₁ f₂) (η' : f₂ f₃) (θ : g₁ g₂) (θ' : g₂ g₃) :

                    The interchange law for horizontal and vertical composition of 2-cells in a bicategory.

                    theorem CategoryTheory.CatEnrichedOrdinary.hComp_assoc {C : Type u} [Category.{v, u} C] [EnrichedOrdinaryCategory Cat C] {a b c d : CatEnrichedOrdinary C} {f f' : a b} {g g' : b c} {h h' : c d} (η : f f') (θ : g g') (κ : h h') :
                    theorem CategoryTheory.CatEnrichedOrdinary.hComp_assoc_heq {C : Type u} [Category.{v, u} C] [EnrichedOrdinaryCategory Cat C] {a b c d : CatEnrichedOrdinary C} {f f' : a b} {g g' : b c} {h h' : c d} (η : f f') (θ : g g') (κ : h h') :
                    hComp (hComp η θ) κ hComp η (hComp θ κ)
                    Equations
                    • One or more equations did not get rendered due to their size.

                    As the associator and left and right unitors are defined as eqToIso of category axioms, the bicategory structure on CatEnrichedOrdinary C is strict.