Documentation

Mathlib.Algebra.Polynomial.Degree.Support

Degree and support of univariate polynomials #

Main results #

theorem Polynomial.as_sum_support {R : Type u} [Semiring R] (p : Polynomial R) :
p = ip.support, (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, C (p.coeff i) * 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, (monomial i) (p.coeff i)
theorem Polynomial.as_sum_range {R : Type u} [Semiring R] (p : Polynomial R) :
p = iFinset.range (p.natDegree + 1), (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), C (p.coeff i) * X ^ i
theorem Polynomial.mem_support_C_mul_X_pow {R : Type u} [Semiring R] {n a : } {c : R} (h : a (C c * X ^ n).support) :
a = n
theorem Polynomial.le_degree_of_mem_supp {R : Type u} [Semiring R] {p : Polynomial R} (a : ) :
a p.supporta p.degree