Documentation

Mathlib.Order.Filter.AtTopBot.Basic

Basic results on Filter.atTop and Filter.atBot filters #

In this file we prove many lemmas like “if f → +∞, then f ± c → +∞”.

theorem Filter.eventually_forall_ge_atTop {α : Type u_3} [Preorder α] {p : αProp} :
(∀ᶠ (x : α) in atTop, ∀ (y : α), x yp y) ∀ᶠ (x : α) in atTop, p x
theorem Filter.eventually_forall_le_atBot {α : Type u_3} [Preorder α] {p : αProp} :
(∀ᶠ (x : α) in atBot, yx, p y) ∀ᶠ (x : α) in atBot, p x
theorem Filter.Tendsto.eventually_forall_ge_atTop {α : Type u_3} {β : Type u_4} [Preorder β] {l : Filter α} {p : βProp} {f : αβ} (hf : Tendsto f l atTop) (h_evtl : ∀ᶠ (x : β) in atTop, p x) :
∀ᶠ (x : α) in l, ∀ (y : β), f x yp y
theorem Filter.Tendsto.eventually_forall_le_atBot {α : Type u_3} {β : Type u_4} [Preorder β] {l : Filter α} {p : βProp} {f : αβ} (hf : Tendsto f l atBot) (h_evtl : ∀ᶠ (x : β) in atBot, p x) :
∀ᶠ (x : α) in l, yf x, p y
theorem Filter.hasAntitoneBasis_atTop {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] :
theorem Filter.atTop_basis {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] :
atTop.HasBasis (fun (x : α) => True) Set.Ici
theorem Filter.atTop_basis_Ioi {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] [NoMaxOrder α] :
atTop.HasBasis (fun (x : α) => True) Set.Ioi
theorem Filter.atTop_basis_Ioi' {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [NoMaxOrder α] (a : α) :
atTop.HasBasis (fun (x : α) => a < x) Set.Ioi
theorem Filter.atTop_basis' {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] (a : α) :
atTop.HasBasis (fun (x : α) => a x) Set.Ici
instance Filter.atTop_neBot {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] :
theorem Filter.atTop_neBot_iff {α : Type u_6} [Preorder α] :
atTop.NeBot Nonempty α IsDirected α fun (x1 x2 : α) => x1 x2
theorem Filter.atBot_neBot_iff {α : Type u_6} [Preorder α] :
atBot.NeBot Nonempty α IsDirected α fun (x1 x2 : α) => x1 x2
@[simp]
theorem Filter.mem_atTop_sets {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] {s : Set α} :
s atTop ∃ (a : α), ba, b s
@[simp]
theorem Filter.eventually_atTop {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : αProp} [Nonempty α] :
(∀ᶠ (x : α) in atTop, p x) ∃ (a : α), ba, p b
theorem Filter.frequently_atTop {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : αProp} [Nonempty α] :
(∃ᶠ (x : α) in atTop, p x) ∀ (a : α), ba, p b
theorem Filter.Eventually.exists_forall_of_atTop {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : αProp} [Nonempty α] :
(∀ᶠ (x : α) in atTop, p x)∃ (a : α), ba, p b

Alias of the forward direction of Filter.eventually_atTop.

