Documentation

Mathlib.Analysis.SpecialFunctions.Exp

Complex and real exponential #

In this file we prove continuity of Complex.exp and Real.exp. We also prove a few facts about limits of Real.exp at infinity.

Tags #

exp

theorem Complex.exp_bound_sq (x : ) (z : ) (hz : z 1) :
(x + z).exp - x.exp - z x.exp x.exp * z ^ 2
theorem Complex.locally_lipschitz_exp {r : } (hr_nonneg : 0 r) (hr_le : r 1) (x : ) (y : ) (hyx : y - x < r) :
y.exp - x.exp (1 + r) * x.exp * y - x
theorem Complex.exp_sub_sum_range_isBigO_pow (n : ) :
(fun (x : ) => x.exp - (Finset.range n).sum fun (i : ) => x ^ i / i.factorial) =O[nhds 0] fun (x : ) => x ^ n
theorem Complex.exp_sub_sum_range_succ_isLittleO_pow (n : ) :
(fun (x : ) => x.exp - (Finset.range (n + 1)).sum fun (i : ) => x ^ i / i.factorial) =o[nhds 0] fun (x : ) => x ^ n
theorem Filter.Tendsto.cexp {α : Type u_1} {l : Filter α} {f : α} {z : } (hf : Filter.Tendsto f l (nhds z)) :
Filter.Tendsto (fun (x : α) => (f x).exp) l (nhds z.exp)
theorem ContinuousWithinAt.cexp {α : Type u_1} [TopologicalSpace α] {f : α} {s : Set α} {x : α} (h : ContinuousWithinAt f s x) :
ContinuousWithinAt (fun (y : α) => (f y).exp) s x
theorem ContinuousAt.cexp {α : Type u_1} [TopologicalSpace α] {f : α} {x : α} (h : ContinuousAt f x) :
ContinuousAt (fun (y : α) => (f y).exp) x
theorem ContinuousOn.cexp {α : Type u_1} [TopologicalSpace α] {f : α} {s : Set α} (h : ContinuousOn f s) :
ContinuousOn (fun (y : α) => (f y).exp) s
theorem Continuous.cexp {α : Type u_1} [TopologicalSpace α] {f : α} (h : Continuous f) :
Continuous fun (y : α) => (f y).exp
theorem Real.exp_sub_sum_range_isBigO_pow (n : ) :
(fun (x : ) => x.exp - (Finset.range n).sum fun (i : ) => x ^ i / i.factorial) =O[nhds 0] fun (x : ) => x ^ n
theorem Real.exp_sub_sum_range_succ_isLittleO_pow (n : ) :
(fun (x : ) => x.exp - (Finset.range (n + 1)).sum fun (i : ) => x ^ i / i.factorial) =o[nhds 0] fun (x : ) => x ^ n
theorem Filter.Tendsto.rexp {α : Type u_1} {l : Filter α} {f : α} {z : } (hf : Filter.Tendsto f l (nhds z)) :
Filter.Tendsto (fun (x : α) => (f x).exp) l (nhds z.exp)
theorem ContinuousWithinAt.rexp {α : Type u_1} [TopologicalSpace α] {f : α} {s : Set α} {x : α} (h : ContinuousWithinAt f s x) :
ContinuousWithinAt (fun (y : α) => (f y).exp) s x
@[deprecated ContinuousWithinAt.rexp]
theorem ContinuousWithinAt.exp {α : Type u_1} [TopologicalSpace α] {f : α} {s : Set α} {x : α} (h : ContinuousWithinAt f s x) :
ContinuousWithinAt (fun (y : α) => (f y).exp) s x

Alias of ContinuousWithinAt.rexp.

theorem ContinuousAt.rexp {α : Type u_1} [TopologicalSpace α] {f : α} {x : α} (h : ContinuousAt f x) :
ContinuousAt (fun (y : α) => (f y).exp) x
@[deprecated ContinuousAt.rexp]
theorem ContinuousAt.exp {α : Type u_1} [TopologicalSpace α] {f : α} {x : α} (h : ContinuousAt f x) :
ContinuousAt (fun (y : α) => (f y).exp) x

Alias of ContinuousAt.rexp.

theorem ContinuousOn.rexp {α : Type u_1} [TopologicalSpace α] {f : α} {s : Set α} (h : ContinuousOn f s) :
ContinuousOn (fun (y : α) => (f y).exp) s
@[deprecated ContinuousOn.rexp]
theorem ContinuousOn.exp {α : Type u_1} [TopologicalSpace α] {f : α} {s : Set α} (h : ContinuousOn f s) :
ContinuousOn (fun (y : α) => (f y).exp) s

