Skip to content

Commit 1251ea6

Browse files
golf
1 parent 45f0066 commit 1251ea6

File tree

1 file changed

+4
-6
lines changed

1 file changed

+4
-6
lines changed

FltRegular/NumberTheory/AuxLemmas.lean

+4-6
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,8 @@ import Mathlib.RingTheory.Valuation.ValuationRing
44
import Mathlib.FieldTheory.Separable
55
import Mathlib.RingTheory.Trace.Defs
66

7+
import Mathlib
8+
79
/-!
810
911
This file contains lemmas that should go somewhere in a file in mathlib.
@@ -132,10 +134,6 @@ lemma IsSeparable_of_isLocalization (R S Rₚ Sₚ) [CommRing R] [CommRing S] [F
132134
[IsScalarTower R S Sₚ] [IsScalarTower R Rₚ Sₚ] (M : Submonoid R) [IsLocalization M Rₚ]
133135
[IsLocalization (Algebra.algebraMapSubmonoid S M) Sₚ] [hRS : Algebra.IsSeparable R S] :
134136
Algebra.IsSeparable Rₚ Sₚ := by
135-
have : algebraMap Rₚ Sₚ = IsLocalization.map (T := Algebra.algebraMapSubmonoid S M) Sₚ
136-
(algebraMap R S) (Submonoid.le_comap_map M) := by
137-
apply IsLocalization.ringHom_ext M
138-
simp only [IsLocalization.map_comp, ← IsScalarTower.algebraMap_eq]
139137
refine ⟨fun x ↦ ?_⟩
140138
obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective (Algebra.algebraMapSubmonoid S M) x
141139
obtain ⟨t, ht, e⟩ := s.prop
@@ -145,8 +143,8 @@ lemma IsSeparable_of_isLocalization (R S Rₚ Sₚ) [CommRing R] [CommRing S] [F
145143
exact isUnit_of_invertible _
146144
· rw [aeval_def]
147145
convert scaleRoots_eval₂_eq_zero _ (r := algebraMap S Sₚ x) _
148-
· rw [this, IsLocalization.map_mk', _root_.map_one, IsLocalization.mk'_eq_mul_mk'_one,
149-
mul_comm]
146+
· rw [IsLocalization.algebraMap_eq_map_map_submonoid M S, IsLocalization.map_mk',
147+
_root_.map_one, IsLocalization.mk'_eq_mul_mk'_one, mul_comm]
150148
congr; ext; exact e.symm
151149
· rw [← aeval_def, ← map_aeval_eq_aeval_map, minpoly.aeval, map_zero]
152150
rw [← IsScalarTower.algebraMap_eq, ← IsScalarTower.algebraMap_eq]

0 commit comments

Comments
 (0)