Documentation

Mathlib.Analysis.SpecialFunctions.Trigonometric.Chebyshev.Extremal

Chebyshev polynomials over the reals: some extremal properties #

Chebyshev polynomials have largest leading coefficient, following proof in https://math.stackexchange.com/a/978145/1277

Main statements #

Implementation #

By monotonicity of 2 ^ (n - 1), we can assume that P has degree exactly n. Using Lagrange interpolation, we can give a formula for the leading coefficient of P as a linear combination of the values of P on the Chebyshev nodes (sumNodes_eq_coeff). The Chebyshev polynomial T_n has value ±1 on the nodes, with the same signs as the coefficients of the linear combination (negOnePow_mul_leadingCoeffC_pos). Since |P (x)| ≤ 1 on the nodes, this implies that the leading coefficient of P is bounded by that of T_n, which is known to equal 2 ^ (n - 1). Moreover, equality holds iff P and T_n agree on the nodes, which implies that they coincide.

noncomputable def Polynomial.Chebyshev.node (n i : ) :

For n ≠ 0 and i ≤ n, node n i is one of the extremal points of the Chebyhsev T polynomial over the interval [-1, 1].

Equations
Instances For
    theorem Polynomial.Chebyshev.node_eq_neg_one {n : } (hn : n 0) :
    node n n = -1
    theorem Polynomial.Chebyshev.eval_T_real_node {n i : } (hi : i Finset.Iic n) :
    eval (node n i) (T n) = (-1) ^ i
    theorem Polynomial.Chebyshev.node_lt {n i j : } (hj : j n) (hij : i < j) :
    node n j < node n i
    theorem Polynomial.Chebyshev.zero_lt_prod_node_sub_node {n i : } (hi : i n) :
    0 < (-1) ^ i * j(Finset.range (n + 1)).erase i, (node n i - node n j)
    noncomputable def Polynomial.Chebyshev.sumNodes (n : ) (c : ) (P : Polynomial ) :

    For a polynomial P and coefficient function c, sumNodes n c P is a linear combination of P evaluated at the n'th order Chebyshev nodes, with coefficients taken from c.

    Equations
    Instances For
      theorem Polynomial.Chebyshev.sumNodes_le_sumNodes_T {n : } {c : } (hcnonneg : in, 0 (-1) ^ i * c i) {P : Polynomial } (hPbnd : xSet.Icc (-1) 1, |eval x P| 1) :
      sumNodes n c P sumNodes n c (T n)
      theorem Polynomial.Chebyshev.sumNodes_eq_sumNodes_T_iff {n : } {c : } (hcpos : in, 0 < (-1) ^ i * c i) {P : Polynomial } (hPdeg : P.degree n) (hPbnd : xSet.Icc (-1) 1, |eval x P| 1) :
      sumNodes n c P = sumNodes n c (T n) P = T n
      theorem Polynomial.Chebyshev.coeff_le_of_forall_abs_le_one {n : } {P : Polynomial } (hPdeg : P.degree n) (hPbnd : xSet.Icc (-1) 1, |eval x P| 1) :
      P.coeff n 2 ^ (n - 1)
      theorem Polynomial.Chebyshev.leadingCoeff_le_of_forall_abs_le_one {n : } {P : Polynomial } (hPdeg : P.degree n) (hPbnd : xSet.Icc (-1) 1, |eval x P| 1) :
      P.leadingCoeff 2 ^ (n - 1)
      theorem Polynomial.Chebyshev.coeff_eq_iff_of_forall_abs_le_one {n : } {P : Polynomial } (hPdeg : P.degree n) (hPbnd : xSet.Icc (-1) 1, |eval x P| 1) :
      P.coeff n = 2 ^ (n - 1) P = T n
      theorem Polynomial.Chebyshev.leadingCoeff_eq_iff_of_forall_abs_le_one {n : } {P : Polynomial } (hn : 2 n) (hPdeg : P.degree n) (hPbnd : xSet.Icc (-1) 1, |eval x P| 1) :
      P.leadingCoeff = 2 ^ (n - 1) P = T n