Documentation

Mathlib.Algebra.Polynomial.Reverse

Reverse of a univariate polynomial #

The main definition is reverse. Applying reverse to a polynomial f : R[X] produces the polynomial with a reversed list of coefficients, equivalent to X^f.natDegree * f(1/X).

The main result is that reverse (f * g) = reverse f * reverse g, provided the leading coefficients of f and g do not multiply to zero.

def Polynomial.revAtFun (N : ) (i : ) :

If i ≤ N, then revAtFun N i returns N - i, otherwise it returns i. This is the map used by the embedding revAt.

Equations
Instances For

    If i ≤ N, then revAt N i returns N - i, otherwise it returns i. Essentially, this embedding is only used for i ≤ N. The advantage of revAt N i over N - i is that revAt is an involution.

    Equations
    Instances For
      @[simp]

      We prefer to use the bundled revAt over unbundled revAtFun.

      @[simp]
      @[simp]
      theorem Polynomial.revAt_le {N : } {i : } (H : i N) :
      (Polynomial.revAt N) i = N - i
      theorem Polynomial.revAt_eq_self_of_lt {N : } {i : } (h : N < i) :
      theorem Polynomial.revAt_add {N : } {O : } {n : } {o : } (hn : n N) (ho : o O) :
      noncomputable def Polynomial.reflect {R : Type u_1} [Semiring R] (N : ) :

      reflect N f is the polynomial such that (reflect N f).coeff i = f.coeff (revAt N i). In other words, the terms with exponent [0, ..., N] now have exponent [N, ..., 0].

      In practice, reflect is only used when N is at least as large as the degree of f.

      Eventually, it will be used with N exactly equal to the degree of f.

      Equations
      Instances For
        theorem Polynomial.reflect_support {R : Type u_1} [Semiring R] (N : ) (f : Polynomial R) :
        (Polynomial.reflect N f).support = Finset.image ((Polynomial.revAt N)) f.support
        @[simp]
        theorem Polynomial.coeff_reflect {R : Type u_1} [Semiring R] (N : ) (f : Polynomial R) (i : ) :
        (Polynomial.reflect N f).coeff i = f.coeff ((Polynomial.revAt N) i)
        @[simp]
        theorem Polynomial.reflect_zero {R : Type u_1} [Semiring R] {N : } :
        @[simp]
        theorem Polynomial.reflect_eq_zero_iff {R : Type u_1} [Semiring R] {N : } {f : Polynomial R} :
        @[simp]
        @[simp]
        theorem Polynomial.reflect_C_mul {R : Type u_1} [Semiring R] (f : Polynomial R) (r : R) (N : ) :
        Polynomial.reflect N (Polynomial.C r * f) = Polynomial.C r * Polynomial.reflect N f
        theorem Polynomial.reflect_C_mul_X_pow {R : Type u_1} [Semiring R] (N : ) (n : ) {c : R} :
        Polynomial.reflect N (Polynomial.C c * Polynomial.X ^ n) = Polynomial.C c * Polynomial.X ^ (Polynomial.revAt N) n
        @[simp]
        theorem Polynomial.reflect_C {R : Type u_1} [Semiring R] (r : R) (N : ) :
        Polynomial.reflect N (Polynomial.C r) = Polynomial.C r * Polynomial.X ^ N
        @[simp]
        theorem Polynomial.reflect_monomial {R : Type u_1} [Semiring R] (N : ) (n : ) :
        Polynomial.reflect N (Polynomial.X ^ n) = Polynomial.X ^ (Polynomial.revAt N) n
        @[simp]
        theorem Polynomial.reflect_one_X {R : Type u_1} [Semiring R] :
        Polynomial.reflect 1 Polynomial.X = 1
        theorem Polynomial.reflect_mul_induction {R : Type u_1} [Semiring R] (cf : ) (cg : ) (N : ) (O : ) (f : Polynomial R) (g : Polynomial R) :
        f.support.card cf.succg.support.card cg.succf.natDegree Ng.natDegree OPolynomial.reflect (N + O) (f * g) = Polynomial.reflect N f * Polynomial.reflect O g
        @[simp]
        theorem Polynomial.reflect_mul {R : Type u_1} [Semiring R] (f : Polynomial R) (g : Polynomial R) {F : } {G : } (Ff : f.natDegree F) (Gg : g.natDegree G) :
        theorem Polynomial.eval₂_reflect_mul_pow {R : Type u_1} [Semiring R] {S : Type u_2} [CommSemiring S] (i : R →+* S) (x : S) [Invertible x] (N : ) (f : Polynomial R) (hf : f.natDegree N) :
        theorem Polynomial.eval₂_reflect_eq_zero_iff {R : Type u_1} [Semiring R] {S : Type u_2} [CommSemiring S] (i : R →+* S) (x : S) [Invertible x] (N : ) (f : Polynomial R) (hf : f.natDegree N) :
        noncomputable def Polynomial.reverse {R : Type u_1} [Semiring R] (f : Polynomial R) :

        The reverse of a polynomial f is the polynomial obtained by "reading f backwards". Even though this is not the actual definition, reverse f = f (1/X) * X ^ f.natDegree.

        Equations
        Instances For
          theorem Polynomial.coeff_reverse {R : Type u_1} [Semiring R] (f : Polynomial R) (n : ) :
          f.reverse.coeff n = f.coeff ((Polynomial.revAt f.natDegree) n)
          @[simp]
          theorem Polynomial.coeff_zero_reverse {R : Type u_1} [Semiring R] (f : Polynomial R) :
          f.reverse.coeff 0 = f.leadingCoeff
          @[simp]
          theorem Polynomial.reverse_eq_zero {R : Type u_1} [Semiring R] {f : Polynomial R} :
          f.reverse = 0 f = 0
          theorem Polynomial.reverse_natDegree_le {R : Type u_1} [Semiring R] (f : Polynomial R) :
          f.reverse.natDegree f.natDegree
          theorem Polynomial.natDegree_eq_reverse_natDegree_add_natTrailingDegree {R : Type u_1} [Semiring R] (f : Polynomial R) :
          f.natDegree = f.reverse.natDegree + f.natTrailingDegree
          theorem Polynomial.reverse_natDegree {R : Type u_1} [Semiring R] (f : Polynomial R) :
          f.reverse.natDegree = f.natDegree - f.natTrailingDegree
          theorem Polynomial.reverse_leadingCoeff {R : Type u_1} [Semiring R] (f : Polynomial R) :
          f.reverse.leadingCoeff = f.trailingCoeff
          theorem Polynomial.natTrailingDegree_reverse {R : Type u_1} [Semiring R] (f : Polynomial R) :
          f.reverse.natTrailingDegree = 0
          theorem Polynomial.reverse_trailingCoeff {R : Type u_1} [Semiring R] (f : Polynomial R) :
          f.reverse.trailingCoeff = f.leadingCoeff
          theorem Polynomial.reverse_mul {R : Type u_1} [Semiring R] {f : Polynomial R} {g : Polynomial R} (fg : f.leadingCoeff * g.leadingCoeff 0) :
          (f * g).reverse = f.reverse * g.reverse
          @[simp]
          theorem Polynomial.reverse_mul_of_domain {R : Type u_2} [Ring R] [NoZeroDivisors R] (f : Polynomial R) (g : Polynomial R) :
          (f * g).reverse = f.reverse * g.reverse
          theorem Polynomial.trailingCoeff_mul {R : Type u_2} [Ring R] [NoZeroDivisors R] (p : Polynomial R) (q : Polynomial R) :
          (p * q).trailingCoeff = p.trailingCoeff * q.trailingCoeff
          @[simp]
          theorem Polynomial.coeff_one_reverse {R : Type u_1} [Semiring R] (f : Polynomial R) :
          f.reverse.coeff 1 = f.nextCoeff
          @[simp]
          theorem Polynomial.reverse_C {R : Type u_1} [Semiring R] (t : R) :
          (Polynomial.C t).reverse = Polynomial.C t
          @[simp]
          theorem Polynomial.reverse_mul_X {R : Type u_1} [Semiring R] (p : Polynomial R) :
          (p * Polynomial.X).reverse = p.reverse
          @[simp]
          theorem Polynomial.reverse_X_mul {R : Type u_1} [Semiring R] (p : Polynomial R) :
          (Polynomial.X * p).reverse = p.reverse
          @[simp]
          theorem Polynomial.reverse_mul_X_pow {R : Type u_1} [Semiring R] (p : Polynomial R) (n : ) :
          (p * Polynomial.X ^ n).reverse = p.reverse
          @[simp]
          theorem Polynomial.reverse_X_pow_mul {R : Type u_1} [Semiring R] (p : Polynomial R) (n : ) :
          (Polynomial.X ^ n * p).reverse = p.reverse
          @[simp]
          theorem Polynomial.reverse_add_C {R : Type u_1} [Semiring R] (p : Polynomial R) (t : R) :
          (p + Polynomial.C t).reverse = p.reverse + Polynomial.C t * Polynomial.X ^ p.natDegree
          @[simp]
          theorem Polynomial.reverse_C_add {R : Type u_1} [Semiring R] (p : Polynomial R) (t : R) :
          (Polynomial.C t + p).reverse = Polynomial.C t * Polynomial.X ^ p.natDegree + p.reverse
          theorem Polynomial.eval₂_reverse_mul_pow {R : Type u_1} [Semiring R] {S : Type u_2} [CommSemiring S] (i : R →+* S) (x : S) [Invertible x] (f : Polynomial R) :
          Polynomial.eval₂ i (x) f.reverse * x ^ f.natDegree = Polynomial.eval₂ i x f
          @[simp]
          theorem Polynomial.eval₂_reverse_eq_zero_iff {R : Type u_1} [Semiring R] {S : Type u_2} [CommSemiring S] (i : R →+* S) (x : S) [Invertible x] (f : Polynomial R) :
          Polynomial.eval₂ i (x) f.reverse = 0 Polynomial.eval₂ i x f = 0
          @[simp]
          @[simp]
          @[simp]
          theorem Polynomial.reverse_neg {R : Type u_1} [Ring R] (f : Polynomial R) :
          (-f).reverse = -f.reverse