mathlib documentation

category_theory.​functor_category

category_theory.​functor_category

@[instance]

functor.category C D gives the category structure on functors and natural transformations between categories C and D.

Notice that if C and D are both small categories at the same universe level, this is another small category at that level. However if C and D are both large categories at the same universe level, this is a small category at the next higher level.

Equations
@[simp]
theorem category_theory.​nat_trans.​vcomp_eq_comp {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {F G H : C D} (α : F G) (β : G H) :

theorem category_theory.​nat_trans.​vcomp_app' {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {F G H : C D} (α : F G) (β : G H) (X : C) :
β).app X = α.app X β.app X

theorem category_theory.​nat_trans.​congr_app {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {F G : C D} {α β : F G} (h : α = β) (X : C) :
α.app X = β.app X

@[simp]
theorem category_theory.​nat_trans.​id_app {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (F : C D) (X : C) :
(𝟙 F).app X = 𝟙 (F.obj X)

@[simp]
theorem category_theory.​nat_trans.​comp_app {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {F G H : C D} (α : F G) (β : G H) (X : C) :
β).app X = α.app X β.app X

theorem category_theory.​nat_trans.​app_naturality {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {E : Type u₃} [category_theory.category E] {F G : C D E} (T : F G) (X : C) {Y Z : D} (f : Y Z) :
(F.obj X).map f (T.app X).app Z = (T.app X).app Y (G.obj X).map f

theorem category_theory.​nat_trans.​naturality_app {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {E : Type u₃} [category_theory.category E] {F G : C D E} (T : F G) (Z : D) {X Y : C} (f : X Y) :
(F.map f).app Z (T.app Y).app Z = (T.app X).app Z (G.map f).app Z

def category_theory.​nat_trans.​hcomp {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {E : Type u₃} [category_theory.category E] {F G : C D} {H I : D E} :
(F G)(H I)(F H G I)

hcomp α β is the horizontal composition of natural transformations.

Equations
@[simp]
theorem category_theory.​nat_trans.​hcomp_app {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {E : Type u₃} [category_theory.category E] {F G : C D} {H I : D E} (α : F G) (β : H I) (X : C) :
β).app X = β.app (F.obj X) I.map (α.app X)

@[simp]
theorem category_theory.​nat_trans.​hcomp_id_app {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {E : Type u₃} [category_theory.category E] {F G : C D} {H : D E} (α : F G) (X : C) :
𝟙 H).app X = H.map (α.app X)

theorem category_theory.​nat_trans.​id_hcomp_app {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {E : Type u₃} [category_theory.category E] {F G : C D} {H : E C} (α : F G) (X : E) :
(𝟙 H α).app X = α.app (H.obj X)

theorem category_theory.​nat_trans.​exchange {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {E : Type u₃} [category_theory.category E] {F G H : C D} {I J K : D E} (α : F G) (β : G H) (γ : I J) (δ : J K) :
β) δ) = α γ β δ

def category_theory.​functor.​flip {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {E : Type u₃} [category_theory.category E] :
C D ED C E

Flip the arguments of a bifunctor. See also currying.lean.

Equations
@[simp]
theorem category_theory.​functor.​flip_obj_obj {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {E : Type u₃} [category_theory.category E] (F : C D E) (c : C) (d : D) :
(F.flip.obj d).obj c = (F.obj c).obj d

@[simp]
theorem category_theory.​functor.​flip_obj_map {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {E : Type u₃} [category_theory.category E] (F : C D E) {c c' : C} (f : c c') (d : D) :
(F.flip.obj d).map f = (F.map f).app d

@[simp]
theorem category_theory.​functor.​flip_map_app {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {E : Type u₃} [category_theory.category E] (F : C D E) {d d' : D} (f : d d') (c : C) :
(F.flip.map f).app c = (F.obj c).map f