Skip to content

Commit 9419c03

Browse files
smorel394smorel394
andcommitted
chore(CategoryTheory/Adjunction/Opposites): redefine the opposite of an adjunction (#20375)
Change the definition of `Adjunction.adjointOfOpAdjointOp` and `Adjunction.opAdjointOpOfAdjoint` to be more in line with the new definition of adjunctions from #16317 (which removes the `homEquiv` field), and rename these to `Adjunction.op` and `Adjunction.unop` to be consistent with `Equivalence.op` and `Equivalence.unop`. The 6 other variants of opposite adjunctions are all defined from `Adjunction.op` or `Adjunction.unop` and should not be needed anymore, they have been marked as `deprecated`, as well as `Adjunction.adjointOfOpAdjointOp` and `Adjunction.opAdjointOpOfAdjoint`. Co-authored-by: smorel394 <67864981+smorel394@users.noreply.github.com>
1 parent bd681bc commit 9419c03

File tree

1 file changed

+23
-59
lines changed

1 file changed

+23
-59
lines changed

Mathlib/CategoryTheory/Adjunction/Opposites.lean

Lines changed: 23 additions & 59 deletions
Original file line numberDiff line numberDiff line change
@@ -29,66 +29,31 @@ namespace CategoryTheory.Adjunction
2929

3030
attribute [local simp] homEquiv_unit homEquiv_counit
3131

32-
/-- If `G.op` is adjoint to `F.op` then `F` is adjoint to `G`. -/
33-
@[simps! unit_app counit_app]
34-
def adjointOfOpAdjointOp (F : C ⥤ D) (G : D ⥤ C) (h : G.op ⊣ F.op) : F ⊣ G :=
35-
Adjunction.mkOfHomEquiv {
36-
homEquiv := fun {X Y} =>
37-
((h.homEquiv (Opposite.op Y) (Opposite.op X)).trans (opEquiv _ _)).symm.trans
38-
(opEquiv _ _)
39-
homEquiv_naturality_left_symm := by
40-
-- Porting note: This proof was handled by `obviously` in mathlib3. The only obstruction to
41-
-- automation fully kicking in here is that the `@[simps]` lemmas of `opEquiv` and
42-
-- `homEquiv` aren't firing.
43-
intros
44-
simp [opEquiv, homEquiv]
45-
homEquiv_naturality_right := by
46-
-- Porting note: This proof was handled by `obviously` in mathlib3. The only obstruction to
47-
-- automation fully kicking in here is that the `@[simps]` lemmas of `opEquiv` and
48-
-- `homEquiv` aren't firing.
49-
intros
50-
simp [opEquiv, homEquiv] }
51-
52-
/-- If `G` is adjoint to `F.op` then `F` is adjoint to `G.unop`. -/
53-
def adjointUnopOfAdjointOp (F : C ⥤ D) (G : Dᵒᵖ ⥤ Cᵒᵖ) (h : G ⊣ F.op) : F ⊣ G.unop :=
54-
adjointOfOpAdjointOp F G.unop (h.ofNatIsoLeft G.opUnopIso.symm)
55-
56-
/-- If `G.op` is adjoint to `F` then `F.unop` is adjoint to `G`. -/
57-
def unopAdjointOfOpAdjoint (F : Cᵒᵖ ⥤ Dᵒᵖ) (G : D ⥤ C) (h : G.op ⊣ F) : F.unop ⊣ G :=
58-
adjointOfOpAdjointOp _ _ (h.ofNatIsoRight F.opUnopIso.symm)
59-
6032
/-- If `G` is adjoint to `F` then `F.unop` is adjoint to `G.unop`. -/
61-
def unopAdjointUnopOfAdjoint (F : Cᵒᵖ ⥤ Dᵒᵖ) (G : Dᵒᵖ ⥤ Cᵒᵖ) (h : G ⊣ F) : F.unop ⊣ G.unop :=
62-
adjointUnopOfAdjointOp F.unop G (h.ofNatIsoRight F.opUnopIso.symm)
33+
@[simps]
34+
def unop {F : Cᵒᵖ ⥤ Dᵒᵖ} {G : Dᵒᵖ ⥤ Cᵒᵖ} (h : G ⊣ F) : F.unop ⊣ G.unop where
35+
unit := NatTrans.unop h.counit
36+
counit := NatTrans.unop h.unit
37+
left_triangle_components _ := Quiver.Hom.op_inj (h.right_triangle_components _)
38+
right_triangle_components _ := Quiver.Hom.op_inj (h.left_triangle_components _)
39+
40+
@[deprecated (since := "2025-01-01")] alias adjointOfOpAdjointOp := unop
41+
@[deprecated (since := "2025-01-01")] alias adjointUnopOfAdjointOp := unop
42+
@[deprecated (since := "2025-01-01")] alias unopAdjointOfOpAdjoint := unop
43+
@[deprecated (since := "2025-01-01")] alias unopAdjointUnopOfAdjoint := unop
6344

6445
/-- If `G` is adjoint to `F` then `F.op` is adjoint to `G.op`. -/
65-
@[simps! unit_app counit_app]
66-
def opAdjointOpOfAdjoint (F : C ⥤ D) (G : D ⥤ C) (h : G ⊣ F) : F.op ⊣ G.op :=
67-
Adjunction.mkOfHomEquiv {
68-
homEquiv := fun X Y =>
69-
(opEquiv _ Y).trans ((h.homEquiv _ _).symm.trans (opEquiv X (Opposite.op _)).symm)
70-
homEquiv_naturality_left_symm := by
71-
-- Porting note: This proof was handled by `obviously` in mathlib3. The only obstruction to
72-
-- automation fully kicking in here is that the `@[simps]` lemmas of `opEquiv` aren't firing.
73-
intros
74-
simp [opEquiv]
75-
homEquiv_naturality_right := by
76-
-- Porting note: This proof was handled by `obviously` in mathlib3. The only obstruction to
77-
-- automation fully kicking in here is that the `@[simps]` lemmas of `opEquiv` aren't firing.
78-
intros
79-
simp [opEquiv] }
80-
81-
/-- If `G` is adjoint to `F.unop` then `F` is adjoint to `G.op`. -/
82-
def adjointOpOfAdjointUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) (G : D ⥤ C) (h : G ⊣ F.unop) : F ⊣ G.op :=
83-
(opAdjointOpOfAdjoint F.unop _ h).ofNatIsoLeft F.opUnopIso
84-
85-
/-- If `G.unop` is adjoint to `F` then `F.op` is adjoint to `G`. -/
86-
def opAdjointOfUnopAdjoint (F : C ⥤ D) (G : Dᵒᵖ ⥤ Cᵒᵖ) (h : G.unop ⊣ F) : F.op ⊣ G :=
87-
(opAdjointOpOfAdjoint _ G.unop h).ofNatIsoRight G.opUnopIso
88-
89-
/-- If `G.unop` is adjoint to `F.unop` then `F` is adjoint to `G`. -/
90-
def adjointOfUnopAdjointUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) (G : Dᵒᵖ ⥤ Cᵒᵖ) (h : G.unop ⊣ F.unop) : F ⊣ G :=
91-
(adjointOpOfAdjointUnop _ _ h).ofNatIsoRight G.opUnopIso
46+
@[simps]
47+
def op {F : C ⥤ D} {G : D ⥤ C} (h : G ⊣ F) : F.op ⊣ G.op where
48+
unit := NatTrans.op h.counit
49+
counit := NatTrans.op h.unit
50+
left_triangle_components _ := Quiver.Hom.unop_inj (by simp)
51+
right_triangle_components _ := Quiver.Hom.unop_inj (by simp)
52+
53+
@[deprecated (since := "2025-01-01")] alias opAdjointOpOfAdjoint := op
54+
@[deprecated (since := "2025-01-01")] alias adjointOpOfAdjointUnop := op
55+
@[deprecated (since := "2025-01-01")] alias opAdjointOfUnopAdjoint := op
56+
@[deprecated (since := "2025-01-01")] alias adjointOfUnopAdjointUnop := op
9257

9358
/-- If `F` and `F'` are both adjoint to `G`, there is a natural isomorphism
9459
`F.op ⋙ coyoneda ≅ F'.op ⋙ coyoneda`.
@@ -119,7 +84,6 @@ Note: it is generally better to use `Adjunction.natIsoEquiv`, see the file `Adju
11984
-/
12085
def natIsoOfLeftAdjointNatIso {F F' : C ⥤ D} {G G' : D ⥤ C}
12186
(adj1 : F ⊣ G) (adj2 : F' ⊣ G') (l : F ≅ F') : G ≅ G' :=
122-
NatIso.removeOp (natIsoOfRightAdjointNatIso (opAdjointOpOfAdjoint _ F' adj2)
123-
(opAdjointOpOfAdjoint _ _ adj1) (NatIso.op l))
87+
NatIso.removeOp (natIsoOfRightAdjointNatIso (op adj2) (op adj1) (NatIso.op l))
12488

12589
end CategoryTheory.Adjunction

0 commit comments

Comments
 (0)