Skip to content

Commit 675c28e

Browse files
committed
feat(Polynomial): new lemmas and noncommutative generalizations (#19665)
In `Algebra/Polynomial/Eval/Degree.lean`: two new lemmas about the coefficients and zero-ness of the composition of a polynomial with `C r * X`. Noncommutative generalizations: `Algebra/Polynomial/Degree/Lemmas.lean`: two lemmas `RingTheory/Polynomial/Tower.lean`: one lemma `RingTheory/Polynomial/IntegralNormalization.lean`: add `of_commute` versions of two lemmas similar to [Polynomial.scaleRoots_eval₂_mul_of_commute](https://leanprover-community.github.io/mathlib4_docs/Mathlib/RingTheory/Polynomial/ScaleRoots.html#Polynomial.scaleRoots_eval%E2%82%82_mul_of_commute) and generalize one lemma.
1 parent 720f178 commit 675c28e

File tree

4 files changed

+46
-30
lines changed

4 files changed

+46
-30
lines changed

Mathlib/Algebra/Polynomial/Degree/Lemmas.lean

Lines changed: 4 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -369,20 +369,16 @@ theorem leadingCoeff_comp (hq : natDegree q ≠ 0) :
369369

370370
end NoZeroDivisors
371371

372-
section CommRing
373-
variable [CommRing R] {p q : R[X]}
374-
375-
@[simp] lemma comp_neg_X_leadingCoeff_eq (p : R[X]) :
372+
@[simp] lemma comp_neg_X_leadingCoeff_eq [Ring R] (p : R[X]) :
376373
(p.comp (-X)).leadingCoeff = (-1) ^ p.natDegree * p.leadingCoeff := by
377374
nontriviality R
378375
by_cases h : p = 0
379376
· simp [h]
380377
rw [Polynomial.leadingCoeff, natDegree_comp_eq_of_mul_ne_zero, coeff_comp_degree_mul_degree] <;>
381-
simp [mul_comm, h]
382-
383-
variable [IsDomain R]
378+
simp [((Commute.neg_one_left _).pow_left _).eq, h]
384379

385-
lemma comp_eq_zero_iff : p.comp q = 0 ↔ p = 0 ∨ p.eval (q.coeff 0) = 0 ∧ q = C (q.coeff 0) := by
380+
lemma comp_eq_zero_iff [Semiring R] [NoZeroDivisors R] {p q : R[X]} :
381+
p.comp q = 0 ↔ p = 0 ∨ p.eval (q.coeff 0) = 0 ∧ q = C (q.coeff 0) := by
386382
refine ⟨fun h ↦ ?_, Or.rec (fun h ↦ by simp [h]) fun h ↦ by rw [h.2, comp_C, h.1, C_0]⟩
387383
have key : p.natDegree = 0 ∨ q.natDegree = 0 := by
388384
rw [← mul_eq_zero, ← natDegree_comp, h, natDegree_zero]
@@ -392,8 +388,6 @@ lemma comp_eq_zero_iff : p.comp q = 0 ↔ p = 0 ∨ p.eval (q.coeff 0) = 0 ∧ q
392388
· rw [key, comp_C, C_eq_zero] at h
393389
exact Or.inr ⟨h, key⟩
394390

395-
end CommRing
396-
397391
section DivisionRing
398392

399393
variable {K : Type*} [DivisionRing K]

Mathlib/Algebra/Polynomial/Eval/Degree.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -3,6 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker
55
-/
6+
import Mathlib.Algebra.GroupWithZero.NonZeroDivisors
67
import Mathlib.Algebra.Polynomial.Degree.Support
78
import Mathlib.Algebra.Polynomial.Degree.Units
89
import Mathlib.Algebra.Polynomial.Eval.Coeff
@@ -107,6 +108,21 @@ theorem coeff_comp_degree_mul_degree (hqd0 : natDegree q ≠ 0) :
107108
case h₁ =>
108109
simp +contextual
109110

111+
@[simp] lemma comp_C_mul_X_coeff {r : R} {n : ℕ} :
112+
(p.comp <| C r * X).coeff n = p.coeff n * r ^ n := by
113+
simp_rw [comp, eval₂_eq_sum_range, (commute_X _).symm.mul_pow,
114+
← C_pow, finset_sum_coeff, coeff_C_mul, coeff_X_pow]
115+
rw [Finset.sum_eq_single n _ fun h ↦ ?_, if_pos rfl, mul_one]
116+
· intro b _ h; simp_rw [if_neg h.symm, mul_zero]
117+
· rw [coeff_eq_zero_of_natDegree_lt, zero_mul]
118+
rwa [Finset.mem_range_succ_iff, not_le] at h
119+
120+
lemma comp_C_mul_X_eq_zero_iff {r : R} (hr : r ∈ nonZeroDivisors R) :
121+
p.comp (C r * X) = 0 ↔ p = 0 := by
122+
simp_rw [ext_iff]
123+
refine forall_congr' fun n ↦ ?_
124+
rw [comp_C_mul_X_coeff, coeff_zero, mul_right_mem_nonZeroDivisors_eq_zero_iff (pow_mem hr _)]
125+
110126
end Comp
111127

112128
section Map

Mathlib/RingTheory/Polynomial/IntegralNormalization.lean

Lines changed: 23 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -115,24 +115,30 @@ theorem integralNormalization_degree : (integralNormalization p).degree = p.degr
115115
· rw [← degree_scaleRoots, ← integralNormalization_mul_C_leadingCoeff]
116116
exact (degree_mul_le _ _).trans (add_le_of_nonpos_right degree_C_le)
117117

118-
variable [CommSemiring S]
118+
variable {A : Type*} [CommSemiring S] [Semiring A]
119119

120120
theorem leadingCoeff_smul_integralNormalization (p : S[X]) :
121121
p.leadingCoeff • integralNormalization p = scaleRoots p p.leadingCoeff := by
122122
rw [Algebra.smul_def, algebraMap_eq, mul_comm, integralNormalization_mul_C_leadingCoeff]
123123

124-
theorem integralNormalization_eval₂_leadingCoeff_mul (h : 1 ≤ p.natDegree) (f : R →+* S) (x : S) :
124+
theorem integralNormalization_eval₂_leadingCoeff_mul_of_commute (h : 1 ≤ p.natDegree) (f : R →+* A)
125+
(x : A) (h₁ : Commute (f p.leadingCoeff) x) (h₂ : ∀ {r r'}, Commute (f r) (f r')) :
125126
(integralNormalization p).eval₂ f (f p.leadingCoeff * x) =
126127
f p.leadingCoeff ^ (p.natDegree - 1) * p.eval₂ f x := by
127128
rw [eval₂_eq_sum_range, eval₂_eq_sum_range, Finset.mul_sum]
128129
apply Finset.sum_congr
129130
· rw [natDegree_eq_of_degree_eq p.integralNormalization_degree]
130131
intro n _hn
131-
rw [mul_pow, ← mul_assoc, ← f.map_pow, ← f.map_mul,
132-
integralNormalization_coeff_mul_leadingCoeff_pow _ h, f.map_mul, f.map_pow]
133-
ring
132+
rw [h₁.mul_pow, ← mul_assoc, ← f.map_pow, ← f.map_mul,
133+
integralNormalization_coeff_mul_leadingCoeff_pow _ h, f.map_mul, h₂.eq, f.map_pow, mul_assoc]
134134

135-
theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S} (hz : eval₂ f z p = 0)
135+
theorem integralNormalization_eval₂_leadingCoeff_mul (h : 1 ≤ p.natDegree) (f : R →+* S) (x : S) :
136+
(integralNormalization p).eval₂ f (f p.leadingCoeff * x) =
137+
f p.leadingCoeff ^ (p.natDegree - 1) * p.eval₂ f x :=
138+
integralNormalization_eval₂_leadingCoeff_mul_of_commute h _ _ (.all _ _) (.all _ _)
139+
140+
theorem integralNormalization_eval₂_eq_zero_of_commute {p : R[X]} (f : R →+* A) {z : A}
141+
(hz : eval₂ f z p = 0) (h₁ : Commute (f p.leadingCoeff) z) (h₂ : ∀ {r r'}, Commute (f r) (f r'))
136142
(inj : ∀ x : R, f x = 0 → x = 0) :
137143
eval₂ f (f p.leadingCoeff * z) (integralNormalization p) = 0 := by
138144
obtain (h | h) := p.natDegree.eq_zero_or_pos
@@ -141,20 +147,20 @@ theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S}
141147
simp [h0]
142148
· rw [eq_C_of_natDegree_eq_zero h, eval₂_C] at hz
143149
exact absurd (inj _ hz) h0
144-
· rw [integralNormalization_eval₂_leadingCoeff_mul h, hz, mul_zero]
150+
· rw [integralNormalization_eval₂_leadingCoeff_mul_of_commute h _ _ h₁ h₂, hz, mul_zero]
145151

