mathlib documentation

category_theory.​opposites

category_theory.​opposites

@[instance]

The hom types of the opposite of a category (or graph).

As with the objects, we'll make this irreducible below. Use f.op and f.unop to convert between morphisms of C and morphisms of Cᵒᵖ.

Equations
def category_theory.​has_hom.​hom.​op {C : Type u₁} [category_theory.has_hom C] {X Y : C} :
(X Y)(opposite.op Y opposite.op X)

Equations

Equations
@[simp]
theorem category_theory.​has_hom.​hom.​unop_op {C : Type u₁} [category_theory.has_hom C] {X Y : C} {f : X Y} :
f.op.unop = f

@[simp]
theorem category_theory.​has_hom.​hom.​op_unop {C : Type u₁} [category_theory.has_hom C] {X Y : Cᵒᵖ} {f : X Y} :
f.unop.op = f

@[simp]
theorem category_theory.​op_comp {C : Type u₁} [category_theory.category C] {X Y Z : C} {f : X Y} {g : Y Z} :
(f g).op = g.op f.op

@[simp]
theorem category_theory.​op_id {C : Type u₁} [category_theory.category C] {X : C} :

@[simp]
theorem category_theory.​unop_comp {C : Type u₁} [category_theory.category C] {X Y Z : Cᵒᵖ} {f : X Y} {g : Y Z} :
(f g).unop = g.unop f.unop

@[simp]
theorem category_theory.​unop_id {C : Type u₁} [category_theory.category C] {X : Cᵒᵖ} :

@[simp]
theorem category_theory.​unop_id_op {C : Type u₁} [category_theory.category C] {X : C} :

@[simp]

@[simp]

The functor from the double-opposite of a category to the underlying category.

Equations
@[simp]
theorem category_theory.​unop_unop_map {C : Type u₁} [category_theory.category C] (X Y : C) (f : X Y) :

The functor from a category to its double-opposite.

Equations
@[simp]
theorem category_theory.​functor.​op_map {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (F : C D) (X Y : Cᵒᵖ) (f : X Y) :
F.op.map f = (F.map f.unop).op

Equations
@[simp]
theorem category_theory.​functor.​op_obj {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (F : C D) (X : Cᵒᵖ) :

@[simp]
theorem category_theory.​functor.​unop_map {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (F : Cᵒᵖ Dᵒᵖ) (X Y : C) (f : X Y) :
F.unop.map f = (F.map f.op).unop

Equations
@[simp]
theorem category_theory.​functor.​unop_obj {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (F : Cᵒᵖ Dᵒᵖ) (X : C) :

def category_theory.​functor.​op_unop_iso {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (F : C D) :
F.op.unop F

The isomorphism between F.op.unop and F.

Equations

The isomorphism between F.unop.op and F.

Equations

Equations
@[simp]
theorem category_theory.​functor.​op_hom_map_app (C : Type u₁) [category_theory.category C] (D : Type u₂) [category_theory.category D] (F G : (C D)ᵒᵖ) (α : F G) (X : Cᵒᵖ) :

Equations

Equations
@[simp]

@[simp]
theorem category_theory.​functor.​left_op_map {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (F : C Dᵒᵖ) (X Y : Cᵒᵖ) (f : X Y) :
F.left_op.map f = (F.map f.unop).unop

@[simp]
theorem category_theory.​functor.​right_op_map {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (F : Cᵒᵖ D) (X Y : C) (f : X Y) :
F.right_op.map f = (F.map f.op).op

@[simp]
theorem category_theory.​functor.​right_op_obj {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (F : Cᵒᵖ D) (X : C) :

Equations
@[instance]

If F is faithful then the right_op of F is also faithful.

Equations
@[instance]

If F is faithful then the left_op of F is also faithful.

Equations
def category_theory.​nat_trans.​op {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {F G : C D} :
(F G)(G.op F.op)

Equations
@[simp]
theorem category_theory.​nat_trans.​op_app {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {F G : C D} (α : F G) (X : Cᵒᵖ) :

def category_theory.​nat_trans.​unop {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {F G : C D} :
(F.op G.op)(G F)

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

def category_theory.​nat_trans.​left_op {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {F G : C Dᵒᵖ} :
(F G)(G.left_op F.left_op)

Equations
@[simp]

def category_theory.​nat_trans.​right_op {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {F G : C Dᵒᵖ} :
(F.left_op G.left_op)(G F)

Equations
@[simp]

def category_theory.​iso.​op {C : Type u₁} [category_theory.category C] {X Y : C} :
(X Y)(opposite.op Y opposite.op X)

Equations
@[simp]
theorem category_theory.​iso.​op_hom {C : Type u₁} [category_theory.category C] {X Y : C} {α : X Y} :
α.op.hom = α.hom.op

@[simp]
theorem category_theory.​iso.​op_inv {C : Type u₁} [category_theory.category C] {X Y : C} {α : X Y} :
α.op.inv = α.inv.op

def category_theory.​nat_iso.​op {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {F G : C D} :
(F G)(G.op F.op)

The natural isomorphism between opposite functors G.op ≅ F.op induced by a natural isomorphism between the original functors F ≅ G.

Equations
def category_theory.​nat_iso.​unop {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] {F G : C D} :
(F.op G.op)(G F)

The natural isomorphism between functors G ≅ F induced by a natural isomorphism between the opposite functors F.op ≅ G.op.

Equations

The equivalence between arrows of the form A ⟶ B and B.unop ⟶ A.unop. Useful for building adjunctions. Note that this (definitionally) gives variants

def op_equiv' (A : C) (B : Cᵒᵖ) : (opposite.op A  B)  (B.unop  A) :=
op_equiv _ _

def op_equiv'' (A : Cᵒᵖ) (B : C) : (A  opposite.op B)  (B  A.unop) :=
op_equiv _ _

def op_equiv''' (A B : C) : (opposite.op A  opposite.op B)  (B  A) :=
op_equiv _ _
Equations
@[simp]
theorem category_theory.​op_equiv_apply {C : Type u₁} [category_theory.category C] (A B : Cᵒᵖ) (f : A B) :