Documentation

Mathlib.Analysis.SpecialFunctions.CompareExp

Growth estimates on x ^ y for complex x, y #

Let l be a filter on such that Complex.re tends to infinity along l and Complex.im z grows at a subexponential rate compared to Complex.re z. Then

We use these assumptions on l for two reasons. First, these are the assumptions that naturally appear in the proof. Second, in some applications (e.g., in Ilyashenko's proof of the individual finiteness theorem for limit cycles of polynomial ODEs with hyperbolic singularities only) natural stronger assumptions (e.g., im z is bounded from below and from above) are not available.

We say that l : Filter is an exponential comparison filter if the real part tends to infinity along l and the imaginary part grows subexponentially compared to the real part. These properties guarantee that (fun z ↦ z ^ a₁ * exp (b₁ * z)) =o[l] (fun z ↦ z ^ a₂ * exp (b₂ * z)) for any complex a₁, a₂ and real b₁ < b₂.

In particular, the second property is automatically satisfied if the imaginary part is bounded along l.

Instances For

    Alternative constructors #

    theorem Complex.IsExpCmpFilter.of_isBigO_im_re_rpow {l : Filter } (hre : Filter.Tendsto Complex.re l Filter.atTop) (r : ) (hr : Complex.im =O[l] fun (z : ) => z.re ^ r) :
    theorem Complex.IsExpCmpFilter.of_isBigO_im_re_pow {l : Filter } (hre : Filter.Tendsto Complex.re l Filter.atTop) (n : ) (hr : Complex.im =O[l] fun (z : ) => z.re ^ n) :
    theorem Complex.IsExpCmpFilter.of_boundedUnder_abs_im {l : Filter } (hre : Filter.Tendsto Complex.re l Filter.atTop) (him : Filter.IsBoundedUnder (fun (x x_1 : ) => x x_1) l fun (z : ) => |z.im|) :
    theorem Complex.IsExpCmpFilter.of_boundedUnder_im {l : Filter } (hre : Filter.Tendsto Complex.re l Filter.atTop) (him_le : Filter.IsBoundedUnder (fun (x x_1 : ) => x x_1) l Complex.im) (him_ge : Filter.IsBoundedUnder (fun (x x_1 : ) => x x_1) l Complex.im) :

    Preliminary lemmas #

    theorem Complex.IsExpCmpFilter.tendsto_abs_re {l : Filter } (hl : Complex.IsExpCmpFilter l) :
    Filter.Tendsto (fun (z : ) => |z.re|) l Filter.atTop
    theorem Complex.IsExpCmpFilter.isLittleO_im_pow_exp_re {l : Filter } (hl : Complex.IsExpCmpFilter l) (n : ) :
    (fun (z : ) => z.im ^ n) =o[l] fun (z : ) => Real.exp z.re
    theorem Complex.IsExpCmpFilter.abs_im_pow_eventuallyLE_exp_re {l : Filter } (hl : Complex.IsExpCmpFilter l) (n : ) :
    (fun (z : ) => |z.im| ^ n) ≤ᶠ[l] fun (z : ) => Real.exp z.re

    If l : Filter is an "exponential comparison filter", then $\log |z| =o(ℜ z)$ along l. This is the main lemma in the proof of Complex.IsExpCmpFilter.isLittleO_cpow_exp below.

    Main results #

    theorem Complex.IsExpCmpFilter.isTheta_cpow_exp_re_mul_log {l : Filter } (hl : Complex.IsExpCmpFilter l) (a : ) :
    (fun (x : ) => x ^ a) =Θ[l] fun (z : ) => Real.exp (a.re * Real.log (Complex.abs z))
    theorem Complex.IsExpCmpFilter.isLittleO_cpow_exp {l : Filter } (hl : Complex.IsExpCmpFilter l) (a : ) {b : } (hb : 0 < b) :
    (fun (z : ) => z ^ a) =o[l] fun (z : ) => Complex.exp (b * z)

    If l : Filter is an "exponential comparison filter", then for any complex a and any positive real b, we have (fun z ↦ z ^ a) =o[l] (fun z ↦ exp (b * z)).

    theorem Complex.IsExpCmpFilter.isLittleO_cpow_mul_exp {l : Filter } {b₁ : } {b₂ : } (hl : Complex.IsExpCmpFilter l) (hb : b₁ < b₂) (a₁ : ) (a₂ : ) :
    (fun (z : ) => z ^ a₁ * Complex.exp (b₁ * z)) =o[l] fun (z : ) => z ^ a₂ * Complex.exp (b₂ * z)

    If l : Filter is an "exponential comparison filter", then for any complex a₁, a₂ and any real b₁ < b₂, we have (fun z ↦ z ^ a₁ * exp (b₁ * z)) =o[l] (fun z ↦ z ^ a₂ * exp (b₂ * z)).

    theorem Complex.IsExpCmpFilter.isLittleO_exp_cpow {l : Filter } (hl : Complex.IsExpCmpFilter l) (a : ) {b : } (hb : b < 0) :
    (fun (z : ) => Complex.exp (b * z)) =o[l] fun (z : ) => z ^ a

    If l : Filter is an "exponential comparison filter", then for any complex a and any negative real b, we have (fun z ↦ exp (b * z)) =o[l] (fun z ↦ z ^ a).

    theorem Complex.IsExpCmpFilter.isLittleO_pow_mul_exp {l : Filter } {b₁ : } {b₂ : } (hl : Complex.IsExpCmpFilter l) (hb : b₁ < b₂) (m : ) (n : ) :
    (fun (z : ) => z ^ m * Complex.exp (b₁ * z)) =o[l] fun (z : ) => z ^ n * Complex.exp (b₂ * z)

    If l : Filter is an "exponential comparison filter", then for any complex a₁, a₂ and any natural b₁ < b₂, we have (fun z ↦ z ^ a₁ * exp (b₁ * z)) =o[l] (fun z ↦ z ^ a₂ * exp (b₂ * z)).

    theorem Complex.IsExpCmpFilter.isLittleO_zpow_mul_exp {l : Filter } {b₁ : } {b₂ : } (hl : Complex.IsExpCmpFilter l) (hb : b₁ < b₂) (m : ) (n : ) :
    (fun (z : ) => z ^ m * Complex.exp (b₁ * z)) =o[l] fun (z : ) => z ^ n * Complex.exp (b₂ * z)

    If l : Filter is an "exponential comparison filter", then for any complex a₁, a₂ and any integer b₁ < b₂, we have (fun z ↦ z ^ a₁ * exp (b₁ * z)) =o[l] (fun z ↦ z ^ a₂ * exp (b₂ * z)).