Documentation

Mathlib.Analysis.Polynomial.CauchyBound

Cauchy's bound on polynomial roots. #

The bound is given by Polynomial.cauchyBound, which for a_n x^n + a_(n-1) x^(n - 1) + ⋯ + a_0 is is 1 + max_(0 ≤ i < n) a_i / a_n.

The theorem that this gives a bound to polynomial roots is Polynomial.IsRoot.norm_lt_cauchyBound.

noncomputable def Polynomial.cauchyBound {K : Type u_1} [NormedDivisionRing K] (p : Polynomial K) :

Cauchy's bound on the roots of a given polynomial. See IsRoot.norm_lt_cauchyBound for the proof that the roots satisfy this bound.

Equations
Instances For
    @[simp]
    theorem Polynomial.one_le_cauchyBound {K : Type u_1} [NormedDivisionRing K] (p : Polynomial K) :
    1 p.cauchyBound
    @[simp]
    theorem Polynomial.cauchyBound_C {K : Type u_1} [NormedDivisionRing K] (x : K) :
    (Polynomial.C x).cauchyBound = 1
    @[simp]
    theorem Polynomial.cauchyBound_X {K : Type u_1} [NormedDivisionRing K] :
    Polynomial.X.cauchyBound = 1
    @[simp]
    theorem Polynomial.cauchyBound_X_add_C {K : Type u_1} [NormedDivisionRing K] (x : K) :
    (Polynomial.X + Polynomial.C x).cauchyBound = x‖₊ + 1
    @[simp]
    theorem Polynomial.cauchyBound_X_sub_C {K : Type u_1} [NormedDivisionRing K] (x : K) :
    (Polynomial.X - Polynomial.C x).cauchyBound = x‖₊ + 1
    @[simp]
    theorem Polynomial.cauchyBound_smul {K : Type u_1} [NormedDivisionRing K] {x : K} (hx : x 0) (p : Polynomial K) :
    (x p).cauchyBound = p.cauchyBound
    theorem Polynomial.IsRoot.norm_lt_cauchyBound {K : Type u_1} [NormedDivisionRing K] {p : Polynomial K} (hp : p 0) {a : K} (h : p.IsRoot a) :
    a‖₊ < p.cauchyBound

    cauchyBound is a bound on the norm of polynomial roots.