Skip to content

Commit 294ff6a

Browse files
committed
perf(BundledCats): more explicit universe annotations (#12741)
Testing for leanprover/lean4#4085 exposed possible regressions typeclass synthesis which can be ameliorated using explicit universes. We do this and we add references to the issue #12737 tracking these changes.
1 parent 4c7c71d commit 294ff6a

File tree

10 files changed

+143
-116
lines changed

10 files changed

+143
-116
lines changed

Mathlib/AlgebraicGeometry/AffineScheme.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -32,6 +32,8 @@ We also define predicates about affine schemes and affine open sets.
3232
3333
-/
3434

35+
-- Explicit universe annotations were used in this file to improve perfomance #12737
36+
3537
set_option linter.uppercaseLean3 false
3638

3739
noncomputable section

Mathlib/AlgebraicGeometry/FunctionField.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -19,6 +19,8 @@ This is a field when the scheme is integral.
1919
field. This map is injective.
2020
-/
2121

22+
-- Explicit universe annotations were used in this file to improve perfomance #12737
23+
2224
set_option linter.uppercaseLean3 false
2325

2426
universe u v

Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -32,6 +32,8 @@ case the unit and the counit would switch to each other.
3232
3333
-/
3434

35+
-- Explicit universe annotations were used in this file to improve perfomance #12737
36+
3537
set_option linter.uppercaseLean3 false
3638

3739
noncomputable section

Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean

Lines changed: 4 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -34,6 +34,8 @@ Further more, these properties are stable under compositions (resp. base change)
3434
3535
-/
3636

37+
-- Explicit universe annotations were used in this file to improve perfomance #12737
38+
3739
universe u
3840

3941
open CategoryTheory Opposite TopologicalSpace CategoryTheory.Limits AlgebraicGeometry
@@ -68,7 +70,7 @@ theorem RespectsIso.basicOpen_iff (hP : RespectsIso @P) {X Y : Scheme.{u}} [IsAf
6870
congr 1
6971
#align ring_hom.respects_iso.basic_open_iff RingHom.RespectsIso.basicOpen_iff
7072

71-
theorem RespectsIso.basicOpen_iff_localization (hP : RespectsIso @P) {X Y : Scheme} [IsAffine X]
73+
theorem RespectsIso.basicOpen_iff_localization (hP : RespectsIso @P) {X Y : Scheme.{u}} [IsAffine X]
7274
[IsAffine Y] (f : X ⟶ Y) (r : Y.presheaf.obj (Opposite.op ⊤)) :
7375
P (Scheme.Γ.map (f ∣_ Y.basicOpen r).op) ↔ P (Localization.awayMap (Scheme.Γ.map f.op) r) := by
7476
refine (hP.basicOpen_iff _ _).trans ?_
@@ -134,7 +136,7 @@ def sourceAffineLocally : AffineTargetMorphismProperty := fun X _ f _ =>
134136
/-- For `P` a property of ring homomorphisms, `affineLocally P` holds for `f : X ⟶ Y` if for each
135137
affine open `U = Spec A ⊆ Y` and `V = Spec B ⊆ f ⁻¹' U`, the ring hom `A ⟶ B` satisfies `P`.
136138
Also see `affineLocally_iff_affineOpens_le`. -/
137-
abbrev affineLocally : MorphismProperty Scheme :=
139+
abbrev affineLocally : MorphismProperty Scheme.{u} :=
138140
targetAffineLocally (sourceAffineLocally @P)
139141
#align algebraic_geometry.affine_locally AlgebraicGeometry.affineLocally
140142

Mathlib/AlgebraicGeometry/OpenImmersion.lean

Lines changed: 44 additions & 42 deletions
Large diffs are not rendered by default.

Mathlib/AlgebraicGeometry/Properties.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -23,8 +23,10 @@ We provide some basic properties of schemes
2323
are reduced.
2424
-/
2525

26-
universe u
2726

27+
-- Explicit universe annotations were used in this file to improve perfomance #12737
28+
29+
universe u
2830

2931
open TopologicalSpace Opposite CategoryTheory CategoryTheory.Limits TopCat
3032

Mathlib/AlgebraicGeometry/Restrict.lean

Lines changed: 39 additions & 36 deletions
Large diffs are not rendered by default.

Mathlib/AlgebraicGeometry/Scheme.lean

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -19,7 +19,12 @@ A morphism of schemes is just a morphism of the underlying locally ringed spaces
1919
2020
-/
2121

22+
-- Explicit universe annotations were used in this file to improve perfomance #12737
23+
2224
set_option linter.uppercaseLean3 false
25+
26+
universe u
27+
2328
noncomputable section
2429

2530
open TopologicalSpace
@@ -245,7 +250,7 @@ def Spec : CommRingCatᵒᵖ ⥤ Scheme where
245250
/-- The empty scheme.
246251
-/
247252
@[simps]
248-
def empty.{u} : Scheme.{u} where
253+
def empty : Scheme where
249254
carrier := TopCat.of PEmpty
250255
presheaf := (CategoryTheory.Functor.const _).obj (CommRingCat.of PUnit)
251256
IsSheaf := Presheaf.isSheaf_of_isTerminal _ CommRingCat.punitIsTerminal

Mathlib/AlgebraicGeometry/Spec.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -34,6 +34,9 @@ The adjunction `Γ ⊣ Spec` is constructed in `Mathlib/AlgebraicGeometry/GammaS
3434
3535
-/
3636

37+
38+
-- Explicit universe annotations were used in this file to improve perfomance #12737
39+
3740
noncomputable section
3841

3942
universe u v

Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean

Lines changed: 38 additions & 34 deletions
Original file line numberDiff line numberDiff line change
@@ -16,6 +16,10 @@ stalks are local rings), and morphisms between these (morphisms in `SheafedSpace
1616
`is_local_ring_hom` on the stalk maps).
1717
-/
1818

19+
-- Explicit universe annotations were used in this file to improve perfomance #12737
20+
21+
universe u
22+
1923
open CategoryTheory
2024

2125
open TopCat
@@ -33,7 +37,7 @@ such that all the stalks are local rings.
3337
3438
A morphism of locally ringed spaces is a morphism of ringed spaces
3539
such that the morphisms induced on stalks are local ring homomorphisms. -/
36-
structure LocallyRingedSpace extends SheafedSpace CommRingCat where
40+
structure LocallyRingedSpace extends SheafedSpace CommRingCat.{u} where
3741
/-- Stalks of a locally ringed space are local rings. -/
3842
localRing : ∀ x, LocalRing (presheaf.stalk x)
3943
set_option linter.uppercaseLean3 false in
@@ -43,7 +47,7 @@ attribute [instance] LocallyRingedSpace.localRing
4347

4448
namespace LocallyRingedSpace
4549

46-
variable (X : LocallyRingedSpace)
50+
variable (X : LocallyRingedSpace.{u})
4751

4852
/-- An alias for `to_SheafedSpace`, where the result type is a `RingedSpace`.
4953
This allows us to use dot-notation for the `RingedSpace` namespace.
@@ -59,7 +63,7 @@ def toTopCat : TopCat :=
5963
set_option linter.uppercaseLean3 false in
6064
#align algebraic_geometry.LocallyRingedSpace.to_Top AlgebraicGeometry.LocallyRingedSpace.toTopCat
6165

62-
instance : CoeSort LocallyRingedSpace (Type*) :=
66+
instance : CoeSort LocallyRingedSpace (Type u) :=
6367
fun X : LocallyRingedSpace => (X.toTopCat : Type _)⟩
6468

6569
instance (x : X) : LocalRing (X.stalk x) :=
@@ -76,7 +80,7 @@ set_option linter.uppercaseLean3 false in
7680
/-- A morphism of locally ringed spaces is a morphism of ringed spaces
7781
such that the morphisms induced on stalks are local ring homomorphisms. -/
7882
@[ext]
79-
structure Hom (X Y : LocallyRingedSpace) : Type _ where
83+
structure Hom (X Y : LocallyRingedSpace.{u}) : Type _ where
8084
/-- the underlying morphism between ringed spaces -/
8185
val : X.toSheafedSpace ⟶ Y.toSheafedSpace
8286
/-- the underlying morphism induces a local ring homomorphism on stalks -/
@@ -87,14 +91,14 @@ set_option linter.uppercaseLean3 false in
8791
instance : Quiver LocallyRingedSpace :=
8892
⟨Hom⟩
8993

90-
@[ext] lemma Hom.ext' (X Y : LocallyRingedSpace) {f g : X ⟶ Y} (h : f.val = g.val) : f = g :=
94+
@[ext] lemma Hom.ext' (X Y : LocallyRingedSpace.{u}) {f g : X ⟶ Y} (h : f.val = g.val) : f = g :=
9195
Hom.ext _ _ h
9296

9397
-- TODO perhaps we should make a bundled `LocalRing` and return one here?
9498
-- TODO define `sheaf.stalk` so we can write `X.𝒪.stalk` here?
9599
/-- The stalk of a locally ringed space, just as a `CommRing`.
96100
-/
97-
noncomputable def stalk (X : LocallyRingedSpace) (x : X) : CommRingCat :=
101+
noncomputable def stalk (X : LocallyRingedSpace.{u}) (x : X) : CommRingCat :=
98102
X.presheaf.stalk x
99103
set_option linter.uppercaseLean3 false in
100104
#align algebraic_geometry.LocallyRingedSpace.stalk AlgebraicGeometry.LocallyRingedSpace.stalk
@@ -106,39 +110,39 @@ instance stalkLocal (x : X) : LocalRing <| X.stalk x := X.localRing x
106110
/-- A morphism of locally ringed spaces `f : X ⟶ Y` induces
107111
a local ring homomorphism from `Y.stalk (f x)` to `X.stalk x` for any `x : X`.
108112
-/
109-
noncomputable def stalkMap {X Y : LocallyRingedSpace} (f : X ⟶ Y) (x : X) :
113+
noncomputable def stalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) :
110114
Y.stalk (f.1.1 x) ⟶ X.stalk x :=
111115
PresheafedSpace.stalkMap f.1 x
112116
set_option linter.uppercaseLean3 false in
113117
#align algebraic_geometry.LocallyRingedSpace.stalk_map AlgebraicGeometry.LocallyRingedSpace.stalkMap
114118

