Skip to content

Commit 44692e2

Browse files
committed
chore: more universe generalisations / fixes (#5659)
There are changes of two types: first I add some `.{w}` in some declarations to ensure that universes are in the right order. Secondly I generalise some results from `Category.{max u1 v1, u2}` to `Category.{v2, u2}`. From the Copenhagen workshop.
1 parent dc0a7a9 commit 44692e2

File tree

3 files changed

+52
-30
lines changed

3 files changed

+52
-30
lines changed

Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean

Lines changed: 16 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -171,7 +171,7 @@ structure MultispanIndex (C : Type u) [Category.{v} C] where
171171

172172
namespace MulticospanIndex
173173

174-
variable {C : Type u} [Category.{v} C] (I : MulticospanIndex C)
174+
variable {C : Type u} [Category.{v} C] (I : MulticospanIndex.{w} C)
175175

176176
/-- The multicospan associated to `I : MulticospanIndex`. -/
177177
def multicospan : WalkingMulticospan I.fstTo I.sndTo ⥤ C where
@@ -244,7 +244,7 @@ end MulticospanIndex
244244

245245
namespace MultispanIndex
246246

247-
variable {C : Type u} [Category.{v} C] (I : MultispanIndex C)
247+
variable {C : Type u} [Category.{v} C] (I : MultispanIndex.{w} C)
248248

249249
/-- The multispan associated to `I : MultispanIndex`. -/
250250
def multispan : WalkingMultispan I.fstFrom I.sndFrom ⥤ C where
@@ -319,19 +319,19 @@ variable {C : Type u} [Category.{v} C]
319319

320320
/-- A multifork is a cone over a multicospan. -/
321321
--@[nolint has_nonempty_instance]
322-
abbrev Multifork (I : MulticospanIndex C) :=
322+
abbrev Multifork (I : MulticospanIndex.{w} C) :=
323323
Cone I.multicospan
324324
#align category_theory.limits.multifork CategoryTheory.Limits.Multifork
325325

326326
/-- A multicofork is a cocone over a multispan. -/
327327
--@[nolint has_nonempty_instance]
328-
abbrev Multicofork (I : MultispanIndex C) :=
328+
abbrev Multicofork (I : MultispanIndex.{w} C) :=
329329
Cocone I.multispan
330330
#align category_theory.limits.multicofork CategoryTheory.Limits.Multicofork
331331

332332
namespace Multifork
333333

334-
variable {I : MulticospanIndex C} (K : Multifork I)
334+
variable {I : MulticospanIndex.{w} C} (K : Multifork I)
335335

336336
/-- The maps from the cone point of a multifork to the objects on the left. -/
337337
def ι (a : I.L) : K.pt ⟶ I.left a :=
@@ -364,7 +364,7 @@ theorem hom_comp_ι (K₁ K₂ : Multifork I) (f : K₁ ⟶ K₂) (j : I.L) : f.
364364

365365
/-- Construct a multifork using a collection `ι` of morphisms. -/
366366
@[simps]
367-
def ofι (I : MulticospanIndex C) (P : C) (ι : ∀ a, P ⟶ I.left a)
367+
def ofι (I : MulticospanIndex.{w} C) (P : C) (ι : ∀ a, P ⟶ I.left a)
368368
(w : ∀ b, ι (I.fstTo b) ≫ I.fst b = ι (I.sndTo b) ≫ I.snd b) : Multifork I where
369369
pt := P
370370
π :=
@@ -475,7 +475,7 @@ end Multifork
475475

476476
namespace MulticospanIndex
477477

478-
variable (I : MulticospanIndex C) [HasProduct I.left] [HasProduct I.right]
478+
variable (I : MulticospanIndex.{w} C) [HasProduct I.left] [HasProduct I.right]
479479

480480
--attribute [local tidy] tactic.case_bash
481481

@@ -526,7 +526,7 @@ end MulticospanIndex
526526

527527
namespace Multicofork
528528

529-
variable {I : MultispanIndex C} (K : Multicofork I)
529+
variable {I : MultispanIndex.{w} C} (K : Multicofork I)
530530

531531
/-- The maps to the cocone point of a multicofork from the objects on the right. -/
532532
def π (b : I.R) : I.right b ⟶ K.pt :=
@@ -556,7 +556,7 @@ lemma π_comp_hom (K₁ K₂ : Multicofork I) (f : K₁ ⟶ K₂) (b : I.R) : K
556556

557557
/-- Construct a multicofork using a collection `π` of morphisms. -/
558558
@[simps]
559-
def ofπ (I : MultispanIndex C) (P : C) (π : ∀ b, I.right b ⟶ P)
559+
def ofπ (I : MultispanIndex.{w} C) (P : C) (π : ∀ b, I.right b ⟶ P)
560560
(w : ∀ a, I.fst a ≫ π (I.fstFrom a) = I.snd a ≫ π (I.sndFrom a)) : Multicofork I where
561561
pt := P
562562
ι :=
@@ -668,7 +668,7 @@ end Multicofork
668668

669669
namespace MultispanIndex
670670

671-
variable (I : MultispanIndex C) [HasCoproduct I.left] [HasCoproduct I.right]
671+
variable (I : MultispanIndex.{w} C) [HasCoproduct I.left] [HasCoproduct I.right]
672672

673673
--attribute [local tidy] tactic.case_bash
674674

@@ -729,31 +729,31 @@ end MultispanIndex
729729

730730
/-- For `I : MulticospanIndex C`, we say that it has a multiequalizer if the associated
731731
multicospan has a limit. -/
732-
abbrev HasMultiequalizer (I : MulticospanIndex C) :=
732+
abbrev HasMultiequalizer (I : MulticospanIndex.{w} C) :=
733733
HasLimit I.multicospan
734734
#align category_theory.limits.has_multiequalizer CategoryTheory.Limits.HasMultiequalizer
735735

736736
noncomputable section
737737

738738
/-- The multiequalizer of `I : MulticospanIndex C`. -/
739-
abbrev multiequalizer (I : MulticospanIndex C) [HasMultiequalizer I] : C :=
739+
abbrev multiequalizer (I : MulticospanIndex.{w} C) [HasMultiequalizer I] : C :=
740740
limit I.multicospan
741741
#align category_theory.limits.multiequalizer CategoryTheory.Limits.multiequalizer
742742

743743
/-- For `I : MultispanIndex C`, we say that it has a multicoequalizer if
744744
the associated multicospan has a limit. -/
745-
abbrev HasMulticoequalizer (I : MultispanIndex C) :=
745+
abbrev HasMulticoequalizer (I : MultispanIndex.{w} C) :=
746746
HasColimit I.multispan
747747
#align category_theory.limits.has_multicoequalizer CategoryTheory.Limits.HasMulticoequalizer
748748

749749
/-- The multiecoqualizer of `I : MultispanIndex C`. -/
750-
abbrev multicoequalizer (I : MultispanIndex C) [HasMulticoequalizer I] : C :=
750+
abbrev multicoequalizer (I : MultispanIndex.{w} C) [HasMulticoequalizer I] : C :=
751751
colimit I.multispan
752752
#align category_theory.limits.multicoequalizer CategoryTheory.Limits.multicoequalizer
753753

754754
namespace Multiequalizer
755755

756-
variable (I : MulticospanIndex C) [HasMultiequalizer I]
756+
variable (I : MulticospanIndex.{w} C) [HasMultiequalizer I]
757757

758758
/-- The canonical map from the multiequalizer to the objects on the left. -/
759759
abbrev ι (a : I.L) : multiequalizer I ⟶ I.left a :=
@@ -833,7 +833,7 @@ end Multiequalizer
833833

834834
namespace Multicoequalizer
835835

836-
variable (I : MultispanIndex C) [HasMulticoequalizer I]
836+
variable (I : MultispanIndex.{w} C) [HasMulticoequalizer I]
837837

838838
/-- The canonical map from the multiequalizer to the objects on the left. -/
839839
abbrev π (b : I.R) : I.right b ⟶ multicoequalizer I :=

Mathlib/CategoryTheory/Sites/Sheaf.lean

Lines changed: 33 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -37,6 +37,16 @@ and `A` live in the same universe.
3737
Cf https://stacks.math.columbia.edu/tag/0073, which is a weaker version of this statement (it's
3838
only over spaces, not sites) and https://stacks.math.columbia.edu/tag/00YR (a), which
3939
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+
4050
-/
4151

