Documentation

Mathlib.CategoryTheory.Bicategory.Basic

Bicategories #

In this file we define typeclass for bicategories.

A bicategory B consists of

We use u, v, and w as the universe variables for objects, 1-morphisms, and 2-morphisms, respectively.

A typeclass for bicategories extends CategoryTheory.CategoryStruct typeclass. This means that we have

For each object a b : B, the collection of 1-morphisms a ⟶ b has a category structure. The 2-morphisms in the bicategory are implemented as the morphisms in this family of categories.

The composition of 1-morphisms is in fact an object part of a functor (a ⟶ b) ⥤ (b ⟶ c) ⥤ (a ⟶ c). The definition of bicategories in this file does not require this functor directly. Instead, it requires the whiskering functions. For a 1-morphism f : a ⟶ b and a 2-morphism η : g ⟶ h between 1-morphisms g h : b ⟶ c, there is a 2-morphism whiskerLeft f η : f ≫ g ⟶ f ≫ h. Similarly, for a 2-morphism η : f ⟶ g between 1-morphisms f g : a ⟶ b and a 1-morphism f : b ⟶ c, there is a 2-morphism whiskerRight η h : f ≫ h ⟶ g ≫ h. These satisfy the exchange law whiskerLeft f θ ≫ whiskerRight η i = whiskerRight η h ≫ whiskerLeft g θ, which is required as an axiom in the definition here.

class CategoryTheory.Bicategory (B : Type u) extends CategoryTheory.CategoryStruct.{v, u} B :
Type (max (max u (v + 1)) (w + 1))

In a bicategory, we can compose the 1-morphisms f : a ⟶ b and g : b ⟶ c to obtain a 1-morphism f ≫ g : a ⟶ c. This composition does not need to be strictly associative, but there is a specified associator, α_ f g h : (f ≫ g) ≫ h ≅ f ≫ (g ≫ h). There is an identity 1-morphism 𝟙 a : a ⟶ a, with specified left and right unitor isomorphisms λ_ f : 𝟙 a ≫ f ≅ f and ρ_ f : f ≫ 𝟙 a ≅ f. These associators and unitors satisfy the pentagon and triangle equations.

See https://ncatlab.org/nlab/show/bicategory.

