Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 6e59fbe

Browse files
alexjbestjcommelin
andcommitted
feat(field_theory/splitting_field): generalize some results from field to domain (#10726)
This PR does a few things generalizing / golfing facts related to polynomials and splitting fields. - Generalize some results in `data.polynomial.field_division` to division rings - generalize `C_leading_coeff_mul_prod_multiset_X_sub_C` from a field to a domain - same for `prod_multiset_X_sub_C_of_monic_of_roots_card_eq` - add a supporting useful lemma `roots_map_of_injective_card_eq_total_degree` saying that if we already have a full (multi)set of roots over a domain, passing along an injection gives the set of roots of the mapped polynomial Inspired by needs of flt-regular. Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent fec084c commit 6e59fbe

File tree

4 files changed

+145
-33
lines changed

4 files changed

+145
-33
lines changed

src/data/multiset/basic.lean

Lines changed: 17 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -2154,10 +2154,14 @@ begin
21542154
simp only [h, if_false, zero_min],
21552155
end
21562156

2157+
theorem count_map {α β : Type*} (f : α → β) (s : multiset α) [decidable_eq β] (b : β) :
2158+
count b (map f s) = (s.filter (λ a, b = f a)).card :=
2159+
countp_map _ _ _
2160+
21572161
/-- `multiset.map f` preserves `count` if `f` is injective on the set of elements contained in
21582162
the multiset -/
21592163
theorem count_map_eq_count [decidable_eq β] (f : α → β) (s : multiset α)
2160-
(hf : set.inj_on f {x : α | x ∈ s}) (x ∈ s) : (s.map f).count (f x) = s.count x :=
2164+
(hf : set.inj_on f {x : α | x ∈ s}) (x ∈ s) : (s.map f).count (f x) = s.count x :=
21612165
begin
21622166
suffices : (filter (λ (a : α), f x = f a) s).count x = card (filter (λ (a : α), f x = f a) s),
21632167
{ rw [count, countp_map, ← this],
@@ -2166,6 +2170,18 @@ begin
21662170
simp only [count_repeat, eq_self_iff_true, if_true, card_repeat]},
21672171
end
21682172

2173+
/-- `multiset.map f` preserves `count` if `f` is injective -/
2174+
theorem count_map_eq_count' [decidable_eq β] (f : α → β) (s : multiset α)
2175+
(hf : function.injective f) (x : α) : (s.map f).count (f x) = s.count x :=
2176+
begin
2177+
by_cases H : x ∈ s,
2178+
{ exact count_map_eq_count f _ (set.inj_on_of_injective hf _) _ H, },
2179+
{ simp [H, not_exists, count_eq_zero, count_eq_zero_of_not_mem H, hf],
2180+
intros y hy hh,
2181+
apply H,
2182+
rwa [← hf hh], }
2183+
end
2184+
21692185
lemma filter_eq' (s : multiset α) (b : α) : s.filter (= b) = repeat b (count b s) :=
21702186
begin
21712187
ext a,

src/data/polynomial/field_division.lean

Lines changed: 26 additions & 25 deletions
Original file line numberDiff line numberDiff line change
@@ -59,19 +59,8 @@ by simp only [polynomial.coe_norm_unit, normalize_apply, hp.leading_coeff, norm_
5959

6060
end is_domain
6161

62-
section field
63-
variables [field R] {p q : polynomial R}
64-
65-
lemma is_unit_iff_degree_eq_zero : is_unit p ↔ degree p = 0 :=
66-
⟨degree_eq_zero_of_is_unit,
67-
λ h, have degree p ≤ 0, by simp [*, le_refl],
68-
have hc : coeff p 00, from λ hc,
69-
by rw [eq_C_of_degree_le_zero this, hc] at h;
70-
simpa using h,
71-
is_unit_iff_dvd_one.2 ⟨C (coeff p 0)⁻¹, begin
72-
conv in p { rw eq_C_of_degree_le_zero this },
73-
rw [← C_mul, _root_.mul_inv_cancel hc, C_1]
74-
end⟩⟩
62+
section division_ring
63+
variables [division_ring R] {p q : polynomial R}
7564

7665
lemma degree_pos_of_ne_zero_of_nonunit (hp0 : p ≠ 0) (hp : ¬is_unit p) :
7766
0 < degree p :=
@@ -92,6 +81,22 @@ have h₁ : (leading_coeff q)⁻¹ ≠ 0 :=
9281
inv_ne_zero (mt leading_coeff_eq_zero.1 h),
9382
by rw [degree_mul, degree_C h₁, add_zero]
9483

84+
end division_ring
85+
86+
section field
87+
variables [field R] {p q : polynomial R}
88+
89+
lemma is_unit_iff_degree_eq_zero : is_unit p ↔ degree p = 0 :=
90+
⟨degree_eq_zero_of_is_unit,
91+
λ h, have degree p ≤ 0, by simp [*, le_refl],
92+
have hc : coeff p 00, from λ hc,
93+
by rw [eq_C_of_degree_le_zero this, hc] at h;
94+
simpa using h,
95+
is_unit_iff_dvd_one.2 ⟨C (coeff p 0)⁻¹, begin
96+
conv in p { rw eq_C_of_degree_le_zero this },
97+
rw [← C_mul, _root_.mul_inv_cancel hc, C_1]
98+
end⟩⟩
99+
95100
theorem irreducible_of_monic {p : polynomial R} (hp1 : p.monic) (hp2 : p ≠ 1) :
96101
irreducible p ↔ (∀ f g : polynomial R, f.monic → g.monic → f * g = p → f = 1 ∨ g = 1) :=
97102
⟨λ hp3 f g hf hg hfg, or.cases_on (hp3.is_unit_or_is_unit hfg.symm)
@@ -427,7 +432,7 @@ begin
427432
rw [← C_inj, this, C_0],
428433
end
429434

430-
lemma prod_multiset_root_eq_finset_root (p : polynomial R) :
435+
lemma prod_multiset_root_eq_finset_root {R : Type*} [comm_ring R] [is_domain R] {p : polynomial R} :
431436
(multiset.map (λ (a : R), X - C a) p.roots).prod =
432437
∏ a in p.roots.to_finset, (X - C a) ^ root_multiplicity a p :=
433438
by simp only [count_roots, finset.prod_multiset_map_count]
@@ -436,7 +441,7 @@ by simp only [count_roots, finset.prod_multiset_map_count]
436441
lemma prod_multiset_X_sub_C_dvd (p : polynomial R) :
437442
(multiset.map (λ (a : R), X - C a) p.roots).prod ∣ p :=
438443
begin
439-
rw prod_multiset_root_eq_finset_root p,
444+
rw prod_multiset_root_eq_finset_root,
440445
have hcoprime : pairwise (is_coprime on λ (a : R), polynomial.X - C (id a)) :=
441446
pairwise_coprime_X_sub function.injective_id,
442447
have H : pairwise (is_coprime on λ (a : R), (polynomial.X - C (id a)) ^ (root_multiplicity a p)),
@@ -447,7 +452,8 @@ begin
447452
exact pow_root_multiplicity_dvd p a
448453
end
449454

450-
lemma roots_C_mul (p : polynomial R) {a : R} (hzero : a ≠ 0) : (C a * p).roots = p.roots :=
455+
lemma roots_C_mul {R : Type*} [comm_ring R] [is_domain R] (p : polynomial R) {a : R}
456+
(hzero : a ≠ 0) : (C a * p).roots = p.roots :=
451457
begin
452458
by_cases hpzero : p = 0,
453459
{ simp only [hpzero, mul_zero] },
@@ -461,15 +467,10 @@ begin
461467
simp only [mulzero, zero_add]
462468
end
463469

464-
lemma roots_normalize : (normalize p).roots = p.roots :=
465-
begin
466-
by_cases hzero : p = 0,
467-
{ rw [hzero, normalize_zero], },
468-
{ have hcoeff : p.leading_coeff ≠ 0,
469-
{ intro h, exact hzero (leading_coeff_eq_zero.1 h) },
470-
rw [normalize_apply, mul_comm, coe_norm_unit_of_ne_zero hzero,
471-
roots_C_mul _ (inv_ne_zero hcoeff)], },
472-
end
470+
lemma roots_normalize {R : Type*} [comm_ring R] [is_domain R] [normalization_monoid R]
471+
{p : polynomial R} : (normalize p).roots = p.roots :=
472+
by rw [normalize_apply, mul_comm, coe_norm_unit,
473+
roots_C_mul _ (norm_unit (leading_coeff p)).ne_zero]
473474

474475
end field
475476
end polynomial

src/data/polynomial/ring_division.lean

Lines changed: 41 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -392,7 +392,6 @@ end
392392

393393
@[simp] lemma roots_C (x : R) : (C x).roots = 0 :=
394394
if H : x = 0 then by rw [H, C_0, roots_zero] else multiset.ext.mpr $ λ r,
395-
have h : C x ≠ 0, from λ h, H $ C_inj.1 $ h.symm ▸ C_0.symm,
396395
have not_root : ¬ is_root (C x) r := mt (λ (h : eval r (C x) = 0), trans eval_C.symm h) H,
397396
by rw [count_roots, count_zero, root_multiplicity_eq_zero not_root]
398397

@@ -431,6 +430,47 @@ calc ((roots ((X : polynomial R) ^ n - C a)).card : with_bot ℕ)
431430
≤ degree ((X : polynomial R) ^ n - C a) : card_roots (X_pow_sub_C_ne_zero hn a)
432431
... = n : degree_X_pow_sub_C hn a
433432

433+
lemma le_root_multiplicity_map {K L : Type*} [comm_ring K]
434+
[comm_ring L] {p : polynomial K} {f : K →+* L} (hf : function.injective f) (a : K) :
435+
root_multiplicity a p ≤ root_multiplicity (f a) (map f p) :=
436+
begin
437+
by_cases hp0 : p = 0, { simp only [hp0, root_multiplicity_zero, map_zero], },
438+
have hmap : map f p ≠ 0, { simpa only [map_zero] using (map_injective f hf).ne hp0, },
439+
rw [root_multiplicity, root_multiplicity, dif_neg hp0, dif_neg hmap],
440+
simp only [not_not, nat.lt_find_iff, nat.le_find_iff],
441+
intros m hm,
442+
have := ring_hom.map_dvd (map_ring_hom f) (hm m le_rfl),
443+
simpa only [coe_map_ring_hom, map_pow, map_sub, map_X, map_C],
444+
end
445+
446+
lemma count_map_roots {K L : Type*} [comm_ring K] [is_domain K]
447+
[comm_ring L] {p : polynomial K} {f : K →+* L} (hf : function.injective f)
448+
(a : L) :
449+
count a (multiset.map f p.roots) ≤ root_multiplicity a (map f p) :=
450+
begin
451+
by_cases h : ∃ t, f t = a,
452+
{ rcases h with ⟨h_w, rfl⟩,
453+
rw [multiset.count_map_eq_count' f _ hf, count_roots],
454+
exact le_root_multiplicity_map hf h_w },
455+
{ suffices : multiset.count a (multiset.map f p.roots) = 0,
456+
{ rw this, exact zero_le _, },
457+
rw [multiset.count_map, multiset.card_eq_zero, multiset.filter_eq_nil],
458+
rintro k hk rfl,
459+
exact h ⟨k, rfl⟩, },
460+
end
461+
462+
lemma roots_map_of_injective_card_eq_total_degree {K L : Type*} [comm_ring K] [is_domain K]
463+
[comm_ring L] [is_domain L] {p : polynomial K} {f : K →+* L} (hf : function.injective f)
464+
(hroots : p.roots.card = p.nat_degree) :
465+
multiset.map f p.roots = (map f p).roots :=
466+
begin
467+
by_cases hp0 : p = 0, { simp only [hp0, roots_zero, multiset.map_zero, map_zero], },
468+
have hmap : map f p ≠ 0, { simpa only [map_zero] using (map_injective f hf).ne hp0, },
469+
apply multiset.eq_of_le_of_card_le,
470+
{ simpa only [multiset.le_iff_count, count_roots] using count_map_roots hf },
471+
{ simpa only [multiset.card_map, hroots] using (card_roots' _).trans (nat_degree_map_le f p) },
472+
end
473+
434474
section nth_roots
435475

436476
/-- `nth_roots n a` noncomputably returns the solutions to `x ^ n = a`-/

src/field_theory/splitting_field.lean

Lines changed: 61 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -38,7 +38,7 @@ if it is the smallest field extension of `K` such that `f` splits.
3838
`L`, then `algebra.adjoin F S` embeds into `L`.
3939
* `polynomial.is_splitting_field.lift`: An embedding of a splitting field of the polynomial `f` into
4040
another field such that `f` splits.
41-
* `polynomial.is_splitting_field.alg_equiv`: Every splitting field of a polynomial `f` is isomorpic
41+
* `polynomial.is_splitting_field.alg_equiv`: Every splitting field of a polynomial `f` is isomorphic
4242
to `splitting_field f` and thus, being a splitting field is unique up to isomorphism.
4343
4444
-/
@@ -375,12 +375,12 @@ end
375375

376376
/-- A monic polynomial `p` that has as many roots as its degree
377377
can be written `p = ∏(X - a)`, for `a` in `p.roots`. -/
378-
lemma prod_multiset_X_sub_C_of_monic_of_roots_card_eq {p : polynomial K}
378+
lemma prod_multiset_X_sub_C_of_monic_of_roots_card_eq_of_field {p : polynomial K}
379379
(hmonic : p.monic) (hroots : p.roots.card = p.nat_degree) :
380380
(multiset.map (λ (a : K), X - C a) p.roots).prod = p :=
381381
begin
382382
have hprodmonic : (multiset.map (λ (a : K), X - C a) p.roots).prod.monic,
383-
{ simp only [prod_multiset_root_eq_finset_root p,
383+
{ simp only [prod_multiset_root_eq_finset_root,
384384
monic_prod_of_monic, monic_X_sub_C, monic_pow, forall_true_iff] },
385385
have hdegree : (multiset.map (λ (a : K), X - C a) p.roots).prod.nat_degree = p.nat_degree,
386386
{ rw [← hroots, nat_degree_multiset_prod _ (zero_nmem_multiset_map_X_sub_C _ (λ a : K, a))],
@@ -402,11 +402,33 @@ begin
402402
exact eq_of_monic_of_associated hprodmonic hmonic hassoc
403403
end
404404

405+
lemma prod_multiset_X_sub_C_of_monic_of_roots_card_eq {K : Type*} [comm_ring K] [is_domain K]
406+
{p : polynomial K} (hmonic : p.monic) (hroots : p.roots.card = p.nat_degree) :
407+
(multiset.map (λ (a : K), X - C a) p.roots).prod = p :=
408+
begin
409+
apply map_injective _ (is_fraction_ring.injective K (fraction_ring K)),
410+
rw map_multiset_prod,
411+
simp only [map_C, function.comp_app, map_X, multiset.map_map, map_sub],
412+
have : p.roots.map (algebra_map K (fraction_ring K)) =
413+
(map (algebra_map K (fraction_ring K)) p).roots :=
414+
roots_map_of_injective_card_eq_total_degree
415+
(is_fraction_ring.injective K (fraction_ring K)) hroots,
416+
rw ← prod_multiset_X_sub_C_of_monic_of_roots_card_eq_of_field
417+
(monic_map (algebra_map K (fraction_ring K)) hmonic),
418+
{ simp only [map_C, function.comp_app, map_X, map_sub],
419+
congr' 1,
420+
rw ← this,
421+
simp, },
422+
{ rw [nat_degree_map' (is_fraction_ring.injective K (fraction_ring K)), ← this],
423+
simp only [←hroots, multiset.card_map], },
424+
end
425+
405426
/-- A polynomial `p` that has as many roots as its degree
406-
can be written `p = p.leading_coeff * ∏(X - a)`, for `a` in `p.roots`. -/
407-
lemma C_leading_coeff_mul_prod_multiset_X_sub_C {p : polynomial K}
427+
can be written `p = p.leading_coeff * ∏(X - a)`, for `a` in `p.roots`.
428+
Used to prove the more general `C_leading_coeff_mul_prod_multiset_X_sub_C` below. -/
429+
lemma C_leading_coeff_mul_prod_multiset_X_sub_C_of_field {p : polynomial K}
408430
(hroots : p.roots.card = p.nat_degree) :
409-
(C p.leading_coeff) * (multiset.map (λ (a : K), X - C a) p.roots).prod = p :=
431+
C p.leading_coeff * (multiset.map (λ (a : K), X - C a) p.roots).prod = p :=
410432
begin
411433
by_cases hzero : p = 0,
412434
{ rw [hzero, leading_coeff_zero, ring_hom.map_zero, zero_mul], },
@@ -425,6 +447,39 @@ begin
425447
... = p : by simp only [mul_one, ring_hom.map_one], },
426448
end
427449

450+
/-- A polynomial `p` that has as many roots as its degree
451+
can be written `p = p.leading_coeff * ∏(X - a)`, for `a` in `p.roots`. -/
452+
lemma C_leading_coeff_mul_prod_multiset_X_sub_C {K : Type*} [comm_ring K] [is_domain K]
453+
{p : polynomial K} (hroots : p.roots.card = p.nat_degree) :
454+
C p.leading_coeff * (multiset.map (λ (a : K), X - C a) p.roots).prod = p :=
455+
begin
456+
by_cases hzero : p = 0,
457+
{ rw [hzero, leading_coeff_zero, ring_hom.map_zero, zero_mul], },
458+
have hcoeff : p.leading_coeff ≠ 0,
459+
{ intro h, exact hzero (leading_coeff_eq_zero.1 h) },
460+
apply map_injective _ (is_fraction_ring.injective K (fraction_ring K)),
461+
rw [map_mul, map_multiset_prod],
462+
simp only [map_C, function.comp_app, map_X, multiset.map_map, map_sub],
463+
have h : p.roots.map (algebra_map K (fraction_ring K)) =
464+
(map (algebra_map K (fraction_ring K)) p).roots :=
465+
roots_map_of_injective_card_eq_total_degree
466+
(is_fraction_ring.injective K (fraction_ring K)) hroots,
467+
have : multiset.card (map (algebra_map K (fraction_ring K)) p).roots =
468+
(map (algebra_map K (fraction_ring K)) p).nat_degree,
469+
{ rw [nat_degree_map' (is_fraction_ring.injective K (fraction_ring K)), ← h],
470+
simp only [←hroots, multiset.card_map], },
471+
rw [← C_leading_coeff_mul_prod_multiset_X_sub_C_of_field this],
472+
simp only [map_C, function.comp_app, map_X, map_sub],
473+
congr' 2,
474+
{ rw leading_coeff_map_of_leading_coeff_ne_zero,
475+
intro hn,
476+
apply hcoeff,
477+
apply is_fraction_ring.injective K (fraction_ring K),
478+
simp [hn], },
479+
rw ← h,
480+
simp,
481+
end
482+
428483
/-- A polynomial splits if and only if it has as many roots as its degree. -/
429484
lemma splits_iff_card_roots {p : polynomial K} :
430485
splits (ring_hom.id K) p ↔ p.roots.card = p.nat_degree :=

0 commit comments

Comments
 (0)