Documentation

Mathlib.Order.Filter.AtTopBot.Tendsto

Limits of Filter.atTop and Filter.atBot #

In this file we prove many lemmas on the combination of Filter.atTop and Filter.atBot and Tendsto.

theorem Filter.not_tendsto_const_atTop {α : Type u_3} {β : Type u_4} [Preorder α] [NoMaxOrder α] (x : α) (l : Filter β) [l.NeBot] :
¬Tendsto (fun (x_1 : β) => x) l atTop
theorem Filter.not_tendsto_const_atBot {α : Type u_3} {β : Type u_4} [Preorder α] [NoMinOrder α] (x : α) (l : Filter β) [l.NeBot] :
¬Tendsto (fun (x_1 : β) => x) l atBot
theorem Filter.Tendsto.eventually_gt_atTop {α : Type u_3} {β : Type u_4} [Preorder β] [NoMaxOrder β] {f : αβ} {l : Filter α} (hf : Tendsto f l atTop) (c : β) :
∀ᶠ (x : α) in l, c < f x
theorem Filter.Tendsto.eventually_ge_atTop {α : Type u_3} {β : Type u_4} [Preorder β] {f : αβ} {l : Filter α} (hf : Tendsto f l atTop) (c : β) :
∀ᶠ (x : α) in l, c f x
theorem Filter.Tendsto.eventually_ne_atTop {α : Type u_3} {β : Type u_4} [Preorder β] [NoMaxOrder β] {f : αβ} {l : Filter α} (hf : Tendsto f l atTop) (c : β) :
∀ᶠ (x : α) in l, f x c
theorem Filter.Tendsto.eventually_ne_atTop' {α : Type u_3} {β : Type u_4} [Preorder β] [NoMaxOrder β] {f : αβ} {l : Filter α} (hf : Tendsto f l atTop) (c : α) :
∀ᶠ (x : α) in l, x c
theorem Filter.Tendsto.eventually_lt_atBot {α : Type u_3} {β : Type u_4} [Preorder β] [NoMinOrder β] {f : αβ} {l : Filter α} (hf : Tendsto f l atBot) (c : β) :
∀ᶠ (x : α) in l, f x < c
theorem Filter.Tendsto.eventually_le_atBot {α : Type u_3} {β : Type u_4} [Preorder β] {f : αβ} {l : Filter α} (hf : Tendsto f l atBot) (c : β) :
∀ᶠ (x : α) in l, f x c
theorem Filter.Tendsto.eventually_ne_atBot {α : Type u_3} {β : Type u_4} [Preorder β] [NoMinOrder β] {f : αβ} {l : Filter α} (hf : Tendsto f l atBot) (c : β) :
∀ᶠ (x : α) in l, f x c
theorem Filter.tendsto_atTop_pure {α : Type u_3} {β : Type u_4} [PartialOrder α] [OrderTop α] (f : αβ) :
theorem Filter.tendsto_atBot_pure {α : Type u_3} {β : Type u_4} [PartialOrder α] [OrderBot α] (f : αβ) :
theorem Filter.tendsto_atTop {α : Type u_3} {β : Type u_4} [Preorder β] {m : αβ} {f : Filter α} :
Tendsto m f atTop ∀ (b : β), ∀ᶠ (a : α) in f, b m a
theorem Filter.tendsto_atBot {α : Type u_3} {β : Type u_4} [Preorder β] {m : αβ} {f : Filter α} :
Tendsto m f atBot ∀ (b : β), ∀ᶠ (a : α) in f, m a b
theorem Filter.tendsto_atTop_mono' {α : Type u_3} {β : Type u_4} [Preorder β] (l : Filter α) ⦃f₁ f₂ : αβ (h : f₁ ≤ᶠ[l] f₂) (h₁ : Tendsto f₁ l atTop) :
Tendsto f₂ l atTop
theorem Filter.tendsto_atBot_mono' {α : Type u_3} {β : Type u_4} [Preorder β] (l : Filter α) ⦃f₁ f₂ : αβ (h : f₁ ≤ᶠ[l] f₂) :
Tendsto f₂ l atBotTendsto f₁ l atBot
theorem Filter.tendsto_atTop_mono {α : Type u_3} {β : Type u_4} [Preorder β] {l : Filter α} {f g : αβ} (h : ∀ (n : α), f n g n) :
theorem Filter.tendsto_atBot_mono {α : Type u_3} {β : Type u_4} [Preorder β] {l : Filter α} {f g : αβ} (h : ∀ (n : α), f n g n) :

