Skip to content

Commit 628cf18

Browse files
acmepjzjcommelin
andcommitted
feat(Mathlib/FieldTheory/IsSepClosed) separably closed field and separable closure - basic definition (#6285)
Main changes - Add `IsSepClosed` and basic properties. - Add `IsSepClosure` and basic properties. Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent 886c740 commit 628cf18

File tree

2 files changed

+185
-0
lines changed

2 files changed

+185
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1890,6 +1890,7 @@ import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure
18901890
import Mathlib.FieldTheory.IsAlgClosed.Basic
18911891
import Mathlib.FieldTheory.IsAlgClosed.Classification
18921892
import Mathlib.FieldTheory.IsAlgClosed.Spectrum
1893+
import Mathlib.FieldTheory.IsSepClosed
18931894
import Mathlib.FieldTheory.KrullTopology
18941895
import Mathlib.FieldTheory.Laurent
18951896
import Mathlib.FieldTheory.Minpoly.Basic
Lines changed: 184 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,184 @@
1+
/-
2+
Copyright (c) 2023 Jz Pan. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Jz Pan
5+
-/
6+
import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure
7+
8+
/-!
9+
# Separably Closed Field
10+
11+
In this file we define the typeclass for separably closed fields and separable closures,
12+
and prove some of their properties.
13+
14+
## Main Definitions
15+
16+
- `IsSepClosed k` is the typeclass saying `k` is a separably closed field, i.e. every separable
17+
polynomial in `k` splits.
18+
19+
- `IsSepClosure k K` is the typeclass saying `K` is a separable closure of `k`, where `k` is a
20+
field. This means that `K` is separably closed and separable over `k`.
21+
22+
## Tags
23+
24+
separable closure, separably closed
25+
26+
## TODO
27+
28+
- `IsSepClosed.lift` is a map from a separable extension `L` of `k`, into any separably
29+
closed extension of `k`.
30+
31+
- `IsSepClosed.equiv` is a proof that any two separable closures of the
32+
same field are isomorphic.
33+
34+
- If `K` is a separably closed field (or algebraically closed field) containing `k`, then all
35+
elements of `K` which are separable over `k` form a separable closure of `k`.
36+
37+
- Using the above result, construct a separable closure as a subfield of an algebraic closure.
38+
39+
- If `k` is a perfect field, then its separable closure coincides with its algebraic closure.
40+
41+
- An algebraic extension of a separably closed field is purely inseparable.
42+
43+
- Maximal separable subextension ...
44+
45+
-/
46+
47+
universe u v w
48+
49+
open scoped Classical BigOperators Polynomial
50+
51+
open Polynomial
52+
53+
variable (k : Type u) [Field k] (K : Type v) [Field K]
54+
55+
/-- Typeclass for separably closed fields.
56+
57+
To show `Polynomial.Splits p f` for an arbitrary ring homomorphism `f`,
58+
see `IsSepClosed.splits_codomain` and `IsSepClosed.splits_domain`.
59+
-/
60+
class IsSepClosed : Prop where
61+
splits_of_separable : ∀ p : k[X], p.Separable → (p.Splits <| RingHom.id k)
62+
63+
variable {k} {K}
64+
65+
/-- Every separable polynomial splits in the field extension `f : k →+* K` if `K` is
66+
separably closed.
67+
68+
See also `IsSepClosed.splits_domain` for the case where `k` is separably closed.
69+
-/
70+
theorem IsSepClosed.splits_codomain [IsSepClosed K] {f : k →+* K}
71+
(p : k[X]) (h : p.Separable) : p.Splits f := by
72+
convert IsSepClosed.splits_of_separable (p.map f) (Separable.map h); simp [splits_map_iff]
73+
74+
/-- Every separable polynomial splits in the field extension `f : k →+* K` if `k` is
75+
separably closed.
76+
77+
See also `IsSepClosed.splits_codomain` for the case where `k` is separably closed.
78+
-/
79+
theorem IsSepClosed.splits_domain [IsSepClosed k] {f : k →+* K}
80+
(p : k[X]) (h : p.Separable) : p.Splits f :=
81+
Polynomial.splits_of_splits_id _ <| IsSepClosed.splits_of_separable _ h
82+
83+
namespace IsSepClosed
84+
85+
theorem exists_root [IsSepClosed k] (p : k[X]) (hp : p.degree ≠ 0) (hsep : p.Separable) :
86+
∃ x, IsRoot p x :=
87+
exists_root_of_splits _ (IsSepClosed.splits_of_separable p hsep) hp
88+
89+
theorem exists_pow_nat_eq [IsSepClosed k] (x : k) (n : ℕ) [hn : NeZero (n : k)] :
90+
∃ z, z ^ n = x := by
91+
have hn' : 0 < n := Nat.pos_of_ne_zero <| fun h => by
92+
rw [h, Nat.cast_zero] at hn
93+
exact hn.out rfl
94+
have : degree (X ^ n - C x) ≠ 0 := by
95+
rw [degree_X_pow_sub_C hn' x]
96+
exact (WithBot.coe_lt_coe.2 hn').ne'
97+
by_cases hx : x = 0
98+
· exact ⟨0, by rw [hx, pow_eq_zero_iff hn']⟩
99+
· obtain ⟨z, hz⟩ := exists_root _ this <| separable_X_pow_sub_C x hn.out hx
100+
use z
101+
simpa [eval_C, eval_X, eval_pow, eval_sub, IsRoot.def, sub_eq_zero] using hz
102+
103+
theorem exists_eq_mul_self [IsSepClosed k] (x : k) [h2 : NeZero (2 : k)] : ∃ z, x = z * z := by
104+
rcases exists_pow_nat_eq x 2 with ⟨z, rfl⟩
105+
exact ⟨z, sq z⟩
106+
107+
theorem roots_eq_zero_iff [IsSepClosed k] {p : k[X]} (hsep : p.Separable) :
108+
p.roots = 0 ↔ p = Polynomial.C (p.coeff 0) := by
109+
refine' ⟨fun h => _, fun hp => by rw [hp, roots_C]⟩
110+
cases' le_or_lt (degree p) 0 with hd hd
111+
· exact eq_C_of_degree_le_zero hd
112+
· obtain ⟨z, hz⟩ := IsSepClosed.exists_root p hd.ne' hsep
113+
rw [← mem_roots (ne_zero_of_degree_gt hd), h] at hz
114+
simp at hz
115+
116+
theorem exists_eval₂_eq_zero [IsSepClosed K] (f : k →+* K)
117+
(p : k[X]) (hp : p.degree ≠ 0) (hsep : p.Separable) :
118+
∃ x, p.eval₂ f x = 0 :=
119+
let ⟨x, hx⟩ := exists_root (p.map f) (by rwa [degree_map_eq_of_injective f.injective])
120+
(Separable.map hsep)
121+
⟨x, by rwa [eval₂_eq_eval_map, ← IsRoot]⟩
122+
123+
variable (k)
124+
125+
theorem exists_aeval_eq_zero [IsSepClosed K] [Algebra k K] (p : k[X])
126+
(hp : p.degree ≠ 0) (hsep : p.Separable) : ∃ x : K, aeval x p = 0 :=
127+
exists_eval₂_eq_zero (algebraMap k K) p hp hsep
128+
129+
theorem of_exists_root (H : ∀ p : k[X], p.Monic → Irreducible p → Separable p → ∃ x, p.eval x = 0) :
130+
IsSepClosed k := by
131+
refine ⟨fun p hsep ↦ Or.inr ?_⟩
132+
intro q hq hdvd
133+
simp only [map_id] at hdvd
134+
have hlc : IsUnit (leadingCoeff q)⁻¹ := IsUnit.inv <| Ne.isUnit <|
135+
leadingCoeff_ne_zero.2 <| Irreducible.ne_zero hq
136+
have hsep' : Separable (q * C (leadingCoeff q)⁻¹) :=
137+
Separable.mul (Separable.of_dvd hsep hdvd) ((separable_C _).2 hlc)
138+
(by simpa only [← isCoprime_mul_unit_right_right (isUnit_C.2 hlc) q 1, one_mul]
139+
using isCoprime_one_right (x := q))
140+
have hirr' := hq
141+
rw [← irreducible_mul_isUnit (isUnit_C.2 hlc)] at hirr'
142+
obtain ⟨x, hx⟩ := H (q * C (leadingCoeff q)⁻¹) (monic_mul_leadingCoeff_inv hq.ne_zero) hirr' hsep'
143+
exact degree_mul_leadingCoeff_inv q hq.ne_zero ▸ degree_eq_one_of_irreducible_of_root hirr' hx
144+
145+
theorem degree_eq_one_of_irreducible [IsSepClosed k] {p : k[X]}
146+
(hp : Irreducible p) (hsep : p.Separable) : p.degree = 1 :=
147+
degree_eq_one_of_irreducible_of_splits hp (IsSepClosed.splits_codomain p hsep)
148+
149+
variable {k}
150+
151+
theorem algebraMap_surjective
152+
[IsSepClosed k] [Algebra k K] [IsSeparable k K] :
153+
Function.Surjective (algebraMap k K) := by
154+
refine fun x => ⟨-(minpoly k x).coeff 0, ?_⟩
155+
have hq : (minpoly k x).leadingCoeff = 1 := minpoly.monic (IsSeparable.isIntegral k x)
156+
have hsep : (minpoly k x).Separable := IsSeparable.separable k x
157+
have h : (minpoly k x).degree = 1 :=
158+
degree_eq_one_of_irreducible k (minpoly.irreducible (IsSeparable.isIntegral k x)) hsep
159+
have : aeval x (minpoly k x) = 0 := minpoly.aeval k x
160+
rw [eq_X_add_C_of_degree_eq_one h, hq, C_1, one_mul, aeval_add, aeval_X, aeval_C,
161+
add_eq_zero_iff_eq_neg] at this
162+
exact (RingHom.map_neg (algebraMap k K) ((minpoly k x).coeff 0)).symm ▸ this.symm
163+
164+
end IsSepClosed
165+
166+
variable (k) (K)
167+
168+
/-- Typeclass for an extension being a separable closure. -/
169+
class IsSepClosure [Algebra k K] : Prop where
170+
sep_closed : IsSepClosed K
171+
separable : IsSeparable k K
172+
173+
variable {k} {K}
174+
175+
theorem isSepClosure_iff [Algebra k K] :
176+
IsSepClosure k K ↔ IsSepClosed K ∧ IsSeparable k K :=
177+
fun h => ⟨h.1, h.2⟩, fun h => ⟨h.1, h.2⟩⟩
178+
179+
instance (priority := 100) IsSepClosure.normal [Algebra k K]
180+
[IsSepClosure k K] : Normal k K :=
181+
fun x => by apply IsIntegral.isAlgebraic; exact IsSepClosure.separable.isIntegral' x,
182+
fun x => @IsSepClosed.splits_codomain _ _ _ _ (IsSepClosure.sep_closed k) _ _ (by
183+
have : IsSeparable k K := IsSepClosure.separable
184+
exact IsSeparable.separable k x)⟩

0 commit comments

Comments
 (0)