mathlib3 documentation

analysis.specific_limits.normed

A collection of specific limit computations #

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

This file contains important specific limit computations in (semi-)normed groups/rings/spaces, as as well as such computations in when the natural proof passes through a fact about normed spaces.

Powers #

theorem normed_field.tendsto_norm_zpow_nhds_within_0_at_top {𝕜 : Type u_1} [normed_field 𝕜] {m : } (hm : m < 0) :
filter.tendsto (λ (x : 𝕜), x ^ m) (nhds_within 0 {0}) filter.at_top
theorem normed_field.tendsto_zero_smul_of_tendsto_zero_of_bounded {ι : Type u_1} {𝕜 : Type u_2} {𝔸 : Type u_3} [normed_field 𝕜] [normed_add_comm_group 𝔸] [normed_space 𝕜 𝔸] {l : filter ι} {ε : ι 𝕜} {f : ι 𝔸} (hε : filter.tendsto ε l (nhds 0)) (hf : filter.is_bounded_under has_le.le l (has_norm.norm f)) :
filter.tendsto f) l (nhds 0)

The (scalar) product of a sequence that tends to zero with a bounded one also tends to zero.

@[simp]
theorem normed_field.continuous_at_zpow {𝕜 : Type u_1} [nontrivially_normed_field 𝕜] {m : } {x : 𝕜} :
continuous_at (λ (x : 𝕜), x ^ m) x x 0 0 m
@[simp]
theorem is_o_pow_pow_of_lt_left {r₁ r₂ : } (h₁ : 0 r₁) (h₂ : r₁ < r₂) :
(λ (n : ), r₁ ^ n) =o[filter.at_top] λ (n : ), r₂ ^ n
theorem is_O_pow_pow_of_le_left {r₁ r₂ : } (h₁ : 0 r₁) (h₂ : r₁ r₂) :
(λ (n : ), r₁ ^ n) =O[filter.at_top] λ (n : ), r₂ ^ n
theorem is_o_pow_pow_of_abs_lt_left {r₁ r₂ : } (h : |r₁| < |r₂|) :
(λ (n : ), r₁ ^ n) =o[filter.at_top] λ (n : ), r₂ ^ n
theorem tfae_exists_lt_is_o_pow (f : ) (R : ) :
[ (a : ) (H : a set.Ioo (-R) R), f =o[filter.at_top] has_pow.pow a, (a : ) (H : a set.Ioo 0 R), f =o[filter.at_top] has_pow.pow a, (a : ) (H : a set.Ioo (-R) R), f =O[filter.at_top] has_pow.pow a, (a : ) (H : a set.Ioo 0 R), f =O[filter.at_top] has_pow.pow a, (a : ) (H : a < R) (C : ) (h₀ : 0 < C 0 < R), (n : ), |f n| C * a ^ n, (a : ) (H : a set.Ioo 0 R) (C : ) (H : C > 0), (n : ), |f n| C * a ^ n, (a : ) (H : a < R), ∀ᶠ (n : ) in filter.at_top, |f n| a ^ n, (a : ) (H : a set.Ioo 0 R), ∀ᶠ (n : ) in filter.at_top, |f n| a ^ n].tfae

Various statements equivalent to the fact that f n grows exponentially slower than R ^ n.

  • 0: $f n = o(a ^ n)$ for some $-R < a < R$;
  • 1: $f n = o(a ^ n)$ for some $0 < a < R$;
  • 2: $f n = O(a ^ n)$ for some $-R < a < R$;
  • 3: $f n = O(a ^ n)$ for some $0 < a < R$;
  • 4: there exist a < R and C such that one of C and R is positive and $|f n| ≤ Ca^n$ for all n;
  • 5: there exists 0 < a < R and a positive C such that $|f n| ≤ Ca^n$ for all n;
  • 6: there exists a < R such that $|f n| ≤ a ^ n$ for sufficiently large n;
  • 7: there exists 0 < a < R such that $|f n| ≤ a ^ n$ for sufficiently large n.

NB: For backwards compatibility, if you add more items to the list, please append them at the end of the list.

theorem is_o_pow_const_const_pow_of_one_lt {R : Type u_1} [normed_ring R] (k : ) {r : } (hr : 1 < r) :
(λ (n : ), n ^ k) =o[filter.at_top] λ (n : ), r ^ n

For any natural k and a real r > 1 we have n ^ k = o(r ^ n) as n → ∞.

theorem is_o_coe_const_pow_of_one_lt {R : Type u_1} [normed_ring R] {r : } (hr : 1 < r) :
coe =o[filter.at_top] λ (n : ), r ^ n

For a real r > 1 we have n = o(r ^ n) as n → ∞.