theorem Filter.exists_eventually_atTop {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] {r : αβProp} :
(∃ (b : β), ∀ᶠ (a : α) in atTop, r a b) ∀ᶠ (a₀ : α) in atTop, ∃ (b : β), aa₀, r a b
theorem Filter.map_atTop_eq {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] {f : αβ} :
map f atTop = ⨅ (a : α), principal (f '' {a' : α | a a'})
theorem Filter.frequently_atTop' {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : αProp} [Nonempty α] [NoMaxOrder α] :
(∃ᶠ (x : α) in atTop, p x) ∀ (a : α), b > a, p b
theorem Filter.atBot_basis_Iio {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] [NoMinOrder α] :
atBot.HasBasis (fun (x : α) => True) Set.Iio
theorem Filter.atBot_basis_Iio' {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [NoMinOrder α] (a : α) :
atBot.HasBasis (fun (x : α) => x < a) Set.Iio
theorem Filter.atBot_basis' {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] (a : α) :
atBot.HasBasis (fun (x : α) => x a) Set.Iic
theorem Filter.atBot_basis {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] :
atBot.HasBasis (fun (x : α) => True) Set.Iic
instance Filter.atBot_neBot {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] :
@[simp]
theorem Filter.mem_atBot_sets {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] {s : Set α} :
s atBot ∃ (a : α), ba, b s
@[simp]
theorem Filter.eventually_atBot {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : αProp} [Nonempty α] :
(∀ᶠ (x : α) in atBot, p x) ∃ (a : α), ba, p b
theorem Filter.frequently_atBot {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : αProp} [Nonempty α] :
(∃ᶠ (x : α) in atBot, p x) ∀ (a : α), ba, p b
theorem Filter.Eventually.exists_forall_of_atBot {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : αProp} [Nonempty α] :
(∀ᶠ (x : α) in atBot, p x)∃ (a : α), ba, p b

Alias of the forward direction of Filter.eventually_atBot.

theorem Filter.exists_eventually_atBot {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] {r : αβProp} :
(∃ (b : β), ∀ᶠ (a : α) in atBot, r a b) ∀ᶠ (a₀ : α) in atBot, ∃ (b : β), aa₀, r a b
theorem Filter.map_atBot_eq {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Nonempty α] {f : αβ} :
map f atBot = ⨅ (a : α), principal (f '' {a' : α | a' a})
theorem Filter.frequently_atBot' {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : αProp} [Nonempty α] [NoMinOrder α] :
(∃ᶠ (x : α) in atBot, p x) ∀ (a : α), b < a, p b

Sequences #

