@@ -14,6 +14,7 @@ import Mathlib.Data.MvPolynomial.CommRing
14
14
import Mathlib.Data.MvPolynomial.Equiv
15
15
import Mathlib.RingTheory.Polynomial.Content
16
16
import Mathlib.RingTheory.UniqueFactorizationDomain
17
+ import Mathlib.RingTheory.Ideal.QuotientOperations
17
18
18
19
/-!
19
20
# Ring-theoretic supplement of Data.Polynomial.
@@ -256,6 +257,13 @@ theorem monic_geom_sum_X {n : ℕ} (hn : n ≠ 0) : (∑ i in range n, (X : R[X]
256
257
set_option linter.uppercaseLean3 false in
257
258
#align polynomial.monic_geom_sum_X Polynomial.monic_geom_sum_X
258
259
260
+ theorem IsNilpotent.C_mul_X_pow_isNilpotent {r : R} (n : ℕ) (hnil : IsNilpotent r) :
261
+ IsNilpotent ((C r) * X ^ n) := by
262
+ refine' Commute.isNilpotent_mul_left (commute_X_pow _ _).symm _
263
+ obtain ⟨m, hm⟩ := hnil
264
+ refine' ⟨m, _⟩
265
+ rw [← C_pow, hm, C_0]
266
+
259
267
end Semiring
260
268
261
269
section Ring
@@ -460,6 +468,58 @@ theorem ker_modByMonicHom (hq : q.Monic) :
460
468
461
469
end ModByMonic
462
470
471
+ /-- Let `P` be a polynomial over `R`. If its constant term is a unit and its other coefficients are
472
+ nilpotent, then `P` is a unit. -/
473
+ theorem isUnit_of_coeff_isUnit_isNilpotent {P : Polynomial R} (hunit : IsUnit (P.coeff 0 ))
474
+ (hnil : ∀ i, i ≠ 0 → IsNilpotent (P.coeff i)) : IsUnit P := by
475
+ induction' h : P.natDegree using Nat.strong_induction_on with k hind generalizing P
476
+ by_cases hdeg : P.natDegree = 0
477
+ { rw [eq_C_of_natDegree_eq_zero hdeg]
478
+ exact hunit.map C }
479
+ set P₁ := P.eraseLead with hP₁
480
+ suffices IsUnit P₁ by
481
+ rw [← eraseLead_add_monomial_natDegree_leadingCoeff P, ← C_mul_X_pow_eq_monomial]
482
+ obtain ⟨Q, hQ⟩ := this
483
+ rw [← hP₁, ← hQ]
484
+ refine' Commute.IsNilpotent.add_isUnit (IsNilpotent.C_mul_X_pow_isNilpotent _ (hnil _ hdeg))
485
+ ((Commute.all _ _).mul_left (Commute.all _ _))
486
+ have hdeg₂ := lt_of_le_of_lt P.eraseLead_natDegree_le (Nat.sub_lt
487
+ (Nat.pos_of_ne_zero hdeg) zero_lt_one)
488
+ refine' hind P₁.natDegree _ _ (fun i hi => _) rfl
489
+ · simp_rw [← h, hdeg₂]
490
+ · simp_rw [eraseLead_coeff_of_ne _ (Ne.symm hdeg), hunit]
491
+ · by_cases H : i ≤ P₁.natDegree
492
+ simp_rw [eraseLead_coeff_of_ne _ (ne_of_lt (lt_of_le_of_lt H hdeg₂)), hnil i hi]
493
+ simp_rw [coeff_eq_zero_of_natDegree_lt (lt_of_not_ge H), IsNilpotent.zero]
494
+
495
+ /-- Let `P` be a polynomial over `R`. If `P` is a unit, then all its coefficients are nilpotent,
496
+ except its constant term which is a unit. -/
497
+ theorem coeff_isUnit_isNilpotent_of_isUnit {P : Polynomial R} (hunit : IsUnit P) :
498
+ IsUnit (P.coeff 0 ) ∧ (∀ i, i ≠ 0 → IsNilpotent (P.coeff i)) := by
499
+ obtain ⟨Q, hQ⟩ := IsUnit.exists_right_inv hunit
500
+ constructor
501
+ . refine' isUnit_of_mul_eq_one _ (Q.coeff 0 ) _
502
+ have h := (mul_coeff_zero P Q).symm
503
+ rwa [hQ, coeff_one_zero] at h
504
+ . intros n hn
505
+ rw [nilpotent_iff_mem_prime]
506
+ intros I hI
507
+ let f := mapRingHom (Ideal.Quotient.mk I)
508
+ have hPQ : degree (f P) = 0 ∧ degree (f Q) = 0 := by
509
+ rw [← Nat.WithBot.add_eq_zero_iff, ← degree_mul, ← _root_.map_mul, hQ, map_one, degree_one]
510
+ have hcoeff : (f P).coeff n = 0 := by
511
+ refine' coeff_eq_zero_of_degree_lt _
512
+ rw [hPQ.1 ]
513
+ exact (@WithBot.coe_pos _ _ _ n).2 (Ne.bot_lt hn)
514
+ rw [coe_mapRingHom, coeff_map, ← RingHom.mem_ker, Ideal.mk_ker] at hcoeff
515
+ exact hcoeff
516
+
517
+ /-- Let `P` be a polynomial over `R`. `P` is a unit if and only if all its coefficients are
518
+ nilpotent, except its constant term which is a unit. -/
519
+ theorem isUnit_iff_coeff_isUnit_isNilpotent (P : Polynomial R) :
520
+ IsUnit P ↔ IsUnit (P.coeff 0 ) ∧ (∀ i, i ≠ 0 → IsNilpotent (P.coeff i)) :=
521
+ ⟨coeff_isUnit_isNilpotent_of_isUnit, fun H => isUnit_of_coeff_isUnit_isNilpotent H.1 H.2 ⟩
522
+
463
523
end CommRing
464
524
465
525
end Polynomial
0 commit comments