Documentation

Mathlib.Algebra.Polynomial.EraseLead

Erase the leading term of a univariate polynomial #

Definition #

eraseLead serves as reduction step in an induction, shaving off one monomial from a polynomial. The definition is set up so that it does not mention subtraction in the definition, and thus works for polynomials over semirings as well as rings.

eraseLead f for a polynomial f is the polynomial obtained by subtracting from f the leading term of f.

Equations
Instances For
    @[simp]
    @[simp]
    @[simp]
    theorem Polynomial.eraseLead_C {R : Type u_1} [Semiring R] (r : R) :
    Polynomial.eraseLead (Polynomial.C r) = 0
    @[simp]
    theorem Polynomial.eraseLead_X {R : Type u_1} [Semiring R] :
    Polynomial.eraseLead Polynomial.X = 0
    @[simp]
    theorem Polynomial.eraseLead_X_pow {R : Type u_1} [Semiring R] (n : ) :
    Polynomial.eraseLead (Polynomial.X ^ n) = 0
    @[simp]
    theorem Polynomial.eraseLead_C_mul_X_pow {R : Type u_1} [Semiring R] (r : R) (n : ) :
    Polynomial.eraseLead (Polynomial.C r * Polynomial.X ^ n) = 0
    @[simp]
    theorem Polynomial.eraseLead_C_mul_X {R : Type u_1} [Semiring R] (r : R) :
    Polynomial.eraseLead (Polynomial.C r * Polynomial.X) = 0
    theorem Polynomial.induction_with_natDegree_le {R : Type u_1} [Semiring R] (P : Polynomial RProp) (N : ) (P_0 : P 0) (P_C_mul_pow : ∀ (n : ) (r : R), r 0n NP (Polynomial.C r * Polynomial.X ^ n)) (P_C_add : ∀ (f g : Polynomial R), Polynomial.natDegree f < Polynomial.natDegree gPolynomial.natDegree g NP fP gP (f + g)) (f : Polynomial R) :

    An induction lemma for polynomials. It takes a natural number N as a parameter, that is required to be at least as big as the nat_degree of the polynomial. This is useful to prove results where you want to change each term in a polynomial to something else depending on the nat_degree of the polynomial itself and not on the specific nat_degree of each term.

    theorem Polynomial.mono_map_natDegree_eq {R : Type u_1} [Semiring R] {S : Type u_2} {F : Type u_3} [Semiring S] [FunLike F (Polynomial R) (Polynomial S)] [AddMonoidHomClass F (Polynomial R) (Polynomial S)] {φ : F} {p : Polynomial R} (k : ) (fu : ) (fu0 : ∀ {n : }, n kfu n = 0) (fc : ∀ {n m : }, k nn < mfu n < fu m) (φ_k : ∀ {f : Polynomial R}, Polynomial.natDegree f < kφ f = 0) (φ_mon_nat : ∀ (n : ) (c : R), c 0Polynomial.natDegree (φ ((Polynomial.monomial n) c)) = fu n) :

    Let φ : R[x] → S[x] be an additive map, k : ℕ a bound, and fu : ℕ → ℕ a "sufficiently monotone" map. Assume also that

    • φ maps to 0 all monomials of degree less than k,
    • φ maps each monomial m in R[x] to a polynomial φ m of degree fu (deg m). Then, φ maps each polynomial p in R[x] to a polynomial of degree fu (deg p).
    theorem Polynomial.map_natDegree_eq_sub {R : Type u_1} [Semiring R] {S : Type u_2} {F : Type u_3} [Semiring S] [FunLike F (Polynomial R) (Polynomial S)] [AddMonoidHomClass F (Polynomial R) (Polynomial S)] {φ : F} {p : Polynomial R} {k : } (φ_k : ∀ (f : Polynomial R), Polynomial.natDegree f < kφ f = 0) (φ_mon : ∀ (n : ) (c : R), c 0Polynomial.natDegree (φ ((Polynomial.monomial n) c)) = n - k) :
    theorem Polynomial.map_natDegree_eq_natDegree {R : Type u_1} [Semiring R] {S : Type u_2} {F : Type u_3} [Semiring S] [FunLike F (Polynomial R) (Polynomial S)] [AddMonoidHomClass F (Polynomial R) (Polynomial S)] {φ : F} (p : Polynomial R) (φ_mon_nat : ∀ (n : ) (c : R), c 0Polynomial.natDegree (φ ((Polynomial.monomial n) c)) = n) :
    theorem Polynomial.card_support_eq' {R : Type u_1} [Semiring R] {n : } (k : Fin n) (x : Fin nR) (hk : Function.Injective k) (hx : ∀ (i : Fin n), x i 0) :
    (Polynomial.support (Finset.sum Finset.univ fun (i : Fin n) => Polynomial.C (x i) * Polynomial.X ^ k i)).card = n
    theorem Polynomial.card_support_eq {R : Type u_1} [Semiring R] {f : Polynomial R} {n : } :
    (Polynomial.support f).card = n ∃ (k : Fin n) (x : Fin nR) (_ : StrictMono k) (_ : ∀ (i : Fin n), x i 0), f = Finset.sum Finset.univ fun (i : Fin n) => Polynomial.C (x i) * Polynomial.X ^ k i
    theorem Polynomial.card_support_eq_one {R : Type u_1} [Semiring R] {f : Polynomial R} :
    (Polynomial.support f).card = 1 ∃ (k : ) (x : R) (_ : x 0), f = Polynomial.C x * Polynomial.X ^ k
    theorem Polynomial.card_support_eq_two {R : Type u_1} [Semiring R] {f : Polynomial R} :
    (Polynomial.support f).card = 2 ∃ (k : ) (m : ) (_ : k < m) (x : R) (y : R) (_ : x 0) (_ : y 0), f = Polynomial.C x * Polynomial.X ^ k + Polynomial.C y * Polynomial.X ^ m
    theorem Polynomial.card_support_eq_three {R : Type u_1} [Semiring R] {f : Polynomial R} :
    (Polynomial.support f).card = 3 ∃ (k : ) (m : ) (n : ) (_ : k < m) (_ : m < n) (x : R) (y : R) (z : R) (_ : x 0) (_ : y 0) (_ : z 0), f = Polynomial.C x * Polynomial.X ^ k + Polynomial.C y * Polynomial.X ^ m + Polynomial.C z * Polynomial.X ^ n