Documentation

Mathlib.CategoryTheory.Bicategory.NaturalTransformation.Lax

Transformations between lax functors #

Just as there are natural transformations between functors, there are transformations between lax functors. The equality in the naturality condition of a natural transformation gets replaced by a specified 2-morphism. Now, there are three possible types of transformations (between lax functors):

Main definitions #

Using these, we define three CategoryStruct (scoped) instances on B ⥤ᴸ C, in the Lax.LaxTrans, Lax.Oplax, and Lax.StrongTrans namespaces. The arrows in these CategoryStruct's are given by lax transformations, oplax transformations, and strong transformations respectively.

We also provide API for going between lax transformations and strong transformations:

References #

structure CategoryTheory.Lax.LaxTrans {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F G : LaxFunctor B C) :
Type (max (max (max u₁ v₁) v₂) w₂)

If η is a lax transformation between F and G, we have a 1-morphism η.app a : F.obj a ⟶ G.obj a for each object a : B. We also have a 2-morphism η.naturality f : app a ≫ G.map f ⟶ F.map f ≫ app b for each 1-morphism f : a ⟶ b. These 2-morphisms satisfy the naturality condition, and preserve the identities and the compositions modulo some adjustments of domains and codomains of 2-morphisms.

Instances For
    @[simp]
    theorem CategoryTheory.Lax.LaxTrans.naturality_naturality_assoc {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : LaxFunctor B C} (self : LaxTrans F G) {a b : B} {f g : a b} (η : f g) {Z : F.obj a G.obj b} (h : CategoryStruct.comp (F.map g) (self.app b) Z) :
    def CategoryTheory.Lax.LaxTrans.id {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : LaxFunctor B C) :

    The identity lax transformation.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[reducible, inline]
      abbrev CategoryTheory.Lax.LaxTrans.vCompApp {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : LaxTrans F G) (θ : LaxTrans G H) (a : B) :
      F.obj a H.obj a

      Auxiliary definition for vComp.

      Equations
      Instances For
        @[reducible, inline]
        abbrev CategoryTheory.Lax.LaxTrans.vCompNaturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : LaxTrans F G) (θ : LaxTrans G H) {a b : B} (f : a b) :

        Auxiliary definition for vComp.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          theorem CategoryTheory.Lax.LaxTrans.vComp_naturality_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : LaxTrans F G) (θ : LaxTrans G H) {a b : B} {f g : a b} (β : f g) :
          def CategoryTheory.Lax.LaxTrans.vComp {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : LaxTrans F G) (θ : LaxTrans G H) :

          Vertical composition of lax transformations.

          Equations
          • η.vComp θ = { app := fun (a : B) => η.vCompApp θ a, naturality := fun {a b : B} => η.vCompNaturality θ, naturality_naturality := , naturality_id := , naturality_comp := }
          Instances For
            @[simp]
            theorem CategoryTheory.Lax.LaxTrans.vComp_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : LaxTrans F G) (θ : LaxTrans G H) (a : B) :
            (η.vComp θ).app a = η.vCompApp θ a
            @[simp]
            theorem CategoryTheory.Lax.LaxTrans.vComp_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : LaxTrans F G) (θ : LaxTrans G H) {a✝ b✝ : B} (f : a✝ b✝) :
            (η.vComp θ).naturality f = η.vCompNaturality θ f

            CategoryStruct on B ⥤ᴸ C where the (1-)morphisms are given by lax transformations.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem CategoryTheory.Lax.LaxTrans.id_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : LaxFunctor B C) (a : B) :
              @[simp]
              theorem CategoryTheory.Lax.LaxTrans.comp_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : LaxFunctor B C} (η : LaxTrans X✝ Y✝) (θ : LaxTrans Y✝ Z✝) {a✝ b✝ : B} (f : a✝ b✝) :
              @[simp]
              theorem CategoryTheory.Lax.LaxTrans.comp_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : LaxFunctor B C} (η : LaxTrans X✝ Y✝) (θ : LaxTrans Y✝ Z✝) (a : B) :
              (CategoryStruct.comp η θ).app a = η.vCompApp θ a
              @[simp]
              theorem CategoryTheory.Lax.LaxTrans.id_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : LaxFunctor B C) {x✝ x✝¹ : B} (f : x✝ x✝¹) :
              structure CategoryTheory.Lax.OplaxTrans {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F G : LaxFunctor B C) :
              Type (max (max (max u₁ v₁) v₂) w₂)

              If η is an oplax transformation between F and G, we have a 1-morphism η.app a : F.obj a ⟶ G.obj a for each object a : B. We also have a 2-morphism η.naturality f : F.map f ≫ app b ⟶ app a ≫ G.map f for each 1-morphism f : a ⟶ b. These 2-morphisms satisfies the naturality condition, and preserve the identities and the compositions modulo some adjustments of domains and codomains of 2-morphisms.

              Instances For
                @[simp]
                theorem CategoryTheory.Lax.OplaxTrans.naturality_naturality_assoc {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : LaxFunctor B C} (self : OplaxTrans F G) {a b : B} {f g : a b} (η : f g) {Z : F.obj a G.obj b} (h : CategoryStruct.comp (self.app a) (G.map g) Z) :
                def CategoryTheory.Lax.OplaxTrans.id {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : LaxFunctor B C) :

                The identity oplax transformation.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  @[reducible, inline]
                  abbrev CategoryTheory.Lax.OplaxTrans.vCompApp {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : OplaxTrans F G) (θ : OplaxTrans G H) (a : B) :
                  F.obj a H.obj a

                  Auxiliary definition for vComp.

                  Equations
                  Instances For
                    @[reducible, inline]
                    abbrev CategoryTheory.Lax.OplaxTrans.vCompNaturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : OplaxTrans F G) (θ : OplaxTrans G H) {a b : B} (f : a b) :

                    Auxiliary definition for vComp.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      theorem CategoryTheory.Lax.OplaxTrans.vComp_naturality_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : OplaxTrans F G) (θ : OplaxTrans G H) {a b : B} {f g : a b} (β : f g) :
                      def CategoryTheory.Lax.OplaxTrans.vComp {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : OplaxTrans F G) (θ : OplaxTrans G H) :

                      Vertical composition of oplax transformations.

                      Equations
                      • η.vComp θ = { app := η.vCompApp θ, naturality := fun {a b : B} => η.vCompNaturality θ, naturality_naturality := , naturality_id := , naturality_comp := }
                      Instances For

                        CategoryStruct on B ⥤ᴸ C where the (1-)morphisms are given by oplax transformations.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          @[simp]
                          theorem CategoryTheory.Lax.OplaxTrans.id_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : LaxFunctor B C) (a : B) :
                          @[simp]
                          theorem CategoryTheory.Lax.OplaxTrans.comp_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : LaxFunctor B C} (η : OplaxTrans X✝ Y✝) (θ : OplaxTrans Y✝ Z✝) (a : B) :
                          (CategoryStruct.comp η θ).app a = η.vCompApp θ a
                          @[simp]
                          @[simp]
                          theorem CategoryTheory.Lax.OplaxTrans.comp_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : LaxFunctor B C} (η : OplaxTrans X✝ Y✝) (θ : OplaxTrans Y✝ Z✝) {a✝ b✝ : B} (f : a✝ b✝) :
                          structure CategoryTheory.Lax.StrongTrans {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F G : LaxFunctor B C) :
                          Type (max (max (max u₁ v₁) v₂) w₂)

                          A strong natural transformation between lax functors F and G is a natural transformation that is "natural up to 2-isomorphisms".

                          More precisely, it consists of the following:

                          • a 1-morphism η.app a : F.obj a ⟶ G.obj a for each object a : B.
                          • a 2-isomorphism η.naturality f : app a ≫ G.map f ≅ F.map f ≫ app b for each 1-morphism f : a ⟶ b.
                          • These 2-isomorphisms satisfy the naturality condition, and preserve the identities and the compositions modulo some adjustments of domains and codomains of 2-morphisms.
                          Instances For
                            @[simp]
                            theorem CategoryTheory.Lax.StrongTrans.naturality_naturality_assoc {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : LaxFunctor B C} (self : StrongTrans F G) {a b : B} {f g : a b} (η : f g) {Z : F.obj a G.obj b} (h : CategoryStruct.comp (F.map g) (self.app b) Z) :
                            structure CategoryTheory.Lax.LaxTrans.StrongCore {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : LaxFunctor B C} (η : F G) :
                            Type (max (max u₁ v₁) w₂)

                            A structure on an lax transformation that promotes it to a strong transformation.

                            See Pseudofunctor.StrongTrans.mkOfLax.

                            Instances For
                              def CategoryTheory.Lax.StrongTrans.toLax {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : LaxFunctor B C} (η : StrongTrans F G) :

                              The underlying lax natural transformation of a strong natural transformation.

                              Equations
                              • η.toLax = { app := η.app, naturality := fun {a b : B} (f : a b) => (η.naturality f).hom, naturality_naturality := , naturality_id := , naturality_comp := }
                              Instances For
                                @[simp]
                                theorem CategoryTheory.Lax.StrongTrans.toLax_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : LaxFunctor B C} (η : StrongTrans F G) (a : B) :
                                η.toLax.app a = η.app a
                                @[simp]
                                theorem CategoryTheory.Lax.StrongTrans.toLax_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : LaxFunctor B C} (η : StrongTrans F G) {a✝ b✝ : B} (f : a✝ b✝) :
                                def CategoryTheory.Lax.StrongTrans.mkOfLax {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : LaxFunctor B C} (η : LaxTrans F G) (η' : LaxTrans.StrongCore η) :

                                Construct a strong natural transformation from an lax natural transformation whose naturality 2-morphism is an isomorphism.

                                Equations
                                Instances For
                                  noncomputable def CategoryTheory.Lax.StrongTrans.mkOfLax' {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : LaxFunctor B C} (η : LaxTrans F G) [∀ (a b : B) (f : a b), IsIso (η.naturality f)] :

                                  Construct a strong natural transformation from an lax natural transformation whose naturality 2-morphism is an isomorphism.

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    def CategoryTheory.Lax.StrongTrans.id {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : LaxFunctor B C) :

                                    The identity strong natural transformation.

                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    Instances For
                                      @[simp]
                                      theorem CategoryTheory.Lax.StrongTrans.id_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : LaxFunctor B C) (a : B) :
                                      (id F).app a = (LaxTrans.id F).app a
                                      @[simp]
                                      theorem CategoryTheory.Lax.StrongTrans.id_naturality_inv {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : LaxFunctor B C) {a✝ b✝ : B} (f : a✝ b✝) :
                                      @[simp]
                                      theorem CategoryTheory.Lax.StrongTrans.id_naturality_hom {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : LaxFunctor B C) {a✝ b✝ : B} (f : a✝ b✝) :
                                      @[simp]
                                      theorem CategoryTheory.Lax.StrongTrans.id.toLax {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : LaxFunctor B C) :
                                      def CategoryTheory.Lax.StrongTrans.vComp {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : StrongTrans F G) (θ : StrongTrans G H) :

                                      Vertical composition of strong natural transformations.

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For
                                        @[simp]
                                        theorem CategoryTheory.Lax.StrongTrans.vComp_naturality_hom {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : StrongTrans F G) (θ : StrongTrans G H) {a✝ b✝ : B} (f : a✝ b✝) :
                                        @[simp]
                                        theorem CategoryTheory.Lax.StrongTrans.vComp_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : StrongTrans F G) (θ : StrongTrans G H) (a : B) :
                                        (η.vComp θ).app a = η.toLax.vCompApp θ.toLax a
                                        @[simp]
                                        theorem CategoryTheory.Lax.StrongTrans.vComp_naturality_inv {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : LaxFunctor B C} (η : StrongTrans F G) (θ : StrongTrans G H) {a✝ b✝ : B} (f : a✝ b✝) :

                                        CategoryStruct on B ⥤ᴸ C where the (1-)morphisms are given by strong transformations.

                                        Equations
                                        • One or more equations did not get rendered due to their size.
                                        Instances For
                                          @[simp]
                                          theorem CategoryTheory.Lax.StrongTrans.categoryStruct_comp_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : LaxFunctor B C} (η : StrongTrans X✝ Y✝) (θ : StrongTrans Y✝ Z✝) {a✝ b✝ : B} (f : a✝ b✝) :
                                          (CategoryStruct.comp η θ).naturality f = Bicategory.associator (η.app a✝) (θ.app a✝) (Z✝.map f) ≪≫ Bicategory.whiskerLeftIso (η.app a✝) (θ.naturality f) ≪≫ (Bicategory.associator (η.app a✝) (Y✝.map f) (θ.app b✝)).symm ≪≫ Bicategory.whiskerRightIso (η.naturality f) (θ.app b✝) ≪≫ Bicategory.associator (X✝.map f) (η.app b✝) (θ.app b✝)
                                          @[simp]
                                          theorem CategoryTheory.Lax.StrongTrans.categoryStruct_comp_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : LaxFunctor B C} (η : StrongTrans X✝ Y✝) (θ : StrongTrans Y✝ Z✝) (a : B) :