theorem Filter.extraction_of_frequently_atTop' {P : Prop} (h : ∀ (N : ), n > N, P n) :
∃ (φ : ), StrictMono φ ∀ (n : ), P (φ n)
theorem Filter.extraction_of_frequently_atTop {P : Prop} (h : ∃ᶠ (n : ) in atTop, P n) :
∃ (φ : ), StrictMono φ ∀ (n : ), P (φ n)
theorem Filter.extraction_of_eventually_atTop {P : Prop} (h : ∀ᶠ (n : ) in atTop, P n) :
∃ (φ : ), StrictMono φ ∀ (n : ), P (φ n)
theorem Filter.extraction_forall_of_frequently {P : Prop} (h : ∀ (n : ), ∃ᶠ (k : ) in atTop, P n k) :
∃ (φ : ), StrictMono φ ∀ (n : ), P n (φ n)
theorem Filter.extraction_forall_of_eventually {P : Prop} (h : ∀ (n : ), ∀ᶠ (k : ) in atTop, P n k) :
∃ (φ : ), StrictMono φ ∀ (n : ), P n (φ n)
theorem Filter.extraction_forall_of_eventually' {P : Prop} (h : ∀ (n : ), ∃ (N : ), kN, P n k) :
∃ (φ : ), StrictMono φ ∀ (n : ), P n (φ n)
theorem Filter.Eventually.atTop_of_arithmetic {p : Prop} {n : } (hn : n 0) (hp : k < n, ∀ᶠ (a : ) in atTop, p (n * a + k)) :
∀ᶠ (a : ) in atTop, p a
theorem Filter.inf_map_atTop_neBot_iff {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {F : Filter β} {u : αβ} [Nonempty α] :
(F map u atTop).NeBot UF, ∀ (N : α), nN, u n U
theorem Filter.exists_le_of_tendsto_atTop {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {u : αβ} [Preorder β] (h : Tendsto u atTop atTop) (a : α) (b : β) :
a'a, b u a'
theorem Filter.exists_le_of_tendsto_atBot {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {u : αβ} [Preorder β] (h : Tendsto u atTop atBot) (a : α) (b : β) :
a'a, u a' b
theorem Filter.exists_lt_of_tendsto_atTop {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {u : αβ} [Preorder β] [NoMaxOrder β] (h : Tendsto u atTop atTop) (a : α) (b : β) :
a'a, b < u a'
theorem Filter.exists_lt_of_tendsto_atBot {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {u : αβ} [Preorder β] [NoMinOrder β] (h : Tendsto u atTop atBot) (a : α) (b : β) :
a'a, u a' < b
theorem Filter.inf_map_atBot_neBot_iff {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {F : Filter β} {u : αβ} :
(F map u atBot).NeBot UF, ∀ (N : α), nN, u n U
theorem Filter.high_scores {β : Type u_4} [LinearOrder β] [NoMaxOrder β] {u : β} (hu : Tendsto u atTop atTop) (N : ) :
nN, k < n, u k < u n

If u is a sequence which is unbounded above, then after any point, it reaches a value strictly greater than all previous values.

theorem Filter.low_scores {β : Type u_4} [LinearOrder β] [NoMinOrder β] {u : β} (hu : Tendsto u atTop atBot) (N : ) :
nN, k < n, u n < u k

If u is a sequence which is unbounded below, then after any point, it reaches a value strictly smaller than all previous values.

theorem Filter.frequently_high_scores {β : Type u_4} [LinearOrder β] [NoMaxOrder β] {u : β} (hu : Tendsto u atTop atTop) :
∃ᶠ (n : ) in atTop, k < n, u k < u n

If u is a sequence which is unbounded above, then it Frequently reaches a value strictly greater than all previous values.

theorem Filter.frequently_low_scores {β : Type u_4} [LinearOrder β] [NoMinOrder β] {u : β} (hu : Tendsto u atTop atBot) :
∃ᶠ (n : ) in atTop, k < n, u n < u k

If u is a sequence which is unbounded below, then it Frequently reaches a value strictly smaller than all previous values.

theorem Filter.strictMono_subseq_of_tendsto_atTop {β : Type u_4} [LinearOrder β] [NoMaxOrder β] {u : β} (hu : Tendsto u atTop atTop) :
∃ (φ : ), StrictMono φ StrictMono (u φ)
theorem Filter.strictMono_subseq_of_id_le {u : } (hu : ∀ (n : ), n u n) :
∃ (φ : ), StrictMono φ StrictMono (u φ)
theorem Filter.tendsto_atTop' {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} {l : Filter β} :
Tendsto f atTop l sl, ∃ (a : α), ba, f b s
theorem Filter.tendsto_atTop_principal {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} {s : Set β} :
Tendsto f atTop (principal s) ∃ (N : α), nN, f n s
theorem Filter.tendsto_atTop_atTop {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} [Preorder β] :
Tendsto f atTop atTop ∀ (b : β), ∃ (i : α), ∀ (a : α), i ab f a

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

theorem Filter.tendsto_atTop_atBot {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} [Preorder β] :
Tendsto f atTop atBot ∀ (b : β), ∃ (i : α), ∀ (a : α), i af a b
theorem Filter.tendsto_atTop_atTop_iff_of_monotone {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} [Preorder β] (hf : Monotone f) :
Tendsto f atTop atTop ∀ (b : β), ∃ (a : α), b f a
theorem Filter.tendsto_atTop_atBot_iff_of_antitone {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} [Preorder β] (hf : Antitone f) :
Tendsto f atTop atBot ∀ (b : β), ∃ (a : α), f a b
theorem Filter.tendsto_atBot' {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} {l : Filter β} :
Tendsto f atBot l sl, ∃ (a : α), ba, f b s
theorem Filter.tendsto_atBot_principal {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} {s : Set β} :
Tendsto f atBot (principal s) ∃ (N : α), nN, f n s
theorem Filter.tendsto_atBot_atTop {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} [Preorder β] :
Tendsto f atBot atTop ∀ (b : β), ∃ (i : α), ai, b f a
theorem Filter.tendsto_atBot_atBot {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} [Preorder β] :
Tendsto f atBot atBot ∀ (b : β), ∃ (i : α), ai, f a b
theorem Filter.tendsto_atBot_atBot_iff_of_monotone {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} [Preorder β] (hf : Monotone f) :
Tendsto f atBot atBot ∀ (b : β), ∃ (a : α), f a b
theorem Filter.tendsto_atBot_atTop_iff_of_antitone {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} [Preorder β] (hf : Antitone f) :
Tendsto f atBot atTop ∀ (b : β), ∃ (a : α), b f a
theorem Monotone.tendsto_atTop_atTop_iff {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} [Preorder β] (hf : Monotone f) :
Filter.Tendsto f Filter.atTop Filter.atTop ∀ (b : β), ∃ (a : α), b f a

Alias of Filter.tendsto_atTop_atTop_iff_of_monotone.

theorem Monotone.tendsto_atBot_atBot_iff {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {f : αβ} [Preorder β] (hf : Monotone f) :
Filter.Tendsto f Filter.atBot Filter.atBot ∀ (b : β), ∃ (a : α), f a b

Alias of Filter.tendsto_atBot_atBot_iff_of_monotone.

theorem Filter.Tendsto.subseq_mem {α : Type u_3} {F : Filter α} {V : Set α} (h : ∀ (n : ), V n F) {u : α} (hu : Tendsto u atTop F) :
∃ (φ : ), StrictMono φ ∀ (n : ), u (φ n) V n
theorem Filter.eventually_atBot_prod_self {α : Type u_3} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : α × αProp} :
(∀ᶠ (x : α × α) in atBot, p x) ∃ (a : α), ∀ (k l : α), k al ap (k, l)
theorem Filter.eventually_atTop_prod_self {α : Type u_3} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : α × αProp} :
(∀ᶠ (x : α × α) in atTop, p x) ∃ (a : α), ∀ (k l : α), a ka lp (k, l)
theorem Filter.eventually_atBot_prod_self' {α : Type u_3} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : α × αProp} :
(∀ᶠ (x : α × α) in atBot, p x) ∃ (a : α), ka, la, p (k, l)
theorem Filter.eventually_atTop_prod_self' {α : Type u_3} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {p : α × αProp} :
(∀ᶠ (x : α × α) in atTop, p x) ∃ (a : α), ka, la, p (k, l)
theorem Filter.eventually_atTop_curry {α : Type u_3} {β : Type u_4} [Preorder α] [Preorder β] {p : α × βProp} (hp : ∀ᶠ (x : α × β) in atTop, p x) :
∀ᶠ (k : α) in atTop, ∀ᶠ (l : β) in atTop, p (k, l)
theorem Filter.eventually_atBot_curry {α : Type u_3} {β : Type u_4} [Preorder α] [Preorder β] {p : α × βProp} (hp : ∀ᶠ (x : α × β) in atBot, p x) :
∀ᶠ (k : α) in atBot, ∀ᶠ (l : β) in atBot, p (k, l)
theorem Filter.map_atTop_eq_of_gc_preorder {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Preorder β] [IsDirected β fun (x1 x2 : β) => x1 x2] {f : αβ} (hf : Monotone f) (b : β) (hgi : cb, ∃ (x : α), f x = c ∀ (a : α), f a c a x) :

A function f maps upwards closed sets (atTop sets) to upwards closed sets when it is a Galois insertion. The Galois "insertion" and "connection" is weakened to only require it to be an insertion and a connection above b.

theorem Filter.map_atTop_eq_of_gc {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [PartialOrder β] [IsDirected β fun (x1 x2 : β) => x1 x2] {f : αβ} (g : βα) (b : β) (hf : Monotone f) (gc : ∀ (a : α), cb, f a c a g c) (hgi : cb, c f (g c)) :

A function f maps upwards closed sets (atTop sets) to upwards closed sets when it is a Galois insertion. The Galois "insertion" and "connection" is weakened to only require it to be an insertion and a connection above b.

theorem Filter.map_atBot_eq_of_gc_preorder {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Preorder β] [IsDirected β fun (x1 x2 : β) => x1 x2] {f : αβ} (hf : Monotone f) (b : β) (hgi : cb, ∃ (x : α), f x = c ∀ (a : α), c f a x a) :
theorem Filter.map_atBot_eq_of_gc {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [PartialOrder β] [IsDirected β fun (x1 x2 : β) => x1 x2] {f : αβ} (g : βα) (b' : β) (hf : Monotone f) (gc : ∀ (a : α), bb', b f a g b a) (hgi : bb', f (g b) b) :
theorem Filter.map_val_atTop_of_Ici_subset {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {a : α} {s : Set α} (h : Set.Ici a s) :
@[simp]
theorem Filter.map_val_Ici_atTop {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] (a : α) :

The image of the filter atTop on Ici a under the coercion equals atTop.

@[simp]
theorem Filter.map_val_Ioi_atTop {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [NoMaxOrder α] (a : α) :

The image of the filter atTop on Ioi a under the coercion equals atTop.

theorem Filter.atTop_Ioi_eq {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] (a : α) :

The atTop filter for an open interval Ioi a comes from the atTop filter in the ambient order.

theorem Filter.atTop_Ici_eq {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] (a : α) :

The atTop filter for an open interval Ici a comes from the atTop filter in the ambient order.

@[simp]
theorem Filter.map_val_Iio_atBot {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [NoMinOrder α] (a : α) :

The atBot filter for an open interval Iio a comes from the atBot filter in the ambient order.

theorem Filter.atBot_Iio_eq {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] (a : α) :

The atBot filter for an open interval Iio a comes from the atBot filter in the ambient order.

@[simp]
theorem Filter.map_val_Iic_atBot {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] (a : α) :

The atBot filter for an open interval Iic a comes from the atBot filter in the ambient order.

theorem Filter.atBot_Iic_eq {α : Type u_3} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] (a : α) :

The atBot filter for an open interval Iic a comes from the atBot filter in the ambient order.

theorem Filter.tendsto_Ioi_atTop {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {a : α} {f : β(Set.Ioi a)} {l : Filter β} :
Tendsto f l atTop Tendsto (fun (x : β) => (f x)) l atTop
theorem Filter.tendsto_Iio_atBot {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {a : α} {f : β(Set.Iio a)} {l : Filter β} :
Tendsto f l atBot Tendsto (fun (x : β) => (f x)) l atBot
theorem Filter.tendsto_Ici_atTop {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {a : α} {f : β(Set.Ici a)} {l : Filter β} :
Tendsto f l atTop Tendsto (fun (x : β) => (f x)) l atTop
theorem Filter.tendsto_Iic_atBot {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {a : α} {f : β(Set.Iic a)} {l : Filter β} :
Tendsto f l atBot Tendsto (fun (x : β) => (f x)) l atBot
@[simp]
theorem Filter.tendsto_comp_val_Ioi_atTop {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [NoMaxOrder α] {a : α} {f : αβ} {l : Filter β} :
Tendsto (fun (x : (Set.Ioi a)) => f x) atTop l Tendsto f atTop l
@[simp]
theorem Filter.tendsto_comp_val_Ici_atTop {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {a : α} {f : αβ} {l : Filter β} :
Tendsto (fun (x : (Set.Ici a)) => f x) atTop l Tendsto f atTop l
@[simp]
theorem Filter.tendsto_comp_val_Iio_atBot {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [NoMinOrder α] {a : α} {f : αβ} {l : Filter β} :
Tendsto (fun (x : (Set.Iio a)) => f x) atBot l Tendsto f atBot l
@[simp]
theorem Filter.tendsto_comp_val_Iic_atBot {α : Type u_3} {β : Type u_4} [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] {a : α} {f : αβ} {l : Filter β} :
Tendsto (fun (x : (Set.Iic a)) => f x) atBot l Tendsto f atBot l
theorem Filter.map_add_atTop_eq_nat (k : ) :
map (fun (a : ) => a + k) atTop = atTop
theorem Filter.map_sub_atTop_eq_nat (k : ) :
map (fun (a : ) => a - k) atTop = atTop
theorem Filter.tendsto_add_atTop_nat (k : ) :
Tendsto (fun (a : ) => a + k) atTop atTop
theorem Filter.tendsto_sub_atTop_nat (k : ) :
Tendsto (fun (a : ) => a - k) atTop atTop
theorem Filter.tendsto_add_atTop_iff_nat {α : Type u_3} {f : α} {l : Filter α} (k : ) :
Tendsto (fun (n : ) => f (n + k)) atTop l Tendsto f atTop l
theorem Filter.map_div_atTop_eq_nat (k : ) (hk : 0 < k) :
map (fun (a : ) => a / k) atTop = atTop
theorem Filter.unbounded_of_tendsto_atTop {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Preorder β] {f : αβ} [NoMaxOrder β] (h : Tendsto f atTop atTop) :
theorem Filter.unbounded_of_tendsto_atBot {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Preorder β] {f : αβ} [NoMinOrder β] (h : Tendsto f atTop atBot) :
theorem Filter.unbounded_of_tendsto_atTop' {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Preorder β] {f : αβ} [NoMaxOrder β] (h : Tendsto f atBot atTop) :
theorem Filter.unbounded_of_tendsto_atBot' {α : Type u_3} {β : Type u_4} [Nonempty α] [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 x2] [Preorder β] {f : αβ} [NoMinOrder β] (h : Tendsto f atBot atBot) :
theorem Filter.HasAntitoneBasis.eventually_subset {ι : Type u_1} {α : Type u_3} [Preorder ι] {l : Filter α} {s : ιSet α} (hl : l.HasAntitoneBasis s) {t : Set α} (ht : t l) :
∀ᶠ (i : ι) in atTop, s i t
theorem Filter.HasAntitoneBasis.tendsto {ι : Type u_1} {α : Type u_3} [Preorder ι] {l : Filter α} {s : ιSet α} (hl : l.HasAntitoneBasis s) {φ : ια} (h : ∀ (i : ι), φ i s i) :
theorem Filter.HasAntitoneBasis.comp_mono {ι : Type u_1} {ι' : Type u_2} {α : Type u_3} [Nonempty ι] [Preorder ι] [IsDirected ι fun (x1 x2 : ι) => x1 x2] [Preorder ι'] {l : Filter α} {s : ι'Set α} (hs : l.HasAntitoneBasis s) {φ : ιι'} (φ_mono : Monotone φ) (hφ : Tendsto φ atTop atTop) :
theorem Filter.HasAntitoneBasis.comp_strictMono {α : Type u_3} {l : Filter α} {s : Set α} (hs : l.HasAntitoneBasis s) {φ : } (hφ : StrictMono φ) :
theorem Filter.HasAntitoneBasis.subbasis_with_rel {α : Type u_3} {f : Filter α} {s : Set α} (hs : f.HasAntitoneBasis s) {r : Prop} (hr : ∀ (m : ), ∀ᶠ (n : ) in atTop, r m n) :
∃ (φ : ), StrictMono φ (∀ ⦃m n : ⦄, m < nr (φ m) (φ n)) f.HasAntitoneBasis (s φ)

Given an antitone basis s : ℕ → Set α of a filter, extract an antitone subbasis s ∘ φ, φ : ℕ → ℕ, such that m < n implies r (φ m) (φ n). This lemma can be used to extract an antitone basis with basis sets decreasing "sufficiently fast".

theorem Filter.subseq_forall_of_frequently {ι : Type u_6} {x : ι} {p : ιProp} {l : Filter ι} (h_tendsto : Tendsto x atTop l) (h : ∃ᶠ (n : ) in atTop, p (x n)) :
∃ (ns : ), Tendsto (fun (n : ) => x (ns n)) atTop l ∀ (n : ), p (x (ns n))
@[deprecated Nat.eventually_pow_lt_factorial_sub (since := "2024-09-25")]
theorem Nat.exists_pow_lt_factorial (c : ) :
n0 > 1, nn0, c ^ n < (n - 1).factorial
@[deprecated Nat.eventually_mul_pow_lt_factorial_sub (since := "2024-09-25")]
theorem Nat.exists_mul_pow_lt_factorial (a c : ) :
∃ (n0 : ), nn0, a * c ^ n < (n - 1).factorial