Skip to content

Commit bd4da70

Browse files
committed
chore(AlgebraicGeometry): remove unnecessary uses of Hom.base (#30957)
The locations were found by a linter, that flagged every occurrence of the chain `DFunLike.coe`, `CategoryTheory.ConcreteCategory.hom` , `AlgebraicGeometry.PresheafedSpace.Hom.base`, `AlgebraicGeometry.LocallyRingedSpace.Hom.toHom` , `AlgebraicGeometry.Scheme.Hom.toLRSHom'` and an explicit `.base` projection.
1 parent eba35bf commit bd4da70

30 files changed

+188
-189
lines changed

Mathlib/AlgebraicGeometry/AffineTransitionLimit.lean

Lines changed: 11 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -43,7 +43,7 @@ lemma Scheme.nonempty_of_isLimit [IsCofilteredOrEmpty I]
4343
classical
4444
cases isEmpty_or_nonempty I
4545
· have e := (isLimitEquivIsTerminalOfIsEmpty _ _ hc).uniqueUpToIso specULiftZIsTerminal
46-
exact Nonempty.map e.inv.base inferInstance
46+
exact Nonempty.map e.inv inferInstance
4747
· have i := Nonempty.some ‹Nonempty I›
4848
have : IsCofiltered I := ⟨⟩
4949
let 𝒰 := (D.obj i).affineCover.finiteSubcover
@@ -62,7 +62,7 @@ lemma Scheme.nonempty_of_isLimit [IsCofilteredOrEmpty I]
6262
· simp [g]
6363
· simp
6464
· exact Finset.mem_image_of_mem _ (Finset.mem_univ _)) (by simp)
65-
exact Function.isEmpty F.base
65+
exact Function.isEmpty F
6666
obtain ⟨x, -⟩ :=
6767
Cover.covers (𝒰.pullback₁ (D.map (g i (by simp)))) (Nonempty.some inferInstance)
6868
exact (this _).elim x
@@ -87,7 +87,7 @@ lemma Scheme.nonempty_of_isLimit [IsCofilteredOrEmpty I]
8787
let α : F ⟶ Over.forget _ ⋙ D := Functor.whiskerRight
8888
(Functor.whiskerLeft (Over.post D) (Over.mapPullbackAdj (𝒰.f j)).counit) (Over.forget _)
8989
exact this.map (((Functor.Initial.isLimitWhiskerEquiv (Over.forget i) c).symm hc).lift
90-
((Cones.postcompose α).obj c'.1)).base
90+
((Cones.postcompose α).obj c'.1))
9191

9292
include hc in
9393
open Scheme.IdealSheafData in
@@ -102,8 +102,8 @@ lemma exists_mem_of_isClosed_of_nonempty
102102
(hZc : ∀ (i : I), IsClosed (Z i))
103103
(hZne : ∀ i, (Z i).Nonempty)
104104
(hZcpt : ∀ i, IsCompact (Z i))
105-
(hmapsTo : ∀ {i i' : I} (f : i ⟶ i'), Set.MapsTo (D.map f).base (Z i) (Z i')) :
106-
∃ (s : c.pt), ∀ i, (c.π.app i).base s ∈ Z i := by
105+
(hmapsTo : ∀ {i i' : I} (f : i ⟶ i'), Set.MapsTo (D.map f) (Z i) (Z i')) :
106+
∃ (s : c.pt), ∀ i, c.π.app i s ∈ Z i := by
107107
let D' : I ⥤ Scheme :=
108108
{ obj i := (vanishingIdeal ⟨Z i, hZc i⟩).subscheme
109109
map {X Y} f := subschemeMap _ _ (D.map f) (by
@@ -153,8 +153,8 @@ lemma exists_mem_of_isClosed_of_nonempty'
153153
(hZne : ∀ i hij, (Z i hij).Nonempty)
154154
(hZcpt : ∀ i hij, IsCompact (Z i hij))
155155
(hstab : ∀ (i i' : I) (hi'i : i' ⟶ i) (hij : i ⟶ j),
156-
Set.MapsTo (D.map hi'i).base (Z i' (hi'i ≫ hij)) (Z i hij)) :
157-
∃ (s : c.pt), ∀ i hij, (c.π.app i).base s ∈ Z i hij := by
156+
Set.MapsTo (D.map hi'i) (Z i' (hi'i ≫ hij)) (Z i hij)) :
157+
∃ (s : c.pt), ∀ i hij, c.π.app i s ∈ Z i hij := by
158158
have {i₁ i₂ : Over j} (f : i₁ ⟶ i₂) : IsAffineHom ((Over.forget j ⋙ D).map f) := by
159159
dsimp; infer_instance
160160
simpa [Over.forall_iff] using exists_mem_of_isClosed_of_nonempty (Over.forget j ⋙ D) _
@@ -383,19 +383,19 @@ variable (A : ExistsHomHomCompEqCompAux D t f)
383383

384384
omit [LocallyOfFiniteType f] in
385385
lemma exists_index : ∃ (i' : I) (hii' : i' ⟶ A.i),
386-
((D.map hii' ≫ pullback.lift A.a A.b (A.ha.symm.trans A.hb)).base ⁻¹'
386+
((D.map hii' ≫ pullback.lift A.a A.b (A.ha.symm.trans A.hb)) ⁻¹'
387387
((Scheme.Pullback.diagonalCoverDiagonalRange f A.𝒰S A.𝒰X : Set <|
388388
↑(pullback f f))ᶜ)) = ∅ := by
389389
let W := Scheme.Pullback.diagonalCoverDiagonalRange f A.𝒰S A.𝒰X
390390
by_contra! h
391391
let Z (i' : I) (hii' : i' ⟶ A.i) :=
392-
(D.map hii' ≫ pullback.lift A.a A.b (A.ha.symm.trans A.hb)).base ⁻¹' Wᶜ
392+
(D.map hii' ≫ pullback.lift A.a A.b (A.ha.symm.trans A.hb)) ⁻¹' Wᶜ
393393
have hZ (i') (hii' : i' ⟶ A.i) : IsClosed (Z i' hii') :=
394394
(W.isOpen.isClosed_compl).preimage <| Scheme.Hom.continuous _
395395
obtain ⟨s, hs⟩ := exists_mem_of_isClosed_of_nonempty' D A.c A.hc Z hZ h
396396
(fun _ _ ↦ (hZ _ _).isCompact) (fun i i' hii' hij ↦ by simp [Z, Set.MapsTo])
397397
refine hs A.i (𝟙 A.i) (Scheme.Pullback.range_diagonal_subset_diagonalCoverDiagonalRange _ _ _ ?_)
398-
use (A.c.π.app A.i ≫ A.a).base s
398+
use (A.c.π.app A.i ≫ A.a) s
399399
have H : A.c.π.app A.i ≫ A.a ≫ pullback.diagonal f =
400400
A.c.π.app A.i ≫ pullback.lift A.a A.b (A.ha.symm.trans A.hb) := by ext <;> simp [hab]
401401
simp [← Scheme.Hom.comp_apply, - Scheme.Hom.comp_base, H]
@@ -416,7 +416,7 @@ def g : D.obj A.i' ⟶ pullback f f :=
416416

417417
omit [LocallyOfFiniteType f] in
418418
lemma range_g_subset :
419-
Set.range A.g.base ⊆ Scheme.Pullback.diagonalCoverDiagonalRange f A.𝒰S A.𝒰X := by
419+
Set.range A.g ⊆ Scheme.Pullback.diagonalCoverDiagonalRange f A.𝒰S A.𝒰X := by
420420
simpa [ExistsHomHomCompEqCompAux.hii', g] using A.exists_index.choose_spec.choose_spec
421421

422422
/-- (Implementation)

Mathlib/AlgebraicGeometry/Cover/Directed.lean

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -45,7 +45,7 @@ class LocallyDirected (𝒰 : X.Cover (precoverage P)) [Category 𝒰.I₀] wher
4545
w {i j : 𝒰.I₀} (hij : i ⟶ j) : trans hij ≫ 𝒰.f j = 𝒰.f i := by aesop_cat
4646
directed {i j : 𝒰.I₀} (x : (pullback (𝒰.f i) (𝒰.f j)).carrier) :
4747
∃ (k : 𝒰.I₀) (hki : k ⟶ i) (hkj : k ⟶ j) (y : 𝒰.X k),
48-
(pullback.lift (trans hki) (trans hkj) (by simp [w])).base y = x
48+
pullback.lift (trans hki) (trans hkj) (by simp [w]) y = x
4949
property_trans {i j : 𝒰.I₀} (hij : i ⟶ j) : P (trans hij) := by infer_instance
5050

5151
variable (𝒰 : X.Cover (precoverage P)) [Category 𝒰.I₀] [𝒰.LocallyDirected]
@@ -66,7 +66,7 @@ lemma trans_comp {i j k : 𝒰.I₀} (hij : i ⟶ j) (hjk : j ⟶ k) :
6666

6767
lemma exists_lift_trans_eq {i j : 𝒰.I₀} (x : (pullback (𝒰.f i) (𝒰.f j)).carrier) :
6868
∃ (k : 𝒰.I₀) (hki : k ⟶ i) (hkj : k ⟶ j) (y : 𝒰.X k),
69-
(pullback.lift (𝒰.trans hki) (𝒰.trans hkj) (by simp)).base y = x :=
69+
pullback.lift (𝒰.trans hki) (𝒰.trans hkj) (by simp) y = x :=
7070
LocallyDirected.directed x
7171

7272
lemma property_trans {i j : 𝒰.I₀} (hij : i ⟶ j) : P (𝒰.trans hij) :=
@@ -100,7 +100,7 @@ instance : (𝒰.functorOfLocallyDirected ⋙ Scheme.forget).IsLocallyDirected w
100100
cond {i j k} fi fj xi xj hxij := by
101101
simp only [Functor.comp_obj, Cover.functorOfLocallyDirected_obj, forget_obj, Functor.comp_map,
102102
Cover.functorOfLocallyDirected_map, forget_map] at hxij
103-
have : (𝒰.f i).base xi = (𝒰.f j).base xj := by
103+
have : 𝒰.f i xi = 𝒰.f j xj := by
104104
rw [← 𝒰.trans_map fi, ← 𝒰.trans_map fj, Hom.comp_base, Hom.comp_base,
105105
ConcreteCategory.comp_apply, hxij, ConcreteCategory.comp_apply]
106106
obtain ⟨z, rfl, rfl⟩ := Scheme.Pullback.exists_preimage_pullback xi xj this
@@ -142,10 +142,10 @@ instance locallyDirectedPullbackCover : Cover.LocallyDirected (𝒰.pullback₁
142142
pullback.map _ _ _ _ (𝟙 Y) (pullback.lift (𝒰.trans hki) (𝒰.trans hkj) (by simp)) (𝟙 X)
143143
(by simp) (by simp) ≫ iso.inv := by
144144
apply pullback.hom_ext <;> apply pullback.hom_ext <;> simp [iso]
145-
obtain ⟨k, hki, hkj, yk, hyk⟩ := 𝒰.exists_lift_trans_eq ((iso.hom ≫ pullback.snd _ _).base x)
145+
obtain ⟨k, hki, hkj, yk, hyk⟩ := 𝒰.exists_lift_trans_eq ((iso.hom ≫ pullback.snd _ _) x)
146146
refine ⟨k, hki, hkj, show x ∈ Set.range _ from ?_⟩
147147
rw [this, Scheme.Hom.comp_base, TopCat.coe_comp, Set.range_comp, Pullback.range_map]
148-
use iso.hom.base x
148+
use iso.hom x
149149
simp only [Hom.id_base, TopCat.hom_id, ContinuousMap.coe_id, Set.range_id, Set.preimage_univ,
150150
Set.univ_inter, Set.mem_preimage, Set.mem_range, hom_inv_apply, and_true]
151151
exact ⟨yk, hyk⟩
@@ -232,7 +232,7 @@ def Cover.LocallyDirected.ofIsBasisOpensRange {𝒰 : X.OpenCover} [Preorder
232232
trans_id i := by rw [← cancel_mono (𝒰.f i)]; simp
233233
trans_comp hij hjk := by rw [← cancel_mono (𝒰.f _)]; simp
234234
directed {i j} x := by
235-
have : (pullback.fst (𝒰.f i) (𝒰.f j) ≫ 𝒰.f i).base x ∈
235+
have : (pullback.fst (𝒰.f i) (𝒰.f j) ≫ 𝒰.f i) x ∈
236236
(pullback.fst (𝒰.f i) (𝒰.f j) ≫ 𝒰.f i).opensRange := ⟨x, rfl⟩
237237
obtain ⟨k, ⟨k, rfl⟩, ⟨y, hy⟩, h⟩ := TopologicalSpace.Opens.isBasis_iff_nbhd.mp H this
238238
refine ⟨k, homOfLE <| hle.mpr <| le_trans h ?_, homOfLE <| hle.mpr <| le_trans h ?_, y, ?_⟩

Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean

Lines changed: 9 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -40,7 +40,7 @@ variable (K : Precoverage Scheme.{u})
4040
is jointly surjective. -/
4141
class JointlySurjective (K : Precoverage Scheme.{u}) : Prop where
4242
exists_eq {X : Scheme.{u}} (S : Presieve X) (hS : S ∈ K X) (x : X) :
43-
∃ (Y : Scheme.{u}) (g : Y ⟶ X), S g ∧ x ∈ Set.range g.base
43+
∃ (Y : Scheme.{u}) (g : Y ⟶ X), S g ∧ x ∈ Set.range g
4444

4545
/-- A cover of `X` in the coverage `K` is a `0`-hypercover for `K`. -/
4646
abbrev Cover (K : Precoverage Scheme.{u}) := Precoverage.ZeroHypercover.{v} K
@@ -51,7 +51,7 @@ variable {X Y Z : Scheme.{u}} (𝒰 : X.Cover K) (f : X ⟶ Z) (g : Y ⟶ Z)
5151
variable [∀ x, HasPullback (𝒰.f x ≫ f) g]
5252

5353
lemma Cover.exists_eq [JointlySurjective K] (𝒰 : X.Cover K) (x : X) :
54-
∃ i y, (𝒰.f i).base y = x := by
54+
∃ i y, 𝒰.f i y = x := by
5555
obtain ⟨Y, g, ⟨i⟩, y, hy⟩ := JointlySurjective.exists_eq 𝒰.presieve₀ 𝒰.mem₀ x
5656
use i, y
5757

@@ -60,11 +60,11 @@ def Cover.idx [JointlySurjective K] (𝒰 : X.Cover K) (x : X) : 𝒰.I₀ :=
6060
(𝒰.exists_eq x).choose
6161

6262
lemma Cover.covers [JointlySurjective K] (𝒰 : X.Cover K) (x : X) :
63-
x ∈ Set.range (𝒰.f (𝒰.idx x)).base :=
63+
x ∈ Set.range (𝒰.f (𝒰.idx x)) :=
6464
(𝒰.exists_eq x).choose_spec
6565

6666
theorem Cover.iUnion_range [JointlySurjective K] {X : Scheme.{u}} (𝒰 : X.Cover K) :
67-
⋃ i, Set.range (𝒰.f i).base = Set.univ := by
67+
⋃ i, Set.range (𝒰.f i) = Set.univ := by
6868
rw [Set.eq_univ_iff_forall]
6969
intro x
7070
rw [Set.mem_iUnion]
@@ -80,7 +80,7 @@ section MorphismProperty
8080
variable {P Q : MorphismProperty Scheme.{u}}
8181

8282
lemma presieve₀_mem_precoverage_iff (E : PreZeroHypercover X) :
83-
E.presieve₀ ∈ precoverage P X ↔ (∀ x, ∃ i, x ∈ Set.range (E.f i).base) ∧ ∀ i, P (E.f i) := by
83+
E.presieve₀ ∈ precoverage P X ↔ (∀ x, ∃ i, x ∈ Set.range (E.f i)) ∧ ∀ i, P (E.f i) := by
8484
simp
8585

8686
@[grind ←]
@@ -91,7 +91,7 @@ lemma Cover.map_prop (𝒰 : X.Cover (precoverage P)) (i : 𝒰.I₀) : P (𝒰.
9191
cover `X`, `Cover.mkOfCovers` is an associated `P`-cover of `X`. -/
9292
@[simps!]
9393
def Cover.mkOfCovers (J : Type*) (obj : J → Scheme.{u}) (map : (j : J) → obj j ⟶ X)
94-
(covers : ∀ x, ∃ j y, (map j).base y = x)
94+
(covers : ∀ x, ∃ j y, map j y = x)
9595
(map_prop : ∀ j, P (map j) := by infer_instance) : X.Cover (precoverage P) where
9696
I₀ := J
9797
X := obj
@@ -106,7 +106,7 @@ def coverOfIsIso [P.ContainsIdentities] [P.RespectsIso] {X Y : Scheme.{u}} (f :
106106
[IsIso f] : Cover.{v} (precoverage P) Y :=
107107
.mkOfCovers PUnit (fun _ ↦ X)
108108
(fun _ ↦ f)
109-
(fun x ↦ ⟨⟨⟩, (inv f).base x, by simp [← Hom.comp_apply]⟩)
109+
(fun x ↦ ⟨⟨⟩, inv f x, by simp [← Hom.comp_apply]⟩)
110110
(fun _ ↦ P.of_isIso f)
111111

112112
instance : JointlySurjective (precoverage P) where
@@ -141,7 +141,7 @@ def Cover.copy [P.RespectsIso] {X : Scheme.{u}} (𝒰 : X.Cover (precoverage P))
141141
refine ⟨fun x ↦ ?_, ?_⟩
142142
· obtain ⟨i, y, rfl⟩ := 𝒰.exists_eq x
143143
obtain ⟨i, rfl⟩ := e₁.surjective i
144-
use i, (e₂ i).inv.base y
144+
use i, (e₂ i).inv y
145145
simp [h]
146146
· simp_rw [h, MorphismProperty.cancel_left_of_respectsIso]
147147
intro i
@@ -201,7 +201,7 @@ structure AffineCover (P : MorphismProperty Scheme.{u}) (S : Scheme.{u}) where
201201
/-- given a point of `x : S`, `idx x` is the index of the component which contains `x` -/
202202
idx (x : S) : I₀
203203
/-- the components cover `S` -/
204-
covers (x : S) : x ∈ Set.range (f (idx x)).base
204+
covers (x : S) : x ∈ Set.range (f (idx x))
205205
/-- the component maps satisfy `P` -/
206206
map_prop (j : I₀) : P (f j) := by infer_instance
207207

Mathlib/AlgebraicGeometry/Cover/Open.lean

Lines changed: 7 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -77,10 +77,10 @@ lemma OpenCover.isOpenCover_opensRange {X : Scheme.{u}} (𝒰 : X.OpenCover) :
7777
def OpenCover.finiteSubcover {X : Scheme.{u}} (𝒰 : OpenCover X) [H : CompactSpace X] :
7878
OpenCover X := by
7979
have :=
80-
@CompactSpace.elim_nhds_subcover _ _ H (fun x : X => Set.range (𝒰.f (𝒰.idx x)).base)
80+
@CompactSpace.elim_nhds_subcover _ _ H (fun x : X => Set.range (𝒰.f (𝒰.idx x)))
8181
fun x => (IsOpenImmersion.isOpen_range (𝒰.f (𝒰.idx x))).mem_nhds (𝒰.covers x)
8282
let t := this.choose
83-
have h : ∀ x : X, ∃ y : t, x ∈ Set.range (𝒰.f (𝒰.idx y)).base := by
83+
have h : ∀ x : X, ∃ y : t, x ∈ Set.range (𝒰.f (𝒰.idx y)) := by
8484
intro x
8585
have h' : x ∈ (⊤ : Set X) := trivial
8686
rw [← Classical.choose_spec this, Set.mem_iUnion] at h'
@@ -283,8 +283,8 @@ theorem affineBasisCover_obj (X : Scheme.{u}) (i : X.affineBasisCover.I₀) :
283283

284284
theorem affineBasisCover_map_range (X : Scheme.{u}) (x : X)
285285
(r : (X.local_affine x).choose_spec.choose) :
286-
Set.range (X.affineBasisCover.f ⟨x, r⟩).base =
287-
(X.affineCover.f x).base '' (PrimeSpectrum.basicOpen r).1 := by
286+
Set.range (X.affineBasisCover.f ⟨x, r⟩) =
287+
(X.affineCover.f x) '' (PrimeSpectrum.basicOpen r).1 := by
288288
simp only [affineBasisCover, Precoverage.ZeroHypercover.bind_toPreZeroHypercover, Set.range_comp,
289289
PreZeroHypercover.bind_f, Hom.comp_base, TopCat.hom_comp, ContinuousMap.coe_comp]
290290
congr
@@ -293,16 +293,16 @@ theorem affineBasisCover_map_range (X : Scheme.{u}) (x : X)
293293
theorem affineBasisCover_is_basis (X : Scheme.{u}) :
294294
TopologicalSpace.IsTopologicalBasis
295295
{x : Set X |
296-
∃ a : X.affineBasisCover.I₀, x = Set.range (X.affineBasisCover.f a).base} := by
296+
∃ a : X.affineBasisCover.I₀, x = Set.range (X.affineBasisCover.f a)} := by
297297
apply TopologicalSpace.isTopologicalBasis_of_isOpen_of_nhds
298298
· rintro _ ⟨a, rfl⟩
299299
exact IsOpenImmersion.isOpen_range (X.affineBasisCover.f a)
300300
· rintro a U haU hU
301301
rcases X.affineCover.covers a with ⟨x, e⟩
302-
let U' := (X.affineCover.f (X.affineCover.idx a)).base ⁻¹' U
302+
let U' := (X.affineCover.f (X.affineCover.idx a)) ⁻¹' U
303303
have hxU' : x ∈ U' := by rw [← e] at haU; exact haU
304304
rcases PrimeSpectrum.isBasis_basic_opens.exists_subset_of_mem_open hxU'
305-
((X.affineCover.f (X.affineCover.idx a)).base.hom.continuous_toFun.isOpen_preimage _
305+
((X.affineCover.f (X.affineCover.idx a)).continuous.isOpen_preimage _
306306
hU) with
307307
⟨_, ⟨_, ⟨s, rfl⟩, rfl⟩, hxV, hVU⟩
308308
refine ⟨_, ⟨⟨_, s⟩, rfl⟩, ?_, ?_⟩ <;> rw [affineBasisCover_map_range]

Mathlib/AlgebraicGeometry/Cover/Sigma.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -31,7 +31,7 @@ noncomputable def sigma (𝒰 : Cover.{v} (precoverage P) S) : S.Cover (precover
3131
rw [presieve₀_mem_precoverage_iff]
3232
refine ⟨fun s ↦ ?_, fun _ ↦ IsZariskiLocalAtSource.sigmaDesc 𝒰.map_prop⟩
3333
obtain ⟨i, y, rfl⟩ := 𝒰.exists_eq s
34-
refine ⟨default, (Sigma.ι 𝒰.X i).base y, by simp [← Scheme.Hom.comp_apply]⟩
34+
refine ⟨default, Sigma.ι 𝒰.X i y, by simp [← Scheme.Hom.comp_apply]⟩
3535

3636
variable [P.IsMultiplicative] {𝒰 𝒱 : Scheme.Cover.{v} (precoverage P) S}
3737

0 commit comments

Comments
 (0)