Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit f347147

Browse files
b-mehtaScott Morrison
andcommitted
feat(category_theory): creation of limits and reflection of isomorphisms (#2426)
Define creation of limits and reflection of isomorphisms. Part 2 of a sequence of PRs aiming to resolve my TODO [here](https://github.com/leanprover-community/mathlib/blob/cf89963e14cf535783cefba471247ae4470fa8c3/src/category_theory/limits/over.lean#L143) - that the forgetful functor from the over category creates connected limits. Remaining: - [x] Add an instance or def which gives that if `F` creates limits and `K ⋙ F` `has_limit` then `has_limit K` as well. - [x] Move the forget creates limits proof to limits/over, and remove the existing proof since this one is strictly stronger - make sure to keep the statement there though using the previous point. - [x] Add more docstrings Probably relevant to @semorrison and @TwoFX. Co-authored-by: Scott Morrison <scott.morrison@anu.edu.au>
1 parent 0d3e546 commit f347147

File tree

8 files changed

+635
-119
lines changed

8 files changed

+635
-119
lines changed

src/category_theory/adjunction/limits.lean

Lines changed: 22 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -25,19 +25,19 @@ section preservation_colimits
2525
variables {J : Type v} [small_category J] (K : J ⥤ C)
2626

2727
def functoriality_right_adjoint : cocone (K ⋙ F) ⥤ cocone K :=
28-
(cocones.functoriality G) ⋙
28+
(cocones.functoriality _ G) ⋙
2929
(cocones.precompose (K.right_unitor.inv ≫ (whisker_left K adj.unit) ≫ (associator _ _ _).inv))
3030

3131
local attribute [reducible] functoriality_right_adjoint
3232

33-
@[simps] def functoriality_unit : 𝟭 (cocone K) ⟶ cocones.functoriality F ⋙ functoriality_right_adjoint adj K :=
33+
@[simps] def functoriality_unit : 𝟭 (cocone K) ⟶ cocones.functoriality _ F ⋙ functoriality_right_adjoint adj K :=
3434
{ app := λ c, { hom := adj.unit.app c.X } }
3535

36-
@[simps] def functoriality_counit : functoriality_right_adjoint adj K ⋙ cocones.functoriality F ⟶ 𝟭 (cocone (K ⋙ F)) :=
36+
@[simps] def functoriality_counit : functoriality_right_adjoint adj K ⋙ cocones.functoriality _ F ⟶ 𝟭 (cocone (K ⋙ F)) :=
3737
{ app := λ c, { hom := adj.counit.app c.X } }
3838

3939
def functoriality_is_left_adjoint :
40-
is_left_adjoint (@cocones.functoriality _ _ _ _ K _ _ F) :=
40+
is_left_adjoint (cocones.functoriality K F) :=
4141
{ right := functoriality_right_adjoint adj K,
4242
adj := mk_of_unit_counit
4343
{ unit := functoriality_unit adj K,
@@ -80,19 +80,19 @@ section preservation_limits
8080
variables {J : Type v} [small_category J] (K : J ⥤ D)
8181

8282
def functoriality_left_adjoint : cone (K ⋙ G) ⥤ cone K :=
83-
(cones.functoriality F) ⋙ (cones.postcompose
83+
(cones.functoriality _ F) ⋙ (cones.postcompose
8484
((associator _ _ _).hom ≫ (whisker_left K adj.counit) ≫ K.right_unitor.hom))
8585

8686
local attribute [reducible] functoriality_left_adjoint
8787

88-
@[simps] def functoriality_unit' : 𝟭 (cone (K ⋙ G)) ⟶ functoriality_left_adjoint adj K ⋙ cones.functoriality G :=
88+
@[simps] def functoriality_unit' : 𝟭 (cone (K ⋙ G)) ⟶ functoriality_left_adjoint adj K ⋙ cones.functoriality _ G :=
8989
{ app := λ c, { hom := adj.unit.app c.X, } }
9090

91-
@[simps] def functoriality_counit' : cones.functoriality G ⋙ functoriality_left_adjoint adj K ⟶ 𝟭 (cone K) :=
91+
@[simps] def functoriality_counit' : cones.functoriality _ G ⋙ functoriality_left_adjoint adj K ⟶ 𝟭 (cone K) :=
9292
{ app := λ c, { hom := adj.counit.app c.X, } }
9393

9494
def functoriality_is_right_adjoint :
95-
is_right_adjoint (@cones.functoriality _ _ _ _ K _ _ G) :=
95+
is_right_adjoint (cones.functoriality K G) :=
9696
{ left := functoriality_left_adjoint adj K,
9797
adj := mk_of_unit_counit
9898
{ unit := functoriality_unit' adj K,
@@ -113,6 +113,20 @@ omit adj
113113
instance is_equivalence_preserves_limits (E : D ⥤ C) [is_equivalence E] : preserves_limits E :=
114114
right_adjoint_preserves_limits E.inv.adjunction
115115

116+
@[priority 100] -- see Note [lower instance priority]
117+
instance is_equivalence_reflects_limits (E : D ⥤ C) [is_equivalence E] : reflects_limits E :=
118+
{ reflects_limits_of_shape := λ J 𝒥, by exactI
119+
{ reflects_limit := λ K,
120+
{ reflects := λ c t,
121+
begin
122+
have l: is_limit (E.inv.map_cone (E.map_cone c)) := preserves_limit.preserves t,
123+
convert is_limit.map_cone_equiv E.fun_inv_id l,
124+
{ rw functor.comp_id },
125+
{ cases c,
126+
cases c_π,
127+
congr; rw functor.comp_id }
128+
end } } }
129+
116130
-- verify the preserve_limits instance works as expected:
117131
example (E : D ⥤ C) [is_equivalence E]
118132
(c : cone K) [h : is_limit c] : is_limit (E.map_cone c) :=

src/category_theory/limits/cones.lean

Lines changed: 22 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -213,9 +213,7 @@ variable (F)
213213
@[simps]
214214
def forget : cone F ⥤ C :=
215215
{ obj := λ t, t.X, map := λ s t f, f.hom }
216-
end
217216

218-
section
219217
variables {D : Type u'} [𝒟 : category.{v} D]
220218
include 𝒟
221219

@@ -276,9 +274,7 @@ variable (F)
276274
@[simps]
277275
def forget : cocone F ⥤ C :=
278276
{ obj := λ t, t.X, map := λ s t f, f.hom }
279-
end
280277

281-
section
282278
variables {D : Type u'} [𝒟 : category.{v} D]
283279
include 𝒟
284280

@@ -302,13 +298,14 @@ variables {F : J ⥤ C} {G : J ⥤ C} (H : C ⥤ D)
302298
open category_theory.limits
303299

304300
/-- The image of a cone in C under a functor G : C ⥤ D is a cone in D. -/
305-
def map_cone (c : cone F) : cone (F ⋙ H) := (cones.functoriality H).obj c
301+
def map_cone (c : cone F) : cone (F ⋙ H) := (cones.functoriality F H).obj c
306302
/-- The image of a cocone in C under a functor G : C ⥤ D is a cocone in D. -/
307-
def map_cocone (c : cocone F) : cocone (F ⋙ H) := (cocones.functoriality H).obj c
303+
def map_cocone (c : cocone F) : cocone (F ⋙ H) := (cocones.functoriality F H).obj c
308304

309305
@[simp] lemma map_cone_X (c : cone F) : (H.map_cone c).X = H.obj c.X := rfl
310306
@[simp] lemma map_cocone_X (c : cocone F) : (H.map_cocone c).X = H.obj c.X := rfl
311307

308+
@[simps]
312309
def map_cone_inv [is_equivalence H]
313310
(c : cone (F ⋙ H)) : cone F :=
314311
let t := (inv H).map_cone c in
@@ -317,18 +314,31 @@ let α : (F ⋙ H) ⋙ inv H ⟶ F :=
317314
{ X := t.X,
318315
π := ((category_theory.cones J C).map α).app (op t.X) t.π }
319316

320-
@[simp] lemma map_cone_inv_X [is_equivalence H] (c : cone (F ⋙ H)) : (H.map_cone_inv c).X = (inv H).obj c.X := rfl
321-
322-
def map_cone_morphism {c c' : cone F} (f : cone_morphism c c') :
323-
cone_morphism (H.map_cone c) (H.map_cone c') := (cones.functoriality H).map f
324-
def map_cocone_morphism {c c' : cocone F} (f : cocone_morphism c c') :
325-
cocone_morphism (H.map_cocone c) (H.map_cocone c') := (cocones.functoriality H).map f
317+
def map_cone_morphism {c c' : cone F} (f : c ⟶ c') :
318+
(H.map_cone c) ⟶ (H.map_cone c') := (cones.functoriality F H).map f
319+
def map_cocone_morphism {c c' : cocone F} (f : c ⟶ c') :
320+
(H.map_cocone c) ⟶ (H.map_cocone c') := (cocones.functoriality F H).map f
326321

327322
@[simp] lemma map_cone_π (c : cone F) (j : J) :
328323
(map_cone H c).π.app j = H.map (c.π.app j) := rfl
329324
@[simp] lemma map_cocone_ι (c : cocone F) (j : J) :
330325
(map_cocone H c).ι.app j = H.map (c.ι.app j) := rfl
331326

327+
/-- `map_cone` is the left inverse to `map_cone_inv`. -/
328+
def map_cone_map_cone_inv {F : J ⥤ D} (H : D ⥤ C) [is_equivalence H] (c : cone (F ⋙ H)) :
329+
map_cone H (map_cone_inv H c) ≅ c :=
330+
begin
331+
apply cones.ext _ (λ j, _),
332+
{ exact H.inv_fun_id.app c.X },
333+
{ dsimp,
334+
erw [comp_id, ← H.inv_fun_id.hom.naturality (c.π.app j), comp_map, H.map_comp],
335+
congr' 1,
336+
erw [← cancel_epi (H.inv_fun_id.inv.app (H.obj (F.obj j))), nat_iso.inv_hom_id_app,
337+
← (functor.as_equivalence H).functor_unit _, ← H.map_comp, nat_iso.hom_inv_id_app,
338+
H.map_id],
339+
refl }
340+
end
341+
332342
end functor
333343

334344
end category_theory

0 commit comments

Comments
 (0)