@@ -44,7 +44,7 @@ def cech_nerve : simplicial_object C :=
44
44
{ obj := λ n, wide_pullback f.right
45
45
(λ i : ulift (fin (n.unop.len + 1 )), f.left) (λ i, f.hom),
46
46
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) }
48
48
49
49
/-- The augmented Čech nerve associated to an arrow. -/
50
50
@[simps]
@@ -55,6 +55,7 @@ def augmented_cech_nerve : simplicial_object.augmented C :=
55
55
56
56
end category_theory.arrow
57
57
58
+ namespace category_theory
58
59
namespace simplicial_object
59
60
60
61
variables [∀ (n : ℕ) (f : arrow C),
@@ -87,4 +88,238 @@ def augmented_cech_nerve : arrow C ⥤ simplicial_object.augmented C :=
87
88
{ left := cech_nerve.map F,
88
89
right := F.right } }
89
90
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
+
90
170
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
0 commit comments