Documentation

Mathlib.Algebra.Polynomial.Homogenize

Homogenize a univariate polynomial #

In this file we define a function Polynomial.homogenize p n that takes a polynomial p and a natural number n and returns a homogeneous bivariate polynomial of degree n.

If n is at least the degree of p, then (homogenize p n).eval ![x, 1] = p.eval x.

We use MvPolynomial (Fin 2) R to represent bivariate polynomials instead of R[X][Y] (i.e., Polynomial (Polynomial R)), because Mathlib has a theory about homogeneous multivariate polynomials, but not about homogeneous bivariate polynomials encoded as R[X][Y].

noncomputable def Polynomial.homogenize {R : Type u_1} [CommSemiring R] (p : Polynomial R) (n : ) :

Given a polynomial p and a number n ≥ natDegree p, returns a homogeneous bivariate polynomial q of degree n such that q(x, 1) = p(x).

It is defined as ∑ k + l = n, a_k X_0^k X_1^l, where a_k is the kth coefficient of p.

Equations
Instances For
    @[simp]
    theorem Polynomial.homogenize_zero {R : Type u_1} [CommSemiring R] (n : ) :
    @[simp]
    theorem Polynomial.homogenize_add {R : Type u_1} [CommSemiring R] (p q : Polynomial R) (n : ) :
    @[simp]
    theorem Polynomial.homogenize_smul {R : Type u_1} [CommSemiring R] {S : Type u_2} [Semiring S] [Module S R] (c : S) (p : Polynomial R) (n : ) :
    (c p).homogenize n = c p.homogenize n
    noncomputable def Polynomial.homogenizeLM {R : Type u_1} [CommSemiring R] (n : ) :

    homogenize as a bundled linear map.

    Equations
    Instances For
      @[simp]
      theorem Polynomial.homogenizeLM_apply {R : Type u_1} [CommSemiring R] (n : ) (p : Polynomial R) :
      @[simp]
      theorem Polynomial.homogenize_finsetSum {R : Type u_1} [CommSemiring R] {ι : Type u_2} (s : Finset ι) (p : ιPolynomial R) (n : ) :
      (∑ is, p i).homogenize n = is, (p i).homogenize n
      theorem Polynomial.homogenize_map {R : Type u_1} [CommSemiring R] {S : Type u_2} [CommSemiring S] (f : R →+* S) (p : Polynomial R) (n : ) :
      @[simp]
      theorem Polynomial.homogenize_C_mul {R : Type u_1} [CommSemiring R] (c : R) (p : Polynomial R) (n : ) :
      @[simp]
      theorem Polynomial.homogenize_monomial {R : Type u_1} [CommSemiring R] {m n : } (h : m n) (r : R) :
      ((monomial m) r).homogenize n = (MvPolynomial.monomial fun₀ | 0 => m | 1 => n - m) r
      theorem Polynomial.homogenize_monomial_of_lt {R : Type u_1} [CommSemiring R] {m n : } (h : n < m) (r : R) :
      ((monomial m) r).homogenize n = 0
      @[simp]
      theorem Polynomial.homogenize_X_pow {R : Type u_1} [CommSemiring R] {m n : } (h : m n) :
      @[simp]
      theorem Polynomial.homogenize_X {R : Type u_1} [CommSemiring R] {n : } (hn : n 0) :
      @[simp]
      theorem Polynomial.homogenize_C {R : Type u_1} [CommSemiring R] (c : R) (n : ) :
      @[simp]
      theorem Polynomial.coeff_homogenize {R : Type u_1} [CommSemiring R] (p : Polynomial R) (n : ) (m : Fin 2 →₀ ) :
      MvPolynomial.coeff m (p.homogenize n) = if m 0 + m 1 = n then p.coeff (m 0) else 0
      theorem Polynomial.eval₂_homogenize_of_eq_one {R : Type u_1} [CommSemiring R] {S : Type u_2} [CommSemiring S] {p : Polynomial R} {n : } (hn : p.natDegree n) (f : R →+* S) (g : Fin 2S) (hg : g 1 = 1) :
      theorem Polynomial.aeval_homogenize_of_eq_one {R : Type u_1} [CommSemiring R] {A : Type u_2} [CommSemiring A] [Algebra R A] {p : Polynomial R} {n : } (hn : p.natDegree n) (g : Fin 2A) (hg : g 1 = 1) :
      (MvPolynomial.aeval g) (p.homogenize n) = (aeval (g 0)) p
      @[simp]

      If deg p ≤ n, then homogenize p n (x, 1) = p x.

      theorem Polynomial.homogenize_eq_of_isHomogeneous {R : Type u_1} [CommSemiring R] {p : Polynomial R} {n : } {q : MvPolynomial (Fin 2) R} (hq : q.IsHomogeneous n) (hpq : (MvPolynomial.aeval ![X, 1]) q = p) :
      theorem Polynomial.homogenize_mul {R : Type u_1} [CommSemiring R] (p q : Polynomial R) {m n : } (hm : p.natDegree m) (hn : q.natDegree n) :
      (p * q).homogenize (m + n) = p.homogenize m * q.homogenize n
      theorem Polynomial.homogenize_finsetProd {R : Type u_1} [CommSemiring R] {ι : Type u_2} {s : Finset ι} {p : ιPolynomial R} {n : ι} (h : is, (p i).natDegree n i) :
      (∏ is, p i).homogenize (∑ is, n i) = is, (p i).homogenize (n i)
      @[simp]
      theorem Polynomial.homogenize_neg {R : Type u_1} [CommRing R] (p : Polynomial R) (n : ) :
      @[simp]
      theorem Polynomial.homogenize_sub {R : Type u_1} [CommRing R] (p q : Polynomial R) (n : ) :
      theorem Polynomial.eval_homogenize {K : Type u_1} [Semifield K] {p : Polynomial K} {n : } (hn : p.natDegree n) (x : Fin 2K) (hx : x 1 0) :
      (MvPolynomial.eval x) (p.homogenize n) = eval (x 0 / x 1) p * x 1 ^ n