Skip to content

Commit 141971f

Browse files
feat: units of polynomial rings (#4691)
We proved that a polynomial is a unit if and only if all of its coefficients are nilpotent, except the constant term which is a unit. Co-authored-by: Cyprien Chauveau cyprien.chauveau@etu.u-paris.fr Co-authored-by: Lucas Pouillart lucas.pouillart@etu.u-paris.fr Co-authored-by: EmilieUthaiwat <102412311+EmilieUthaiwat@users.noreply.github.com>
1 parent 30a855d commit 141971f

File tree

3 files changed

+82
-0
lines changed

3 files changed

+82
-0
lines changed

Mathlib/Data/Polynomial/RingDivision.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -228,6 +228,8 @@ theorem degree_coe_units [Nontrivial R] (u : R[X]ˣ) : degree (u : R[X]) = 0 :=
228228
degree_eq_zero_of_isUnit ⟨u, rfl⟩
229229
#align polynomial.degree_coe_units Polynomial.degree_coe_units
230230

231+
/-- Characterization of a unit of a polynomial ring over an integral domain `R`.
232+
See `Polynomial.isUnit_iff_coeff_isUnit_isNilpotent` when `R` is a commutative ring. -/
231233
theorem isUnit_iff : IsUnit p ↔ ∃ r : R, IsUnit r ∧ C r = p :=
232234
fun hp =>
233235
⟨p.coeff 0,

Mathlib/RingTheory/Nilpotent.lean

Lines changed: 20 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -11,6 +11,7 @@ Authors: Oliver Nash
1111
import Mathlib.Data.Nat.Choose.Sum
1212
import Mathlib.Algebra.Algebra.Bilinear
1313
import Mathlib.RingTheory.Ideal.Operations
14+
import Mathlib.Algebra.GeomSum
1415

1516
/-!
1617
# Nilpotent elements
@@ -28,6 +29,8 @@ import Mathlib.RingTheory.Ideal.Operations
2829

2930
universe u v
3031

32+
open BigOperators
33+
3134
variable {R S : Type u} {x y : R}
3235

3336
/-- An element is said to be nilpotent if some natural-number-power of it equals zero.
@@ -64,6 +67,23 @@ theorem IsNilpotent.map [MonoidWithZero R] [MonoidWithZero S] {r : R} {F : Type
6467
rw [← map_pow, hr.choose_spec, map_zero]
6568
#align is_nilpotent.map IsNilpotent.map
6669

70+
theorem IsNilpotent.sub_one_isUnit [Ring R] {r : R} (hnil : IsNilpotent r) : IsUnit (r - 1) := by
71+
obtain ⟨n, hn⟩ := hnil
72+
refine' ⟨⟨r - 1, -∑ i in Finset.range n, r ^ i, _, _⟩, rfl⟩
73+
· rw [mul_neg, mul_geom_sum, hn]
74+
simp
75+
· rw [neg_mul, geom_sum_mul, hn]
76+
simp
77+
78+
theorem Commute.IsNilpotent.add_isUnit [Ring R] {r : R} {u : Rˣ} (hnil : IsNilpotent r)
79+
(hru : Commute r (↑u⁻¹ : R)) : IsUnit (u + r) := by
80+
rw [← Units.isUnit_mul_units _ u⁻¹, add_mul, Units.mul_inv, ← IsUnit.neg_iff, add_comm, neg_add,
81+
← sub_eq_add_neg]
82+
obtain ⟨n, hn⟩ := hnil
83+
refine' IsNilpotent.sub_one_isUnit ⟨n, _⟩
84+
rw [neg_pow, hru.mul_pow, hn]
85+
simp
86+
6787
/-- A structure that has zero and pow is reduced if it has no nonzero nilpotent elements. -/
6888
@[mk_iff isReduced_iff]
6989
class IsReduced (R : Type _) [Zero R] [Pow R ℕ] : Prop where

Mathlib/RingTheory/Polynomial/Basic.lean

Lines changed: 60 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -14,6 +14,7 @@ import Mathlib.Data.MvPolynomial.CommRing
1414
import Mathlib.Data.MvPolynomial.Equiv
1515
import Mathlib.RingTheory.Polynomial.Content
1616
import Mathlib.RingTheory.UniqueFactorizationDomain
17+
import Mathlib.RingTheory.Ideal.QuotientOperations
1718

1819
/-!
1920
# 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]
256257
set_option linter.uppercaseLean3 false in
257258
#align polynomial.monic_geom_sum_X Polynomial.monic_geom_sum_X
258259

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+
259267
end Semiring
260268

261269
section Ring
@@ -460,6 +468,58 @@ theorem ker_modByMonicHom (hq : q.Monic) :
460468

461469
end ModByMonic
462470

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+
463523
end CommRing
464524

465525
end Polynomial

0 commit comments

Comments
 (0)