Skip to content

Commit c187e36

Browse files
committed
chore(AlgebraicGeometry): minor cleanup (#13080)
Motivated by trying to clean up `CommRingCat` and friends, but this is a downstream fix that seems fine independently (and helpful later).
1 parent af03206 commit c187e36

File tree

3 files changed

+9
-12
lines changed

3 files changed

+9
-12
lines changed

Mathlib/AlgebraicGeometry/Spec.lean

Lines changed: 6 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -70,13 +70,14 @@ set_option linter.uppercaseLean3 false in
7070

7171
@[simp]
7272
theorem Spec.topMap_id (R : CommRingCat.{u}) : Spec.topMap (𝟙 R) = 𝟙 (Spec.topObj R) :=
73-
PrimeSpectrum.comap_id
73+
rfl
7474
set_option linter.uppercaseLean3 false in
7575
#align algebraic_geometry.Spec.Top_map_id AlgebraicGeometry.Spec.topMap_id
7676

77+
@[simp]
7778
theorem Spec.topMap_comp {R S T : CommRingCat.{u}} (f : R ⟶ S) (g : S ⟶ T) :
7879
Spec.topMap (f ≫ g) = Spec.topMap g ≫ Spec.topMap f :=
79-
PrimeSpectrum.comap_comp _ _
80+
rfl
8081
set_option linter.uppercaseLean3 false in
8182
#align algebraic_geometry.Spec.Top_map_comp AlgebraicGeometry.Spec.topMap_comp
8283

@@ -88,8 +89,6 @@ set_option linter.uppercaseLean3 false in
8889
def Spec.toTop : CommRingCat.{u}ᵒᵖ ⥤ TopCat where
8990
obj R := Spec.topObj (unop R)
9091
map {R S} f := Spec.topMap f.unop
91-
map_id R := by dsimp; rw [Spec.topMap_id]
92-
map_comp {R S T} f g := by dsimp; rw [Spec.topMap_comp]
9392
set_option linter.uppercaseLean3 false in
9493
#align algebraic_geometry.Spec.to_Top AlgebraicGeometry.Spec.toTop
9594

@@ -120,11 +119,10 @@ set_option linter.uppercaseLean3 false in
120119
theorem Spec.sheafedSpaceMap_id {R : CommRingCat.{u}} :
121120
Spec.sheafedSpaceMap (𝟙 R) = 𝟙 (Spec.sheafedSpaceObj R) :=
122121
AlgebraicGeometry.PresheafedSpace.Hom.ext _ _ (Spec.topMap_id R) <| by
123-
ext U
122+
ext
124123
dsimp
125-
erw [PresheafedSpace.id_c_app, comap_id]; swap
126-
· rw [Spec.topMap_id, TopologicalSpace.Opens.map_id_obj_unop]
127-
simp [eqToHom_map]
124+
erw [comap_id (by simp)]
125+
simp
128126
set_option linter.uppercaseLean3 false in
129127
#align algebraic_geometry.Spec.SheafedSpace_map_id AlgebraicGeometry.Spec.sheafedSpaceMap_id
130128

Mathlib/AlgebraicGeometry/StructureSheaf.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1188,7 +1188,7 @@ The comap of the identity is the identity. In this variant of the lemma, two ope
11881188
`V` are given as arguments, together with a proof that `U = V`. This is useful when `U` and `V`
11891189
are not definitionally equal.
11901190
-/
1191-
theorem comap_id (U V : Opens (PrimeSpectrum.Top R)) (hUV : U = V) :
1191+
theorem comap_id {U V : Opens (PrimeSpectrum.Top R)} (hUV : U = V) :
11921192
(comap (RingHom.id R) U V fun p hpV => by rwa [hUV, PrimeSpectrum.comap_id]) =
11931193
eqToHom (show (structureSheaf R).1.obj (op U) = _ by rw [hUV]) :=
11941194
by erw [comap_id_eq_map U V (eqToHom hUV.symm), eqToHom_op, eqToHom_map]
@@ -1197,7 +1197,7 @@ theorem comap_id (U V : Opens (PrimeSpectrum.Top R)) (hUV : U = V) :
11971197
@[simp]
11981198
theorem comap_id' (U : Opens (PrimeSpectrum.Top R)) :
11991199
(comap (RingHom.id R) U U fun p hpU => by rwa [PrimeSpectrum.comap_id]) = RingHom.id _ := by
1200-
rw [comap_id U U rfl]; rfl
1200+
rw [comap_id rfl]; rfl
12011201
#align algebraic_geometry.structure_sheaf.comap_id' AlgebraicGeometry.StructureSheaf.comap_id'
12021202

12031203
theorem comap_comp (f : R →+* S) (g : S →+* P) (U : Opens (PrimeSpectrum.Top R))

Mathlib/Geometry/RingedSpace/SheafedSpace.lean

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -143,8 +143,7 @@ set_option linter.uppercaseLean3 false in
143143

144144
@[simp]
145145
theorem id_c_app (X : SheafedSpace C) (U) :
146-
(𝟙 X : X ⟶ X).c.app U = eqToHom (by aesop_cat) := by
147-
aesop_cat
146+
(𝟙 X : X ⟶ X).c.app U = 𝟙 _ := rfl
148147
set_option linter.uppercaseLean3 false in
149148
#align algebraic_geometry.SheafedSpace.id_c_app AlgebraicGeometry.SheafedSpace.id_c_app
150149

0 commit comments

Comments
 (0)