Skip to content

Commit 6df2c84

Browse files
adamtopazjcommelin
andcommitted
feat: port CategoryTheory.Limits.KanExtension (#2601)
Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent a223779 commit 6df2c84

File tree

2 files changed

+373
-0
lines changed

2 files changed

+373
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -304,6 +304,7 @@ import Mathlib.CategoryTheory.Limits.FullSubcategory
304304
import Mathlib.CategoryTheory.Limits.FunctorCategory
305305
import Mathlib.CategoryTheory.Limits.HasLimits
306306
import Mathlib.CategoryTheory.Limits.IsLimit
307+
import Mathlib.CategoryTheory.Limits.KanExtension
307308
import Mathlib.CategoryTheory.Limits.Preserves.Basic
308309
import Mathlib.CategoryTheory.Limits.Preserves.Finite
309310
import Mathlib.CategoryTheory.Limits.Preserves.Limits
Lines changed: 372 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,372 @@
1+
/-
2+
Copyright (c) 2021 Adam Topaz. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Bhavik Mehta, Adam Topaz
5+
6+
! This file was ported from Lean 3 source module category_theory.limits.kan_extension
7+
! leanprover-community/mathlib commit c9c9fa15fec7ca18e9ec97306fb8764bfe988a7e
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.CategoryTheory.Limits.Shapes.Terminal
12+
import Mathlib.CategoryTheory.PUnit
13+
import Mathlib.CategoryTheory.StructuredArrow
14+
15+
/-!
16+
17+
# Kan extensions
18+
19+
This file defines the right and left Kan extensions of a functor.
20+
They exist under the assumption that the target category has enough limits
21+
resp. colimits.
22+
23+
The main definitions are `Ran ι` and `Lan ι`, where `ι : S ⥤ L` is a functor.
24+
Namely, `Ran ι` is the right Kan extension, while `Lan ι` is the left Kan extension,
25+
both as functors `(S ⥤ D) ⥤ (L ⥤ D)`.
26+
27+
To access the right resp. left adjunction associated to these, use `Ran.adjunction`
28+
resp. `Lan.adjunction`.
29+
30+
# Projects
31+
32+
A lot of boilerplate could be generalized by defining and working with pseudofunctors.
33+
34+
-/
35+
36+
37+
noncomputable section
38+
39+
namespace CategoryTheory
40+
41+
open Limits
42+
43+
universe v v₁ v₂ v₃ u₁ u₂ u₃
44+
45+
variable {S : Type u₁} {L : Type u₂} {D : Type u₃}
46+
47+
variable [Category.{v₁} S] [Category.{v₂} L] [Category.{v₃} D]
48+
49+
variable (ι : S ⥤ L)
50+
51+
namespace Ran
52+
53+
attribute [local simp] StructuredArrow.proj
54+
55+
/-- The diagram indexed by `Ran.index ι x` used to define `Ran`. -/
56+
abbrev diagram (F : S ⥤ D) (x : L) : StructuredArrow x ι ⥤ D :=
57+
StructuredArrow.proj x ι ⋙ F
58+
set_option linter.uppercaseLean3 false in
59+
#align category_theory.Ran.diagram CategoryTheory.Ran.diagram
60+
61+
variable {ι}
62+
63+
/-- A cone over `Ran.diagram ι F x` used to define `Ran`. -/
64+
@[simp]
65+
def cone {F : S ⥤ D} {G : L ⥤ D} (x : L) (f : ι ⋙ G ⟶ F) : Cone (diagram ι F x)
66+
where
67+
pt := G.obj x
68+
π :=
69+
{ app := fun i => G.map i.hom ≫ f.app i.right
70+
naturality := by
71+
rintro ⟨⟨il⟩, ir, i⟩ ⟨⟨jl⟩, jr, j⟩ ⟨⟨⟨fl⟩⟩, fr, ff⟩
72+
dsimp at *
73+
dsimp at ff
74+
simp only [Category.id_comp, Category.assoc] at *
75+
rw [ff]
76+
have := f.naturality
77+
aesop_cat }
78+
set_option linter.uppercaseLean3 false in
79+
#align category_theory.Ran.cone CategoryTheory.Ran.cone
80+
81+
variable (ι)
82+
83+
/-- An auxiliary definition used to define `Ran`. -/
84+
@[simps]
85+
def loc (F : S ⥤ D) [h : ∀ x, HasLimit (diagram ι F x)] : L ⥤ D
86+
where
87+
obj x := limit (diagram ι F x)
88+
map {X Y} f :=
89+
haveI : HasLimit <| StructuredArrow.map f ⋙ diagram ι F X := h Y
90+
limit.pre (diagram ι F X) (StructuredArrow.map f)
91+
map_id := by
92+
intro l
93+
haveI : HasLimit (StructuredArrow.map (𝟙 _) ⋙ diagram ι F l) := h _
94+
dsimp
95+
ext j
96+
simp only [Category.id_comp, limit.pre_π]
97+
congr 1
98+
simp
99+
map_comp := by
100+
intro x y z f g
101+
apply limit.hom_ext
102+
intro j
103+
-- Porting note: The fact that we need to add these instances all over the place
104+
-- is certainly not ideal.
105+
haveI : HasLimit (StructuredArrow.map f ⋙ diagram ι F _) := h _
106+
haveI : HasLimit (StructuredArrow.map g ⋙ diagram ι F _) := h _
107+
haveI : HasLimit (StructuredArrow.map (f ≫ g) ⋙ diagram ι F _) := h _
108+
haveI : HasLimit (StructuredArrow.map g ⋙ StructuredArrow.map f ⋙ diagram ι F _) := h _
109+
haveI : HasLimit ((StructuredArrow.map g ⋙ StructuredArrow.map f) ⋙ diagram ι F _) := h _
110+
erw [limit.pre_pre, limit.pre_π, limit.pre_π]
111+
congr 1
112+
aesop_cat
113+
set_option linter.uppercaseLean3 false in
114+
#align category_theory.Ran.loc CategoryTheory.Ran.loc
115+
116+
/-- An auxiliary definition used to define `Ran` and `Ran.adjunction`. -/
117+
@[simps]
118+
def equiv (F : S ⥤ D) [h : ∀ x, HasLimit (diagram ι F x)] (G : L ⥤ D) :
119+
(G ⟶ loc ι F) ≃ (((whiskeringLeft _ _ _).obj ι).obj G ⟶ F)
120+
where
121+
toFun f :=
122+
{ app := fun x => f.app _ ≫ limit.π (diagram ι F (ι.obj x)) (StructuredArrow.mk (𝟙 _))
123+
naturality := by
124+
intro x y ff
125+
dsimp only [whiskeringLeft]
126+
simp only [Functor.comp_map, NatTrans.naturality_assoc, loc_map, Category.assoc]
127+
congr 1
128+
haveI : HasLimit (StructuredArrow.map (ι.map ff) ⋙ diagram ι F (ι.obj x)) := h _
129+
erw [limit.pre_π]
130+
let t : StructuredArrow.mk (𝟙 (ι.obj x)) ⟶
131+
(StructuredArrow.map (ι.map ff)).obj (StructuredArrow.mk (𝟙 (ι.obj y))) :=
132+
StructuredArrow.homMk ff ?_
133+
convert (limit.w (diagram ι F (ι.obj x)) t).symm using 1
134+
simp }
135+
invFun f :=
136+
{ app := fun x => limit.lift (diagram ι F x) (cone _ f)
137+
naturality := by
138+
intro x y ff
139+
apply limit.hom_ext
140+
intros j
141+
haveI : HasLimit (StructuredArrow.map ff ⋙ diagram ι F x) := h _
142+
erw [limit.lift_pre, limit.lift_π, Category.assoc, limit.lift_π (cone _ f) j]
143+
simp }
144+
left_inv := by
145+
intro x
146+
ext k
147+
apply limit.hom_ext
148+
intros j
149+
dsimp only [cone]
150+
rw [limit.lift_π]
151+
simp only [NatTrans.naturality_assoc, loc_map]
152+
haveI : HasLimit (StructuredArrow.map j.hom ⋙ diagram ι F k) := h _
153+
erw [limit.pre_π]
154+
congr
155+
rcases j with ⟨⟨⟩, _, _⟩
156+
aesop_cat
157+
right_inv := by aesop_cat
158+
set_option linter.uppercaseLean3 false in
159+
#align category_theory.Ran.equiv CategoryTheory.Ran.equiv
160+
161+
end Ran
162+
163+
/-- The right Kan extension of a functor. -/
164+
@[simps!]
165+
def ran [∀ X, HasLimitsOfShape (StructuredArrow X ι) D] : (S ⥤ D) ⥤ L ⥤ D :=
166+
Adjunction.rightAdjointOfEquiv (fun F G => (Ran.equiv ι G F).symm) (by {
167+
-- Porting note: was `tidy`
168+
intros X' X Y f g
169+
ext t
170+
apply limit.hom_ext
171+
intros j
172+
dsimp [Ran.equiv]
173+
simp })
174+
set_option linter.uppercaseLean3 false in
175+
#align category_theory.Ran CategoryTheory.ran
176+
177+
namespace Ran
178+
179+
variable (D)
180+
181+
/-- The adjunction associated to `Ran`. -/
182+
def adjunction [∀ X, HasLimitsOfShape (StructuredArrow X ι) D] :
183+
(whiskeringLeft _ _ D).obj ι ⊣ ran ι :=
184+
Adjunction.adjunctionOfEquivRight _ _
185+
set_option linter.uppercaseLean3 false in
186+
#align category_theory.Ran.adjunction CategoryTheory.Ran.adjunction
187+
188+
theorem reflective [Full ι] [Faithful ι] [∀ X, HasLimitsOfShape (StructuredArrow X ι) D] :
189+
IsIso (adjunction D ι).counit := by
190+
suffices : ∀ (X : S ⥤ D), IsIso (NatTrans.app (adjunction D ι).counit X)
191+
· apply NatIso.isIso_of_isIso_app
192+
intro F
193+
suffices : ∀ (X : S), IsIso (NatTrans.app (NatTrans.app (adjunction D ι).counit F) X)
194+
· apply NatIso.isIso_of_isIso_app
195+
intro X
196+
dsimp [adjunction, equiv]
197+
simp only [Category.id_comp]
198+
exact
199+
IsIso.of_iso
200+
((limit.isLimit _).conePointUniqueUpToIso
201+
(limitOfDiagramInitial StructuredArrow.mkIdInitial _))
202+
set_option linter.uppercaseLean3 false in
203+
#align category_theory.Ran.reflective CategoryTheory.Ran.reflective
204+
205+
end Ran
206+
207+
namespace Lan
208+
209+
attribute [local simp] CostructuredArrow.proj
210+
211+
/-- The diagram indexed by `Lan.index ι x` used to define `Lan`. -/
212+
abbrev diagram (F : S ⥤ D) (x : L) : CostructuredArrow ι x ⥤ D :=
213+
CostructuredArrow.proj ι x ⋙ F
214+
set_option linter.uppercaseLean3 false in
215+
#align category_theory.Lan.diagram CategoryTheory.Lan.diagram
216+
217+
variable {ι}
218+
219+
/-- A cocone over `Lan.diagram ι F x` used to define `Lan`. -/
220+
@[simp]
221+
def cocone {F : S ⥤ D} {G : L ⥤ D} (x : L) (f : F ⟶ ι ⋙ G) : Cocone (diagram ι F x)
222+
where
223+
pt := G.obj x
224+
ι :=
225+
{ app := fun i => f.app i.left ≫ G.map i.hom
226+
naturality := by
227+
rintro ⟨ir, ⟨il⟩, i⟩ ⟨jl, ⟨jr⟩, j⟩ ⟨fl, ⟨⟨fl⟩⟩, ff⟩
228+
dsimp at *
229+
simp only [Functor.comp_map, Category.comp_id, NatTrans.naturality_assoc]
230+
rw [← G.map_comp, ff]
231+
aesop_cat }
232+
set_option linter.uppercaseLean3 false in
233+
#align category_theory.Lan.cocone CategoryTheory.Lan.cocone
234+
235+
variable (ι)
236+
237+
/-- An auxiliary definition used to define `Lan`. -/
238+
@[simps]
239+
def loc (F : S ⥤ D) [I : ∀ x, HasColimit (diagram ι F x)] : L ⥤ D
240+
where
241+
obj x := colimit (diagram ι F x)
242+
map {x y} f :=
243+
haveI : HasColimit (CostructuredArrow.map f ⋙ diagram ι F y) := I _
244+
colimit.pre (diagram ι F y) (CostructuredArrow.map f)
245+
map_id := by
246+
intro l
247+
dsimp
248+
haveI : HasColimit (CostructuredArrow.map (𝟙 l) ⋙ diagram ι F l) := I _
249+
ext j
250+
erw [colimit.ι_pre, Category.comp_id]
251+
congr 1
252+
simp
253+
map_comp := by
254+
intro x y z f g
255+
dsimp
256+
haveI : HasColimit (CostructuredArrow.map (f ≫ g) ⋙ diagram ι F z) := I _
257+
ext j
258+
let ff : CostructuredArrow ι _ ⥤ _ := CostructuredArrow.map f
259+
let gg : CostructuredArrow ι _ ⥤ _ := CostructuredArrow.map g
260+
let dd := diagram ι F z
261+
-- Porting note: It seems that even Lean3 had some trouble with instances in this case.
262+
-- I don't know why lean can't deduce the following three instances...
263+
haveI : HasColimit (ff ⋙ gg ⋙ dd) := I _
264+
haveI : HasColimit ((ff ⋙ gg) ⋙ dd) := I _
265+
haveI : HasColimit (gg ⋙ dd) := I _
266+
change _ = colimit.ι ((ff ⋙ gg) ⋙ dd) j ≫ _ ≫ _
267+
erw [colimit.pre_pre dd gg ff, colimit.ι_pre, colimit.ι_pre]
268+
congr 1
269+
simp
270+
set_option linter.uppercaseLean3 false in
271+
#align category_theory.Lan.loc CategoryTheory.Lan.loc
272+
273+
/-- An auxiliary definition used to define `Lan` and `Lan.adjunction`. -/
274+
@[simps]
275+
def equiv (F : S ⥤ D) [I : ∀ x, HasColimit (diagram ι F x)] (G : L ⥤ D) :
276+
(loc ι F ⟶ G) ≃ (F ⟶ ((whiskeringLeft _ _ _).obj ι).obj G)
277+
where
278+
toFun f :=
279+
{ app := fun x => colimit.ι (diagram ι F (ι.obj x)) (CostructuredArrow.mk (𝟙 _)) ≫ f.app _
280+
naturality := by
281+
intro x y ff
282+
dsimp only [whiskeringLeft]
283+
simp only [Functor.comp_map, Category.assoc]
284+
rw [← f.naturality (ι.map ff), ← Category.assoc, ← Category.assoc]
285+
let fff : CostructuredArrow ι _ ⥤ _ := CostructuredArrow.map (ι.map ff)
286+
-- same issue :-(
287+
haveI : HasColimit (fff ⋙ diagram ι F (ι.obj y)) := I _
288+
erw [colimit.ι_pre (diagram ι F (ι.obj y)) fff (CostructuredArrow.mk (𝟙 _))]
289+
let xx : CostructuredArrow ι (ι.obj y) := CostructuredArrow.mk (ι.map ff)
290+
let yy : CostructuredArrow ι (ι.obj y) := CostructuredArrow.mk (𝟙 _)
291+
let fff : xx ⟶ yy :=
292+
CostructuredArrow.homMk ff
293+
(by
294+
simp only [CostructuredArrow.mk_hom_eq_self]
295+
erw [Category.comp_id])
296+
erw [colimit.w (diagram ι F (ι.obj y)) fff]
297+
congr
298+
simp }
299+
invFun f :=
300+
{ app := fun x => colimit.desc (diagram ι F x) (cocone _ f)
301+
naturality := by
302+
intro x y ff
303+
apply colimit.hom_ext
304+
intros j
305+
haveI : HasColimit (CostructuredArrow.map ff ⋙ diagram ι F y) := I _
306+
erw [colimit.pre_desc, ← Category.assoc, colimit.ι_desc, colimit.ι_desc]
307+
simp }
308+
left_inv := by
309+
intros x
310+
dsimp
311+
ext k
312+
dsimp
313+
apply colimit.hom_ext
314+
intros j
315+
rw [colimit.ι_desc]
316+
dsimp only [cocone]
317+
rw [Category.assoc, ← x.naturality j.hom, ← Category.assoc]
318+
congr 1
319+
dsimp [loc]
320+
haveI : HasColimit (CostructuredArrow.map j.hom ⋙ diagram ι F k) := I _
321+
erw [colimit.ι_pre (diagram ι F k) (CostructuredArrow.map j.hom)]
322+
congr
323+
rcases j with ⟨_, ⟨⟩, _⟩
324+
simp only [CostructuredArrow.map_mk, Category.id_comp]
325+
rfl
326+
right_inv := by aesop_cat
327+
set_option linter.uppercaseLean3 false in
328+
#align category_theory.Lan.equiv CategoryTheory.Lan.equiv
329+
330+
end Lan
331+
332+
/-- The left Kan extension of a functor. -/
333+
@[simps!]
334+
def lan [∀ X, HasColimitsOfShape (CostructuredArrow ι X) D] : (S ⥤ D) ⥤ L ⥤ D :=
335+
Adjunction.leftAdjointOfEquiv (fun F G => Lan.equiv ι F G) (by {
336+
intros X' X Y f g
337+
ext
338+
simp [Lan.equiv] })
339+
set_option linter.uppercaseLean3 false in
340+
#align category_theory.Lan CategoryTheory.lan
341+
342+
namespace Lan
343+
344+
variable (D)
345+
346+
/-- The adjunction associated to `Lan`. -/
347+
def adjunction [∀ X, HasColimitsOfShape (CostructuredArrow ι X) D] :
348+
lan ι ⊣ (whiskeringLeft _ _ D).obj ι :=
349+
Adjunction.adjunctionOfEquivLeft _ _
350+
set_option linter.uppercaseLean3 false in
351+
#align category_theory.Lan.adjunction CategoryTheory.Lan.adjunction
352+
353+
theorem coreflective [Full ι] [Faithful ι] [∀ X, HasColimitsOfShape (CostructuredArrow ι X) D] :
354+
IsIso (adjunction D ι).unit := by
355+
suffices : ∀ (X : S ⥤ D), IsIso (NatTrans.app (adjunction D ι).unit X)
356+
· apply NatIso.isIso_of_isIso_app
357+
intro F
358+
suffices : ∀ (X : S), IsIso (NatTrans.app (NatTrans.app (adjunction D ι).unit F) X)
359+
· apply NatIso.isIso_of_isIso_app
360+
intro X
361+
dsimp [adjunction, equiv]
362+
simp only [Category.comp_id]
363+
exact
364+
IsIso.of_iso
365+
((colimit.isColimit _).coconePointUniqueUpToIso
366+
(colimitOfDiagramTerminal CostructuredArrow.mkIdTerminal _)).symm
367+
set_option linter.uppercaseLean3 false in
368+
#align category_theory.Lan.coreflective CategoryTheory.Lan.coreflective
369+
370+
end Lan
371+
372+
end CategoryTheory

0 commit comments

Comments
 (0)