Documentation

Mathlib.CategoryTheory.Bicategory.NaturalTransformation.Oplax

Transformations between oplax functors #

Just as there are natural transformations between functors, there are transformations between oplax 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 oplax functors):

Main definitions #

Using these, we define three CategoryStruct (scoped) instances on B ⥤ᵒᵖᴸ C, in the Oplax.LaxTrans, Oplax.OplaxTrans, and Oplax.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 oplax transformations and strong transformations:

References #

structure CategoryTheory.Oplax.LaxTrans {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F G : OplaxFunctor 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.Oplax.LaxTrans.naturality_naturality_assoc {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : OplaxFunctor 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.Oplax.LaxTrans.id {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : OplaxFunctor 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.Oplax.LaxTrans.vCompApp {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : OplaxFunctor 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.Oplax.LaxTrans.vCompNaturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : OplaxFunctor 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.Oplax.LaxTrans.vComp_naturality_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : OplaxFunctor B C} (η : LaxTrans F G) (θ : LaxTrans G H) {a b : B} {f g : a b} (β : f g) :
          def CategoryTheory.Oplax.LaxTrans.vComp {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : OplaxFunctor 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

            CategoryStruct on OplaxFunctor 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.Oplax.LaxTrans.comp_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : OplaxFunctor B C} (η : LaxTrans X✝ Y✝) (θ : LaxTrans Y✝ Z✝) {a✝ b✝ : B} (f : a✝ b✝) :
              @[simp]
              @[simp]
              theorem CategoryTheory.Oplax.LaxTrans.comp_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : OplaxFunctor B C} (η : LaxTrans X✝ Y✝) (θ : LaxTrans Y✝ Z✝) (a : B) :
              (CategoryStruct.comp η θ).app a = η.vCompApp θ a
              @[simp]
              theorem CategoryTheory.Oplax.LaxTrans.id_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : OplaxFunctor B C) (a : B) :
              structure CategoryTheory.Oplax.OplaxTrans {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F G : OplaxFunctor 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.Oplax.OplaxTrans.naturality_naturality_assoc {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : OplaxFunctor 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) :
                @[deprecated CategoryTheory.Oplax.OplaxTrans (since := "2025-04-23")]
                def CategoryTheory.OplaxNatTrans {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F G : OplaxFunctor B C) :
                Type (max (max (max u₁ v₁) v₂) w₂)

                Alias of CategoryTheory.Oplax.OplaxTrans.


                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.

                Equations
                Instances For
                  @[simp]
                  theorem CategoryTheory.Oplax.OplaxTrans.whiskerRight_naturality_comp_assoc {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : OplaxFunctor B C} (η : OplaxTrans F G) {a b c : B} {a' : C} (f : a b) (g : b c) (h : G.obj c a') {Z : F.obj a a'} (h✝ : CategoryStruct.comp (η.app a) (CategoryStruct.comp (CategoryStruct.comp (G.map f) (G.map g)) h) Z) :

                  The identity oplax transformation.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    def CategoryTheory.Oplax.OplaxTrans.vcomp {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : OplaxFunctor B C} (η : OplaxTrans F G) (θ : OplaxTrans G H) :

                    Vertical composition of oplax transformations.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    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.Oplax.OplaxTrans.comp_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : OplaxFunctor B C} (η : OplaxTrans X✝ Y✝) (θ : OplaxTrans Y✝ Z✝) {a b : B} (f : a b) :
                        @[simp]
                        @[simp]
                        @[simp]
                        theorem CategoryTheory.Oplax.OplaxTrans.comp_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : OplaxFunctor B C} (η : OplaxTrans X✝ Y✝) (θ : OplaxTrans Y✝ Z✝) (a : B) :
                        structure CategoryTheory.Oplax.StrongTrans {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F G : OplaxFunctor B C) :
                        Type (max (max (max u₁ v₁) v₂) w₂)

                        A strong natural transformation between oplax 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 : F.map f ≫ app b ⟶ app a ≫ G.map f 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
                          @[deprecated CategoryTheory.Oplax.StrongTrans (since := "2025-04-23")]
                          def CategoryTheory.Oplax.StrongOplaxNatTrans {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F G : OplaxFunctor B C) :
                          Type (max (max (max u₁ v₁) v₂) w₂)

                          Alias of CategoryTheory.Oplax.StrongTrans.


                          A strong natural transformation between oplax 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 : F.map f ≫ app b ⟶ app a ≫ G.map f 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.
                          Equations
                          Instances For
                            @[simp]
                            theorem CategoryTheory.Oplax.StrongTrans.naturality_naturality_assoc {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : OplaxFunctor 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 (self.app a) (G.map g) Z) :
                            structure CategoryTheory.Oplax.OplaxTrans.StrongCore {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : OplaxFunctor B C} (η : F G) :
                            Type (max (max u₁ v₁) w₂)

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

                            See Pseudofunctor.StrongTrans.mkOfOplax.

                            Instances For
                              def CategoryTheory.Oplax.StrongTrans.toOplax {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : OplaxFunctor B C} (η : StrongTrans F G) :

                              The underlying oplax natural transformation of a strong natural transformation.

                              Equations
                              • η.toOplax = { app := η.app, naturality := fun {a b : B} (f : a b) => (η.naturality f).hom, naturality_naturality := , naturality_id := , naturality_comp := }
                              Instances For
                                @[simp]
                                theorem CategoryTheory.Oplax.StrongTrans.toOplax_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : OplaxFunctor B C} (η : StrongTrans F G) (a : B) :
                                η.toOplax.app a = η.app a
                                @[simp]
                                theorem CategoryTheory.Oplax.StrongTrans.toOplax_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : OplaxFunctor B C} (η : StrongTrans F G) {a✝ b✝ : B} (f : a✝ b✝) :

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

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

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

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

                                    The identity strong natural transformation.

                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    Instances For
                                      @[simp]
                                      theorem CategoryTheory.Oplax.StrongTrans.id_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : OplaxFunctor B C) (a : B) :
                                      (id F).app a = (OplaxTrans.id F).app a
                                      @[simp]
                                      theorem CategoryTheory.Oplax.StrongTrans.id_naturality_hom {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : OplaxFunctor B C) {a✝ b✝ : B} (f : a✝ b✝) :
                                      @[simp]
                                      theorem CategoryTheory.Oplax.StrongTrans.id_naturality_inv {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] (F : OplaxFunctor B C) {a✝ b✝ : B} (f : a✝ b✝) :
                                      def CategoryTheory.Oplax.StrongTrans.vcomp {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : OplaxFunctor 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.Oplax.StrongTrans.vcomp_naturality_hom {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : OplaxFunctor B C} (η : StrongTrans F G) (θ : StrongTrans G H) {a✝ b✝ : B} (f : a✝ b✝) :
                                        @[simp]
                                        theorem CategoryTheory.Oplax.StrongTrans.vcomp_naturality_inv {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : OplaxFunctor B C} (η : StrongTrans F G) (θ : StrongTrans G H) {a✝ b✝ : B} (f : a✝ b✝) :
                                        @[simp]
                                        theorem CategoryTheory.Oplax.StrongTrans.vcomp_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G H : OplaxFunctor B C} (η : StrongTrans F G) (θ : StrongTrans G H) (a : B) :
                                        (η.vcomp θ).app a = (η.toOplax.vcomp θ.toOplax).app a

                                        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.Oplax.StrongTrans.categoryStruct_comp_naturality {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : OplaxFunctor B C} (η : StrongTrans X✝ Y✝) (θ : StrongTrans Y✝ Z✝) {a✝ b✝ : B} (f : a✝ b✝) :
                                          (CategoryStruct.comp η θ).naturality f = (Bicategory.associator (X✝.map f) (η.app b✝) (θ.app b✝)).symm ≪≫ Bicategory.whiskerRightIso (η.naturality f) (θ.app b✝) ≪≫ Bicategory.associator (η.app a✝) (Y✝.map f) (θ.app b✝) ≪≫ Bicategory.whiskerLeftIso (η.app a✝) (θ.naturality f) ≪≫ (Bicategory.associator (η.app a✝) (θ.app a✝) (Z✝.map f)).symm
                                          @[simp]
                                          theorem CategoryTheory.Oplax.StrongTrans.categoryStruct_comp_app {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {X✝ Y✝ Z✝ : OplaxFunctor B C} (η : StrongTrans X✝ Y✝) (θ : StrongTrans Y✝ Z✝) (a : B) :
                                          theorem CategoryTheory.Oplax.StrongTrans.whiskerLeft_naturality_naturality_app {B : Type u_1} [Bicategory B] {G H : OplaxFunctor B Cat} (θ : StrongTrans G H) {a b : B} {a' : Cat} (f : a' G.obj a) {g h : a b} (β : g h) (X : a') :
                                          CategoryStruct.comp ((θ.app b).map ((G.map₂ β).app (f.obj X))) ((θ.naturality h).hom.app (f.obj X)) = CategoryStruct.comp ((θ.naturality g).hom.app (f.obj X)) ((H.map₂ β).app ((θ.app a).obj (f.obj X)))
                                          theorem CategoryTheory.Oplax.StrongTrans.whiskerRight_naturality_naturality_app {B : Type u_1} [Bicategory B] {F G : OplaxFunctor B Cat} (η : StrongTrans F G) {a b : B} {a' : Cat} {f g : a b} (β : f g) (h : G.obj b a') (X : (F.obj a)) :
                                          CategoryStruct.comp (h.map ((η.app b).map ((F.map₂ β).app X))) (h.map ((η.naturality g).hom.app X)) = CategoryStruct.comp (h.map ((η.naturality f).hom.app X)) (h.map ((G.map₂ β).app ((η.app a).obj X)))
                                          theorem CategoryTheory.Oplax.StrongTrans.whiskerLeft_naturality_comp_app {B : Type u_1} [Bicategory B] {G H : OplaxFunctor B Cat} (θ : StrongTrans G H) {a b c : B} {a' : Cat} (f : a' G.obj a) (g : a b) (h : b c) (X : a') :
                                          CategoryStruct.comp ((θ.naturality (CategoryStruct.comp g h)).hom.app (f.obj X)) ((H.mapComp g h).app ((θ.app a).obj (f.obj X))) = CategoryStruct.comp ((θ.app c).map ((G.mapComp g h).app (f.obj X))) (CategoryStruct.comp ((θ.naturality h).hom.app ((G.map g).obj (f.obj X))) ((H.map h).map ((θ.naturality g).hom.app (f.obj X))))
                                          @[simp]
                                          theorem CategoryTheory.Oplax.StrongTrans.whiskerRight_naturality_comp_assoc {B : Type u₁} [Bicategory B] {C : Type u₂} [Bicategory C] {F G : OplaxFunctor B C} (η : StrongTrans F G) {a b c : B} {a' : C} (f : a b) (g : b c) (h : G.obj c a') {Z : F.obj a a'} (h✝ : CategoryStruct.comp (η.app a) (CategoryStruct.comp (CategoryStruct.comp (G.map f) (G.map g)) h) Z) :
                                          theorem CategoryTheory.Oplax.StrongTrans.whiskerRight_naturality_comp_app {B : Type u_1} [Bicategory B] {F G : OplaxFunctor B Cat} (η : StrongTrans F G) {a b c : B} {a' : Cat} (f : a b) (g : b c) (h : G.obj c a') (X : (F.obj a)) :
                                          CategoryStruct.comp (h.map ((η.naturality (CategoryStruct.comp f g)).hom.app X)) (h.map ((G.mapComp f g).app ((η.app a).obj X))) = CategoryStruct.comp (h.map ((η.app c).map ((F.mapComp f g).app X))) (CategoryStruct.comp (h.map (CategoryStruct.id ((η.app c).obj ((F.map g).obj ((F.map f).obj X))))) (CategoryStruct.comp (h.map ((η.naturality g).hom.app ((F.map f).obj X))) (CategoryStruct.comp (h.map (CategoryStruct.id ((G.map g).obj ((η.app b).obj ((F.map f).obj X))))) (CategoryStruct.comp (h.map ((G.map g).map ((η.naturality f).hom.app X))) (h.map (CategoryStruct.id ((G.map g).obj ((G.map f).obj ((η.app a).obj X)))))))))
                                          theorem CategoryTheory.Oplax.StrongTrans.whiskerLeft_naturality_id_app {B : Type u_1} [Bicategory B] {G H : OplaxFunctor B Cat} (θ : StrongTrans G H) {a : B} {a' : Cat} (f : a' G.obj a) (X : a') :
                                          CategoryStruct.comp ((θ.naturality (CategoryStruct.id a)).hom.app (f.obj X)) ((H.mapId a).app ((θ.app a).obj (f.obj X))) = (θ.app a).map ((G.mapId a).app (f.obj X))
                                          theorem CategoryTheory.Oplax.StrongTrans.whiskerRight_naturality_id_app {B : Type u_1} [Bicategory B] {F G : OplaxFunctor B Cat} (η : StrongTrans F G) {a : B} {a' : Cat} (f : G.obj a a') (X : (F.obj a)) :
                                          CategoryStruct.comp (f.map ((η.naturality (CategoryStruct.id a)).hom.app X)) (f.map ((G.mapId a).app ((η.app a).obj X))) = CategoryStruct.comp (f.map ((η.app a).map ((F.mapId a).app X))) (CategoryStruct.comp (f.map (CategoryStruct.id ((η.app a).obj ((CategoryStruct.id (F.obj a)).obj X)))) (f.map (CategoryStruct.id ((η.app a).obj X))))