Alias of ContinuousOn.rexp.

theorem Continuous.rexp {α : Type u_1} [TopologicalSpace α] {f : α} (h : Continuous f) :
Continuous fun (y : α) => (f y).exp
@[deprecated Continuous.rexp]
theorem Continuous.exp {α : Type u_1} [TopologicalSpace α] {f : α} (h : Continuous f) :
Continuous fun (y : α) => (f y).exp

Alias of Continuous.rexp.

theorem Real.exp_half (x : ) :
(x / 2).exp = x.exp
theorem Real.tendsto_exp_atTop :
Filter.Tendsto Real.exp Filter.atTop Filter.atTop

The real exponential function tends to +∞ at +∞.

theorem Real.tendsto_exp_neg_atTop_nhds_zero :
Filter.Tendsto (fun (x : ) => (-x).exp) Filter.atTop (nhds 0)

The real exponential function tends to 0 at -∞ or, equivalently, exp(-x) tends to 0 at +∞

@[deprecated Real.tendsto_exp_neg_atTop_nhds_zero]
theorem Real.tendsto_exp_neg_atTop_nhds_0 :
Filter.Tendsto (fun (x : ) => (-x).exp) Filter.atTop (nhds 0)

Alias of Real.tendsto_exp_neg_atTop_nhds_zero.


The real exponential function tends to 0 at -∞ or, equivalently, exp(-x) tends to 0 at +∞

The real exponential function tends to 1 at 0.

@[deprecated Real.tendsto_exp_nhds_zero_nhds_one]

Alias of Real.tendsto_exp_nhds_zero_nhds_one.


The real exponential function tends to 1 at 0.

@[simp]
theorem Real.isBoundedUnder_ge_exp_comp {α : Type u_1} (l : Filter α) (f : α) :
Filter.IsBoundedUnder (fun (x x_1 : ) => x x_1) l fun (x : α) => (f x).exp
@[simp]
theorem Real.isBoundedUnder_le_exp_comp {α : Type u_1} {l : Filter α} {f : α} :
(Filter.IsBoundedUnder (fun (x x_1 : ) => x x_1) l fun (x : α) => (f x).exp) Filter.IsBoundedUnder (fun (x x_1 : ) => x x_1) l f
theorem Real.tendsto_exp_div_pow_atTop (n : ) :
Filter.Tendsto (fun (x : ) => x.exp / x ^ n) Filter.atTop Filter.atTop

The function exp(x)/x^n tends to +∞ at +∞, for any natural number n

theorem Real.tendsto_pow_mul_exp_neg_atTop_nhds_zero (n : ) :
Filter.Tendsto (fun (x : ) => x ^ n * (-x).exp) Filter.atTop (nhds 0)

The function x^n * exp(-x) tends to 0 at +∞, for any natural number n.

@[deprecated Real.tendsto_pow_mul_exp_neg_atTop_nhds_zero]
theorem Real.tendsto_pow_mul_exp_neg_atTop_nhds_0 (n : ) :
Filter.Tendsto (fun (x : ) => x ^ n * (-x).exp) Filter.atTop (nhds 0)

Alias of Real.tendsto_pow_mul_exp_neg_atTop_nhds_zero.


The function x^n * exp(-x) tends to 0 at +∞, for any natural number n.

theorem Real.tendsto_mul_exp_add_div_pow_atTop (b : ) (c : ) (n : ) (hb : 0 < b) :
Filter.Tendsto (fun (x : ) => (b * x.exp + c) / x ^ n) Filter.atTop Filter.atTop

The function (b * exp x + c) / (x ^ n) tends to +∞ at +∞, for any natural number n and any real numbers b and c such that b is positive.

theorem Real.tendsto_div_pow_mul_exp_add_atTop (b : ) (c : ) (n : ) (hb : 0 b) :
Filter.Tendsto (fun (x : ) => x ^ n / (b * x.exp + c)) Filter.atTop (nhds 0)

The function (x ^ n) / (b * exp x + c) tends to 0 at +∞, for any natural number n and any real numbers b and c such that b is nonzero.

