Documentation

Mathlib.Algebra.Polynomial.Degree.Support

Degree and support of univariate polynomials #

Main results #

theorem Polynomial.supDegree_eq_natDegree {R : Type u} [Semiring R] (p : Polynomial R) :
AddMonoidAlgebra.supDegree id p.toFinsupp = p.natDegree
theorem Polynomial.le_natDegree_of_mem_supp {R : Type u} [Semiring R] {p : Polynomial R} (a : ) :
a p.supporta p.natDegree
theorem Polynomial.supp_subset_range {R : Type u} {m : } [Semiring R] {p : Polynomial R} (h : p.natDegree < m) :
p.support Finset.range m
theorem Polynomial.supp_subset_range_natDegree_succ {R : Type u} [Semiring R] {p : Polynomial R} :
p.support Finset.range (p.natDegree + 1)
theorem Polynomial.as_sum_support {R : Type u} [Semiring R] (p : Polynomial R) :
p = ip.support, (Polynomial.monomial i) (p.coeff i)
theorem Polynomial.as_sum_support_C_mul_X_pow {R : Type u} [Semiring R] (p : Polynomial R) :
p = ip.support, Polynomial.C (p.coeff i) * Polynomial.X ^ i
theorem Polynomial.sum_over_range' {R : Type u} {S : Type v} [Semiring R] [AddCommMonoid S] (p : Polynomial R) {f : RS} (h : ∀ (n : ), f n 0 = 0) (n : ) (w : p.natDegree < n) :
p.sum f = aFinset.range n, f a (p.coeff a)

We can reexpress a sum over p.support as a sum over range n, for any n satisfying p.natDegree < n.

theorem Polynomial.sum_over_range {R : Type u} {S : Type v} [Semiring R] [AddCommMonoid S] (p : Polynomial R) {f : RS} (h : ∀ (n : ), f n 0 = 0) :
p.sum f = aFinset.range (p.natDegree + 1), f a (p.coeff a)

We can reexpress a sum over p.support as a sum over range (p.natDegree + 1).

theorem Polynomial.sum_fin {R : Type u} {S : Type v} [Semiring R] [AddCommMonoid S] (f : RS) (hf : ∀ (i : ), f i 0 = 0) {n : } {p : Polynomial R} (hn : p.degree < n) :
i : Fin n, f (↑i) (p.coeff i) = p.sum f
theorem Polynomial.as_sum_range' {R : Type u} [Semiring R] (p : Polynomial R) (n : ) (w : p.natDegree < n) :
p = iFinset.range n, (Polynomial.monomial i) (p.coeff i)
theorem Polynomial.as_sum_range {R : Type u} [Semiring R] (p : Polynomial R) :
p = iFinset.range (p.natDegree + 1), (Polynomial.monomial i) (p.coeff i)
theorem Polynomial.as_sum_range_C_mul_X_pow {R : Type u} [Semiring R] (p : Polynomial R) :
p = iFinset.range (p.natDegree + 1), Polynomial.C (p.coeff i) * Polynomial.X ^ i
theorem Polynomial.mem_support_C_mul_X_pow {R : Type u} [Semiring R] {n a : } {c : R} (h : a (Polynomial.C c * Polynomial.X ^ n).support) :
a = n
theorem Polynomial.card_support_C_mul_X_pow_le_one {R : Type u} [Semiring R] {c : R} {n : } :
(Polynomial.C c * Polynomial.X ^ n).support.card 1
theorem Polynomial.card_supp_le_succ_natDegree {R : Type u} [Semiring R] (p : Polynomial R) :
p.support.card p.natDegree + 1
theorem Polynomial.le_degree_of_mem_supp {R : Type u} [Semiring R] {p : Polynomial R} (a : ) :
a p.supporta p.degree
theorem Polynomial.nonempty_support_iff {R : Type u} [Semiring R] {p : Polynomial R} :
p.support.Nonempty p 0
theorem Polynomial.natDegree_mem_support_of_nonzero {R : Type u} [Semiring R] {p : Polynomial R} (H : p 0) :
p.natDegree p.support
theorem Polynomial.natDegree_eq_support_max' {R : Type u} [Semiring R] {p : Polynomial R} (h : p 0) :
p.natDegree = p.support.max'