Induction on polynomials #
This file contains lemmas dealing with different flavours of induction on polynomials.
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
theorem
Polynomial.divX_hom_toFun
{R : Type u}
[Semiring R]
{p : Polynomial R}
 :
Polynomial.divX_hom p = p.divX
theorem
Polynomial.natDegree_divX_eq_natDegree_tsub_one
{R : Type u}
[Semiring R]
{p : Polynomial R}
 :
theorem
Polynomial.natDegree_divX_le
{R : Type u}
[Semiring R]
{p : Polynomial R}
 :
p.divX.natDegree ≤ p.natDegree
theorem
Polynomial.degree_divX_lt
{R : Type u}
[Semiring R]
{p : Polynomial R}
(hp0 : p ≠ 0)
 :
p.divX.degree < p.degree
@[irreducible]
noncomputable def
Polynomial.recOnHorner
{R : Type u}
[Semiring R]
{M : Polynomial R → Sort u_1}
(p : Polynomial R)
(M0 : M 0)
(MC : (p : Polynomial R) → (a : R) → p.coeff 0 = 0 → a ≠ 0 → M p → M (p + Polynomial.C a))
(MX : (p : Polynomial R) → p ≠ 0 → M p → M (p * Polynomial.X))
 :
M p
An induction principle for polynomials, valued in Sort* instead of Prop.
Equations
- One or more equations did not get rendered due to their size.
Instances For
theorem
Polynomial.degree_pos_induction_on
{R : Type u}
[Semiring R]
{P : Polynomial R → Prop}
(p : Polynomial R)
(h0 : 0 < p.degree)
(hC : ∀ {a : R}, a ≠ 0 → P (Polynomial.C a * Polynomial.X))
(hX : ∀ {p : Polynomial R}, 0 < p.degree → P p → P (p * Polynomial.X))
(hadd : ∀ {p : Polynomial R} {a : R}, 0 < p.degree → P p → P (p + Polynomial.C a))
 :
P p
A property holds for all polynomials of positive degree with coefficients in a semiring R
if it holds for
- a * X, with- a ∈ R,
- p * X, with- p ∈ R[X],
- p + a, with- a ∈ R,- p ∈ R[X], with appropriate restrictions on each term.
See natDegree_ne_zero_induction_on for a similar statement involving no explicit multiplication.
theorem
Polynomial.natDegree_ne_zero_induction_on
{R : Type u}
[Semiring R]
{M : Polynomial R → Prop}
{f : Polynomial R}
(f0 : f.natDegree ≠ 0)
(h_C_add : ∀ {a : R} {p : Polynomial R}, M p → M (Polynomial.C a + p))
(h_add : ∀ {p q : Polynomial R}, M p → M q → M (p + q))
(h_monomial : ∀ {n : ℕ} {a : R}, a ≠ 0 → n ≠ 0 → M ((Polynomial.monomial n) a))
 :
M f
A property holds for all polynomials of non-zero natDegree with coefficients in a
semiring R if it holds for
- p + a, with- a ∈ R,- p ∈ R[X],
- p + q, with- p, q ∈ R[X],
- monomials with nonzero coefficient and non-zero exponent,
with appropriate restrictions on each term.
Note that multiplication is "hidden" in the assumption on monomials, so there is no explicit
multiplication in the statement.
See degree_pos_induction_onfor a similar statement involving more explicit multiplications.