Skip to content

Commit d732a68

Browse files
kim-emdagurtomas
authored andcommitted
feat: if the body constrains universes, make it explicit in the signature. (#14069)
In many places in Mathlib universes in the type signature are invisibly constrained by the RHS of the `def`. This PR makes all these explicit in the type signature. Likely we will change the Lean behaviour to disallow this in leanprover/lean4#4482 (i.e. this is the backport to `master` of the fixes required for that).
1 parent 9bc57dc commit d732a68

File tree

21 files changed

+111
-99
lines changed

21 files changed

+111
-99
lines changed

Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean

+5-5
Original file line numberDiff line numberDiff line change
@@ -50,7 +50,7 @@ open CategoryTheory Limits
5050

5151
namespace ModuleCat
5252

53-
universe v u₁ u₂ u₃
53+
universe v u₁ u₂ u₃ w
5454

5555
namespace RestrictScalars
5656

@@ -618,11 +618,11 @@ def restrictCoextendScalarsAdj {R : Type u₁} {S : Type u₂} [Ring R] [Ring S]
618618
#align category_theory.Module.restrict_coextend_scalars_adj ModuleCat.restrictCoextendScalarsAdj
619619

620620
instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) :
621-
(restrictScalars f).IsLeftAdjoint :=
621+
(restrictScalars.{max u₂ w} f).IsLeftAdjoint :=
622622
(restrictCoextendScalarsAdj f).isLeftAdjoint
623623

624624
instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) :
625-
(coextendScalars f).IsRightAdjoint :=
625+
(coextendScalars.{u₁, u₂, max u₂ w} f).IsRightAdjoint :=
626626
(restrictCoextendScalarsAdj f).isRightAdjoint
627627

