Lemmas for calculating the degree of univariate polynomials #
Main results #
degree_mul
: The degree of the product is the sum of degreesleadingCoeff_add_of_degree_eq
andleadingCoeff_add_of_degree_lt
: The leading_coefficient of a sum is determined by the leading coefficients and degrees
theorem
Polynomial.supDegree_eq_degree
{R : Type u}
[Semiring R]
(p : Polynomial R)
:
AddMonoidAlgebra.supDegree WithBot.some p.toFinsupp = p.degree
theorem
Polynomial.degree_lt_wf
{R : Type u}
[Semiring R]
:
WellFounded fun (p q : Polynomial R) => p.degree < q.degree
Equations
- Polynomial.instWellFoundedRelation = { rel := fun (p q : Polynomial R) => p.degree < q.degree, wf := ⋯ }
theorem
Polynomial.monic_of_subsingleton
{R : Type u}
[Semiring R]
[Subsingleton R]
(p : Polynomial R)
:
p.Monic
theorem
Polynomial.degree_of_subsingleton
{R : Type u}
[Semiring R]
{p : Polynomial R}
[Subsingleton R]
:
theorem
Polynomial.natDegree_of_subsingleton
{R : Type u}
[Semiring R]
{p : Polynomial R}
[Subsingleton R]
:
p.natDegree = 0
theorem
Polynomial.le_natDegree_of_ne_zero
{R : Type u}
{n : ℕ}
[Semiring R]
{p : Polynomial R}
(h : p.coeff n ≠ 0)
:
n ≤ p.natDegree
theorem
Polynomial.degree_eq_of_le_of_coeff_ne_zero
{R : Type u}
{n : ℕ}
[Semiring R]
{p : Polynomial R}
(pn : p.degree ≤ ↑n)
(p1 : p.coeff n ≠ 0)
:
p.degree = ↑n
theorem
Polynomial.natDegree_eq_of_le_of_coeff_ne_zero
{R : Type u}
{n : ℕ}
[Semiring R]
{p : Polynomial R}
(pn : p.natDegree ≤ n)
(p1 : p.coeff n ≠ 0)
:
p.natDegree = n
theorem
Polynomial.natDegree_lt_natDegree
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(hp : p ≠ 0)
(hpq : p.degree < q.degree)
:
p.natDegree < q.natDegree
theorem
Polynomial.coeff_eq_zero_of_degree_lt
{R : Type u}
{n : ℕ}
[Semiring R]
{p : Polynomial R}
(h : p.degree < ↑n)
:
p.coeff n = 0
theorem
Polynomial.coeff_eq_zero_of_natDegree_lt
{R : Type u}
[Semiring R]
{p : Polynomial R}
{n : ℕ}
(h : p.natDegree < n)
:
p.coeff n = 0
@[simp]
theorem
Polynomial.coeff_natDegree_eq_zero_of_degree_lt
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : p.degree < q.degree)
:
p.coeff q.natDegree = 0
theorem
Polynomial.ne_zero_of_degree_gt
{R : Type u}
[Semiring R]
{p : Polynomial R}
{n : WithBot ℕ}
(h : n < p.degree)
:
p ≠ 0
theorem
Polynomial.ne_zero_of_degree_ge_degree
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(hpq : p.degree ≤ q.degree)
(hp : p ≠ 0)
:
q ≠ 0
theorem
Polynomial.ne_zero_of_natDegree_gt
{R : Type u}
[Semiring R]
{p : Polynomial R}
{n : ℕ}
(h : n < p.natDegree)
:
p ≠ 0
theorem
Polynomial.degree_lt_degree
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : p.natDegree < q.natDegree)
:
p.degree < q.degree
theorem
Polynomial.natDegree_lt_natDegree_iff
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(hp : p ≠ 0)
:
theorem
Polynomial.eq_C_of_degree_le_zero
{R : Type u}
[Semiring R]
{p : Polynomial R}
(h : p.degree ≤ 0)
:
p = Polynomial.C (p.coeff 0)
theorem
Polynomial.eq_C_of_degree_eq_zero
{R : Type u}
[Semiring R]
{p : Polynomial R}
(h : p.degree = 0)
:
p = Polynomial.C (p.coeff 0)
theorem
Polynomial.degree_add_eq_left_of_degree_lt
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : q.degree < p.degree)
:
theorem
Polynomial.degree_add_eq_right_of_degree_lt
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : p.degree < q.degree)
:
theorem
Polynomial.natDegree_add_eq_left_of_degree_lt
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : q.degree < p.degree)
:
theorem
Polynomial.natDegree_add_eq_left_of_natDegree_lt
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : q.natDegree < p.natDegree)
:
theorem
Polynomial.natDegree_add_eq_right_of_degree_lt
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : p.degree < q.degree)
:
theorem
Polynomial.natDegree_add_eq_right_of_natDegree_lt
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : p.natDegree < q.natDegree)
:
theorem
Polynomial.degree_add_C
{R : Type u}
{a : R}
[Semiring R]
{p : Polynomial R}
(hp : 0 < p.degree)
:
@[simp]
@[simp]
theorem
Polynomial.degree_add_eq_of_leadingCoeff_add_ne_zero
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : p.leadingCoeff + q.leadingCoeff ≠ 0)
:
theorem
Polynomial.natDegree_eq_of_natDegree_add_lt_left
{R : Type u}
[Semiring R]
(p q : Polynomial R)
(H : (p + q).natDegree < p.natDegree)
:
p.natDegree = q.natDegree
theorem
Polynomial.natDegree_eq_of_natDegree_add_lt_right
{R : Type u}
[Semiring R]
(p q : Polynomial R)
(H : (p + q).natDegree < q.natDegree)
:
p.natDegree = q.natDegree
theorem
Polynomial.natDegree_eq_of_natDegree_add_eq_zero
{R : Type u}
[Semiring R]
(p q : Polynomial R)
(H : (p + q).natDegree = 0)
:
p.natDegree = q.natDegree
theorem
Polynomial.monic_of_natDegree_le_of_coeff_eq_one
{R : Type u}
[Semiring R]
{p : Polynomial R}
(n : ℕ)
(pn : p.natDegree ≤ n)
(p1 : p.coeff n = 1)
:
p.Monic
theorem
Polynomial.monic_of_degree_le_of_coeff_eq_one
{R : Type u}
[Semiring R]
{p : Polynomial R}
(n : ℕ)
(pn : p.degree ≤ ↑n)
(p1 : p.coeff n = 1)
:
p.Monic
theorem
Polynomial.leadingCoeff_add_of_degree_lt
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : p.degree < q.degree)
:
theorem
Polynomial.leadingCoeff_add_of_degree_lt'
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : q.degree < p.degree)
:
@[simp]
theorem
Polynomial.degree_mul'
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : p.leadingCoeff * q.leadingCoeff ≠ 0)
:
theorem
Polynomial.natDegree_mul'
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : p.leadingCoeff * q.leadingCoeff ≠ 0)
:
theorem
Polynomial.leadingCoeff_mul'
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(h : p.leadingCoeff * q.leadingCoeff ≠ 0)
:
theorem
Polynomial.leadingCoeff_monic_mul
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(hp : p.Monic)
:
theorem
Polynomial.leadingCoeff_mul_monic
{R : Type u}
[Semiring R]
{p q : Polynomial R}
(hq : q.Monic)
:
@[simp]
@[simp]
@[simp]
theorem
Polynomial.degree_lt_degree_mul_X
{R : Type u}
[Semiring R]
{p : Polynomial R}
(hp : p ≠ 0)
:
theorem
Polynomial.eq_C_of_natDegree_le_zero
{R : Type u}
[Semiring R]
{p : Polynomial R}
(h : p.natDegree ≤ 0)
:
p = Polynomial.C (p.coeff 0)
theorem
Polynomial.eq_C_of_natDegree_eq_zero
{R : Type u}
[Semiring R]
{p : Polynomial R}
(h : p.natDegree = 0)
:
p = Polynomial.C (p.coeff 0)
theorem
Polynomial.eq_C_coeff_zero_iff_natDegree_eq_zero
{R : Type u}
[Semiring R]
{p : Polynomial R}
:
theorem
Polynomial.eq_one_of_monic_natDegree_zero
{R : Type u}
[Semiring R]
{p : Polynomial R}
(hf : p.Monic)
(hfd : p.natDegree = 0)
:
p = 1
theorem
Polynomial.Monic.natDegree_eq_zero
{R : Type u}
[Semiring R]
{p : Polynomial R}
(hf : p.Monic)
:
@[simp]
theorem
Polynomial.natDegree_mul_X
{R : Type u}
[Semiring R]
[Nontrivial R]
{p : Polynomial R}
(hp : p ≠ 0)
:
@[simp]
theorem
Polynomial.natDegree_X_mul
{R : Type u}
[Semiring R]
[Nontrivial R]
{p : Polynomial R}
(hp : p ≠ 0)
:
@[simp]
theorem
Polynomial.natDegree_mul_X_pow
{R : Type u}
[Semiring R]
[Nontrivial R]
{p : Polynomial R}
(n : ℕ)
(hp : p ≠ 0)
:
@[simp]
theorem
Polynomial.natDegree_X_pow_mul
{R : Type u}
[Semiring R]
[Nontrivial R]
{p : Polynomial R}
(n : ℕ)
(hp : p ≠ 0)
:
@[simp]
@[simp]
theorem
Polynomial.degree_mul_X_pow
{R : Type u}
[Semiring R]
[Nontrivial R]
{p : Polynomial R}
(n : ℕ)
:
theorem
Polynomial.degree_sub_C
{R : Type u}
{a : R}
[Ring R]
{p : Polynomial R}
(hp : 0 < p.degree)
:
@[simp]
theorem
Polynomial.leadingCoeff_sub_of_degree_lt
{R : Type u}
[Ring R]
{p q : Polynomial R}
(h : q.degree < p.degree)
:
theorem
Polynomial.leadingCoeff_sub_of_degree_lt'
{R : Type u}
[Ring R]
{p q : Polynomial R}
(h : p.degree < q.degree)
:
theorem
Polynomial.leadingCoeff_sub_of_degree_eq
{R : Type u}
[Ring R]
{p q : Polynomial R}
(h : p.degree = q.degree)
(hlc : p.leadingCoeff ≠ q.leadingCoeff)
:
theorem
Polynomial.degree_sub_eq_left_of_degree_lt
{R : Type u}
[Ring R]
{p q : Polynomial R}
(h : q.degree < p.degree)
:
theorem
Polynomial.degree_sub_eq_right_of_degree_lt
{R : Type u}
[Ring R]
{p q : Polynomial R}
(h : p.degree < q.degree)
:
theorem
Polynomial.natDegree_sub_eq_left_of_natDegree_lt
{R : Type u}
[Ring R]
{p q : Polynomial R}
(h : q.natDegree < p.natDegree)
:
theorem
Polynomial.natDegree_sub_eq_right_of_natDegree_lt
{R : Type u}
[Ring R]
{p q : Polynomial R}
(h : p.natDegree < q.natDegree)
:
@[simp]
theorem
Polynomial.degree_X_pow_add_C
{R : Type u}
[Nontrivial R]
[Semiring R]
{n : ℕ}
(hn : 0 < n)
(a : R)
:
theorem
Polynomial.X_pow_add_C_ne_zero
{R : Type u}
[Nontrivial R]
[Semiring R]
{n : ℕ}
(hn : 0 < n)
(a : R)
:
theorem
Polynomial.zero_nmem_multiset_map_X_add_C
{R : Type u}
[Nontrivial R]
[Semiring R]
{α : Type u_1}
(m : Multiset α)
(f : α → R)
:
0 ∉ Multiset.map (fun (a : α) => Polynomial.X + Polynomial.C (f a)) m
theorem
Polynomial.X_pow_add_C_ne_one
{R : Type u}
[Nontrivial R]
[Semiring R]
{n : ℕ}
(hn : 0 < n)
(a : R)
:
@[simp]
degree
as a monoid homomorphism between R[X]
and Multiplicative (WithBot ℕ)
.
This is useful to prove results about multiplication and degree.
Equations
- Polynomial.degreeMonoidHom = { toFun := Polynomial.degree, map_one' := ⋯, map_mul' := ⋯ }
Instances For
@[simp]
theorem
Polynomial.degree_pow
{R : Type u}
[Semiring R]
[NoZeroDivisors R]
[Nontrivial R]
(p : Polynomial R)
(n : ℕ)
:
@[simp]
theorem
Polynomial.leadingCoeff_mul
{R : Type u}
[Semiring R]
[NoZeroDivisors R]
(p q : Polynomial R)
:
Polynomial.leadingCoeff
bundled as a MonoidHom
when R
has NoZeroDivisors
, and thus
leadingCoeff
is multiplicative
Equations
- Polynomial.leadingCoeffHom = { toFun := Polynomial.leadingCoeff, map_one' := ⋯, map_mul' := ⋯ }
Instances For
@[simp]
theorem
Polynomial.leadingCoeffHom_apply
{R : Type u}
[Semiring R]
[NoZeroDivisors R]
(p : Polynomial R)
:
Polynomial.leadingCoeffHom p = p.leadingCoeff
@[simp]
theorem
Polynomial.leadingCoeff_pow
{R : Type u}
[Semiring R]
[NoZeroDivisors R]
(p : Polynomial R)
(n : ℕ)
:
theorem
Polynomial.leadingCoeff_dvd_leadingCoeff
{R : Type u}
[Semiring R]
[NoZeroDivisors R]
{a p : Polynomial R}
(hap : a ∣ p)
:
a.leadingCoeff ∣ p.leadingCoeff
theorem
Polynomial.degree_le_mul_left
{R : Type u}
[Semiring R]
[NoZeroDivisors R]
{q : Polynomial R}
(p : Polynomial R)
(hq : q ≠ 0)
:
theorem
Polynomial.Monic.natDegree_pos
{R : Type u}
[CommSemiring R]
{p : Polynomial R}
(hp : p.Monic)
:
theorem
Polynomial.Monic.degree_pos
{R : Type u}
[CommSemiring R]
{p : Polynomial R}
(hp : p.Monic)
:
@[simp]
theorem
Polynomial.degree_X_pow_sub_C
{R : Type u}
[Ring R]
[Nontrivial R]
{n : ℕ}
(hn : 0 < n)
(a : R)
:
theorem
Polynomial.X_pow_sub_C_ne_zero
{R : Type u}
[Ring R]
[Nontrivial R]
{n : ℕ}
(hn : 0 < n)
(a : R)
:
theorem
Polynomial.zero_nmem_multiset_map_X_sub_C
{R : Type u}
[Ring R]
[Nontrivial R]
{α : Type u_1}
(m : Multiset α)
(f : α → R)
:
0 ∉ Multiset.map (fun (a : α) => Polynomial.X - Polynomial.C (f a)) m