4252

@@ -215,15 +225,15 @@ variable {J}
215225
If `P`s a sheaf, `S` is a cover of `X`, and `x` is a collection of morphisms from `E`
216226
to `P` evaluated at terms in the cover which are compatible, then we can amalgamate
217227
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}
219229
(hP : Presheaf.IsSheaf J P) (S : J.Cover X) (x : ∀ I : S.Arrow, E ⟶ P.obj (op I.Y))
220230
(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) :=
221231
(hP _ _ S.condition).amalgamate (fun Y f hf => x ⟨Y, f, hf⟩) fun Y₁ Y₂ Z g₁ g₂ f₁ f₂ h₁ h₂ w =>
222232
hx ⟨Y₁, Y₂, Z, g₁, g₂, f₁, f₂, h₁, h₂, w⟩
223233
#align category_theory.presheaf.is_sheaf.amalgamate CategoryTheory.Presheaf.IsSheaf.amalgamate
224234

225235
@[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}
227237
(hP : Presheaf.IsSheaf J P) (S : J.Cover X) (x : ∀ I : S.Arrow, E ⟶ P.obj (op I.Y))
228238
(hx : ∀ I : S.Relation, x I.fst ≫ P.map I.g₁.op = x I.snd ≫ P.map I.g₂.op) (I : S.Arrow) :
229239
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 :
233243
(fun Y₁ Y₂ Z g₁ g₂ f₁ f₂ h₁ h₂ w => hx ⟨Y₁, Y₂, Z, g₁, g₂, f₁, f₂, h₁, h₂, w⟩) f hf
234244
#align category_theory.presheaf.is_sheaf.amalgamate_map CategoryTheory.Presheaf.IsSheaf.amalgamate_map
235245

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}
237247
(hP : Presheaf.IsSheaf J P) (S : J.Cover X) (e₁ e₂ : E ⟶ P.obj (op X))
238248
(h : ∀ I : S.Arrow, e₁ ≫ P.map I.f.op = e₂ ≫ P.map I.f.op) : e₁ = e₂ :=
239249
(hP _ _ S.condition).isSeparatedFor.ext fun Y f hf => h ⟨Y, f, hf⟩
@@ -454,13 +464,20 @@ namespace Presheaf
454464
-- between 00VQ and 00VR.
455465
variable {C : Type u₁} [Category.{v₁} C]
456466

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']
458475