115-
instance {X Y : LocallyRingedSpace} (f : X ⟶ Y) (x : X) : IsLocalRingHom (stalkMap f x) :=
119+
instance {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : IsLocalRingHom (stalkMap f x) :=
116120
f.2 x
117121

118-
instance {X Y : LocallyRingedSpace} (f : X ⟶ Y) (x : X) :
122+
instance {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) :
119123
IsLocalRingHom (PresheafedSpace.stalkMap f.1 x) :=
120124
f.2 x
121125

122126
/-- The identity morphism on a locally ringed space. -/
123127
@[simps]
124-
def id (X : LocallyRingedSpace) : Hom X X :=
128+
def id (X : LocallyRingedSpace.{u}) : Hom X X :=
125129
⟨𝟙 _, fun x => by erw [PresheafedSpace.stalkMap.id]; apply isLocalRingHom_id⟩
126130
set_option linter.uppercaseLean3 false in
127131
#align algebraic_geometry.LocallyRingedSpace.id AlgebraicGeometry.LocallyRingedSpace.id
128132

129-
instance (X : LocallyRingedSpace) : Inhabited (Hom X X) :=
133+
instance (X : LocallyRingedSpace.{u}) : Inhabited (Hom X X) :=
130134
⟨id X⟩
131135

132136
/-- Composition of morphisms of locally ringed spaces. -/
133-
def comp {X Y Z : LocallyRingedSpace} (f : Hom X Y) (g : Hom Y Z) : Hom X Z :=
137+
def comp {X Y Z : LocallyRingedSpace.{u}} (f : Hom X Y) (g : Hom Y Z) : Hom X Z :=
134138
⟨f.val ≫ g.val, fun x => by
135139
erw [PresheafedSpace.stalkMap.comp]
136140
exact @isLocalRingHom_comp _ _ _ _ _ _ _ _ (f.2 _) (g.2 _)⟩
137141
set_option linter.uppercaseLean3 false in
138142
#align algebraic_geometry.LocallyRingedSpace.comp AlgebraicGeometry.LocallyRingedSpace.comp
139143

140144
/-- The category of locally ringed spaces. -/
141-
instance : Category LocallyRingedSpace where
145+
instance : Category LocallyRingedSpace.{u} where
142146
Hom := Hom
143147
id := id
144148
comp {X Y Z} f g := comp f g
@@ -148,7 +152,7 @@ instance : Category LocallyRingedSpace where
148152

149153
/-- The forgetful functor from `LocallyRingedSpace` to `SheafedSpace CommRing`. -/
150154
@[simps]
151-
def forgetToSheafedSpace : LocallyRingedSpace ⥤ SheafedSpace CommRingCat where
155+
def forgetToSheafedSpace : LocallyRingedSpace.{u} ⥤ SheafedSpace CommRingCat.{u} where
152156
obj X := X.toSheafedSpace
153157
map {X Y} f := f.1
154158
set_option linter.uppercaseLean3 false in
@@ -159,13 +163,13 @@ instance : forgetToSheafedSpace.Faithful where
159163

160164
/-- The forgetful functor from `LocallyRingedSpace` to `Top`. -/
161165
@[simps!]
162-
def forgetToTop : LocallyRingedSpace ⥤ TopCat :=
166+
def forgetToTop : LocallyRingedSpace.{u} ⥤ TopCat.{u} :=
163167
forgetToSheafedSpace ⋙ SheafedSpace.forget _
164168
set_option linter.uppercaseLean3 false in
165169
#align algebraic_geometry.LocallyRingedSpace.forget_to_Top AlgebraicGeometry.LocallyRingedSpace.forgetToTop
166170

167171
@[simp]
168-
theorem comp_val {X Y Z : LocallyRingedSpace} (f : X ⟶ Y) (g : Y ⟶ Z) :
172+
theorem comp_val {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) :
169173
(f ≫ g).val = f.val ≫ g.val :=
170174
rfl
171175
set_option linter.uppercaseLean3 false in
@@ -174,13 +178,13 @@ set_option linter.uppercaseLean3 false in
174178
-- Porting note: complains that `(f ≫ g).val.c` can be further simplified
175179
-- so changed to its simp normal form `(f.val ≫ g.val).c`
176180
@[simp]
177-
theorem comp_val_c {X Y Z : LocallyRingedSpace} (f : X ⟶ Y) (g : Y ⟶ Z) :
181+
theorem comp_val_c {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) :
178182
(f.1 ≫ g.1).c = g.val.c ≫ (Presheaf.pushforward _ g.val.base).map f.val.c :=
179183
rfl
180184
set_option linter.uppercaseLean3 false in
181185
#align algebraic_geometry.LocallyRingedSpace.comp_val_c AlgebraicGeometry.LocallyRingedSpace.comp_val_c
182186

183-
theorem comp_val_c_app {X Y Z : LocallyRingedSpace} (f : X ⟶ Y) (g : Y ⟶ Z) (U : (Opens Z)ᵒᵖ) :
187+
theorem comp_val_c_app {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (U : (Opens Z)ᵒᵖ) :
184188
(f ≫ g).val.c.app U = g.val.c.app U ≫ f.val.c.app (op <| (Opens.map g.val.base).obj U.unop) :=
185189
rfl
186190
set_option linter.uppercaseLean3 false in
@@ -192,8 +196,8 @@ spaces can be lifted to a morphism `X ⟶ Y` as locally ringed spaces.
192196
See also `iso_of_SheafedSpace_iso`.
193197
-/
194198
@[simps]
195-
def homOfSheafedSpaceHomOfIsIso {X Y : LocallyRingedSpace} (f : X.toSheafedSpace ⟶ Y.toSheafedSpace)
196-
[IsIso f] : X ⟶ Y :=
199+
def homOfSheafedSpaceHomOfIsIso {X Y : LocallyRingedSpace.{u}}
200+
(f : X.toSheafedSpace ⟶ Y.toSheafedSpace) [IsIso f] : X ⟶ Y :=
197201
Hom.mk f fun x =>
198202
-- Here we need to see that the stalk maps are really local ring homomorphisms.
199203
-- This can be solved by type class inference, because stalk maps of isomorphisms
@@ -210,7 +214,7 @@ This is related to the property that the functor `forget_to_SheafedSpace` reflec
210214
In fact, it is slightly stronger as we do not require `f` to come from a morphism between
211215
_locally_ ringed spaces.
212216
-/
213-
def isoOfSheafedSpaceIso {X Y : LocallyRingedSpace} (f : X.toSheafedSpace ≅ Y.toSheafedSpace) :
217+
def isoOfSheafedSpaceIso {X Y : LocallyRingedSpace.{u}} (f : X.toSheafedSpace ≅ Y.toSheafedSpace) :
214218
X ≅ Y where
215219
hom := homOfSheafedSpaceHomOfIsIso f.hom
216220
inv := homOfSheafedSpaceHomOfIsIso f.inv
@@ -224,15 +228,15 @@ instance : forgetToSheafedSpace.ReflectsIsomorphisms where reflects {_ _} f i :=
224228
⟨homOfSheafedSpaceHomOfIsIso (CategoryTheory.inv (forgetToSheafedSpace.map f)),
225229
Hom.ext _ _ (IsIso.hom_inv_id (I := i)), Hom.ext _ _ (IsIso.inv_hom_id (I := i))⟩ }
226230

227-
instance is_sheafedSpace_iso {X Y : LocallyRingedSpace} (f : X ⟶ Y) [IsIso f] : IsIso f.1 :=
231+
instance is_sheafedSpace_iso {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) [IsIso f] : IsIso f.1 :=
228232
LocallyRingedSpace.forgetToSheafedSpace.map_isIso f
229233
set_option linter.uppercaseLean3 false in
230234
#align algebraic_geometry.LocallyRingedSpace.is_SheafedSpace_iso AlgebraicGeometry.LocallyRingedSpace.is_sheafedSpace_iso
231235

232236
/-- The restriction of a locally ringed space along an open embedding.
233237
-/
234238
@[simps!]
235-
def restrict {U : TopCat} (X : LocallyRingedSpace) {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) :
239+
def restrict {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) :
236240
LocallyRingedSpace where
237241
localRing := by
238242
intro x
@@ -244,23 +248,23 @@ set_option linter.uppercaseLean3 false in
244248
#align algebraic_geometry.LocallyRingedSpace.restrict AlgebraicGeometry.LocallyRingedSpace.restrict
245249

246250
/-- The canonical map from the restriction to the subspace. -/
247-
def ofRestrict {U : TopCat} (X : LocallyRingedSpace) {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) :
248-
X.restrict h ⟶ X :=
251+
def ofRestrict {U : TopCat} (X : LocallyRingedSpace.{u})
252+
{f : U ⟶ X.toTopCat} (h : OpenEmbedding f) : X.restrict h ⟶ X :=
249253
⟨X.toPresheafedSpace.ofRestrict h, fun _ => inferInstance⟩
250254
set_option linter.uppercaseLean3 false in
251255
#align algebraic_geometry.LocallyRingedSpace.of_restrict AlgebraicGeometry.LocallyRingedSpace.ofRestrict
252256

253257
/-- The restriction of a locally ringed space `X` to the top subspace is isomorphic to `X` itself.
254258
-/
255-
def restrictTopIso (X : LocallyRingedSpace) :
259+
def restrictTopIso (X : LocallyRingedSpace.{u}) :
256260
X.restrict (Opens.openEmbedding ⊤) ≅ X :=
257261
isoOfSheafedSpaceIso X.toSheafedSpace.restrictTopIso
258262
set_option linter.uppercaseLean3 false in
259263
#align algebraic_geometry.LocallyRingedSpace.restrict_top_iso AlgebraicGeometry.LocallyRingedSpace.restrictTopIso
260264

261265
/-- The global sections, notated Gamma.
262266
-/
263-
def Γ : LocallyRingedSpaceᵒᵖ ⥤ CommRingCat :=
267+
def Γ : LocallyRingedSpace.{u}ᵒᵖ ⥤ CommRingCat.{u} :=
264268
forgetToSheafedSpace.op ⋙ SheafedSpace.Γ
265269
set_option linter.uppercaseLean3 false in
266270
#align algebraic_geometry.LocallyRingedSpace.Γ AlgebraicGeometry.LocallyRingedSpace.Γ
@@ -271,28 +275,28 @@ set_option linter.uppercaseLean3 false in
271275
#align algebraic_geometry.LocallyRingedSpace.Γ_def AlgebraicGeometry.LocallyRingedSpace.Γ_def
272276

273277
@[simp]
274-
theorem Γ_obj (X : LocallyRingedSpaceᵒᵖ) : Γ.obj X = X.unop.presheaf.obj (op ⊤) :=
278+
theorem Γ_obj (X : LocallyRingedSpace.{u}ᵒᵖ) : Γ.obj X = X.unop.presheaf.obj (op ⊤) :=
275279
rfl
276280
set_option linter.uppercaseLean3 false in
277281
#align algebraic_geometry.LocallyRingedSpace.Γ_obj AlgebraicGeometry.LocallyRingedSpace.Γ_obj
278282

279-
theorem Γ_obj_op (X : LocallyRingedSpace) : Γ.obj (op X) = X.presheaf.obj (op ⊤) :=
283+
theorem Γ_obj_op (X : LocallyRingedSpace.{u}) : Γ.obj (op X) = X.presheaf.obj (op ⊤) :=
280284
rfl
281285
set_option linter.uppercaseLean3 false in
282286
#align algebraic_geometry.LocallyRingedSpace.Γ_obj_op AlgebraicGeometry.LocallyRingedSpace.Γ_obj_op
283287

284288
@[simp]
285-
theorem Γ_map {X Y : LocallyRingedSpaceᵒᵖ} (f : X ⟶ Y) : Γ.map f = f.unop.1.c.app (op ⊤) :=
289+
theorem Γ_map {X Y : LocallyRingedSpace.{u}ᵒᵖ} (f : X ⟶ Y) : Γ.map f = f.unop.1.c.app (op ⊤) :=
286290
rfl
287291
set_option linter.uppercaseLean3 false in
288292
#align algebraic_geometry.LocallyRingedSpace.Γ_map AlgebraicGeometry.LocallyRingedSpace.Γ_map
289293

290-
theorem Γ_map_op {X Y : LocallyRingedSpace} (f : X ⟶ Y) : Γ.map f.op = f.1.c.app (op ⊤) :=
294+
theorem Γ_map_op {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) : Γ.map f.op = f.1.c.app (op ⊤) :=
291295
rfl
292296
set_option linter.uppercaseLean3 false in
293297
#align algebraic_geometry.LocallyRingedSpace.Γ_map_op AlgebraicGeometry.LocallyRingedSpace.Γ_map_op
294298

295-
theorem preimage_basicOpen {X Y : LocallyRingedSpace} (f : X ⟶ Y) {U : Opens Y}
299+
theorem preimage_basicOpen {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) {U : Opens Y}
296300
(s : Y.presheaf.obj (op U)) :
297301
(Opens.map f.1.base).obj (Y.toRingedSpace.basicOpen s) =
298302
@RingedSpace.basicOpen X.toRingedSpace ((Opens.map f.1.base).obj U) (f.1.c.app _ s) := by
@@ -311,7 +315,7 @@ set_option linter.uppercaseLean3 false in
311315

312316
-- This actually holds for all ringed spaces with nontrivial stalks.
313317
@[simp]
314-
theorem basicOpen_zero (X : LocallyRingedSpace) (U : Opens X.carrier) :
318+
theorem basicOpen_zero (X : LocallyRingedSpace.{u}) (U : Opens X.carrier) :
315319
X.toRingedSpace.basicOpen (0 : X.presheaf.obj <| op U) = ⊥ := by
316320
ext x
317321
simp only [RingedSpace.basicOpen, Opens.coe_mk, Set.mem_image, Set.mem_setOf_eq, Subtype.exists,
@@ -324,7 +328,7 @@ theorem basicOpen_zero (X : LocallyRingedSpace) (U : Opens X.carrier) :
324328
set_option linter.uppercaseLean3 false in
325329
#align algebraic_geometry.LocallyRingedSpace.basic_open_zero AlgebraicGeometry.LocallyRingedSpace.basicOpen_zero
326330

327-
instance component_nontrivial (X : LocallyRingedSpace) (U : Opens X.carrier) [hU : Nonempty U] :
331+
instance component_nontrivial (X : LocallyRingedSpace.{u}) (U : Opens X.carrier) [hU : Nonempty U] :
328332
Nontrivial (X.presheaf.obj <| op U) :=
329333
(X.presheaf.germ hU.some).domain_nontrivial
330334
set_option linter.uppercaseLean3 false in

0 commit comments

Comments
 (0)