Skip to content

Commit bfdf7d0

Browse files
committed
chore: golf IsSplittingField.algEquiv (#8142)
Also golfs `Normal.of_algEquiv` and `Algebra.IsIntegral.of_finite` and refactors `Algebra.IsAlgebraic.bijective_of_isScalarTower`.
1 parent b3e4ef9 commit bfdf7d0

File tree

5 files changed

+28
-43
lines changed

5 files changed

+28
-43
lines changed

Mathlib/FieldTheory/Normal.lean

Lines changed: 6 additions & 15 deletions
Original file line numberDiff line numberDiff line change
@@ -110,24 +110,15 @@ theorem AlgHom.normal_bijective [h : Normal F E] (ϕ : E →ₐ[F] K) : Function
110110
-- Porting note: `[Field F] [Field E] [Algebra F E]` added by hand.
111111
variable {F E} {E' : Type*} [Field F] [Field E] [Algebra F E] [Field E'] [Algebra F E']
112112

113-
theorem Normal.of_algEquiv [h : Normal F E] (f : E ≃ₐ[F] E') : Normal F E' :=
114-
normal_iff.2 fun x => by
115-
cases' h.out (f.symm x) with hx hhx
116-
have H := map_isIntegral f.toAlgHom hx
117-
simp [AlgEquiv.toAlgHom_eq_coe] at H
118-
use H
119-
apply Polynomial.splits_of_splits_of_dvd (algebraMap F E') (minpoly.ne_zero hx)
120-
· rw [← AlgHom.comp_algebraMap f.toAlgHom]
121-
exact Polynomial.splits_comp_of_splits (algebraMap F E) f.toAlgHom.toRingHom hhx
122-
· apply minpoly.dvd _ _
123-
rw [← AddEquiv.map_eq_zero_iff f.symm.toAddEquiv]
124-
exact
125-
Eq.trans (Polynomial.aeval_algHom_apply f.symm.toAlgHom x (minpoly F (f.symm x))).symm
126-
(minpoly.aeval _ _)
113+
theorem Normal.of_algEquiv [h : Normal F E] (f : E ≃ₐ[F] E') : Normal F E' := by
114+
rw [normal_iff] at h ⊢
115+
intro x; specialize h (f.symm x)
116+
rw [← f.apply_symm_apply x, minpoly.algEquiv_eq, ← f.toAlgHom.comp_algebraMap]
117+
exact ⟨map_isIntegral f h.1, splits_comp_of_splits _ _ h.2
127118
#align normal.of_alg_equiv Normal.of_algEquiv
128119

129120
theorem AlgEquiv.transfer_normal (f : E ≃ₐ[F] E') : Normal F E ↔ Normal F E' :=
130-
fun _ => Normal.of_algEquiv f, fun _ => Normal.of_algEquiv f.symm⟩
121+
fun _ Normal.of_algEquiv f, fun _ Normal.of_algEquiv f.symm⟩
131122
#align alg_equiv.transfer_normal AlgEquiv.transfer_normal
132123

133124
open IntermediateField

Mathlib/FieldTheory/SplittingField/Construction.lean

Lines changed: 4 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -359,23 +359,10 @@ instance (f : K[X]) : NoZeroSMulDivisors K f.SplittingField :=
359359
inferInstance
360360

361361
/-- Any splitting field is isomorphic to `SplittingFieldAux f`. -/
362-
def algEquiv (f : K[X]) [IsSplittingField K L f] : L ≃ₐ[K] SplittingField f := by
363-
refine'
364-
AlgEquiv.ofBijective (lift L f <| splits (SplittingField f) f)
365-
⟨RingHom.injective (lift L f <| splits (SplittingField f) f).toRingHom, _⟩
366-
haveI := finiteDimensional (SplittingField f) f
367-
haveI := finiteDimensional L f
368-
have : FiniteDimensional.finrank K L = FiniteDimensional.finrank K (SplittingField f) :=
369-
le_antisymm
370-
(LinearMap.finrank_le_finrank_of_injective
371-
(show Function.Injective (lift L f <| splits (SplittingField f) f).toLinearMap from
372-
RingHom.injective (lift L f <| splits (SplittingField f) f : L →+* f.SplittingField)))
373-
(LinearMap.finrank_le_finrank_of_injective
374-
(show Function.Injective (lift (SplittingField f) f <| splits L f).toLinearMap from
375-
RingHom.injective (lift (SplittingField f) f <| splits L f : f.SplittingField →+* L)))
376-
change Function.Surjective (lift L f <| splits (SplittingField f) f).toLinearMap
377-
refine' (LinearMap.injective_iff_surjective_of_finrank_eq_finrank this).1 _
378-
exact RingHom.injective (lift L f <| splits (SplittingField f) f : L →+* f.SplittingField)
362+
def algEquiv (f : K[X]) [h : IsSplittingField K L f] : L ≃ₐ[K] SplittingField f :=
363+
AlgEquiv.ofBijective (lift L f <| splits (SplittingField f) f) <|
364+
have := finiteDimensional L f
365+
((Algebra.isAlgebraic_of_finite K L).algHom_bijective₂ _ <| lift _ f h.1).1
379366
#align polynomial.is_splitting_field.alg_equiv Polynomial.IsSplittingField.algEquiv
380367

381368
end IsSplittingField

Mathlib/Logic/Function/Basic.lean

Lines changed: 9 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -127,7 +127,7 @@ theorem Injective.of_comp {g : γ → α} (I : Injective (f ∘ g)) : Injective
127127
#align function.injective.of_comp Function.Injective.of_comp
128128

129129
@[simp]
130-
theorem Injective.of_comp_iff {f : α → β} (hf : Injective f) (g : γ → α) :
130+
theorem Injective.of_comp_iff (hf : Injective f) (g : γ → α) :
131131
Injective (f ∘ g) ↔ Injective g :=
132132
⟨Injective.of_comp, hf.comp⟩
133133
#align function.injective.of_comp_iff Function.Injective.of_comp_iff
@@ -138,6 +138,10 @@ theorem Injective.of_comp_right {g : γ → α} (I : Injective (f ∘ g)) (hg :
138138
obtain ⟨y, rfl⟩ := hg y
139139
exact congr_arg g (I h)
140140

141+
theorem Surjective.bijective₂_of_injective {g : γ → α} (hf : Surjective f) (hg : Surjective g)
142+
(I : Injective (f ∘ g)) : Bijective f ∧ Bijective g :=
143+
⟨⟨I.of_comp_right hg, hf⟩, I.of_comp, hg⟩
144+
141145
@[simp]
142146
theorem Injective.of_comp_iff' (f : α → β) {g : γ → α} (hg : Bijective g) :
143147
Injective (f ∘ g) ↔ Injective f :=
@@ -182,6 +186,10 @@ theorem Surjective.of_comp_iff (f : α → β) {g : γ → α} (hg : Surjective
182186
theorem Surjective.of_comp_left {g : γ → α} (S : Surjective (f ∘ g)) (hf : Injective f) :
183187
Surjective g := fun a ↦ let ⟨c, hc⟩ := S (f a); ⟨c, hf hc⟩
184188

189+
theorem Injective.bijective₂_of_surjective {g : γ → α} (hf : Injective f) (hg : Injective g)
190+
(S : Surjective (f ∘ g)) : Bijective f ∧ Bijective g :=
191+
⟨⟨hf, S.of_comp⟩, hg, S.of_comp_left hf⟩
192+
185193
@[simp]
186194
theorem Surjective.of_comp_iff' (hf : Bijective f) (g : γ → α) :
187195
Surjective (f ∘ g) ↔ Surjective g :=

Mathlib/RingTheory/Algebraic.lean

Lines changed: 7 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -298,17 +298,20 @@ theorem Algebra.IsAlgebraic.algHom_bijective (ha : Algebra.IsAlgebraic K L) (f :
298298
exact ⟨a, Subtype.ext_iff.1 ha⟩
299299
#align algebra.is_algebraic.alg_hom_bijective Algebra.IsAlgebraic.algHom_bijective
300300

301+
theorem Algebra.IsAlgebraic.algHom_bijective₂ [Field R] [Algebra K R]
302+
(ha : Algebra.IsAlgebraic K L) (f : L →ₐ[K] R) (g : R →ₐ[K] L) :
303+
Function.Bijective f ∧ Function.Bijective g :=
304+
(g.injective.bijective₂_of_surjective f.injective (ha.algHom_bijective <| g.comp f).2).symm
305+
301306
theorem Algebra.IsAlgebraic.bijective_of_isScalarTower (ha : Algebra.IsAlgebraic K L)
302307
[Field R] [Algebra K R] [Algebra L R] [IsScalarTower K L R] (f : R →ₐ[K] L) :
303308
Function.Bijective f :=
304-
⟨f.injective, Function.Surjective.of_comp <|
305-
(ha.algHom_bijective <| f.comp <| IsScalarTower.toAlgHom K L R).2
309+
(ha.algHom_bijective₂ (IsScalarTower.toAlgHom K L R) f).2
306310

307311
theorem Algebra.IsAlgebraic.bijective_of_isScalarTower' [Field R] [Algebra K R]
308312
(ha : Algebra.IsAlgebraic K R) [Algebra L R] [IsScalarTower K L R] (f : R →ₐ[K] L) :
309313
Function.Bijective f :=
310-
⟨f.injective, Function.Surjective.of_comp_left
311-
(ha.algHom_bijective <| (IsScalarTower.toAlgHom K L R).comp f).2 (RingHom.injective _)⟩
314+
(ha.algHom_bijective₂ f (IsScalarTower.toAlgHom K L R)).1
312315

313316
theorem AlgHom.bijective [FiniteDimensional K L] (ϕ : L →ₐ[K] L) : Function.Bijective ϕ :=
314317
(Algebra.isAlgebraic_of_finite K L).algHom_bijective ϕ

Mathlib/RingTheory/IntegralClosure.lean

Lines changed: 2 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -458,12 +458,8 @@ theorem Algebra.IsIntegral.finite (h : Algebra.IsIntegral R A) [h' : Algebra.Fin
458458
exact Algebra.smul_def _ _
459459
#align algebra.is_integral.finite Algebra.IsIntegral.finite
460460

461-
theorem Algebra.IsIntegral.of_finite [h : Module.Finite R A] : Algebra.IsIntegral R A := by
462-
apply RingHom.Finite.to_isIntegral
463-
rw [RingHom.Finite]
464-
convert h
465-
ext
466-
exact (Algebra.smul_def _ _).symm
461+
theorem Algebra.IsIntegral.of_finite [h : Module.Finite R A] : Algebra.IsIntegral R A :=
462+
fun _ ↦ isIntegral_of_mem_of_FG ⊤ h.1 _ trivial
467463
#align algebra.is_integral.of_finite Algebra.IsIntegral.of_finite
468464

469465
/-- finite = integral + finite type -/

0 commit comments

Comments
 (0)