@[simp]
theorem Real.coe_expOrderIso_apply (x : ) :
(Real.expOrderIso x) = x.exp
@[simp]
theorem Real.map_exp_atTop :
Filter.map Real.exp Filter.atTop = Filter.atTop
@[simp]
theorem Real.comap_exp_atTop :
Filter.comap Real.exp Filter.atTop = Filter.atTop
@[simp]
theorem Real.tendsto_exp_comp_atTop {α : Type u_1} {l : Filter α} {f : α} :
Filter.Tendsto (fun (x : α) => (f x).exp) l Filter.atTop Filter.Tendsto f l Filter.atTop
theorem Real.tendsto_comp_exp_atTop {α : Type u_1} {l : Filter α} {f : α} :
Filter.Tendsto (fun (x : ) => f x.exp) Filter.atTop l Filter.Tendsto f Filter.atTop l
@[simp]
theorem Real.tendsto_comp_exp_atBot {α : Type u_1} {l : Filter α} {f : α} :
Filter.Tendsto (fun (x : ) => f x.exp) Filter.atBot l Filter.Tendsto f (nhdsWithin 0 (Set.Ioi 0)) l
@[simp]
@[simp]
theorem Real.tendsto_exp_comp_nhds_zero {α : Type u_1} {l : Filter α} {f : α} :
Filter.Tendsto (fun (x : α) => (f x).exp) l (nhds 0) Filter.Tendsto f l Filter.atBot
theorem Real.isLittleO_pow_exp_atTop {n : } :
(fun (x : ) => x ^ n) =o[Filter.atTop] Real.exp
@[simp]
theorem Real.isBigO_exp_comp_exp_comp {α : Type u_1} {l : Filter α} {f : α} {g : α} :
((fun (x : α) => (f x).exp) =O[l] fun (x : α) => (g x).exp) Filter.IsBoundedUnder (fun (x x_1 : ) => x x_1) l (f - g)
@[simp]
theorem Real.isTheta_exp_comp_exp_comp {α : Type u_1} {l : Filter α} {f : α} {g : α} :
((fun (x : α) => (f x).exp) =Θ[l] fun (x : α) => (g x).exp) Filter.IsBoundedUnder (fun (x x_1 : ) => x x_1) l fun (x : α) => |f x - g x|
@[simp]
theorem Real.isLittleO_exp_comp_exp_comp {α : Type u_1} {l : Filter α} {f : α} {g : α} :
((fun (x : α) => (f x).exp) =o[l] fun (x : α) => (g x).exp) Filter.Tendsto (fun (x : α) => g x - f x) l Filter.atTop
theorem Real.isLittleO_one_exp_comp {α : Type u_1} {l : Filter α} {f : α} :
((fun (x : α) => 1) =o[l] fun (x : α) => (f x).exp) Filter.Tendsto f l Filter.atTop
@[simp]
theorem Real.isBigO_one_exp_comp {α : Type u_1} {l : Filter α} {f : α} :
((fun (x : α) => 1) =O[l] fun (x : α) => (f x).exp) Filter.IsBoundedUnder (fun (x x_1 : ) => x x_1) l f

Real.exp (f x) is bounded away from zero along a filter if and only if this filter is bounded from below under f.

theorem Real.isBigO_exp_comp_one {α : Type u_1} {l : Filter α} {f : α} :
((fun (x : α) => (f x).exp) =O[l] fun (x : α) => 1) Filter.IsBoundedUnder (fun (x x_1 : ) => x x_1) l f

Real.exp (f x) is bounded away from zero along a filter if and only if this filter is bounded from below under f.

@[simp]
theorem Real.isTheta_exp_comp_one {α : Type u_1} {l : Filter α} {f : α} :
((fun (x : α) => (f x).exp) =Θ[l] fun (x : α) => 1) Filter.IsBoundedUnder (fun (x x_1 : ) => x x_1) l fun (x : α) => |f x|

Real.exp (f x) is bounded away from zero and infinity along a filter l if and only if |f x| is bounded from above along this filter.

theorem Real.summable_exp_nat_mul_iff {a : } :
(Summable fun (n : ) => (n * a).exp) a < 0
theorem Real.summable_exp_neg_nat :
Summable fun (n : ) => (-n).exp
theorem Real.summable_pow_mul_exp_neg_nat_mul (k : ) {r : } (hr : 0 < r) :
Summable fun (n : ) => n ^ k * (-r * n).exp
theorem HasSum.rexp {ι : Type u_1} {f : ι} {a : } (h : HasSum f a) :
HasProd (Real.exp f) a.exp

If f has sum a, then exp ∘ f has product exp a.

theorem Complex.tendsto_exp_nhds_zero_iff {α : Type u_1} {l : Filter α} {f : α} :
Filter.Tendsto (fun (x : α) => (f x).exp) l (nhds 0) Filter.Tendsto (fun (x : α) => (f x).re) l Filter.atBot
theorem HasSum.cexp {ι : Type u_1} {f : ι} {a : } (h : HasSum f a) :

If f has sum a, then exp ∘ f has product exp a.