@@ -173,7 +173,7 @@ section RegularSheaves
173
173
174
174
variable {C}
175
175
176
- open Opposite
176
+ open Opposite Presieve
177
177
178
178
/-- A presieve is *regular* if it consists of a single effective epimorphism. -/
179
179
class Presieve.regular {X : C} (R : Presieve X) : Prop where
@@ -186,10 +186,8 @@ The map to the explicit equalizer used in the sheaf condition.
186
186
-/
187
187
def MapToEqualizer (P : Cᵒᵖ ⥤ Type (max u v)) {W X B : C} (f : X ⟶ B)
188
188
(g₁ g₂ : W ⟶ X) (w : g₁ ≫ f = g₂ ≫ f) :
189
- P.obj (op B) → { x : P.obj (op X) | P.map g₁.op x = P.map g₂.op x } :=
190
- fun t ↦ ⟨P.map f.op t, by
191
- change (P.map _ ≫ P.map _) _ = (P.map _ ≫ P.map _) _;
192
- simp_rw [← P.map_comp, ← op_comp, w] ⟩
189
+ P.obj (op B) → { x : P.obj (op X) | P.map g₁.op x = P.map g₂.op x } := fun t ↦
190
+ ⟨P.map f.op t, by simp only [Set.mem_setOf_eq, ← FunctorToTypes.map_comp_apply, ← op_comp, w]⟩
193
191
194
192
/--
195
193
The sheaf condition with respect to regular presieves, given the existence of the relavant pullback.
@@ -199,164 +197,54 @@ def EqualizerCondition (P : Cᵒᵖ ⥤ Type (max u v)) : Prop :=
199
197
(MapToEqualizer P π (pullback.fst (f := π) (g := π)) (pullback.snd (f := π) (g := π))
200
198
pullback.condition)
201
199
202
- /--
203
- The `FirstObj` in the sheaf condition diagram is isomorphic to `F` applied to `X`.
204
- -/
205
- noncomputable
206
- def EqualizerFirstObjIso (F : Cᵒᵖ ⥤ Type (max u v)) {B X : C} (π : X ⟶ B) :
207
- Equalizer.FirstObj F (Presieve.singleton π) ≅ F.obj (op X) :=
208
- CategoryTheory.Equalizer.firstObjEqFamily F (Presieve.singleton π) ≪≫
209
- { hom := fun e ↦ e π (Presieve.singleton_self π)
210
- inv := fun e _ _ h ↦ by
211
- induction h with
212
- | mk => exact e
213
- hom_inv_id := by
214
- funext _ _ _ h
215
- induction h with
216
- | mk => rfl
217
- inv_hom_id := by aesop }
218
-
219
- instance {B X : C} (π : X ⟶ B) [HasPullback π π] :
220
- (Presieve.singleton π).hasPullbacks where
221
- has_pullbacks hf _ hg := by
222
- cases hf
223
- cases hg
224
- infer_instance
225
-
226
- /--
227
- The `SecondObj` in the sheaf condition diagram is isomorphic to `F` applied to the pullback of `π`
228
- with itself
229
- -/
230
- noncomputable
231
- def EqualizerSecondObjIso (F : Cᵒᵖ ⥤ Type (max u v)) {B X : C} (π : X ⟶ B) [HasPullback π π] :
232
- Equalizer.Presieve.SecondObj F (Presieve.singleton π) ≅ F.obj (op (Limits.pullback π π)) :=
233
- Types.productIso.{max u v, max u v} _ ≪≫
234
- { hom := fun e ↦ e (⟨X, ⟨π, Presieve.singleton_self π⟩⟩, ⟨X, ⟨π, Presieve.singleton_self π⟩⟩)
235
- inv := fun x ⟨⟨_, ⟨_, h₁⟩⟩ , ⟨_, ⟨_, h₂⟩⟩⟩ ↦ by
236
- induction h₁
237
- induction h₂
238
- exact x
239
- hom_inv_id := by
240
- funext _ ⟨⟨_, ⟨_, h₁⟩⟩ , ⟨_, ⟨_, h₂⟩⟩⟩
241
- induction h₁
242
- induction h₂
243
- rfl
244
- inv_hom_id := by aesop }
245
-
246
200
lemma EqualizerCondition.isSheafFor {B : C} {S : Presieve B} [S.regular] [S.hasPullbacks]
247
201
{F : Cᵒᵖ ⥤ Type (max u v)}
248
- (hFecs : EqualizerCondition F) : S.IsSheafFor F := by
249
- obtain ⟨X, π, ⟨hS, πsurj⟩⟩ := Presieve.regular.single_epi (R := S)
250
- rw [Presieve.ofArrows_pUnit] at hS
251
- haveI : (Presieve.singleton π).hasPullbacks := by rw [← hS]; infer_instance
252
- haveI : HasPullback π π :=
253
- Presieve.hasPullbacks.has_pullbacks (Presieve.singleton.mk) (Presieve.singleton.mk)
202
+ (hF : EqualizerCondition F) : S.IsSheafFor F := by
203
+ obtain ⟨X, π, hS, πsurj⟩ := Presieve.regular.single_epi (R := S)
254
204
subst hS
255
- rw [Equalizer.Presieve.sheaf_condition, Limits.Types.type_equalizer_iff_unique ]
205
+ rw [isSheafFor_arrows_iff_pullbacks ]
256
206
intro y h
257
- specialize hFecs X B π
258
- have hy : F.map (pullback.fst (f := π) (g := π)).op ((EqualizerFirstObjIso F π).hom y) =
259
- F.map (pullback.snd (f := π) (g := π)).op ((EqualizerFirstObjIso F π).hom y) :=
260
- calc
261
- _ = (Equalizer.Presieve.firstMap F (Presieve.singleton π) ≫
262
- (EqualizerSecondObjIso F π).hom) y := by
263
- simp [EqualizerSecondObjIso, EqualizerFirstObjIso, Equalizer.Presieve.firstMap]
264
- _ = (Equalizer.Presieve.secondMap F (Presieve.singleton π) ≫ (EqualizerSecondObjIso F π).hom)
265
- y := by simp only [Equalizer.Presieve.SecondObj, types_comp_apply]; rw [h]
266
- _ = _ := by
267
- simp [EqualizerSecondObjIso, EqualizerFirstObjIso, Equalizer.Presieve.secondMap]
268
- obtain ⟨x, ⟨hx₁, hx₂⟩⟩ : ∃! x, F.map π.op x = (EqualizerFirstObjIso F π).hom y
269
- · rw [Function.bijective_iff_existsUnique] at hFecs
270
- specialize hFecs ⟨(EqualizerFirstObjIso F π).hom y, hy⟩
271
- obtain ⟨x, ⟨hx₁, hx₂⟩⟩ := hFecs
272
- refine ⟨x, ⟨Subtype.ext_iff.mp hx₁, ?_⟩⟩
273
- intros
274
- apply hx₂
275
- rwa [Subtype.ext_iff]
276
- have fork_comp : Equalizer.forkMap F (Presieve.singleton π) ≫ (EqualizerFirstObjIso F π).hom =
277
- F.map π.op := by ext; simp [EqualizerFirstObjIso, Equalizer.forkMap]
278
- rw [← fork_comp] at hx₁ hx₂
279
- refine ⟨x, ⟨?_, ?_⟩⟩
280
- · apply_fun (EqualizerFirstObjIso F π).hom using injective_of_mono (EqualizerFirstObjIso F π).hom
281
- exact hx₁
282
- · intro z hz
283
- apply_fun (EqualizerFirstObjIso F π).hom at hz
284
- exact hx₂ z hz
285
-
286
- lemma IsSheafForRegular.equalizerCondition {F : Cᵒᵖ ⥤ Type (max u v)}
207
+ have : (Presieve.singleton π).hasPullbacks := by rw [← ofArrows_pUnit]; infer_instance
208
+ have : HasPullback π π := hasPullbacks.has_pullbacks Presieve.singleton.mk Presieve.singleton.mk
209
+ specialize hF X B π
210
+ rw [Function.bijective_iff_existsUnique] at hF
211
+ obtain ⟨t, ht, ht'⟩ := hF ⟨y (), h () ()⟩
212
+ refine ⟨t, fun _ ↦ ?_, fun x h ↦ ht' x ?_⟩
213
+ · simpa [MapToEqualizer] using ht
214
+ · simpa [MapToEqualizer] using h ()
215
+
216
+ lemma equalizerCondition_of_regular {F : Cᵒᵖ ⥤ Type (max u v)}
287
217
(hSF : ∀ {B : C} (S : Presieve B) [S.regular] [S.hasPullbacks], S.IsSheafFor F) :
288
218
EqualizerCondition F := by
289
219
intro X B π _ _
290
- haveI : (Presieve.singleton π).regular :=
291
- ⟨X, π, ⟨(Presieve.ofArrows_pUnit π).symm, inferInstance⟩⟩
292
- specialize hSF (Presieve.singleton π)
293
- rw [Equalizer.Presieve.sheaf_condition, Limits.Types.type_equalizer_iff_unique] at hSF
220
+ have : (ofArrows (fun _ ↦ X) (fun _ ↦ π)).regular := ⟨X, π, rfl, inferInstance⟩
221
+ have : (ofArrows (fun () ↦ X) (fun _ ↦ π)).hasPullbacks := ⟨
222
+ fun hf _ hg ↦ (by cases hf; cases hg; infer_instance)⟩
223
+ specialize hSF (ofArrows (fun () ↦ X) (fun _ ↦ π))
224
+ rw [isSheafFor_arrows_iff_pullbacks] at hSF
294
225
rw [Function.bijective_iff_existsUnique]
295
- intro ⟨b, hb⟩
296
- specialize hSF ((EqualizerFirstObjIso F π).inv b) ?_
297
- · apply_fun (EqualizerSecondObjIso F π).hom using injective_of_mono _
298
- calc
299
- _ = F.map (pullback.fst (f := π) (g := π)).op b := by
300
- simp only [Equalizer.Presieve.SecondObj, EqualizerSecondObjIso, Equalizer.Presieve.firstMap,
301
- EqualizerFirstObjIso, Iso.trans_inv, types_comp_apply, Equalizer.firstObjEqFamily_inv,
302
- Iso.trans_hom, Types.productIso_hom_comp_eval_apply, Types.Limit.lift_π_apply', Fan.mk_pt,
303
- Fan.mk_π_app]; rfl
304
- _ = F.map (pullback.snd (f := π) (g := π)).op b := hb
305
- _ = ((EqualizerFirstObjIso F π).inv ≫ Equalizer.Presieve.secondMap F (Presieve.singleton π) ≫
306
- (EqualizerSecondObjIso F π).hom) b := by
307
- simp only [EqualizerFirstObjIso, Iso.trans_inv, Equalizer.Presieve.SecondObj,
308
- Equalizer.Presieve.secondMap, EqualizerSecondObjIso, Iso.trans_hom,
309
- Types.productIso_hom_comp_eval, limit.lift_π, Fan.mk_pt, Fan.mk_π_app, types_comp_apply,
310
- Equalizer.firstObjEqFamily_inv, Types.Limit.lift_π_apply']; rfl
311
- · obtain ⟨a, ⟨ha₁, ha₂⟩⟩ := hSF
312
- refine ⟨a, ⟨?_, ?_⟩⟩
313
- · ext
314
- dsimp [MapToEqualizer]
315
- apply_fun (EqualizerFirstObjIso F π).hom at ha₁
316
- simp only [inv_hom_id_apply] at ha₁
317
- rw [← ha₁]
318
- simp only [EqualizerFirstObjIso, Equalizer.forkMap, Iso.trans_hom, types_comp_apply,
319
- Equalizer.firstObjEqFamily_hom, Types.pi_lift_π_apply]
320
- · intro y hy
321
- apply ha₂
322
- apply_fun (EqualizerFirstObjIso F π).hom using injective_of_mono _
323
- simp only [inv_hom_id_apply]
324
- simp only [MapToEqualizer, Set.mem_setOf_eq, Subtype.mk.injEq] at hy
325
- rw [← hy]
326
- simp only [EqualizerFirstObjIso, Equalizer.forkMap, Iso.trans_hom, types_comp_apply,
327
- Equalizer.firstObjEqFamily_hom, Types.pi_lift_π_apply]
226
+ intro ⟨x, hx⟩
227
+ obtain ⟨t, ht, ht'⟩ := hSF (fun _ ↦ x) (fun _ _ ↦ hx)
228
+ refine ⟨t, ?_, fun y h ↦ ht' y ?_⟩
229
+ · simpa [MapToEqualizer] using ht ()
230
+ · simpa [MapToEqualizer] using h
328
231
329
232
lemma isSheafFor_regular_of_projective {X : C} (S : Presieve X) [S.regular] [Projective X]
330
233
(F : Cᵒᵖ ⥤ Type (max u v)) : S.IsSheafFor F := by
331
234
obtain ⟨Y, f, rfl, hf⟩ := Presieve.regular.single_epi (R := S)
332
- let g := Projective.factorThru (𝟙 _) f
333
- have hfg : g ≫ f = 𝟙 _ := by
334
- simp only [Projective.factorThru_comp]
335
- intro y hy
336
- refine' ⟨F.map g.op <| y f <| Presieve.ofArrows.mk (), fun Z h hZ => _, fun z hz => _⟩
337
- · cases' hZ with u
338
- have := hy (f₁ := f) (f₂ := f) (𝟙 Y) (f ≫ g) (Presieve.ofArrows.mk ())
339
- (Presieve.ofArrows.mk ()) ?_
340
- · rw [op_id, F.map_id, types_id_apply] at this
341
- rw [← types_comp_apply (F.map g.op) (F.map f.op), ← F.map_comp, ← op_comp]
342
- exact this.symm
343
- · rw [Category.id_comp, Category.assoc, hfg, Category.comp_id]
344
- · have := congr_arg (F.map g.op) <| hz f (Presieve.ofArrows.mk ())
345
- rwa [← types_comp_apply (F.map f.op) (F.map g.op), ← F.map_comp, ← op_comp, hfg, op_id,
346
- F.map_id, types_id_apply] at this
235
+ rw [isSheafFor_arrows_iff]
236
+ refine fun x hx ↦ ⟨F.map (Projective.factorThru (𝟙 _) f).op <| x (), fun _ ↦ ?_, fun y h ↦ ?_⟩
237
+ · simpa using (hx () () Y (𝟙 Y) (f ≫ (Projective.factorThru (𝟙 _) f)) (by simp)).symm
238
+ · simp only [← h (), ← FunctorToTypes.map_comp_apply, ← op_comp, Projective.factorThru_comp,
239
+ op_id, FunctorToTypes.map_id_apply]
347
240
348
241
lemma isSheaf_iff_equalizerCondition (F : Cᵒᵖ ⥤ Type (max u v)) [Preregular C] [HasPullbacks C] :
349
242
Presieve.IsSheaf (regularCoverage C).toGrothendieck F ↔ EqualizerCondition F := by
350
243
rw [Presieve.isSheaf_coverage]
351
- refine ⟨?_, ?_⟩
352
- · intro h
353
- apply IsSheafForRegular.equalizerCondition
354
- intro B S _ _
355
- apply h S
356
- obtain ⟨Y, f, rfl, _⟩ := Presieve.regular.single_epi (R := S)
357
- use Y, f
358
- · intro h X S ⟨Y, f, hh⟩
359
- haveI : S.regular := ⟨Y, f, hh⟩
244
+ refine ⟨fun h ↦ equalizerCondition_of_regular fun S _ _ ↦ h S ?_, fun h X S ⟨Y, f, hh⟩ ↦ ?_⟩
245
+ · obtain ⟨Y, f, rfl, _⟩ := Presieve.regular.single_epi (R := S)
246
+ exact ⟨Y, f, rfl, inferInstance⟩
247
+ · have : S.regular := ⟨Y, f, hh⟩
360
248
exact h.isSheafFor
361
249
362
250
lemma isSheaf_of_projective (F : Cᵒᵖ ⥤ Type (max u v)) [Preregular C] [∀ (X : C), Projective X] :
0 commit comments