Instances

    Left whiskering for morphisms

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For

      Right whiskering for morphisms

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For

        The associator isomorphism: (f ≫ g) ≫ h ≅ f ≫ g ≫ h

        Equations
        Instances For

          The left unitor: 𝟙 a ≫ f ≅ f

          Equations
          Instances For

            The right unitor: f ≫ 𝟙 b ≅ f

            Equations
            Instances For

              Simp-normal form for 2-morphisms #

              Rewriting involving associators and unitors could be very complicated. We try to ease this complexity by putting carefully chosen simp lemmas that rewrite any 2-morphisms into simp-normal form defined below. Rewriting into simp-normal form is also useful when applying (forthcoming) coherence tactic.

              The simp-normal form of 2-morphisms is defined to be an expression that has the minimal number of parentheses. More precisely,

              1. it is a composition of 2-morphisms like η₁ ≫ η₂ ≫ η₃ ≫ η₄ ≫ η₅ such that each ηᵢ is either a structural 2-morphisms (2-morphisms made up only of identities, associators, unitors) or non-structural 2-morphisms, and
              2. each non-structural 2-morphism in the composition is of the form f₁ ◁ f₂ ◁ f₃ ◁ η ▷ f₄ ▷ f₅, where each fᵢ is a 1-morphism that is not the identity or a composite and η is a non-structural 2-morphisms that is also not the identity or a composite.

              Note that f₁ ◁ f₂ ◁ f₃ ◁ η ▷ f₄ ▷ f₅ is actually f₁ ◁ (f₂ ◁ (f₃ ◁ ((η ▷ f₄) ▷ f₅))).

              theorem CategoryTheory.Bicategory.whisker_assoc_assoc {B : Type u} [self : Bicategory B] {a b c d : B} (f : a b) {g g' : b c} (η : g g') (h : c d) {Z : a d} (h✝ : CategoryStruct.comp (CategoryStruct.comp f g') h Z) :
              theorem CategoryTheory.Bicategory.comp_whiskerRight_assoc {B : Type u} [self : Bicategory B] {a b c : B} {f g h : a b} (η : f g) (θ : g h) (i : b c) {Z : a c} (h✝ : CategoryStruct.comp h i Z) :
              theorem CategoryTheory.Bicategory.whisker_exchange_assoc {B : Type u} [self : Bicategory B] {a b c : B} {f g : a b} {h i : b c} (η : f g) (θ : h i) {Z : a c} (h✝ : CategoryStruct.comp g i Z) :
              theorem CategoryTheory.Bicategory.comp_whiskerLeft_assoc {B : Type u} [self : Bicategory B] {a b c d : B} (f : a b) (g : b c) {h h' : c d} (η : h h') {Z : a d} (h✝ : CategoryStruct.comp (CategoryStruct.comp f g) h' Z) :
              theorem CategoryTheory.Bicategory.whiskerRight_comp_assoc {B : Type u} [self : Bicategory B] {a b c d : B} {f f' : a b} (η : f f') (g : b c) (h : c d) {Z : a d} (h✝ : CategoryStruct.comp f' (CategoryStruct.comp g h) Z) :
              theorem CategoryTheory.Bicategory.whiskerLeft_comp_assoc {B : Type u} [self : Bicategory B] {a b c : B} (f : a b) {g h i : b c} (η : g h) (θ : h i) {Z : a c} (h✝ : CategoryStruct.comp f i Z) :
              @[simp]
              @[simp]
              theorem CategoryTheory.Bicategory.whiskerLeft_hom_inv_assoc {B : Type u} [Bicategory B] {a b c : B} (f : a b) {g h : b c} (η : g h) {Z : a c} (h✝ : CategoryStruct.comp f g Z) :
              @[simp]
              theorem CategoryTheory.Bicategory.hom_inv_whiskerRight_assoc {B : Type u} [Bicategory B] {a b c : B} {f g : a b} (η : f g) (h : b c) {Z : a c} (h✝ : CategoryStruct.comp f h Z) :
              @[simp]
              @[simp]
              theorem CategoryTheory.Bicategory.whiskerLeft_inv_hom_assoc {B : Type u} [Bicategory B] {a b c : B} (f : a b) {g h : b c} (η : g h) {Z : a c} (h✝ : CategoryStruct.comp f h Z) :
              @[simp]
              theorem CategoryTheory.Bicategory.inv_hom_whiskerRight_assoc {B : Type u} [Bicategory B] {a b c : B} {f g : a b} (η : f g) (h : b c) {Z : a c} (h✝ : CategoryStruct.comp g h Z) :
              def CategoryTheory.Bicategory.whiskerLeftIso {B : Type u} [Bicategory B] {a b c : B} (f : a b) {g h : b c} (η : g h) :

              The left whiskering of a 2-isomorphism is a 2-isomorphism.

              Equations
              Instances For
                @[simp]
                theorem CategoryTheory.Bicategory.whiskerLeftIso_hom {B : Type u} [Bicategory B] {a b c : B} (f : a b) {g h : b c} (η : g h) :
                @[simp]
                theorem CategoryTheory.Bicategory.whiskerLeftIso_inv {B : Type u} [Bicategory B] {a b c : B} (f : a b) {g h : b c} (η : g h) :
                instance CategoryTheory.Bicategory.whiskerLeft_isIso {B : Type u} [Bicategory B] {a b c : B} (f : a b) {g h : b c} (η : g h) [IsIso η] :
                @[simp]
                theorem CategoryTheory.Bicategory.inv_whiskerLeft {B : Type u} [Bicategory B] {a b c : B} (f : a b) {g h : b c} (η : g h) [IsIso η] :
                def CategoryTheory.Bicategory.whiskerRightIso {B : Type u} [Bicategory B] {a b c : B} {f g : a b} (η : f g) (h : b c) :

                The right whiskering of a 2-isomorphism is a 2-isomorphism.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  @[simp]
                  theorem CategoryTheory.Bicategory.whiskerRightIso_inv {B : Type u} [Bicategory B] {a b c : B} {f g : a b} (η : f g) (h : b c) :
                  @[simp]
                  theorem CategoryTheory.Bicategory.whiskerRightIso_hom {B : Type u} [Bicategory B] {a b c : B} {f g : a b} (η : f g) (h : b c) :
                  instance CategoryTheory.Bicategory.whiskerRight_isIso {B : Type u} [Bicategory B] {a b c : B} {f g : a b} (η : f g) (h : b c) [IsIso η] :
                  @[simp]
                  theorem CategoryTheory.Bicategory.inv_whiskerRight {B : Type u} [Bicategory B] {a b c : B} {f g : a b} (η : f g) (h : b c) [IsIso η] :
                  theorem CategoryTheory.Bicategory.associator_naturality_left {B : Type u} [Bicategory B] {a b c d : B} {f f' : a b} (η : f f') (g : b c) (h : c d) :
                  theorem CategoryTheory.Bicategory.whiskerRight_comp_symm {B : Type u} [Bicategory B] {a b c d : B} {f f' : a b} (η : f f') (g : b c) (h : c d) :
                  theorem CategoryTheory.Bicategory.associator_naturality_middle {B : Type u} [Bicategory B] {a b c d : B} (f : a b) {g g' : b c} (η : g g') (h : c d) :
                  theorem CategoryTheory.Bicategory.associator_inv_naturality_middle {B : Type u} [Bicategory B] {a b c d : B} (f : a b) {g g' : b c} (η : g g') (h : c d) :
                  theorem CategoryTheory.Bicategory.whisker_assoc_symm {B : Type u} [Bicategory B] {a b c d : B} (f : a b) {g g' : b c} (η : g g') (h : c d) :
                  theorem CategoryTheory.Bicategory.whisker_assoc_symm_assoc {B : Type u} [Bicategory B] {a b c d : B} (f : a b) {g g' : b c} (η : g g') (h : c d) {Z : a d} (h✝ : CategoryStruct.comp f (CategoryStruct.comp g' h) Z) :
                  theorem CategoryTheory.Bicategory.associator_naturality_right {B : Type u} [Bicategory B] {a b c d : B} (f : a b) (g : b c) {h h' : c d} (η : h h') :
                  theorem CategoryTheory.Bicategory.comp_whiskerLeft_symm {B : Type u} [Bicategory B] {a b c d : B} (f : a b) (g : b c) {h h' : c d} (η : h h') :
                  theorem CategoryTheory.Bicategory.whiskerLeft_iff {B : Type u} [Bicategory B] {a b : B} {f g : a b} (η θ : f g) :
                  theorem CategoryTheory.Bicategory.whiskerRight_iff {B : Type u} [Bicategory B] {a b : B} {f g : a b} (η θ : f g) :
                  @[simp]

                  We state it as a simp lemma, which is regarded as an involved version of id_whiskerRight f g : 𝟙 f ▷ g = 𝟙 (f ≫ g).

                  def CategoryTheory.Bicategory.precomp {B : Type u} [Bicategory B] {a b : B} (c : B) (f : a b) :
                  Functor (b c) (a c)

                  Precomposition of a 1-morphism as a functor.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem CategoryTheory.Bicategory.precomp_obj {B : Type u} [Bicategory B] {a b : B} (c : B) (f : a b) (x✝ : b c) :
                    (precomp c f).obj x✝ = CategoryStruct.comp f x✝
                    @[simp]
                    theorem CategoryTheory.Bicategory.precomp_map {B : Type u} [Bicategory B] {a b : B} (c : B) (f : a b) {X✝ Y✝ : b c} (x✝ : X✝ Y✝) :
                    (precomp c f).map x✝ = whiskerLeft f x✝
                    def CategoryTheory.Bicategory.precomposing {B : Type u} [Bicategory B] (a b c : B) :
                    Functor (a b) (Functor (b c) (a c))

                    Precomposition of a 1-morphism as a functor from the category of 1-morphisms a ⟶ b into the category of functors (b ⟶ c) ⥤ (a ⟶ c).

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[simp]
                      theorem CategoryTheory.Bicategory.precomposing_obj {B : Type u} [Bicategory B] (a b c : B) (f : a b) :
                      (precomposing a b c).obj f = precomp c f
                      @[simp]
                      theorem CategoryTheory.Bicategory.precomposing_map_app {B : Type u} [Bicategory B] (a b c : B) {X✝ Y✝ : a b} (η : X✝ Y✝) (x✝ : b c) :
                      ((precomposing a b c).map η).app x✝ = whiskerRight η x✝
                      def CategoryTheory.Bicategory.postcomp {B : Type u} [Bicategory B] {b c : B} (a : B) (f : b c) :
                      Functor (a b) (a c)

                      Postcomposition of a 1-morphism as a functor.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        @[simp]
                        theorem CategoryTheory.Bicategory.postcomp_obj {B : Type u} [Bicategory B] {b c : B} (a : B) (f : b c) (x✝ : a b) :
                        (postcomp a f).obj x✝ = CategoryStruct.comp x✝ f
                        @[simp]
                        theorem CategoryTheory.Bicategory.postcomp_map {B : Type u} [Bicategory B] {b c : B} (a : B) (f : b c) {X✝ Y✝ : a b} (x✝ : X✝ Y✝) :
                        (postcomp a f).map x✝ = whiskerRight x✝ f
                        def CategoryTheory.Bicategory.postcomposing {B : Type u} [Bicategory B] (a b c : B) :
                        Functor (b c) (Functor (a b) (a c))

                        Postcomposition of a 1-morphism as a functor from the category of 1-morphisms b ⟶ c into the category of functors (a ⟶ b) ⥤ (a ⟶ c).

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          @[simp]
                          theorem CategoryTheory.Bicategory.postcomposing_map_app {B : Type u} [Bicategory B] (a b c : B) {X✝ Y✝ : b c} (η : X✝ Y✝) (x✝ : a b) :
                          ((postcomposing a b c).map η).app x✝ = whiskerLeft x✝ η
                          @[simp]
                          theorem CategoryTheory.Bicategory.postcomposing_obj {B : Type u} [Bicategory B] (a b c : B) (f : b c) :
                          (postcomposing a b c).obj f = postcomp a f
                          def CategoryTheory.Bicategory.associatorNatIsoLeft {B : Type u} [Bicategory B] {b c d : B} (a : B) (g : b c) (h : c d) :

                          Left component of the associator as a natural isomorphism.

                          Equations
                          Instances For
                            @[simp]
                            theorem CategoryTheory.Bicategory.associatorNatIsoLeft_hom_app {B : Type u} [Bicategory B] {b c d : B} (a : B) (g : b c) (h : c d) (X : a b) :
                            @[simp]
                            theorem CategoryTheory.Bicategory.associatorNatIsoLeft_inv_app {B : Type u} [Bicategory B] {b c d : B} (a : B) (g : b c) (h : c d) (X : a b) :
                            def CategoryTheory.Bicategory.associatorNatIsoMiddle {B : Type u} [Bicategory B] {a b c d : B} (f : a b) (h : c d) :
                            ((precomposing a b c).obj f).comp ((postcomposing a c d).obj h) ((postcomposing b c d).obj h).comp ((precomposing a b d).obj f)

                            Middle component of the associator as a natural isomorphism.

                            Equations
                            Instances For
                              @[simp]
                              theorem CategoryTheory.Bicategory.associatorNatIsoMiddle_inv_app {B : Type u} [Bicategory B] {a b c d : B} (f : a b) (h : c d) (X : b c) :
                              @[simp]
                              theorem CategoryTheory.Bicategory.associatorNatIsoMiddle_hom_app {B : Type u} [Bicategory B] {a b c d : B} (f : a b) (h : c d) (X : b c) :
                              def CategoryTheory.Bicategory.associatorNatIsoRight {B : Type u} [Bicategory B] {a b c : B} (f : a b) (g : b c) (d : B) :

                              Right component of the associator as a natural isomorphism.

                              Equations
                              Instances For
                                @[simp]
                                theorem CategoryTheory.Bicategory.associatorNatIsoRight_inv_app {B : Type u} [Bicategory B] {a b c : B} (f : a b) (g : b c) (d : B) (X : c d) :
                                @[simp]
                                theorem CategoryTheory.Bicategory.associatorNatIsoRight_hom_app {B : Type u} [Bicategory B] {a b c : B} (f : a b) (g : b c) (d : B) (X : c d) :