mathlib3 documentation

analysis.special_functions.compare_exp

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

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

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.

structure complex.is_exp_cmp_filter (l : filter ) :
Prop

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 (λ z, z ^ a₁ * exp (b₁ * z)) =o[l] (λ 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.

Alternative constructors #

Preliminary lemmas #

theorem complex.is_exp_cmp_filter.is_o_im_pow_exp_re {l : filter } (hl : complex.is_exp_cmp_filter l) (n : ) :
(λ (z : ), z.im ^ n) =o[l] λ (z : ), rexp 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.is_exp_cmp_filter.is_o_cpow_exp below.

Main results #

theorem complex.is_exp_cmp_filter.is_o_cpow_exp {l : filter } (hl : complex.is_exp_cmp_filter l) (a : ) {b : } (hb : 0 < b) :
(λ (z : ), z ^ a) =o[l] λ (z : ), cexp (b * z)

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

theorem complex.is_exp_cmp_filter.is_o_cpow_mul_exp {l : filter } {b₁ b₂ : } (hl : complex.is_exp_cmp_filter l) (hb : b₁ < b₂) (a₁ a₂ : ) :
(λ (z : ), z ^ a₁ * cexp (b₁ * z)) =o[l] λ (z : ), z ^ a₂ * cexp (b₂ * z)

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

theorem complex.is_exp_cmp_filter.is_o_exp_cpow {l : filter } (hl : complex.is_exp_cmp_filter l) (a : ) {b : } (hb : b < 0) :
(λ (z : ), cexp (b * z)) =o[l] λ (z : ), z ^ a

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

theorem complex.is_exp_cmp_filter.is_o_pow_mul_exp {l : filter } {b₁ b₂ : } (hl : complex.is_exp_cmp_filter l) (hb : b₁ < b₂) (m n : ) :
(λ (z : ), z ^ m * cexp (b₁ * z)) =o[l] λ (z : ), z ^ n * cexp (b₂ * z)

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

theorem complex.is_exp_cmp_filter.is_o_zpow_mul_exp {l : filter } {b₁ b₂ : } (hl : complex.is_exp_cmp_filter l) (hb : b₁ < b₂) (m n : ) :
(λ (z : ), z ^ m * cexp (b₁ * z)) =o[l] λ (z : ), z ^ n * cexp (b₂ * z)

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