@@ -14,7 +14,6 @@ import Mathlib.CategoryTheory.Opposites
14
14
15
15
This file contains constructions to relate adjunctions of functors to adjunctions of their
16
16
opposites.
17
- These constructions are used to show uniqueness of adjoints (up to natural isomorphism).
18
17
19
18
## Tags
20
19
adjunction, opposite, uniqueness
@@ -128,213 +127,27 @@ def leftAdjointsCoyonedaEquiv {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (a
128
127
((adj1.homEquiv X.unop Y).trans (adj2.homEquiv X.unop Y).symm).toIso
129
128
#align category_theory.adjunction.left_adjoints_coyoneda_equiv CategoryTheory.Adjunction.leftAdjointsCoyonedaEquiv
130
129
131
- /-- If `F` and `F'` are both left adjoint to `G`, then they are naturally isomorphic. -/
132
- def leftAdjointUniq {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G) :
133
- F ≅ F' :=
134
- NatIso.removeOp (Coyoneda.preimageNatIso (leftAdjointsCoyonedaEquiv adj2 adj1))
135
- #align category_theory.adjunction.left_adjoint_uniq CategoryTheory.Adjunction.leftAdjointUniq
136
-
137
- -- Porting note (#10618): removed simp as simp can prove this
138
- theorem homEquiv_leftAdjointUniq_hom_app {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G)
139
- (x : C) : adj1.homEquiv _ _ ((leftAdjointUniq adj1 adj2).hom.app x) = adj2.unit.app x := by
140
- apply (adj1.homEquiv _ _).symm.injective
141
- apply Quiver.Hom.op_inj
142
- apply coyoneda.map_injective
143
- --swap; infer_instance
144
- ext
145
- simp [leftAdjointUniq, leftAdjointsCoyonedaEquiv]
146
- #align category_theory.adjunction.hom_equiv_left_adjoint_uniq_hom_app CategoryTheory.Adjunction.homEquiv_leftAdjointUniq_hom_app
147
-
148
- @[reassoc (attr := simp)]
149
- theorem unit_leftAdjointUniq_hom {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G) :
150
- adj1.unit ≫ whiskerRight (leftAdjointUniq adj1 adj2).hom G = adj2.unit := by
151
- ext x
152
- rw [NatTrans.comp_app, ← homEquiv_leftAdjointUniq_hom_app adj1 adj2]
153
- simp [← G.map_comp]
154
- #align category_theory.adjunction.unit_left_adjoint_uniq_hom CategoryTheory.Adjunction.unit_leftAdjointUniq_hom
155
-
156
- @[reassoc (attr := simp)]
157
- theorem unit_leftAdjointUniq_hom_app {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G)
158
- (x : C) : adj1.unit.app x ≫ G.map ((leftAdjointUniq adj1 adj2).hom.app x) = adj2.unit.app x :=
159
- by rw [← unit_leftAdjointUniq_hom adj1 adj2]; rfl
160
- #align category_theory.adjunction.unit_left_adjoint_uniq_hom_app CategoryTheory.Adjunction.unit_leftAdjointUniq_hom_app
161
-
162
- @[reassoc (attr := simp)]
163
- theorem leftAdjointUniq_hom_counit {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G) :
164
- whiskerLeft G (leftAdjointUniq adj1 adj2).hom ≫ adj2.counit = adj1.counit := by
165
- ext x
166
- apply Quiver.Hom.op_inj
167
- apply coyoneda.map_injective
168
- ext
169
- simp [leftAdjointUniq, leftAdjointsCoyonedaEquiv]
170
- #align category_theory.adjunction.left_adjoint_uniq_hom_counit CategoryTheory.Adjunction.leftAdjointUniq_hom_counit
171
-
172
- @[reassoc (attr := simp)]
173
- theorem leftAdjointUniq_hom_app_counit {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G)
174
- (x : D) :
175
- (leftAdjointUniq adj1 adj2).hom.app (G.obj x) ≫ adj2.counit.app x = adj1.counit.app x := by
176
- rw [← leftAdjointUniq_hom_counit adj1 adj2]
177
- rfl
178
- #align category_theory.adjunction.left_adjoint_uniq_hom_app_counit CategoryTheory.Adjunction.leftAdjointUniq_hom_app_counit
179
-
180
- @[simp]
181
- theorem leftAdjointUniq_inv_app {F F' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G) (x : C) :
182
- (leftAdjointUniq adj1 adj2).inv.app x = (leftAdjointUniq adj2 adj1).hom.app x :=
183
- rfl
184
- #align category_theory.adjunction.left_adjoint_uniq_inv_app CategoryTheory.Adjunction.leftAdjointUniq_inv_app
185
-
186
- @[reassoc (attr := simp)]
187
- theorem leftAdjointUniq_trans {F F' F'' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G)
188
- (adj3 : F'' ⊣ G) :
189
- (leftAdjointUniq adj1 adj2).hom ≫ (leftAdjointUniq adj2 adj3).hom =
190
- (leftAdjointUniq adj1 adj3).hom := by
191
- ext
192
- apply Quiver.Hom.op_inj
193
- apply coyoneda.map_injective
194
- ext
195
- simp [leftAdjointsCoyonedaEquiv, leftAdjointUniq]
196
- #align category_theory.adjunction.left_adjoint_uniq_trans CategoryTheory.Adjunction.leftAdjointUniq_trans
197
-
198
- @[reassoc (attr := simp)]
199
- theorem leftAdjointUniq_trans_app {F F' F'' : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F' ⊣ G)
200
- (adj3 : F'' ⊣ G) (x : C) :
201
- (leftAdjointUniq adj1 adj2).hom.app x ≫ (leftAdjointUniq adj2 adj3).hom.app x =
202
- (leftAdjointUniq adj1 adj3).hom.app x := by
203
- rw [← leftAdjointUniq_trans adj1 adj2 adj3]
204
- rfl
205
- #align category_theory.adjunction.left_adjoint_uniq_trans_app CategoryTheory.Adjunction.leftAdjointUniq_trans_app
206
-
207
- @[simp]
208
- theorem leftAdjointUniq_refl {F : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) :
209
- (leftAdjointUniq adj1 adj1).hom = 𝟙 _ := by
210
- ext
211
- apply Quiver.Hom.op_inj
212
- apply coyoneda.map_injective
213
- ext
214
- simp [leftAdjointsCoyonedaEquiv, leftAdjointUniq]
215
- #align category_theory.adjunction.left_adjoint_uniq_refl CategoryTheory.Adjunction.leftAdjointUniq_refl
216
-
217
- /-- If `G` and `G'` are both right adjoint to `F`, then they are naturally isomorphic. -/
218
- def rightAdjointUniq {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G') :
219
- G ≅ G' :=
220
- NatIso.removeOp (leftAdjointUniq (opAdjointOpOfAdjoint _ F adj2) (opAdjointOpOfAdjoint _ _ adj1))
221
- #align category_theory.adjunction.right_adjoint_uniq CategoryTheory.Adjunction.rightAdjointUniq
222
-
223
- -- Porting note (#10618): simp can prove this
224
- theorem homEquiv_symm_rightAdjointUniq_hom_app {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G)
225
- (adj2 : F ⊣ G') (x : D) :
226
- (adj2.homEquiv _ _).symm ((rightAdjointUniq adj1 adj2).hom.app x) = adj1.counit.app x := by
227
- apply Quiver.Hom.op_inj
228
- convert homEquiv_leftAdjointUniq_hom_app (opAdjointOpOfAdjoint _ F adj2)
229
- (opAdjointOpOfAdjoint _ _ adj1) (Opposite.op x)
230
- -- Porting note: was `simpa`
231
- simp only [opAdjointOpOfAdjoint, Functor.op_obj, Opposite.unop_op, mkOfHomEquiv_unit_app,
232
- Equiv.trans_apply, homEquiv_counit, Functor.id_obj]
233
- -- Porting note: Yet another `erw`...
234
- -- https://github.com/leanprover-community/mathlib4/issues/5164
235
- erw [F.map_id]
236
- rw [Category.id_comp]
237
- rfl
238
- #align category_theory.adjunction.hom_equiv_symm_right_adjoint_uniq_hom_app CategoryTheory.Adjunction.homEquiv_symm_rightAdjointUniq_hom_app
239
-
240
- @[reassoc (attr := simp)]
241
- theorem unit_rightAdjointUniq_hom_app {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G')
242
- (x : C) : adj1.unit.app x ≫ (rightAdjointUniq adj1 adj2).hom.app (F.obj x) =
243
- adj2.unit.app x := by
244
- apply Quiver.Hom.op_inj
245
- convert
246
- leftAdjointUniq_hom_app_counit (opAdjointOpOfAdjoint _ _ adj2)
247
- (opAdjointOpOfAdjoint _ _ adj1) (Opposite.op x) using 1
248
- --all_goals simp
249
- all_goals {
250
- -- Porting note: Again, something seems wrong here... Some `simp` lemmas are not firing!
251
- simp only [Functor.id_obj, Functor.comp_obj, op_comp, Functor.op_obj, Opposite.unop_op,
252
- opAdjointOpOfAdjoint, mkOfHomEquiv_counit_app, Equiv.invFun_as_coe, Equiv.symm_trans_apply,
253
- Equiv.symm_symm, homEquiv_unit]
254
- erw [Functor.map_id]
255
- rw [Category.comp_id]
256
- rfl }
257
- #align category_theory.adjunction.unit_right_adjoint_uniq_hom_app CategoryTheory.Adjunction.unit_rightAdjointUniq_hom_app
258
-
259
- @[reassoc (attr := simp)]
260
- theorem unit_rightAdjointUniq_hom {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G') :
261
- adj1.unit ≫ whiskerLeft F (rightAdjointUniq adj1 adj2).hom = adj2.unit := by
262
- ext x
263
- simp
264
- #align category_theory.adjunction.unit_right_adjoint_uniq_hom CategoryTheory.Adjunction.unit_rightAdjointUniq_hom
265
-
266
- @[reassoc (attr := simp)]
267
- theorem rightAdjointUniq_hom_app_counit {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G')
268
- (x : D) :
269
- F.map ((rightAdjointUniq adj1 adj2).hom.app x) ≫ adj2.counit.app x = adj1.counit.app x := by
270
- apply Quiver.Hom.op_inj
271
- convert
272
- unit_leftAdjointUniq_hom_app (opAdjointOpOfAdjoint _ _ adj2)
273
- (opAdjointOpOfAdjoint _ _ adj1) (Opposite.op x) using 1
274
- · simp only [Functor.id_obj, op_comp, Functor.comp_obj, Functor.op_obj, Opposite.unop_op,
275
- opAdjointOpOfAdjoint_unit_app, Functor.op_map]
276
- dsimp [opEquiv]
277
- simp only [← op_comp]
278
- congr 2
279
- simp
280
- · simp only [Functor.id_obj, opAdjointOpOfAdjoint_unit_app, Opposite.unop_op]
281
- erw [Functor.map_id, Category.id_comp]
282
- rfl
283
- #align category_theory.adjunction.right_adjoint_uniq_hom_app_counit CategoryTheory.Adjunction.rightAdjointUniq_hom_app_counit
284
-
285
- @[reassoc (attr := simp)]
286
- theorem rightAdjointUniq_hom_counit {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G') :
287
- whiskerRight (rightAdjointUniq adj1 adj2).hom F ≫ adj2.counit = adj1.counit := by
288
- ext
289
- simp
290
- #align category_theory.adjunction.right_adjoint_uniq_hom_counit CategoryTheory.Adjunction.rightAdjointUniq_hom_counit
291
-
292
- @[simp]
293
- theorem rightAdjointUniq_inv_app {F : C ⥤ D} {G G' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G')
294
- (x : D) : (rightAdjointUniq adj1 adj2).inv.app x = (rightAdjointUniq adj2 adj1).hom.app x :=
295
- rfl
296
- #align category_theory.adjunction.right_adjoint_uniq_inv_app CategoryTheory.Adjunction.rightAdjointUniq_inv_app
297
-
298
- @[reassoc (attr := simp)]
299
- theorem rightAdjointUniq_trans_app {F : C ⥤ D} {G G' G'' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G')
300
- (adj3 : F ⊣ G'') (x : D) :
301
- (rightAdjointUniq adj1 adj2).hom.app x ≫ (rightAdjointUniq adj2 adj3).hom.app x =
302
- (rightAdjointUniq adj1 adj3).hom.app x := by
303
- apply Quiver.Hom.op_inj
304
- dsimp [rightAdjointUniq]
305
- simp
306
- #align category_theory.adjunction.right_adjoint_uniq_trans_app CategoryTheory.Adjunction.rightAdjointUniq_trans_app
307
-
308
- @[reassoc (attr := simp)]
309
- theorem rightAdjointUniq_trans {F : C ⥤ D} {G G' G'' : D ⥤ C} (adj1 : F ⊣ G) (adj2 : F ⊣ G')
310
- (adj3 : F ⊣ G'') :
311
- (rightAdjointUniq adj1 adj2).hom ≫ (rightAdjointUniq adj2 adj3).hom =
312
- (rightAdjointUniq adj1 adj3).hom := by
313
- ext
314
- simp
315
- #align category_theory.adjunction.right_adjoint_uniq_trans CategoryTheory.Adjunction.rightAdjointUniq_trans
130
+ /-- Given two adjunctions, if the right adjoints are naturally isomorphic, then so are the left
131
+ adjoints.
316
132
317
- @[simp]
318
- theorem rightAdjointUniq_refl {F : C ⥤ D} {G : D ⥤ C} (adj1 : F ⊣ G) :
319
- (rightAdjointUniq adj1 adj1).hom = 𝟙 _ := by
320
- delta rightAdjointUniq
321
- simp
322
- #align category_theory.adjunction.right_adjoint_uniq_refl CategoryTheory.Adjunction.rightAdjointUniq_refl
133
+ Note: it is generally better to use `Adjunction.natIsoEquiv`, see the file `Adjunction.Unique`.
134
+ The reason this definition still exists is that apparently `CategoryTheory.extendAlongYonedaYoneda`
135
+ uses its definitional properties (TODO: figure out a way to avoid this).
136
+ -/
137
+ def natIsoOfRightAdjointNatIso {F F' : C ⥤ D} {G G' : D ⥤ C}
138
+ (adj1 : F ⊣ G) (adj2 : F' ⊣ G') (r : G ≅ G') : F ≅ F' :=
139
+ NatIso.removeOp (Coyoneda.preimageNatIso (leftAdjointsCoyonedaEquiv adj2 (adj1.ofNatIsoRight r)))
140
+ #align category_theory.adjunction.nat_iso_of_right_adjoint_nat_iso CategoryTheory.Adjunction.natIsoOfRightAdjointNatIso
323
141
324
142
/-- Given two adjunctions, if the left adjoints are naturally isomorphic, then so are the right
325
143
adjoints.
144
+
145
+ Note: it is generally better to use `Adjunction.natIsoEquiv`, see the file `Adjunction.Unique`.
326
146
-/
327
147
def natIsoOfLeftAdjointNatIso {F F' : C ⥤ D} {G G' : D ⥤ C}
328
148
(adj1 : F ⊣ G) (adj2 : F' ⊣ G') (l : F ≅ F') : G ≅ G' :=
329
- rightAdjointUniq adj1 (adj2.ofNatIsoLeft l.symm)
149
+ NatIso.removeOp (natIsoOfRightAdjointNatIso (opAdjointOpOfAdjoint _ F' adj2)
150
+ (opAdjointOpOfAdjoint _ _ adj1) (NatIso.op l))
330
151
#align category_theory.adjunction.nat_iso_of_left_adjoint_nat_iso CategoryTheory.Adjunction.natIsoOfLeftAdjointNatIso
331
152
332
- /-- Given two adjunctions, if the right adjoints are naturally isomorphic, then so are the left
333
- adjoints.
334
- -/
335
- def natIsoOfRightAdjointNatIso {F F' : C ⥤ D} {G G' : D ⥤ C}
336
- (adj1 : F ⊣ G) (adj2 : F' ⊣ G') (r : G ≅ G') : F ≅ F' :=
337
- leftAdjointUniq adj1 (adj2.ofNatIsoRight r.symm)
338
- #align category_theory.adjunction.nat_iso_of_right_adjoint_nat_iso CategoryTheory.Adjunction.natIsoOfRightAdjointNatIso
339
-
340
153
end CategoryTheory.Adjunction
0 commit comments