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

Commit 31ba155

Browse files
adamtopazkim-em
andcommitted
feat(algebraic_topology/cech_nerve): The Cech nerve is a right adjoint. (#7716)
Also fixes the namespace in the file `algebraic_topology/cech_nerve.lean`. This is needed for LTE Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent 272a930 commit 31ba155

File tree

3 files changed

+279
-2
lines changed

3 files changed

+279
-2
lines changed

src/algebraic_topology/cech_nerve.lean

Lines changed: 236 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -44,7 +44,7 @@ def cech_nerve : simplicial_object C :=
4444
{ obj := λ n, wide_pullback f.right
4545
(λ i : ulift (fin (n.unop.len + 1)), f.left) (λ i, f.hom),
4646
map := λ m n g, wide_pullback.lift (wide_pullback.base _)
47-
(λ i, wide_pullback.π _ $ ulift.up $ g.unop.to_preorder_hom i.down) (by tidy) }
47+
(λ i, wide_pullback.π (λ i, f.hom) $ ulift.up $ g.unop.to_preorder_hom i.down) (by tidy) }
4848

4949
/-- The augmented Čech nerve associated to an arrow. -/
5050
@[simps]
@@ -55,6 +55,7 @@ def augmented_cech_nerve : simplicial_object.augmented C :=
5555

5656
end category_theory.arrow
5757

58+
namespace category_theory
5859
namespace simplicial_object
5960

6061
variables [∀ (n : ℕ) (f : arrow C),
@@ -87,4 +88,238 @@ def augmented_cech_nerve : arrow C ⥤ simplicial_object.augmented C :=
8788
{ left := cech_nerve.map F,
8889
right := F.right } }
8990