146-
end Semiring
147-
148-
section CommSemiring
149-
150-
variable [CommSemiring R] [CommSemiring S]
152+
theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S} (hz : eval₂ f z p = 0)
153+
(inj : ∀ x : R, f x = 0 → x = 0) :
154+
eval₂ f (f p.leadingCoeff * z) (integralNormalization p) = 0 :=
155+
integralNormalization_eval₂_eq_zero_of_commute _ hz (.all _ _) (.all _ _) inj
151156

152-
theorem integralNormalization_aeval_eq_zero [Algebra R S] {f : R[X]} {z : S} (hz : aeval z f = 0)
153-
(inj : ∀ x : R, algebraMap R S x = 0 → x = 0) :
154-
aeval (algebraMap R S f.leadingCoeff * z) (integralNormalization f) = 0 :=
155-
integralNormalization_eval₂_eq_zero (algebraMap R S) hz inj
157+
theorem integralNormalization_aeval_eq_zero [Algebra S A] {f : S[X]} {z : A} (hz : aeval z f = 0)
158+
(inj : ∀ x : S, algebraMap S A x = 0 → x = 0) :
159+
aeval (algebraMap S A f.leadingCoeff * z) (integralNormalization f) = 0 :=
160+
integralNormalization_eval₂_eq_zero_of_commute (algebraMap S A) hz
161+
(Algebra.commute_algebraMap_left _ _) (.map (.all _ _) _) inj
156162

157-
end CommSemiring
163+
end Semiring
158164

159165
section IsCancelMulZero
160166

Mathlib/RingTheory/Polynomial/Tower.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -36,9 +36,9 @@ theorem aeval_map_algebraMap (x : B) (p : R[X]) : aeval x (map (algebraMap R A)
3636
rw [aeval_def, aeval_def, eval₂_map, IsScalarTower.algebraMap_eq R A B]
3737

3838
@[simp]
39-
lemma eval_map_algebraMap (P : R[X]) (a : A) :
40-
(map (algebraMap R A) P).eval a = aeval a P := by
41-
rw [← aeval_map_algebraMap (A := A), coe_aeval_eq_eval]
39+
lemma eval_map_algebraMap (P : R[X]) (b : B) :
40+
(map (algebraMap R B) P).eval b = aeval b P := by
41+
rw [aeval_def, eval_map]
4242

4343
end Semiring
4444

0 commit comments

Comments
 (0)