Convergence to infinity and countably generated filters #
In this file we prove that
Filter.atTop
andFilter.atBot
filters on a countable type are countably generated;Filter.exists_seq_tendsto
: iff
is a nontrivial countably generated filter, then there exists a sequence that converges. tof
;Filter.tendsto_iff_seq_tendsto
: convergence along a countably generated filter is equivalent to convergence along all sequences that converge to this filter.
@[instance 200]
instance
Filter.atTop.isCountablyGenerated
{α : Type u_1}
[Preorder α]
[Countable α]
:
Filter.atTop.IsCountablyGenerated
@[instance 200]
instance
Filter.atBot.isCountablyGenerated
{α : Type u_1}
[Preorder α]
[Countable α]
:
Filter.atBot.IsCountablyGenerated
instance
OrderDual.instIsCountablyGeneratedAtTop
{α : Type u_1}
[Preorder α]
[Filter.atBot.IsCountablyGenerated]
:
Filter.atTop.IsCountablyGenerated
instance
OrderDual.instIsCountablyGeneratedAtBot
{α : Type u_1}
[Preorder α]
[Filter.atTop.IsCountablyGenerated]
:
Filter.atBot.IsCountablyGenerated
theorem
Filter.atTop_countable_basis
{α : Type u_1}
[Preorder α]
[IsDirected α fun (x1 x2 : α) => x1 ≤ x2]
[Nonempty α]
[Countable α]
:
Filter.atTop.HasCountableBasis (fun (x : α) => True) Set.Ici
theorem
Filter.atBot_countable_basis
{α : Type u_1}
[Preorder α]
[IsDirected α fun (x1 x2 : α) => x1 ≥ x2]
[Nonempty α]
[Countable α]
:
Filter.atBot.HasCountableBasis (fun (x : α) => True) Set.Iic
theorem
Filter.exists_seq_tendsto
{α : Type u_1}
(f : Filter α)
[f.IsCountablyGenerated]
[f.NeBot]
:
∃ (x : ℕ → α), Filter.Tendsto x Filter.atTop f
If f
is a nontrivial countably generated filter, then there exists a sequence that converges
to f
.
theorem
Filter.exists_seq_monotone_tendsto_atTop_atTop
(α : Type u_3)
[Preorder α]
[Nonempty α]
[IsDirected α fun (x1 x2 : α) => x1 ≤ x2]
[Filter.atTop.IsCountablyGenerated]
:
∃ (xs : ℕ → α), Monotone xs ∧ Filter.Tendsto xs Filter.atTop Filter.atTop
theorem
Filter.exists_seq_antitone_tendsto_atTop_atBot
(α : Type u_3)
[Preorder α]
[Nonempty α]
[IsDirected α fun (x1 x2 : α) => x1 ≥ x2]
[Filter.atBot.IsCountablyGenerated]
:
∃ (xs : ℕ → α), Antitone xs ∧ Filter.Tendsto xs Filter.atTop Filter.atBot
theorem
Filter.tendsto_iff_seq_tendsto
{α : Type u_1}
{β : Type u_2}
{f : α → β}
{k : Filter α}
{l : Filter β}
[k.IsCountablyGenerated]
:
Filter.Tendsto f k l ↔ ∀ (x : ℕ → α), Filter.Tendsto x Filter.atTop k → Filter.Tendsto (f ∘ x) Filter.atTop l
An abstract version of continuity of sequentially continuous functions on metric spaces:
if a filter k
is countably generated then Tendsto f k l
iff for every sequence u
converging to k
, f ∘ u
tends to l
.
theorem
Filter.tendsto_of_seq_tendsto
{α : Type u_1}
{β : Type u_2}
{f : α → β}
{k : Filter α}
{l : Filter β}
[k.IsCountablyGenerated]
:
(∀ (x : ℕ → α), Filter.Tendsto x Filter.atTop k → Filter.Tendsto (f ∘ x) Filter.atTop l) → Filter.Tendsto f k l
theorem
Filter.eventually_iff_seq_eventually
{ι : Type u_3}
{l : Filter ι}
{p : ι → Prop}
[l.IsCountablyGenerated]
:
(∀ᶠ (n : ι) in l, p n) ↔ ∀ (x : ℕ → ι), Filter.Tendsto x Filter.atTop l → ∀ᶠ (n : ℕ) in Filter.atTop, p (x n)
theorem
Filter.frequently_iff_seq_frequently
{ι : Type u_3}
{l : Filter ι}
{p : ι → Prop}
[l.IsCountablyGenerated]
:
(∃ᶠ (n : ι) in l, p n) ↔ ∃ (x : ℕ → ι), Filter.Tendsto x Filter.atTop l ∧ ∃ᶠ (n : ℕ) in Filter.atTop, p (x n)
theorem
Filter.exists_seq_forall_of_frequently
{ι : Type u_3}
{l : Filter ι}
{p : ι → Prop}
[l.IsCountablyGenerated]
(h : ∃ᶠ (n : ι) in l, p n)
:
∃ (ns : ℕ → ι), Filter.Tendsto ns Filter.atTop l ∧ ∀ (n : ℕ), p (ns n)
theorem
Filter.frequently_iff_seq_forall
{ι : Type u_3}
{l : Filter ι}
{p : ι → Prop}
[l.IsCountablyGenerated]
:
(∃ᶠ (n : ι) in l, p n) ↔ ∃ (ns : ℕ → ι), Filter.Tendsto ns Filter.atTop l ∧ ∀ (n : ℕ), p (ns n)
theorem
Filter.tendsto_of_subseq_tendsto
{α : Type u_1}
{ι : Type u_3}
{x : ι → α}
{f : Filter α}
{l : Filter ι}
[l.IsCountablyGenerated]
(hxy :
∀ (ns : ℕ → ι),
Filter.Tendsto ns Filter.atTop l → ∃ (ms : ℕ → ℕ), Filter.Tendsto (fun (n : ℕ) => x (ns (ms n))) Filter.atTop f)
:
Filter.Tendsto x l f
A sequence converges if every subsequence has a convergent subsequence.
theorem
Filter.subseq_tendsto_of_neBot
{α : Type u_1}
{f : Filter α}
[f.IsCountablyGenerated]
{u : ℕ → α}
(hx : (f ⊓ Filter.map u Filter.atTop).NeBot)
:
∃ (θ : ℕ → ℕ), StrictMono θ ∧ Filter.Tendsto (u ∘ θ) Filter.atTop f