theorem is_o_pow_const_mul_const_pow_const_pow_of_norm_lt {R : Type u_1} [normed_ring R] (k : ) {r₁ : R} {r₂ : } (h : r₁ < r₂) :
(λ (n : ), n ^ k * r₁ ^ n) =o[filter.at_top] λ (n : ), r₂ ^ n

If ‖r₁‖ < r₂, then for any naturak k we have n ^ k r₁ ^ n = o (r₂ ^ n) as n → ∞.

theorem tendsto_pow_const_div_const_pow_of_one_lt (k : ) {r : } (hr : 1 < r) :
filter.tendsto (λ (n : ), n ^ k / r ^ n) filter.at_top (nhds 0)
theorem tendsto_pow_const_mul_const_pow_of_abs_lt_one (k : ) {r : } (hr : |r| < 1) :
filter.tendsto (λ (n : ), n ^ k * r ^ n) filter.at_top (nhds 0)

If |r| < 1, then n ^ k r ^ n tends to zero for any natural k.

theorem tendsto_pow_const_mul_const_pow_of_lt_one (k : ) {r : } (hr : 0 r) (h'r : r < 1) :
filter.tendsto (λ (n : ), n ^ k * r ^ n) filter.at_top (nhds 0)

If 0 ≤ r < 1, then n ^ k r ^ n tends to zero for any natural k. This is a specialized version of tendsto_pow_const_mul_const_pow_of_abs_lt_one, singled out for ease of application.

theorem tendsto_self_mul_const_pow_of_abs_lt_one {r : } (hr : |r| < 1) :
filter.tendsto (λ (n : ), n * r ^ n) filter.at_top (nhds 0)

If |r| < 1, then n * r ^ n tends to zero.

theorem tendsto_self_mul_const_pow_of_lt_one {r : } (hr : 0 r) (h'r : r < 1) :
filter.tendsto (λ (n : ), n * r ^ n) filter.at_top (nhds 0)

If 0 ≤ r < 1, then n * r ^ n tends to zero. This is a specialized version of tendsto_self_mul_const_pow_of_abs_lt_one, singled out for ease of application.

theorem tendsto_pow_at_top_nhds_0_of_norm_lt_1 {R : Type u_1} [normed_ring R] {x : R} (h : x < 1) :
filter.tendsto (λ (n : ), x ^ n) filter.at_top (nhds 0)

In a normed ring, the powers of an element x with ‖x‖ < 1 tend to zero.

theorem tendsto_pow_at_top_nhds_0_of_abs_lt_1 {r : } (h : |r| < 1) :
filter.tendsto (λ (n : ), r ^ n) filter.at_top (nhds 0)

Geometric series #

theorem has_sum_geometric_of_norm_lt_1 {K : Type u_4} [normed_field K] {ξ : K} (h : ξ < 1) :
has_sum (λ (n : ), ξ ^ n) (1 - ξ)⁻¹
theorem summable_geometric_of_norm_lt_1 {K : Type u_4} [normed_field K] {ξ : K} (h : ξ < 1) :
summable (λ (n : ), ξ ^ n)
theorem tsum_geometric_of_norm_lt_1 {K : Type u_4} [normed_field K] {ξ : K} (h : ξ < 1) :
∑' (n : ), ξ ^ n = (1 - ξ)⁻¹
theorem has_sum_geometric_of_abs_lt_1 {r : } (h : |r| < 1) :
has_sum (λ (n : ), r ^ n) (1 - r)⁻¹
theorem summable_geometric_of_abs_lt_1 {r : } (h : |r| < 1) :
summable (λ (n : ), r ^ n)
theorem tsum_geometric_of_abs_lt_1 {r : } (h : |r| < 1) :
∑' (n : ), r ^ n = (1 - r)⁻¹
@[simp]
theorem summable_geometric_iff_norm_lt_1 {K : Type u_4} [normed_field K] {ξ : K} :
summable (λ (n : ), ξ ^ n) ξ < 1

A geometric series in a normed field is summable iff the norm of the common ratio is less than one.

theorem summable_norm_pow_mul_geometric_of_norm_lt_1 {R : Type u_1} [normed_ring R] (k : ) {r : R} (hr : r < 1) :
summable (λ (n : ), n ^ k * r ^ n)
theorem summable_pow_mul_geometric_of_norm_lt_1 {R : Type u_1} [normed_ring R] [complete_space R] (k : ) {r : R} (hr : r < 1) :
summable (λ (n : ), n ^ k * r ^ n)
theorem has_sum_coe_mul_geometric_of_norm_lt_1 {𝕜 : Type u_1} [normed_field 𝕜] [complete_space 𝕜] {r : 𝕜} (hr : r < 1) :
has_sum (λ (n : ), n * r ^ n) (r / (1 - r) ^ 2)

If ‖r‖ < 1, then ∑' n : ℕ, n * r ^ n = r / (1 - r) ^ 2, has_sum version.

theorem tsum_coe_mul_geometric_of_norm_lt_1 {𝕜 : Type u_1} [normed_field 𝕜] [complete_space 𝕜] {r : 𝕜} (hr : r < 1) :
∑' (n : ), n * r ^ n = r / (1 - r) ^ 2

If ‖r‖ < 1, then ∑' n : ℕ, n * r ^ n = r / (1 - r) ^ 2.

theorem seminormed_add_comm_group.cauchy_seq_of_le_geometric {α : Type u_1} [seminormed_add_comm_group α] {C r : } (hr : r < 1) {u : α} (h : (n : ), u n - u (n + 1) C * r ^ n) :
theorem dist_partial_sum_le_of_le_geometric {α : Type u_1} [seminormed_add_comm_group α] {r C : } {f : α} (hf : (n : ), f n C * r ^ n) (n : ) :
has_dist.dist ((finset.range n).sum (λ (i : ), f i)) ((finset.range (n + 1)).sum (λ (i : ), f i)) C * r ^ n
theorem cauchy_seq_finset_of_geometric_bound {α : Type u_1} [seminormed_add_comm_group α] {r C : } {f : α} (hr : r < 1) (hf : (n : ), f n C * r ^ n) :
cauchy_seq (λ (s : finset ), s.sum (λ (x : ), f x))

If ‖f n‖ ≤ C * r ^ n for all n : ℕ and some r < 1, then the partial sums of f form a Cauchy sequence. This lemma does not assume 0 ≤ r or 0 ≤ C.

theorem norm_sub_le_of_geometric_bound_of_has_sum {α : Type u_1} [seminormed_add_comm_group α] {r C : } {f : α} (hr : r < 1) (hf : (n : ), f n C * r ^ n) {a : α} (ha : has_sum f a) (n : ) :
(finset.range n).sum (λ (x : ), f x) - a C * r ^ n / (1 - r)

If ‖f n‖ ≤ C * r ^ n for all n : ℕ and some r < 1, then the partial sums of f are within distance C * r ^ n / (1 - r) of the sum of the series. This lemma does not assume 0 ≤ r or 0 ≤ C.

@[simp]
theorem dist_partial_sum {α : Type u_1} [seminormed_add_comm_group α] (u : α) (n : ) :
has_dist.dist ((finset.range (n + 1)).sum (λ (k : ), u k)) ((finset.range n).sum (λ (k : ), u k)) = u n
@[simp]
theorem dist_partial_sum' {α : Type u_1} [seminormed_add_comm_group α] (u : α) (n : ) :
has_dist.dist ((finset.range n).sum (λ (k : ), u k)) ((finset.range (n + 1)).sum (λ (k : ), u k)) = u n
theorem cauchy_series_of_le_geometric {α : Type u_1} [seminormed_add_comm_group α] {C : } {u : α} {r : } (hr : r < 1) (h : (n : ), u n C * r ^ n) :
cauchy_seq (λ (n : ), (finset.range n).sum (λ (k : ), u k))
theorem normed_add_comm_group.cauchy_series_of_le_geometric' {α : Type u_1} [seminormed_add_comm_group α] {C : } {u : α} {r : } (hr : r < 1) (h : (n : ), u n C * r ^ n) :
cauchy_seq (λ (n : ), (finset.range (n + 1)).sum (λ (k : ), u k))
theorem normed_add_comm_group.cauchy_series_of_le_geometric'' {α : Type u_1} [seminormed_add_comm_group α] {C : } {u : α} {N : } {r : } (hr₀ : 0 < r) (hr₁ : r < 1) (h : (n : ), n N u n C * r ^ n) :
cauchy_seq (λ (n : ), (finset.range (n + 1)).sum (λ (k : ), u k))
theorem normed_ring.summable_geometric_of_norm_lt_1 {R : Type u_4} [normed_ring R] [complete_space R] (x : R) (h : x < 1) :
summable (λ (n : ), x ^ n)

A geometric series in a complete normed ring is summable. Proved above (same name, different namespace) for not-necessarily-complete normed fields.

theorem normed_ring.tsum_geometric_of_norm_lt_1 {R : Type u_4} [normed_ring R] [complete_space R] (x : R) (h : x < 1) :
∑' (n : ), x ^ n 1 - 1 + (1 - x)⁻¹

Bound for the sum of a geometric series in a normed ring. This formula does not assume that the normed ring satisfies the axiom ‖1‖ = 1.

theorem geom_series_mul_neg {R : Type u_4} [normed_ring R] [complete_space R] (x : R) (h : x < 1) :
(∑' (i : ), x ^ i) * (1 - x) = 1
theorem mul_neg_geom_series {R : Type u_4} [normed_ring R] [complete_space R] (x : R) (h : x < 1) :
(1 - x) * ∑' (i : ), x ^ i = 1

Summability tests based on comparison with geometric series #

theorem summable_of_ratio_norm_eventually_le {α : Type u_1} [seminormed_add_comm_group α] [complete_space α] {f : α} {r : } (hr₁ : r < 1) (h : ∀ᶠ (n : ) in filter.at_top, f (n + 1) r * f n) :
theorem summable_of_ratio_test_tendsto_lt_one {α : Type u_1} [normed_add_comm_group α] [complete_space α] {f : α} {l : } (hl₁ : l < 1) (hf : ∀ᶠ (n : ) in filter.at_top, f n 0) (h : filter.tendsto (λ (n : ), f (n + 1) / f n) filter.at_top (nhds l)) :
theorem not_summable_of_ratio_norm_eventually_ge {α : Type u_1} [seminormed_add_comm_group α] {f : α} {r : } (hr : 1 < r) (hf : ∃ᶠ (n : ) in filter.at_top, f n 0) (h : ∀ᶠ (n : ) in filter.at_top, r * f n f (n + 1)) :
theorem not_summable_of_ratio_test_tendsto_gt_one {α : Type u_1} [seminormed_add_comm_group α] {f : α} {l : } (hl : 1 < l) (h : filter.tendsto (λ (n : ), f (n + 1) / f n) filter.at_top (nhds l)) :

Dirichlet and alternating series tests #

theorem monotone.cauchy_seq_series_mul_of_tendsto_zero_of_bounded {E : Type u_4} [normed_add_comm_group E] [normed_space E] {b : } {f : } {z : E} (hfa : monotone f) (hf0 : filter.tendsto f filter.at_top (nhds 0)) (hgb : (n : ), (finset.range n).sum (λ (i : ), z i) b) :
cauchy_seq (λ (n : ), (finset.range (n + 1)).sum (λ (i : ), f i z i))

Dirichlet's Test for monotone sequences.

theorem antitone.cauchy_seq_series_mul_of_tendsto_zero_of_bounded {E : Type u_4} [normed_add_comm_group E] [normed_space E] {b : } {f : } {z : E} (hfa : antitone f) (hf0 : filter.tendsto f filter.at_top (nhds 0)) (hzb : (n : ), (finset.range n).sum (λ (i : ), z i) b) :
cauchy_seq (λ (n : ), (finset.range (n + 1)).sum (λ (i : ), f i z i))

Dirichlet's test for antitone sequences.

theorem norm_sum_neg_one_pow_le (n : ) :
(finset.range n).sum (λ (i : ), (-1) ^ i) 1
theorem monotone.cauchy_seq_alternating_series_of_tendsto_zero {f : } (hfa : monotone f) (hf0 : filter.tendsto f filter.at_top (nhds 0)) :
cauchy_seq (λ (n : ), (finset.range (n + 1)).sum (λ (i : ), (-1) ^ i * f i))

The alternating series test for monotone sequences. See also tendsto_alternating_series_of_monotone_tendsto_zero.

theorem monotone.tendsto_alternating_series_of_tendsto_zero {f : } (hfa : monotone f) (hf0 : filter.tendsto f filter.at_top (nhds 0)) :
(l : ), filter.tendsto (λ (n : ), (finset.range (n + 1)).sum (λ (i : ), (-1) ^ i * f i)) filter.at_top (nhds l)

The alternating series test for monotone sequences.

theorem antitone.cauchy_seq_alternating_series_of_tendsto_zero {f : } (hfa : antitone f) (hf0 : filter.tendsto f filter.at_top (nhds 0)) :
cauchy_seq (λ (n : ), (finset.range (n + 1)).sum (λ (i : ), (-1) ^ i * f i))

The alternating series test for antitone sequences. See also tendsto_alternating_series_of_antitone_tendsto_zero.

theorem antitone.tendsto_alternating_series_of_tendsto_zero {f : } (hfa : antitone f) (hf0 : filter.tendsto f filter.at_top (nhds 0)) :
(l : ), filter.tendsto (λ (n : ), (finset.range (n + 1)).sum (λ (i : ), (-1) ^ i * f i)) filter.at_top (nhds l)

The alternating series test for antitone sequences.

Factorial #

theorem real.summable_pow_div_factorial (x : ) :
summable (λ (n : ), x ^ n / (n.factorial))

The series ∑' n, x ^ n / n! is summable of any x : ℝ. See also exp_series_div_summable for a version that also works in , and exp_series_summable' for a version that works in any normed algebra over or .