628628
namespace ExtendRestrictScalarsAdj
@@ -857,11 +857,11 @@ def extendRestrictScalarsAdj {R : Type u₁} {S : Type u₂} [CommRing R] [CommR
857857
#align category_theory.Module.extend_restrict_scalars_adj ModuleCat.extendRestrictScalarsAdj
858858

859859
instance {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* S) :
860-
(extendScalars f).IsLeftAdjoint :=
860+
(extendScalars.{u₁, u₂, max u₂ w} f).IsLeftAdjoint :=
861861
(extendRestrictScalarsAdj f).isLeftAdjoint
862862

863863
instance {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* S) :
864-
(restrictScalars f).IsRightAdjoint :=
864+
(restrictScalars.{max u₂ w, u₁, u₂} f).IsRightAdjoint :=
865865
(extendRestrictScalarsAdj f).isRightAdjoint
866866

867867
noncomputable instance preservesLimitRestrictScalars

Mathlib/CategoryTheory/Category/Grpd.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -135,7 +135,7 @@ def piLimitFanIsLimit ⦃J : Type u⦄ (F : J → Grpd.{u, u}) : Limits.IsLimit
135135
set_option linter.uppercaseLean3 false in
136136
#align category_theory.Groupoid.pi_limit_fan_is_limit CategoryTheory.Grpd.piLimitFanIsLimit
137137

138-
instance has_pi : Limits.HasProducts Grpd.{u, u} :=
138+
instance has_pi : Limits.HasProducts.{u} Grpd.{u, u} :=
139139
Limits.hasProducts_of_limit_fans (by apply piLimitFan) (by apply piLimitFanIsLimit)
140140
set_option linter.uppercaseLean3 false in
141141
#align category_theory.Groupoid.has_pi CategoryTheory.Grpd.has_pi

Mathlib/CategoryTheory/EffectiveEpi/Preserves.lean

+7-5
Original file line numberDiff line numberDiff line change
@@ -103,11 +103,11 @@ class PreservesEffectiveEpiFamilies (F : C ⥤ D) : Prop where
103103
A functor preserves effective epimorphic families if it maps effective epimorphic families to
104104
effective epimorphic families.
105105
-/
106-
preserves : ∀ {α : Type*} {B : C} (X : α → C) (π : (a : α) → (X a ⟶ B)) [EffectiveEpiFamily X π],
106+
preserves : ∀ {α : Type u} {B : C} (X : α → C) (π : (a : α) → (X a ⟶ B)) [EffectiveEpiFamily X π],
107107
EffectiveEpiFamily (fun a ↦ F.obj (X a)) (fun a ↦ F.map (π a))
108108

109-
instance map_effectiveEpiFamily (F : C ⥤ D) [F.PreservesEffectiveEpiFamilies]
110-
{α : Type*} {B : C} (X : α → C) (π : (a : α) → (X a ⟶ B)) [EffectiveEpiFamily X π] :
109+
instance map_effectiveEpiFamily (F : C ⥤ D) [PreservesEffectiveEpiFamilies.{_, _, u} F]
110+
{α : Type u} {B : C} (X : α → C) (π : (a : α) → (X a ⟶ B)) [EffectiveEpiFamily X π] :
111111
EffectiveEpiFamily (fun a ↦ F.obj (X a)) (fun a ↦ F.map (π a)) :=
112112
PreservesEffectiveEpiFamilies.preserves X π
113113

@@ -128,7 +128,8 @@ instance map_finite_effectiveEpiFamily (F : C ⥤ D) [F.PreservesFiniteEffective
128128
EffectiveEpiFamily (fun a ↦ F.obj (X a)) (fun a ↦ F.map (π a)) :=
129129
PreservesFiniteEffectiveEpiFamilies.preserves X π
130130

131-
instance (F : C ⥤ D) [PreservesEffectiveEpiFamilies F] : PreservesFiniteEffectiveEpiFamilies F where
131+
instance (F : C ⥤ D) [PreservesEffectiveEpiFamilies.{_, _, 0} F] :
132+
PreservesFiniteEffectiveEpiFamilies F where
132133
preserves _ _ := inferInstance
133134

134135
instance (F : C ⥤ D) [PreservesFiniteEffectiveEpiFamilies F] : PreservesEffectiveEpis F where
@@ -191,7 +192,8 @@ lemma finite_effectiveEpiFamily_of_map (F : C ⥤ D) [ReflectsFiniteEffectiveEpi
191192
EffectiveEpiFamily X π :=
192193
ReflectsFiniteEffectiveEpiFamilies.reflects X π h
193194

194-
instance (F : C ⥤ D) [ReflectsEffectiveEpiFamilies F] : ReflectsFiniteEffectiveEpiFamilies F where
195+
instance (F : C ⥤ D) [ReflectsEffectiveEpiFamilies.{_, _, 0} F] :
196+
ReflectsFiniteEffectiveEpiFamilies F where
195197
reflects _ _ h := by
196198
have := F.effectiveEpiFamily_of_map _ _ h
197199
infer_instance

Mathlib/CategoryTheory/EssentiallySmall.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -96,7 +96,7 @@ class LocallySmall (C : Type u) [Category.{v} C] : Prop where
9696
hom_small : ∀ X Y : C, Small.{w} (X ⟶ Y) := by infer_instance
9797
#align category_theory.locally_small CategoryTheory.LocallySmall
9898

99-
instance (C : Type u) [Category.{v} C] [LocallySmall.{w} C] (X Y : C) : Small (X ⟶ Y) :=
99+
instance (C : Type u) [Category.{v} C] [LocallySmall.{w} C] (X Y : C) : Small.{w, v} (X ⟶ Y) :=
100100
LocallySmall.hom_small X Y
101101

102102
theorem locallySmall_of_faithful {C : Type u} [Category.{v} C] {D : Type u'} [Category.{v'} D]

Mathlib/CategoryTheory/Limits/Constructions/Over/Basic.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -51,7 +51,7 @@ instance hasFiniteLimits {B : C} [HasFiniteWidePullbacks C] : HasFiniteLimits (O
5151
exact ConstructProducts.over_binaryProduct_of_pullback
5252
#align category_theory.over.has_finite_limits CategoryTheory.Over.hasFiniteLimits
5353

54-
instance hasLimits {B : C} [HasWidePullbacks.{w} C] : HasLimitsOfSize.{w} (Over B) := by
54+
instance hasLimits {B : C} [HasWidePullbacks.{w} C] : HasLimitsOfSize.{w, w} (Over B) := by
5555
apply @has_limits_of_hasEqualizers_and_products _ _ ?_ ?_
5656
· exact ConstructProducts.over_products_of_widePullbacks
5757
· apply @hasEqualizers_of_hasPullbacks_and_binary_products _ _ ?_ _

Mathlib/CategoryTheory/Sites/Adjunction.lean

+8-11
Original file line numberDiff line numberDiff line change
@@ -20,10 +20,10 @@ namespace CategoryTheory
2020

2121
open GrothendieckTopology CategoryTheory Limits Opposite
2222

23-
universe v u
23+
universe v₁ v₂ u₁ u₂
2424

25-
variable {C : Type u} [Category.{v} C] (J : GrothendieckTopology C)
26-
variable {D : Type*} [Category D]
25+
variable {C : Type u} [Category.{v} C] (J : GrothendieckTopology C)
26+
variable {D : Type u₂} [Category.{v₂} D]
2727
variable {E : Type*} [Category E]
2828
variable {F : D ⥤ E} {G : E ⥤ D}
2929
variable [HasWeakSheafify J D]
@@ -119,21 +119,21 @@ section ForgetToType
119119
variable [ConcreteCategory D] [HasSheafCompose J (forget D)]
120120

121121
/-- This is the functor sending a sheaf of types `X` to the sheafification of `X ⋙ G`. -/
122-
abbrev composeAndSheafifyFromTypes (G : Type max v u ⥤ D) : SheafOfTypes J ⥤ Sheaf J D :=
122+
abbrev composeAndSheafifyFromTypes (G : Type max v₁ u₁ ⥤ D) : SheafOfTypes J ⥤ Sheaf J D :=
123123
(sheafEquivSheafOfTypes J).inverse ⋙ composeAndSheafify _ G
124124
set_option linter.uppercaseLean3 false in
125125
#align category_theory.Sheaf.compose_and_sheafify_from_types CategoryTheory.Sheaf.composeAndSheafifyFromTypes
126126

127127
/-- A variant of the adjunction between sheaf categories, in the case where the right adjoint
128128
is the forgetful functor to sheaves of types. -/
129-
def adjunctionToTypes {G : Type max v u ⥤ D} (adj : G ⊣ forget D) :
129+
def adjunctionToTypes {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D) :
130130
composeAndSheafifyFromTypes J G ⊣ sheafForget J :=
131131
(sheafEquivSheafOfTypes J).symm.toAdjunction.comp (adjunction J adj)
132132
set_option linter.uppercaseLean3 false in
133133
#align category_theory.Sheaf.adjunction_to_types CategoryTheory.Sheaf.adjunctionToTypes
134134

135135
@[simp]
136-
theorem adjunctionToTypes_unit_app_val {G : Type max v u ⥤ D} (adj : G ⊣ forget D)
136+
theorem adjunctionToTypes_unit_app_val {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D)
137137
(Y : SheafOfTypes J) :
138138
((adjunctionToTypes J adj).unit.app Y).val =
139139
(adj.whiskerRight _).unit.app ((sheafOfTypesToPresheaf J).obj Y) ≫
@@ -145,7 +145,7 @@ set_option linter.uppercaseLean3 false in
145145
#align category_theory.Sheaf.adjunction_to_types_unit_app_val CategoryTheory.Sheaf.adjunctionToTypes_unit_app_val
146146

147147
@[simp]
148-
theorem adjunctionToTypes_counit_app_val {G : Type max v u ⥤ D} (adj : G ⊣ forget D)
148+
theorem adjunctionToTypes_counit_app_val {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D)
149149
(X : Sheaf J D) :
150150
((adjunctionToTypes J adj).counit.app X).val =
151151
sheafifyLift J ((Functor.associator _ _ _).hom ≫ (adj.whiskerRight _).counit.app _) X.2 := by
@@ -159,10 +159,7 @@ theorem adjunctionToTypes_counit_app_val {G : Type max v u ⥤ D} (adj : G ⊣ f
159159
NatIso.ofComponents, Adjunction.whiskerRight, Adjunction.mkOfUnitCounit]
160160
simp
161161

162-
set_option linter.uppercaseLean3 false in
163-
#align category_theory.Sheaf.adjunction_to_types_counit_app_val CategoryTheory.Sheaf.adjunctionToTypes_counit_app_val
164-
165-
instance [(forget D).IsRightAdjoint ] : (sheafForget J : Sheaf J D ⥤ _).IsRightAdjoint :=
162+
instance [(forget D).IsRightAdjoint] : (sheafForget (D := D) J).IsRightAdjoint :=
166163
(adjunctionToTypes J (Adjunction.ofIsRightAdjoint (forget D))).isRightAdjoint
167164

168165
end ForgetToType

Mathlib/CategoryTheory/Sites/Equivalence.lean

+7-7
Original file line numberDiff line numberDiff line change
@@ -36,15 +36,15 @@ sufficiently small limits in the sheaf category on the essentially small site.
3636
3737
-/
3838

39-
universe u
39+
universe v₁ v₂ v₃ u₁ u₂ u₃ w
4040

4141
namespace CategoryTheory
4242

4343
open Functor Limits GrothendieckTopology
4444

45-
variable {C : Type*} [Category C] (J : GrothendieckTopology C)
46-
variable {D : Type*} [Category D] (K : GrothendieckTopology D) (e : C ≌ D) (G : D ⥤ C)
47-
variable (A : Type*) [Category A]
45+
variable {C : Type u₁} [Category.{v₁} C] (J : GrothendieckTopology C)
46+
variable {D : Type u₂} [Category.{v₂} D] (K : GrothendieckTopology D) (e : C ≌ D) (G : D ⥤ C)
47+
variable (A : Type u₃) [Category.{v₃} A]
4848

4949
namespace Equivalence
5050

@@ -213,7 +213,7 @@ theorem hasSheafCompose : J.HasSheafCompose F where
213213

214214
end Equivalence
215215

216-
variable [EssentiallySmall C]
216+
variable [EssentiallySmall.{w} C]
217217
variable (B : Type*) [Category B] (F : A ⥤ B)
218218
variable [HasSheafify ((equivSmallModel C).locallyCoverDense J).inducedTopology A]
219219
variable [((equivSmallModel C).locallyCoverDense J).inducedTopology.HasSheafCompose F]
@@ -239,13 +239,13 @@ instance hasSheafComposeEssentiallySmallSite : HasSheafCompose J F :=
239239

240240
instance hasLimitsEssentiallySmallSite
241241
[HasLimits <| Sheaf ((equivSmallModel C).locallyCoverDense J).inducedTopology A] :
242-
HasLimitsOfSize <| Sheaf J A :=
242+
HasLimitsOfSize.{max v₃ w, max v₃ w} <| Sheaf J A :=
243243
Adjunction.has_limits_of_equivalence ((equivSmallModel C).sheafCongr J
244244
((equivSmallModel C).locallyCoverDense J).inducedTopology A).functor
245245

246246
instance hasColimitsEssentiallySmallSite
247247
[HasColimits <| Sheaf ((equivSmallModel C).locallyCoverDense J).inducedTopology A] :
248-
HasColimitsOfSize <| Sheaf J A :=
248+
HasColimitsOfSize.{max v₃ w, max v₃ w} <| Sheaf J A :=
249249
Adjunction.has_colimits_of_equivalence ((equivSmallModel C).sheafCongr J
250250
((equivSmallModel C).locallyCoverDense J).inducedTopology A).functor
251251

Mathlib/Condensed/Limits.lean

+4-2
Original file line numberDiff line numberDiff line change
@@ -20,12 +20,14 @@ instance : HasLimits CondensedSet.{u} := by
2020
change HasLimits (Sheaf _ _)
2121
infer_instance
2222

23-
instance : HasLimitsOfSize.{u} CondensedSet.{u} := hasLimitsOfSizeShrink.{u, u+1, u+1, u} _
23+
instance : HasLimitsOfSize.{u, u + 1} CondensedSet.{u} :=
24+
hasLimitsOfSizeShrink.{u, u+1, u+1, u} _
2425

2526
variable (R : Type (u+1)) [Ring R]
2627

2728
instance : HasLimits (CondensedMod.{u} R) := by
2829
change HasLimits (Sheaf _ _)
2930
infer_instance
3031

31-
instance : HasLimitsOfSize.{u} (CondensedMod.{u} R) := hasLimitsOfSizeShrink.{u, u+1, u+1, u} _
32+
instance : HasLimitsOfSize.{u, u + 1} (CondensedMod.{u} R) :=
33+
hasLimitsOfSizeShrink.{u, u+1, u+1, u} _

Mathlib/Data/Fintype/Prod.lean

+7-5
Original file line numberDiff line numberDiff line change
@@ -76,7 +76,7 @@ theorem infinite_prod : Infinite (α × β) ↔ Infinite α ∧ Nonempty β ∨
7676
exact H'.false
7777
#align infinite_prod infinite_prod
7878

79-
instance Pi.infinite_of_left {ι : Sort*} {π : ι → Sort _} [∀ i, Nontrivial <| π i] [Infinite ι] :
79+
instance Pi.infinite_of_left {ι : Sort*} {π : ι → Type*} [∀ i, Nontrivial <| π i] [Infinite ι] :
8080
Infinite (∀ i : ι, π i) := by
8181
choose m n hm using fun i => exists_pair_ne (π i)
8282
refine Infinite.of_injective (fun i => update m i (n i)) fun x y h => of_not_not fun hne => ?_
@@ -85,25 +85,27 @@ instance Pi.infinite_of_left {ι : Sort*} {π : ι → Sort _} [∀ i, Nontrivia
8585
#align pi.infinite_of_left Pi.infinite_of_left
8686

8787
/-- If at least one `π i` is infinite and the rest nonempty, the pi type of all `π` is infinite. -/
88-
theorem Pi.infinite_of_exists_right {ι : Type*} {π : ι → Type*} (i : ι) [Infinite <| π i]
88+
theorem Pi.infinite_of_exists_right {ι : Sort*} {π : ι → Sort*} (i : ι) [Infinite <| π i]
8989
[∀ i, Nonempty <| π i] : Infinite (∀ i : ι, π i) :=
9090
let ⟨m⟩ := @Pi.instNonempty ι π _
9191
Infinite.of_injective _ (update_injective m i)
9292
#align pi.infinite_of_exists_right Pi.infinite_of_exists_right
9393

9494
/-- See `Pi.infinite_of_exists_right` for the case that only one `π i` is infinite. -/
95-
instance Pi.infinite_of_right {ι : Sort _} {π : ι → Sort _} [∀ i, Infinite <| π i] [Nonempty ι] :
95+
instance Pi.infinite_of_right {ι : Sort*} {π : ι → Type*} [∀ i, Infinite <| π i] [Nonempty ι] :
9696
Infinite (∀ i : ι, π i) :=
9797
Pi.infinite_of_exists_right (Classical.arbitrary ι)
9898
#align pi.infinite_of_right Pi.infinite_of_right
9999

100100
/-- Non-dependent version of `Pi.infinite_of_left`. -/
101-
instance Function.infinite_of_left {ι π : Sort _} [Nontrivial π] [Infinite ι] : Infinite (ι → π) :=
101+
instance Function.infinite_of_left {ι : Sort*} {π : Type*} [Nontrivial π] [Infinite ι] :
102+
Infinite (ι → π) :=
102103
Pi.infinite_of_left
103104
#align function.infinite_of_left Function.infinite_of_left
104105

105106
/-- Non-dependent version of `Pi.infinite_of_exists_right` and `Pi.infinite_of_right`. -/
106-
instance Function.infinite_of_right {ι π : Sort _} [Infinite π] [Nonempty ι] : Infinite (ι → π) :=
107+
instance Function.infinite_of_right {ι : Sort*} {π : Type*} [Infinite π] [Nonempty ι] :
108+
Infinite (ι → π) :=
107109
Pi.infinite_of_right
108110
#align function.infinite_of_right Function.infinite_of_right
109111

Mathlib/Geometry/RingedSpace/Basic.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -37,7 +37,7 @@ namespace AlgebraicGeometry
3737

3838
/-- The type of Ringed spaces, as an abbreviation for `SheafedSpace CommRingCat`. -/
3939
abbrev RingedSpace : TypeMax.{u+1, v+1} :=
40-
SheafedSpace.{_, v, u} CommRingCat.{v}
40+
SheafedSpace.{v+1, v, u} CommRingCat.{v}
4141
set_option linter.uppercaseLean3 false in
4242
#align algebraic_geometry.RingedSpace AlgebraicGeometry.RingedSpace
4343

Mathlib/Geometry/RingedSpace/SheafedSpace.lean

+5-2
Original file line numberDiff line numberDiff line change
@@ -21,7 +21,10 @@ presheaves.
2121
open CategoryTheory TopCat TopologicalSpace Opposite CategoryTheory.Limits CategoryTheory.Category
2222
CategoryTheory.Functor
2323

24-
variable (C : Type*) [Category C]
24+
universe u v
25+
26+
variable (C : Type u) [Category.{v} C]
27+
2528

2629
-- Porting note: removed
2730
-- local attribute [tidy] tactic.op_induction'
@@ -251,7 +254,7 @@ noncomputable instance [HasLimits C] :
251254
limit_isSheaf _ fun j => Sheaf.pushforward_sheaf_of_sheaf _ (K.obj (unop j)).2
252255
(colimit.isoColimitCocone ⟨_, PresheafedSpace.colimitCoconeIsColimit _⟩).symm⟩⟩
253256

254-
instance [HasLimits C] : HasColimits (SheafedSpace C) :=
257+
instance [HasLimits C] : HasColimits.{v} (SheafedSpace C) :=
255258
hasColimits_of_hasColimits_createsColimits forgetToPresheafedSpace
256259

257260
noncomputable instance [HasLimits C] : PreservesColimits (forget C) :=

Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean

+3-2
Original file line numberDiff line numberDiff line change
@@ -296,8 +296,9 @@ instance generators_connected (G) [Groupoid.{u, u} G] [IsConnected G] [IsFreeGro
296296

297297
/-- A vertex group in a free connected groupoid is free. With some work one could drop the
298298
connectedness assumption, by looking at connected components. -/
299-
instance endIsFreeOfConnectedFree {G} [Groupoid G] [IsConnected G] [IsFreeGroupoid G] (r : G) :
300-
IsFreeGroup (End r) :=
299+
instance endIsFreeOfConnectedFree
300+
{G : Type u} [Groupoid G] [IsConnected G] [IsFreeGroupoid G] (r : G) :
301+
IsFreeGroup.{u} (End r) :=
301302
SpanningTree.endIsFree <| geodesicSubtree (symgen r)
302303
#align is_free_groupoid.End_is_free_of_connected_free IsFreeGroupoid.endIsFreeOfConnectedFree
303304

Mathlib/Logic/Equiv/Basic.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -1976,7 +1976,7 @@ end
19761976

19771977
section BinaryOp
19781978

1979-
variable (e : α₁ ≃ β₁) (f : α₁ → α₁ → α₁)
1979+
variable {α₁ β₁ : Type*} (e : α₁ ≃ β₁) (f : α₁ → α₁ → α₁)
19801980

19811981
theorem semiconj_conj (f : α₁ → α₁) : Semiconj e f (e.conj f) := fun x => by simp
19821982
#align equiv.semiconj_conj Equiv.semiconj_conj

0 commit comments

Comments
 (0)