Limits and asymptotics of power functions at +∞
#
This file contains results about the limiting behaviour of power functions at +∞
. For convenience
some results on asymptotics as x → 0
(those which are not just continuity statements) are also
located here.
Limits at +∞
#
theorem
tendsto_rpow_atTop
{y : ℝ}
(hy : 0 < y)
:
Filter.Tendsto (fun (x : ℝ) => x ^ y) Filter.atTop Filter.atTop
The function x ^ y
tends to +∞
at +∞
for any positive real y
.
theorem
tendsto_rpow_neg_atTop
{y : ℝ}
(hy : 0 < y)
:
Filter.Tendsto (fun (x : ℝ) => x ^ (-y)) Filter.atTop (nhds 0)
The function x ^ (-y)
tends to 0
at +∞
for any positive real y
.
theorem
tendsto_rpow_atTop_of_base_lt_one
(b : ℝ)
(hb₀ : -1 < b)
(hb₁ : b < 1)
:
Filter.Tendsto (fun (x : ℝ) => b ^ x) Filter.atTop (nhds 0)
theorem
tendsto_rpow_atTop_of_base_gt_one
(b : ℝ)
(hb : 1 < b)
:
Filter.Tendsto (fun (x : ℝ) => b ^ x) Filter.atBot (nhds 0)
theorem
tendsto_rpow_atBot_of_base_lt_one
(b : ℝ)
(hb₀ : 0 < b)
(hb₁ : b < 1)
:
Filter.Tendsto (fun (x : ℝ) => b ^ x) Filter.atBot Filter.atTop
theorem
tendsto_rpow_atBot_of_base_gt_one
(b : ℝ)
(hb : 1 < b)
:
Filter.Tendsto (fun (x : ℝ) => b ^ x) Filter.atBot (nhds 0)
The function x ^ (1 / x)
tends to 1
at +∞
.
The function x ^ (-1 / x)
tends to 1
at +∞
.
theorem
tendsto_exp_div_rpow_atTop
(s : ℝ)
:
Filter.Tendsto (fun (x : ℝ) => Real.exp x / x ^ s) Filter.atTop Filter.atTop
The function exp(x) / x ^ s
tends to +∞
at +∞
, for any real number s
.
The function exp (b * x) / x ^ s
tends to +∞
at +∞
, for any real s
and b > 0
.
theorem
NNReal.tendsto_rpow_atTop
{y : ℝ}
(hy : 0 < y)
:
Filter.Tendsto (fun (x : NNReal) => x ^ y) Filter.atTop Filter.atTop
Asymptotic results: IsBigO
, IsLittleO
and IsTheta
#
theorem
Complex.isTheta_cpow_rpow
{α : Type u_1}
{l : Filter α}
{f g : α → ℂ}
(hl_im : Filter.IsBoundedUnder (fun (x1 x2 : ℝ) => x1 ≤ x2) l fun (x : α) => |(g x).im|)
(hl : ∀ᶠ (x : α) in l, f x = 0 → (g x).re = 0 → g x = 0)
:
theorem
Asymptotics.IsBigOWith.rpow
{α : Type u_1}
{r c : ℝ}
{l : Filter α}
{f g : α → ℝ}
(h : Asymptotics.IsBigOWith c l f g)
(hc : 0 ≤ c)
(hr : 0 ≤ r)
(hg : 0 ≤ᶠ[l] g)
:
Asymptotics.IsBigOWith (c ^ r) l (fun (x : α) => f x ^ r) fun (x : α) => g x ^ r
theorem
tendsto_log_div_rpow_nhds_zero
{r : ℝ}
(hr : r < 0)
:
Filter.Tendsto (fun (x : ℝ) => Real.log x / x ^ r) (nhdsWithin 0 (Set.Ioi 0)) (nhds 0)
theorem
tendsto_log_mul_rpow_nhds_zero
{r : ℝ}
(hr : 0 < r)
:
Filter.Tendsto (fun (x : ℝ) => Real.log x * x ^ r) (nhdsWithin 0 (Set.Ioi 0)) (nhds 0)
theorem
tendsto_log_mul_self_nhds_zero_left :
Filter.Tendsto (fun (x : ℝ) => Real.log x * x) (nhdsWithin 0 (Set.Iio 0)) (nhds 0)