91+
/-- A helper function used in defining the Čech adjunction. -/
92+
@[simps]
93+
def equivalence_right_to_left (X : simplicial_object.augmented C) (F : arrow C)
94+
(G : X ⟶ F.augmented_cech_nerve) : augmented.to_arrow.obj X ⟶ F :=
95+
{ left := G.left.app _ ≫ wide_pullback.π (λ i, F.hom) ⟨0⟩,
96+
right := G.right,
97+
w' := begin
98+
have := G.w,
99+
apply_fun (λ e, e.app (opposite.op $ simplex_category.mk 0)) at this,
100+
tidy,
101+
end }
102+
103+
/-- A helper function used in defining the Čech adjunction. -/
104+
@[simps]
105+
def equivalence_left_to_right (X : simplicial_object.augmented C) (F : arrow C)
106+
(G : augmented.to_arrow.obj X ⟶ F) : X ⟶ F.augmented_cech_nerve :=
107+
{ left :=
108+
{ app := λ x, limits.wide_pullback.lift (X.hom.app _ ≫ G.right)
109+
(λ i, X.left.map (simplex_category.const x.unop i.down).op ≫ G.left) (by tidy),
110+
naturality' := begin
111+
intros x y f,
112+
ext,
113+
{ dsimp,
114+
simp only [wide_pullback.lift_π, category.assoc],
115+
rw [← category.assoc, ← X.left.map_comp],
116+
refl },
117+
{ dsimp,
118+
simp only [functor.const.obj_map, nat_trans.naturality_assoc,
119+
wide_pullback.lift_base, category.assoc],
120+
erw category.id_comp }
121+
end },
122+
right := G.right }
123+
124+
/-- A helper function used in defining the Čech adjunction. -/
125+
@[simps]
126+
def cech_nerve_equiv (X : simplicial_object.augmented C) (F : arrow C) :
127+
(augmented.to_arrow.obj X ⟶ F) ≃ (X ⟶ F.augmented_cech_nerve) :=
128+
{ to_fun := equivalence_left_to_right _ _,
129+
inv_fun := equivalence_right_to_left _ _,
130+
left_inv := begin
131+
intro A,
132+
dsimp,
133+
ext,
134+
{ dsimp,
135+
erw wide_pullback.lift_π,
136+
nth_rewrite 1 ← category.id_comp A.left,
137+
congr' 1,
138+
convert X.left.map_id _,
139+
rw ← op_id,
140+
congr' 1,
141+
ext ⟨a,ha⟩,
142+
change a < 1 at ha,
143+
change 0 = a,
144+
linarith },
145+
{ refl, }
146+
end,
147+
right_inv := begin
148+
intro A,
149+
dsimp,
150+
ext _ ⟨j⟩,
151+
{ dsimp,
152+
simp only [arrow.cech_nerve_map, wide_pullback.lift_π, nat_trans.naturality_assoc],
153+
erw wide_pullback.lift_π,
154+
refl },
155+
{ dsimp,
156+
erw wide_pullback.lift_base,
157+
have := A.w,
158+
apply_fun (λ e, e.app x) at this,
159+
rw nat_trans.comp_app at this,
160+
erw this,
161+
refl },
162+
{ refl }
163+
end }
164+
165+
/-- The augmented Čech nerve construction is right adjoint to the `to_arrow` functor. -/
166+
abbreviation cech_nerve_adjunction :
167+
(augmented.to_arrow : _ ⥤ arrow C) ⊣ augmented_cech_nerve :=
168+
adjunction.mk_of_hom_equiv { hom_equiv := cech_nerve_equiv }
169+
90170
end simplicial_object
171+
172+
end category_theory
173+
174+
namespace category_theory.arrow
175+
176+
variables (f : arrow C)
177+
variables [∀ n : ℕ, has_wide_pushout f.left (λ i : ulift (fin (n+1)), f.right) (λ i, f.hom)]
178+
179+
/-- The Čech conerve associated to an arrow. -/
180+
@[simps]
181+
def cech_conerve : cosimplicial_object C :=
182+
{ obj := λ n, wide_pushout f.left
183+
(λ i : ulift (fin (n.len + 1)), f.right) (λ i, f.hom),
184+
map := λ m n g, wide_pushout.desc (wide_pushout.head _)
185+
(λ i, wide_pushout.ι (λ i, f.hom) $ ulift.up $ g.to_preorder_hom i.down)
186+
begin
187+
rintros ⟨⟨j⟩⟩,
188+
dsimp,
189+
rw [wide_pushout.arrow_ι (λ i, f.hom)],
190+
end }
191+
192+
/-- The augmented Čech conerve associated to an arrow. -/
193+
@[simps]
194+
def augmented_cech_conerve : cosimplicial_object.augmented C :=
195+
{ left := f.left,
196+
right := f.cech_conerve,
197+
hom := { app := λ i, wide_pushout.head _ } }
198+
199+
end category_theory.arrow
200+
201+
namespace category_theory
202+
namespace cosimplicial_object
203+
204+
variables [∀ (n : ℕ) (f : arrow C),
205+
has_wide_pushout f.left (λ i : ulift (fin (n+1)), f.right) (λ i, f.hom)]
206+
207+
/-- The Čech conerve construction, as a functor from `arrow C`. -/
208+
@[simps]
209+
def cech_conerve : arrow C ⥤ cosimplicial_object C :=
210+
{ obj := λ f, f.cech_conerve,
211+
map := λ f g F,
212+
{ app := λ n, wide_pushout.desc (F.left ≫ wide_pushout.head _)
213+
(λ i, F.right ≫ wide_pushout.ι _ i)
214+
(λ i, by { rw [←arrow.w_assoc F, wide_pushout.arrow_ι (λ i, g.hom)], }) },
215+
-- tidy needs a bit of help here...
216+
map_id' := by { intros i, ext, tidy },
217+
map_comp' := begin
218+
intros f g h F G,
219+
ext,
220+
all_goals {
221+
dsimp,
222+
simp only [category.assoc, limits.wide_pushout.head_desc_assoc,
223+
limits.wide_pushout.ι_desc_assoc, limits.colimit.ι_desc],
224+
simpa only [← category.assoc], },
225+
end }
226+
227+
/-- The augmented Čech conerve construction, as a functor from `arrow C`. -/
228+
@[simps]
229+
def augmented_cech_conerve : arrow C ⥤ cosimplicial_object.augmented C :=
230+
{ obj := λ f, f.augmented_cech_conerve,
231+
map := λ f g F,
232+
{ left := F.left,
233+
right := cech_conerve.map F, } }
234+
235+
/-- A helper function used in defining the Čech conerve adjunction. -/
236+
@[simps]
237+
def equivalence_left_to_right (F : arrow C) (X : cosimplicial_object.augmented C)
238+
(G : F.augmented_cech_conerve ⟶ X) : F ⟶ augmented.to_arrow.obj X :=
239+
{ left := G.left,
240+
right :=
241+
(wide_pushout.ι (λ i, F.hom) (_root_.ulift.up 0) ≫ G.right.app (simplex_category.mk 0) : _),
242+
w' := begin
243+
have := G.w,
244+
apply_fun (λ e, e.app (simplex_category.mk 0)) at this,
245+
dsimp at this,
246+
simpa only [category_theory.functor.id_map, augmented.to_arrow_obj_hom,
247+
wide_pushout.arrow_ι_assoc (λ i, F.hom)],
248+
end }
249+
250+
/-- A helper function used in defining the Čech conerve adjunction. -/
251+
@[simps]
252+
def equivalence_right_to_left (F : arrow C) (X : cosimplicial_object.augmented C)
253+
(G : F ⟶ augmented.to_arrow.obj X) : F.augmented_cech_conerve ⟶ X :=
254+
{ left := G.left,
255+
right := { app := λ x, limits.wide_pushout.desc (G.left ≫ X.hom.app _)
256+
(λ i, G.right ≫ X.right.map (simplex_category.const x i.down))
257+
begin
258+
rintros ⟨j⟩,
259+
rw ←arrow.w_assoc G,
260+
dsimp,
261+
have t := X.hom.naturality (x.const j),
262+
dsimp at t, simp only [category.id_comp] at t,
263+
rw ←t,
264+
end,
265+
naturality' := begin
266+
intros x y f,
267+
ext,
268+
{ dsimp,
269+
simp only [wide_pushout.ι_desc_assoc, wide_pushout.ι_desc],
270+
rw [category.assoc, ←X.right.map_comp],
271+
refl },
272+
{ dsimp,
273+
simp only [functor.const.obj_map, ←nat_trans.naturality,
274+
wide_pushout.head_desc_assoc, wide_pushout.head_desc, category.assoc],
275+
erw category.id_comp }
276+
end }, }
277+
278+
/-- A helper function used in defining the Čech conerve adjunction. -/
279+
@[simps]
280+
def cech_conerve_equiv (F : arrow C) (X : cosimplicial_object.augmented C) :
281+
(F.augmented_cech_conerve ⟶ X) ≃ (F ⟶ augmented.to_arrow.obj X) :=
282+
{ to_fun := equivalence_left_to_right _ _,
283+
inv_fun := equivalence_right_to_left _ _,
284+
left_inv := begin
285+
intro A,
286+
dsimp,
287+
ext,
288+
{ refl, },
289+
{ cases j,
290+
dsimp,
291+
simp only [arrow.cech_conerve_map, wide_pushout.ι_desc, category.assoc,
292+
←nat_trans.naturality, wide_pushout.ι_desc_assoc],
293+
refl },
294+
{ dsimp,
295+
erw wide_pushout.head_desc,
296+
have := A.w,
297+
apply_fun (λ e, e.app x) at this,
298+
rw nat_trans.comp_app at this,
299+
erw this,
300+
refl },
301+
end,
302+
right_inv := begin
303+
intro A,
304+
dsimp,
305+
ext,
306+
{ refl, },
307+
{ dsimp,
308+
erw wide_pushout.ι_desc,
309+
nth_rewrite 1 ← category.comp_id A.right,
310+
congr' 1,
311+
convert X.right.map_id _,
312+
ext ⟨a,ha⟩,
313+
change a < 1 at ha,
314+
change 0 = a,
315+
linarith },
316+
end }
317+
318+
/-- The augmented Čech conerve construction is left adjoint to the `to_arrow` functor. -/
319+
abbreviation cech_conerve_adjunction :
320+
augmented_cech_conerve ⊣ (augmented.to_arrow : _ ⥤ arrow C) :=
321+
adjunction.mk_of_hom_equiv { hom_equiv := cech_conerve_equiv }
322+
323+
end cosimplicial_object
324+
325+
end category_theory