Sequences #

theorem Monotone.upperBounds_range_comp_tendsto_atTop {α : Type u_3} {β : Type u_4} {γ : Type u_5} [Preorder β] [Preorder γ] {l : Filter α} [l.NeBot] {f : βγ} (hf : Monotone f) {g : αβ} (hg : Filter.Tendsto g l Filter.atTop) :

If f is a monotone function and g tends to atTop along a nontrivial filter. then the upper bounds of the range of f ∘ g are the same as the upper bounds of the range of f.

This lemma together with exists_seq_monotone_tendsto_atTop_atTop below is useful to reduce a statement about a monotone family indexed by a type with countably generated atTop (e.g., ) to the case of a family indexed by natural numbers.

theorem Monotone.lowerBounds_range_comp_tendsto_atBot {α : Type u_3} {β : Type u_4} {γ : Type u_5} [Preorder β] [Preorder γ] {l : Filter α} [l.NeBot] {f : βγ} (hf : Monotone f) {g : αβ} (hg : Filter.Tendsto g l Filter.atBot) :

If f is a monotone function and g tends to atBot along a nontrivial filter. then the lower bounds of the range of f ∘ g are the same as the lower bounds of the range of f.

theorem Antitone.lowerBounds_range_comp_tendsto_atTop {α : Type u_3} {β : Type u_4} {γ : Type u_5} [Preorder β] [Preorder γ] {l : Filter α} [l.NeBot] {f : βγ} (hf : Antitone f) {g : αβ} (hg : Filter.Tendsto g l Filter.atTop) :

If f is an antitone function and g tends to atTop along a nontrivial filter. then the upper bounds of the range of f ∘ g are the same as the upper bounds of the range of f.

theorem Antitone.upperBounds_range_comp_tendsto_atBot {α : Type u_3} {β : Type u_4} {γ : Type u_5} [Preorder β] [Preorder γ] {l : Filter α} [l.NeBot] {f : βγ} (hf : Antitone f) {g : αβ} (hg : Filter.Tendsto g l Filter.atBot) :

If f is an antitone function and g tends to atBot along a nontrivial filter. then the upper bounds of the range of f ∘ g are the same as the upper bounds of the range of f.

theorem Filter.tendsto_atTop_atTop_of_monotone {α : Type u_3} {β : Type u_4} [Preorder α] [Preorder β] {f : αβ} (hf : Monotone f) (h : ∀ (b : β), ∃ (a : α), b f a) :
theorem Filter.tendsto_atTop_atBot_of_antitone {α : Type u_3} {β : Type u_4} [Preorder α] [Preorder β] {f : αβ} (hf : Antitone f) (h : ∀ (b : β), ∃ (a : α), f a b) :
theorem Filter.tendsto_atBot_atBot_of_monotone {α : Type u_3} {β : Type u_4} [Preorder α] [Preorder β] {f : αβ} (hf : Monotone f) (h : ∀ (b : β), ∃ (a : α), f a b) :
theorem Filter.tendsto_atBot_atTop_of_antitone {α : Type u_3} {β : Type u_4} [Preorder α] [Preorder β] {f : αβ} (hf : Antitone f) (h : ∀ (b : β), ∃ (a : α), b f a) :
theorem Monotone.tendsto_atTop_atTop {α : Type u_3} {β : Type u_4} [Preorder α] [Preorder β] {f : αβ} (hf : Monotone f) (h : ∀ (b : β), ∃ (a : α), b f a) :

