@@ -37,6 +37,16 @@ and `A` live in the same universe.
37
37
Cf https://stacks.math.columbia.edu/tag/0073, which is a weaker version of this statement (it's
38
38
only over spaces, not sites) and https://stacks.math.columbia.edu/tag/00YR (a), which
39
39
additionally assumes filtered colimits.
40
+
41
+ ## Implementation notes
42
+
43
+ Occasionally we need to take a limit in `A` of a collection of morphisms of `C` indexed
44
+ by a collection of objects in `C`. This turns out to force the morphisms of `A` to be
45
+ in a sufficiently large universe. Rather than use `UnivLE` we prove some results for
46
+ a category `A'` instead, whose morphism universe of `A'` is defined to be `max u₁ v₁`, where
47
+ `u₁, v₁` are the universes for `C`. Perhaps after we get better at handling universe
48
+ inequalities this can be changed.
49
+
40
50
-/
41
51
42
52
@@ -215,15 +225,15 @@ variable {J}
215
225
If `P`s a sheaf, `S` is a cover of `X`, and `x` is a collection of morphisms from `E`
216
226
to `P` evaluated at terms in the cover which are compatible, then we can amalgamate
217
227
the `x`s to obtain a single morphism `E ⟶ P.obj (op X)`. -/
218
- def IsSheaf.amalgamate {A : Type u₂} [Category.{max v₁ u₁ } A] {E : A} {X : C} {P : Cᵒᵖ ⥤ A}
228
+ def IsSheaf.amalgamate {A : Type u₂} [Category.{v₂ } A] {E : A} {X : C} {P : Cᵒᵖ ⥤ A}
219
229
(hP : Presheaf.IsSheaf J P) (S : J.Cover X) (x : ∀ I : S.Arrow, E ⟶ P.obj (op I.Y))
220
230
(hx : ∀ I : S.Relation, x I.fst ≫ P.map I.g₁.op = x I.snd ≫ P.map I.g₂.op) : E ⟶ P.obj (op X) :=
221
231
(hP _ _ S.condition).amalgamate (fun Y f hf => x ⟨Y, f, hf⟩) fun Y₁ Y₂ Z g₁ g₂ f₁ f₂ h₁ h₂ w =>
222
232
hx ⟨Y₁, Y₂, Z, g₁, g₂, f₁, f₂, h₁, h₂, w⟩
223
233
#align category_theory.presheaf.is_sheaf.amalgamate CategoryTheory.Presheaf.IsSheaf.amalgamate
224
234
225
235
@[reassoc (attr := simp)]
226
- theorem IsSheaf.amalgamate_map {A : Type u₂} [Category.{max v₁ u₁ } A] {E : A} {X : C} {P : Cᵒᵖ ⥤ A}
236
+ theorem IsSheaf.amalgamate_map {A : Type u₂} [Category.{v₂ } A] {E : A} {X : C} {P : Cᵒᵖ ⥤ A}
227
237
(hP : Presheaf.IsSheaf J P) (S : J.Cover X) (x : ∀ I : S.Arrow, E ⟶ P.obj (op I.Y))
228
238
(hx : ∀ I : S.Relation, x I.fst ≫ P.map I.g₁.op = x I.snd ≫ P.map I.g₂.op) (I : S.Arrow) :
229
239
hP.amalgamate S x hx ≫ P.map I.f.op = x _ := by
@@ -233,7 +243,7 @@ theorem IsSheaf.amalgamate_map {A : Type u₂} [Category.{max v₁ u₁} A] {E :
233
243
(fun Y₁ Y₂ Z g₁ g₂ f₁ f₂ h₁ h₂ w => hx ⟨Y₁, Y₂, Z, g₁, g₂, f₁, f₂, h₁, h₂, w⟩) f hf
234
244
#align category_theory.presheaf.is_sheaf.amalgamate_map CategoryTheory.Presheaf.IsSheaf.amalgamate_map
235
245
236
- theorem IsSheaf.hom_ext {A : Type u₂} [Category.{max v₁ u₁ } A] {E : A} {X : C} {P : Cᵒᵖ ⥤ A}
246
+ theorem IsSheaf.hom_ext {A : Type u₂} [Category.{v₂ } A] {E : A} {X : C} {P : Cᵒᵖ ⥤ A}
237
247
(hP : Presheaf.IsSheaf J P) (S : J.Cover X) (e₁ e₂ : E ⟶ P.obj (op X))
238
248
(h : ∀ I : S.Arrow, e₁ ≫ P.map I.f.op = e₂ ≫ P.map I.f.op) : e₁ = e₂ :=
239
249
(hP _ _ S.condition).isSeparatedFor.ext fun Y f hf => h ⟨Y, f, hf⟩
@@ -454,13 +464,20 @@ namespace Presheaf
454
464
-- between 00VQ and 00VR.
455
465
variable {C : Type u₁} [Category.{v₁} C]
456
466
457
- variable {A : Type u₂} [Category.{max v₁ u₁} A]
467
+ -- `A` is a general category; `A'` is a variant where the morphisms live in a large enough
468
+ -- universe to guarantee that we can take limits in A of things coming from C.
469
+ -- I would have liked to use something like `UnivLE.{max v₁ u₁, v₂}` as a hypothesis on
470
+ -- `A`'s morphism universe rather than introducing `A'` but I can't get it to work.
471
+ -- So, for now, results which need max v₁ u₁ ≤ v₂ are just stated for `A'` and `P' : Cᵒᵖ ⥤ A'`
472
+ -- instead.
473
+ variable {A : Type u₂} [Category.{v₂} A]
474
+ variable {A' : Type u₂} [Category.{max v₁ u₁} A']
458
475
459
476
variable (J : GrothendieckTopology C)
460
477
461
478
variable {U : C} (R : Presieve U)
462
479
463
- variable (P : Cᵒᵖ ⥤ A)
480
+ variable (P : Cᵒᵖ ⥤ A) (P' : Cᵒᵖ ⥤ A')
464
481
465
482
section MultiequalizerConditions
466
483
@@ -527,6 +544,7 @@ end MultiequalizerConditions
527
544
section
528
545
529
546
variable [HasProducts.{max u₁ v₁} A]
547
+ variable [HasProducts.{max u₁ v₁} A']
530
548
531
549
/--
532
550
The middle object of the fork diagram given in Equation (3) of [ MM92 ] , as well as the fork diagram
@@ -580,6 +598,8 @@ def IsSheaf' (P : Cᵒᵖ ⥤ A) : Prop :=
580
598
∀ (U : C) (R : Presieve U) (_ : generate R ∈ J U), Nonempty (IsLimit (Fork.ofι _ (w R P)))
581
599
#align category_theory.presheaf.is_sheaf' CategoryTheory.Presheaf.IsSheaf'
582
600
601
+ -- Again I wonder whether `UnivLE` can somehow be used to allow `s` to take
602
+ -- values in a more general universe.
583
603
/-- (Implementation). An auxiliary lemma to convert between sheaf conditions. -/
584
604
def isSheafForIsSheafFor' (P : Cᵒᵖ ⥤ A) (s : A ⥤ Type max v₁ u₁)
585
605
[∀ J, PreservesLimitsOfShape (Discrete.{max v₁ u₁} J) s] (U : C) (R : Presieve U) :
@@ -608,14 +628,16 @@ def isSheafForIsSheafFor' (P : Cᵒᵖ ⥤ A) (s : A ⥤ Type max v₁ u₁)
608
628
simp [Fork.ι]
609
629
#align category_theory.presheaf.is_sheaf_for_is_sheaf_for' CategoryTheory.Presheaf.isSheafForIsSheafFor'
610
630
631
+ -- Remark : this lemma and the next use `A'` not `A`; `A'` is `A` but with a universe
632
+ -- restriction. Can they be generalised?
611
633
/-- The equalizer definition of a sheaf given by `isSheaf'` is equivalent to `isSheaf`. -/
612
- theorem isSheaf_iff_isSheaf' : IsSheaf J P ↔ IsSheaf' J P := by
634
+ theorem isSheaf_iff_isSheaf' : IsSheaf J P' ↔ IsSheaf' J P' := by
613
635
constructor
614
636
· intro h U R hR
615
637
refine' ⟨_⟩
616
638
apply coyonedaJointlyReflectsLimits
617
639
intro X
618
- have q : Presieve.IsSheafFor (P ⋙ coyoneda.obj X) _ := h X.unop _ hR
640
+ have q : Presieve.IsSheafFor (P' ⋙ coyoneda.obj X) _ := h X.unop _ hR
619
641
rw [← Presieve.isSheafFor_iff_generate] at q
620
642
rw [Equalizer.Presieve.sheaf_condition] at q
621
643
replace q := Classical.choice q
@@ -645,13 +667,13 @@ Note this lemma applies for "algebraic" categories, eg groups, abelian groups an
645
667
for the category of topological spaces, topological rings, etc since reflecting isomorphisms doesn't
646
668
hold.
647
669
-/
648
- theorem isSheaf_iff_isSheaf_forget (s : A ⥤ Type max v₁ u₁) [HasLimits A] [PreservesLimits s]
649
- [ReflectsIsomorphisms s] : IsSheaf J P ↔ IsSheaf J (P ⋙ s) := by
670
+ theorem isSheaf_iff_isSheaf_forget (s : A' ⥤ Type max v₁ u₁) [HasLimits A' ] [PreservesLimits s]
671
+ [ReflectsIsomorphisms s] : IsSheaf J P' ↔ IsSheaf J (P' ⋙ s) := by
650
672
rw [isSheaf_iff_isSheaf', isSheaf_iff_isSheaf']
651
673
refine' forall_congr' (fun U => ball_congr (fun R _ => _))
652
674
letI : ReflectsLimits s := reflectsLimitsOfReflectsIsomorphisms
653
- have : IsLimit (s.mapCone (Fork.ofι _ (w R P))) ≃ IsLimit (Fork.ofι _ (w R (P ⋙ s))) :=
654
- isSheafForIsSheafFor' P s U R
675
+ have : IsLimit (s.mapCone (Fork.ofι _ (w R P' ))) ≃ IsLimit (Fork.ofι _ (w R (P' ⋙ s))) :=
676
+ isSheafForIsSheafFor' P' s U R
655
677
rw [← Equiv.nonempty_congr this]
656
678
constructor
657
679
· haveI := preservesSmallestLimitsOfPreservesLimits s
0 commit comments