src/algebraic_topology/simplex_category.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -114,6 +114,14 @@ instance small_category : small_category.{u} simplex_category :=
114114
id := λ m, simplex_category.hom.id _,
115115
comp := λ _ _ _ f g, simplex_category.hom.comp g f, }
116116

117+
/-- The constant morphism from [0]. -/
118+
def const (x : simplex_category.{u}) (i : fin (x.len+1)) : [0] ⟶ x :=
119+
hom.mk $ ⟨λ _, i, by tauto⟩
120+
121+
@[simp]
122+
lemma const_comp (x y : simplex_category.{u}) (i : fin (x.len + 1)) (f : x ⟶ y) :
123+
const x i ≫ f = const y (f.to_preorder_hom i) := rfl
124+
117125
/--
118126
Make a morphism `[n] ⟶ [m]` from a monotone map between fin's.
119127
This is useful for constructing morphisms beetween `[n]` directly

src/algebraic_topology/simplicial_object.lean

Lines changed: 35 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -8,7 +8,7 @@ import category_theory.category.ulift
88
import category_theory.limits.functor_category
99
import category_theory.opposites
1010
import category_theory.adjunction.limits
11-
import category_theory.comma
11+
import category_theory.arrow
1212

1313
/-!
1414
# Simplicial objects in a category.
@@ -180,6 +180,23 @@ def drop : augmented C ⥤ simplicial_object C := comma.fst _ _
180180
@[simps]
181181
def point : augmented C ⥤ C := comma.snd _ _
182182

