@@ -38,7 +38,7 @@ if it is the smallest field extension of `K` such that `f` splits.
38
38
`L`, then `algebra.adjoin F S` embeds into `L`.
39
39
* `polynomial.is_splitting_field.lift`: An embedding of a splitting field of the polynomial `f` into
40
40
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
42
42
to `splitting_field f` and thus, being a splitting field is unique up to isomorphism.
43
43
44
44
-/
@@ -375,12 +375,12 @@ end
375
375
376
376
/-- A monic polynomial `p` that has as many roots as its degree
377
377
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}
379
379
(hmonic : p.monic) (hroots : p.roots.card = p.nat_degree) :
380
380
(multiset.map (λ (a : K), X - C a) p.roots).prod = p :=
381
381
begin
382
382
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,
384
384
monic_prod_of_monic, monic_X_sub_C, monic_pow, forall_true_iff] },
385
385
have hdegree : (multiset.map (λ (a : K), X - C a) p.roots).prod.nat_degree = p.nat_degree,
386
386
{ rw [← hroots, nat_degree_multiset_prod _ (zero_nmem_multiset_map_X_sub_C _ (λ a : K, a))],
@@ -402,11 +402,33 @@ begin
402
402
exact eq_of_monic_of_associated hprodmonic hmonic hassoc
403
403
end
404
404
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
+
405
426
/-- 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}
408
430
(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 :=
410
432
begin
411
433
by_cases hzero : p = 0 ,
412
434
{ rw [hzero, leading_coeff_zero, ring_hom.map_zero, zero_mul], },
@@ -425,6 +447,39 @@ begin
425
447
... = p : by simp only [mul_one, ring_hom.map_one], },
426
448
end
427
449
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
+
428
483
/-- A polynomial splits if and only if it has as many roots as its degree. -/
429
484
lemma splits_iff_card_roots {p : polynomial K} :
430
485
splits (ring_hom.id K) p ↔ p.roots.card = p.nat_degree :=
0 commit comments