459476
variable (J : GrothendieckTopology C)
460477

461478
variable {U : C} (R : Presieve U)
462479

463-
variable (P : Cᵒᵖ ⥤ A)
480+
variable (P : Cᵒᵖ ⥤ A) (P' : Cᵒᵖ ⥤ A')
464481

465482
section MultiequalizerConditions
466483

@@ -527,6 +544,7 @@ end MultiequalizerConditions
527544
section
528545

529546
variable [HasProducts.{max u₁ v₁} A]
547+
variable [HasProducts.{max u₁ v₁} A']
530548

531549
/--
532550
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 :=
580598
∀ (U : C) (R : Presieve U) (_ : generate R ∈ J U), Nonempty (IsLimit (Fork.ofι _ (w R P)))
581599
#align category_theory.presheaf.is_sheaf' CategoryTheory.Presheaf.IsSheaf'
582600

601+
-- Again I wonder whether `UnivLE` can somehow be used to allow `s` to take
602+
-- values in a more general universe.
583603
/-- (Implementation). An auxiliary lemma to convert between sheaf conditions. -/
584604
def isSheafForIsSheafFor' (P : Cᵒᵖ ⥤ A) (s : A ⥤ Type max v₁ u₁)
585605
[∀ 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₁)
608628
simp [Fork.ι]
609629
#align category_theory.presheaf.is_sheaf_for_is_sheaf_for' CategoryTheory.Presheaf.isSheafForIsSheafFor'
610630

631+
-- Remark : this lemma and the next use `A'` not `A`; `A'` is `A` but with a universe
632+
-- restriction. Can they be generalised?
611633
/-- 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
613635
constructor
614636
· intro h U R hR
615637
refine' ⟨_⟩
616638
apply coyonedaJointlyReflectsLimits
617639
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
619641
rw [← Presieve.isSheafFor_iff_generate] at q
620642
rw [Equalizer.Presieve.sheaf_condition] at q
621643
replace q := Classical.choice q
@@ -645,13 +667,13 @@ Note this lemma applies for "algebraic" categories, eg groups, abelian groups an
645667
for the category of topological spaces, topological rings, etc since reflecting isomorphisms doesn't
646668
hold.
647669
-/
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
650672
rw [isSheaf_iff_isSheaf', isSheaf_iff_isSheaf']
651673
refine' forall_congr' (fun U => ball_congr (fun R _ => _))
652674
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
655677
rw [← Equiv.nonempty_congr this]
656678
constructor
657679
· haveI := preservesSmallestLimitsOfPreservesLimits s

Mathlib/CategoryTheory/Sites/Whiskering.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -27,13 +27,13 @@ namespace CategoryTheory
2727

2828
open CategoryTheory.Limits
2929

30-
universe v₁ v₂ u₁ u₂ u₃
30+
universe v₁ v₂ v₃ u₁ u₂ u₃
3131

3232
variable {C : Type u₁} [Category.{v₁} C]
3333

34-
variable {A : Type u₂} [Category.{max v₁ u₁} A]
34+
variable {A : Type u₂} [Category.{v₂} A]
3535

36-
variable {B : Type u₃} [Category.{max v₁ u₁} B]
36+
variable {B : Type u₃} [Category.{v₃} B]
3737

3838
variable {J : GrothendieckTopology C}
3939

0 commit comments

Comments
 (0)