183+
/-- The functor from augmented objects to arrows. -/
184+
@[simps]
185+
def to_arrow : augmented C ⥤ arrow C :=
186+
{ obj := λ X,
187+
{ left := (drop.obj X) _[0],
188+
right := (point.obj X),
189+
hom := X.hom.app _ },
190+
map := λ X Y η,
191+
{ left := (drop.map η).app _,
192+
right := (point.map η),
193+
w' := begin
194+
dsimp,
195+
rw ← nat_trans.comp_app,
196+
erw η.w,
197+
refl,
198+
end } }
199+
183200
variable (C)
184201

185202
/-- Functor composition induces a functor on augmented simplicial objects. -/
@@ -366,6 +383,23 @@ def drop : augmented C ⥤ cosimplicial_object C := comma.snd _ _
366383
@[simps]
367384
def point : augmented C ⥤ C := comma.fst _ _
368385

386+
/-- The functor from augmented objects to arrows. -/
387+
@[simps]
388+
def to_arrow : augmented C ⥤ arrow C :=
389+
{ obj := λ X,
390+
{ left := (point.obj X),
391+
right := (drop.obj X) _[0],
392+
hom := X.hom.app _ },
393+
map := λ X Y η,
394+
{ left := (point.map η),
395+
right := (drop.map η).app _,
396+
w' := begin
397+
dsimp,
398+
rw ← nat_trans.comp_app,
399+
erw ← η.w,
400+
refl,
401+
end } }
402+
369403
variable (C)
370404

371405
/-- Functor composition induces a functor on augmented cosimplicial objects. -/

0 commit comments

Comments
 (0)