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):
- Lax natural transformations;
- OpLax natural transformations;
- Strong natural transformations. These differ in the direction (and invertibility) of the 2-morphisms involved in the naturality condition.
Main definitions #
Lax.LaxTrans F G: lax transformations between lax functorsFandG. The naturality condition is given by a 2-morphismapp a ≫ G.map f ⟶ F.map f ≫ app bfor each 1-morphismf : a ⟶ b.Lax.OplaxTrans F G: oplax transformations between lax functorsFandG. The naturality condition is given by a 2-morphismF.map f ≫ app b ⟶ app a ≫ G.map ffor each 1-morphismf : a ⟶ b.Lax.StrongTrans F G: Strong transformations between lax functorsFandG. The naturality condition is given by a 2-morphismapp a ≫ G.map f ≅ F.map f ≫ app bfor each 1-morphismf : a ⟶ b.
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:
Lax.StrongCore F G: a structure on a lax transformation between lax functors that promotes it to a strong transformation.Lax.mkOfLax η η': given an lax transformationηsuch that each component 2-morphism is an isomorphism,mkOfLaxgives the corresponding strong transformation.
References #
- Niles Johnson, Donald Yau, 2-Dimensional Categories, section 4.2.
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.
The component 1-morphisms of a lax transformation.
- naturality {a b : B} (f : a ⟶ b) : CategoryStruct.comp (self.app a) (G.map f) ⟶ CategoryStruct.comp (F.map f) (self.app b)
The 2-morphisms underlying the lax naturality constraint.
- naturality_naturality {a b : B} {f g : a ⟶ b} (η : f ⟶ g) : CategoryStruct.comp (self.naturality f) (Bicategory.whiskerRight (F.map₂ η) (self.app b)) = CategoryStruct.comp (Bicategory.whiskerLeft (self.app a) (G.map₂ η)) (self.naturality g)
Naturality of the lax naturality constraint.
- naturality_id (a : B) : CategoryStruct.comp (Bicategory.whiskerLeft (self.app a) (G.mapId a)) (self.naturality (CategoryStruct.id a)) = CategoryStruct.comp (Bicategory.rightUnitor (self.app a)).hom (CategoryStruct.comp (Bicategory.leftUnitor (self.app a)).inv (Bicategory.whiskerRight (F.mapId a) (self.app a)))
Lax unity.
- naturality_comp {a b c : B} (f : a ⟶ b) (g : b ⟶ c) : CategoryStruct.comp (Bicategory.whiskerLeft (self.app a) (G.mapComp f g)) (self.naturality (CategoryStruct.comp f g)) = CategoryStruct.comp (Bicategory.associator (self.app a) (G.map f) (G.map g)).inv (CategoryStruct.comp (Bicategory.whiskerRight (self.naturality f) (G.map g)) (CategoryStruct.comp (Bicategory.associator (F.map f) (self.app b) (G.map g)).hom (CategoryStruct.comp (Bicategory.whiskerLeft (F.map f) (self.naturality g)) (CategoryStruct.comp (Bicategory.associator (F.map f) (F.map g) (self.app c)).inv (Bicategory.whiskerRight (F.mapComp f g) (self.app c))))))
Lax functoriality.
Instances For
The identity lax transformation.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
Auxiliary definition for vComp.
Equations
- η.vCompApp θ a = CategoryTheory.CategoryStruct.comp (η.app a) (θ.app a)
Instances For
Auxiliary definition for vComp.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Vertical composition of lax transformations.
Equations
Instances For
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
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.
The component 1-morphisms of an oplax transformation.
- naturality {a b : B} (f : a ⟶ b) : CategoryStruct.comp (F.map f) (self.app b) ⟶ CategoryStruct.comp (self.app a) (G.map f)
The 2-morphisms underlying the oplax naturality constraint.
- naturality_naturality {a b : B} {f g : a ⟶ b} (η : f ⟶ g) : CategoryStruct.comp (Bicategory.whiskerRight (F.map₂ η) (self.app b)) (self.naturality g) = CategoryStruct.comp (self.naturality f) (Bicategory.whiskerLeft (self.app a) (G.map₂ η))
Naturality of the oplax naturality constraint.
- naturality_id (a : B) : CategoryStruct.comp (Bicategory.whiskerRight (F.mapId a) (self.app a)) (self.naturality (CategoryStruct.id a)) = CategoryStruct.comp (Bicategory.leftUnitor (self.app a)).hom (CategoryStruct.comp (Bicategory.rightUnitor (self.app a)).inv (Bicategory.whiskerLeft (self.app a) (G.mapId a)))
- naturality_comp {a b c : B} (f : a ⟶ b) (g : b ⟶ c) : CategoryStruct.comp (Bicategory.whiskerRight (F.mapComp f g) (self.app c)) (self.naturality (CategoryStruct.comp f g)) = CategoryStruct.comp (Bicategory.associator (F.map f) (F.map g) (self.app c)).hom (CategoryStruct.comp (Bicategory.whiskerLeft (F.map f) (self.naturality g)) (CategoryStruct.comp (Bicategory.associator (F.map f) (self.app b) (G.map g)).inv (CategoryStruct.comp (Bicategory.whiskerRight (self.naturality f) (G.map g)) (CategoryStruct.comp (Bicategory.associator (self.app a) (G.map f) (G.map g)).hom (Bicategory.whiskerLeft (self.app a) (G.mapComp f g))))))
Instances For
The identity oplax transformation.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
Auxiliary definition for vComp.
Equations
- η.vCompApp θ a = CategoryTheory.CategoryStruct.comp (η.app a) (θ.app a)
Instances For
Auxiliary definition for vComp.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Vertical composition of oplax transformations.
Equations
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
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 afor each objecta : B. - a 2-isomorphism
η.naturality f : app a ≫ G.map f ≅ F.map f ≫ app bfor each 1-morphismf : 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.
- naturality {a b : B} (f : a ⟶ b) : CategoryStruct.comp (self.app a) (G.map f) ≅ CategoryStruct.comp (F.map f) (self.app b)
- naturality_naturality {a b : B} {f g : a ⟶ b} (η : f ⟶ g) : CategoryStruct.comp (self.naturality f).hom (Bicategory.whiskerRight (F.map₂ η) (self.app b)) = CategoryStruct.comp (Bicategory.whiskerLeft (self.app a) (G.map₂ η)) (self.naturality g).hom
- naturality_id (a : B) : CategoryStruct.comp (Bicategory.whiskerLeft (self.app a) (G.mapId a)) (self.naturality (CategoryStruct.id a)).hom = CategoryStruct.comp (Bicategory.rightUnitor (self.app a)).hom (CategoryStruct.comp (Bicategory.leftUnitor (self.app a)).inv (Bicategory.whiskerRight (F.mapId a) (self.app a)))
Lax unity.
- naturality_comp {a b c : B} (f : a ⟶ b) (g : b ⟶ c) : CategoryStruct.comp (Bicategory.whiskerLeft (self.app a) (G.mapComp f g)) (self.naturality (CategoryStruct.comp f g)).hom = CategoryStruct.comp (Bicategory.associator (self.app a) (G.map f) (G.map g)).inv (CategoryStruct.comp (Bicategory.whiskerRight (self.naturality f).hom (G.map g)) (CategoryStruct.comp (Bicategory.associator (F.map f) (self.app b) (G.map g)).hom (CategoryStruct.comp (Bicategory.whiskerLeft (F.map f) (self.naturality g).hom) (CategoryStruct.comp (Bicategory.associator (F.map f) (F.map g) (self.app c)).inv (Bicategory.whiskerRight (F.mapComp f g) (self.app c))))))
Lax functoriality.
Instances For
A structure on an lax transformation that promotes it to a strong transformation.
See Pseudofunctor.StrongTrans.mkOfLax.
- naturality {a b : B} (f : a ⟶ b) : CategoryStruct.comp (η.app a) (G.map f) ≅ CategoryStruct.comp (F.map f) (η.app b)
The underlying 2-isomorphisms of the naturality constraint.
The 2-isomorphisms agree with the underlying 2-morphism of the lax transformation.
Instances For
The underlying lax natural transformation of a strong natural transformation.
Equations
Instances For
Construct a strong natural transformation from an lax natural transformation whose naturality 2-morphism is an isomorphism.
Equations
- CategoryTheory.Lax.StrongTrans.mkOfLax η η' = { app := η.app, naturality := fun {a b : B} => η'.naturality, naturality_naturality := ⋯, naturality_id := ⋯, naturality_comp := ⋯ }
Instances For
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
The identity strong natural transformation.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
Vertical composition of strong natural 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 strong
transformations.
Equations
- One or more equations did not get rendered due to their size.