Alias of Filter.tendsto_atTop_atTop_of_monotone.

theorem Monotone.tendsto_atBot_atBot {α : Type u_3} {β : Type u_4} [Preorder α] [Preorder β] {f : αβ} (hf : Monotone f) (h : ∀ (b : β), ∃ (a : α), f a b) :

Alias of Filter.tendsto_atBot_atBot_of_monotone.

theorem Filter.comap_embedding_atTop {β : Type u_4} {γ : Type u_5} [Preorder β] [Preorder γ] {e : βγ} (hm : ∀ (b₁ b₂ : β), e b₁ e b₂ b₁ b₂) (hu : ∀ (c : γ), ∃ (b : β), c e b) :
theorem Filter.comap_embedding_atBot {β : Type u_4} {γ : Type u_5} [Preorder β] [Preorder γ] {e : βγ} (hm : ∀ (b₁ b₂ : β), e b₁ e b₂ b₁ b₂) (hu : ∀ (c : γ), ∃ (b : β), e b c) :
theorem Filter.tendsto_atTop_embedding {α : Type u_3} {β : Type u_4} {γ : Type u_5} [Preorder β] [Preorder γ] {f : αβ} {e : βγ} {l : Filter α} (hm : ∀ (b₁ b₂ : β), e b₁ e b₂ b₁ b₂) (hu : ∀ (c : γ), ∃ (b : β), c e b) :
theorem Filter.tendsto_atBot_embedding {α : Type u_3} {β : Type u_4} {γ : Type u_5} [Preorder β] [Preorder γ] {f : αβ} {e : βγ} {l : Filter α} (hm : ∀ (b₁ b₂ : β), e b₁ e b₂ b₁ b₂) (hu : ∀ (c : γ), ∃ (b : β), e b c) :

A function f goes to -∞ independent of an order-preserving embedding e.

theorem Filter.tendsto_atTop_atTop_of_monotone' {ι : Type u_1} {α : Type u_3} [Preorder ι] [LinearOrder α] {u : ια} (h : Monotone u) (H : ¬BddAbove (Set.range u)) :

If u is a monotone function with linear ordered codomain and the range of u is not bounded above, then Tendsto u atTop atTop.

theorem Filter.tendsto_atBot_atBot_of_monotone' {ι : Type u_1} {α : Type u_3} [Preorder ι] [LinearOrder α] {u : ια} (h : Monotone u) (H : ¬BddBelow (Set.range u)) :

If u is a monotone function with linear ordered codomain and the range of u is not bounded below, then Tendsto u atBot atBot.

theorem Filter.tendsto_atTop_of_monotone_of_filter {ι : Type u_1} {α : Type u_3} [Preorder ι] [Preorder α] {l : Filter ι} {u : ια} (h : Monotone u) [l.NeBot] (hu : Tendsto u l atTop) :

If a monotone function u : ι → α tends to atTop along some non-trivial filter l, then it tends to atTop along atTop.

theorem Filter.tendsto_atBot_of_monotone_of_filter {ι : Type u_1} {α : Type u_3} [Preorder ι] [Preorder α] {l : Filter ι} {u : ια} (h : Monotone u) [l.NeBot] (hu : Tendsto u l atBot) :

If a monotone function u : ι → α tends to atBot along some non-trivial filter l, then it tends to atBot along atBot.

theorem Filter.tendsto_atTop_of_monotone_of_subseq {ι : Type u_1} {ι' : Type u_2} {α : Type u_3} [Preorder ι] [Preorder α] {u : ια} {φ : ι'ι} (h : Monotone u) {l : Filter ι'} [l.NeBot] (H : Tendsto (u φ) l atTop) :
theorem Filter.tendsto_atBot_of_monotone_of_subseq {ι : Type u_1} {ι' : Type u_2} {α : Type u_3} [Preorder ι] [Preorder α] {u : ια} {φ : ι'ι} (h : Monotone u) {l : Filter ι'} [l.NeBot] (H : Tendsto (u φ) l atBot) :