Documentation

Mathlib.Topology.ClusterPt

Lemmas on cluster and accumulation points #

In this file we prove various lemmas on cluster points (also known as limit points and accumulation points) of a filter and of a sequence.

A filter F on X has x as a cluster point if ClusterPt x F : 𝓝 x ⊓ F ≠ ⊥. A map f : α → X clusters at x along F : Filter α if MapClusterPt x F f : ClusterPt x (map f F). In particular the notion of cluster point of a sequence u is MapClusterPt x atTop u.

theorem ClusterPt.neBot {X : Type u} [TopologicalSpace X] {x : X} {F : Filter X} (h : ClusterPt x F) :
(nhds x F).NeBot
theorem Filter.HasBasis.clusterPt_iff {X : Type u} [TopologicalSpace X] {x : X} {ιX : Sort u_3} {ιF : Sort u_4} {pX : ιXProp} {sX : ιXSet X} {pF : ιFProp} {sF : ιFSet X} {F : Filter X} (hX : (nhds x).HasBasis pX sX) (hF : F.HasBasis pF sF) :
ClusterPt x F ∀ ⦃i : ιX⦄, pX i∀ ⦃j : ιF⦄, pF j(sX i sF j).Nonempty
theorem Filter.HasBasis.clusterPt_iff_frequently {X : Type u} [TopologicalSpace X] {x : X} {ι : Sort u_3} {p : ιProp} {s : ιSet X} {F : Filter X} (hx : (nhds x).HasBasis p s) :
ClusterPt x F ∀ (i : ι), p i∃ᶠ (x : X) in F, x s i
theorem clusterPt_iff_frequently {X : Type u} [TopologicalSpace X] {x : X} {F : Filter X} :
ClusterPt x F snhds x, ∃ᶠ (y : X) in F, y s
theorem ClusterPt.frequently {X : Type u} [TopologicalSpace X] {x : X} {F : Filter X} {p : XProp} (hx : ClusterPt x F) (hp : ∀ᶠ (y : X) in nhds x, p y) :
∃ᶠ (y : X) in F, p y
theorem clusterPt_iff_nonempty {X : Type u} [TopologicalSpace X] {x : X} {F : Filter X} :
ClusterPt x F ∀ ⦃U : Set X⦄, U nhds x∀ ⦃V : Set X⦄, V F(U V).Nonempty
@[deprecated clusterPt_iff_nonempty (since := "2025-03-16")]
theorem clusterPt_iff {X : Type u} [TopologicalSpace X] {x : X} {F : Filter X} :
ClusterPt x F ∀ ⦃U : Set X⦄, U nhds x∀ ⦃V : Set X⦄, V F(U V).Nonempty

Alias of clusterPt_iff_nonempty.

theorem Filter.HasBasis.clusterPt_iff_forall_mem_closure {X : Type u} [TopologicalSpace X] {x : X} {ι : Sort u_3} {p : ιProp} {s : ιSet X} {F : Filter X} (hF : F.HasBasis p s) :
ClusterPt x F ∀ (i : ι), p ix closure (s i)
theorem clusterPt_iff_forall_mem_closure {X : Type u} [TopologicalSpace X] {x : X} {F : Filter X} :
ClusterPt x F sF, x closure s
theorem clusterPt_principal_iff {X : Type u} [TopologicalSpace X] {x : X} {s : Set X} :
ClusterPt x (Filter.principal s) Unhds x, (U s).Nonempty

x is a cluster point of a set s if every neighbourhood of x meets s on a nonempty set. See also mem_closure_iff_clusterPt.

theorem ClusterPt.of_le_nhds {X : Type u} [TopologicalSpace X] {x : X} {f : Filter X} (H : f nhds x) [f.NeBot] :
theorem ClusterPt.of_le_nhds' {X : Type u} [TopologicalSpace X] {x : X} {f : Filter X} (H : f nhds x) (_hf : f.NeBot) :
theorem ClusterPt.of_nhds_le {X : Type u} [TopologicalSpace X] {x : X} {f : Filter X} (H : nhds x f) :
theorem ClusterPt.mono {X : Type u} [TopologicalSpace X] {x : X} {f g : Filter X} (H : ClusterPt x f) (h : f g) :
theorem ClusterPt.of_inf_left {X : Type u} [TopologicalSpace X] {x : X} {f g : Filter X} (H : ClusterPt x (f g)) :
theorem ClusterPt.of_inf_right {X : Type u} [TopologicalSpace X] {x : X} {f g : Filter X} (H : ClusterPt x (f g)) :
theorem mapClusterPt_def {X : Type u} [TopologicalSpace X] {α : Type u_1} {F : Filter α} {u : αX} {x : X} :
theorem MapClusterPt.clusterPt {X : Type u} [TopologicalSpace X] {α : Type u_1} {F : Filter α} {u : αX} {x : X} :

Alias of the forward direction of mapClusterPt_def.

theorem Filter.EventuallyEq.mapClusterPt_iff {X : Type u} [TopologicalSpace X] {α : Type u_1} {F : Filter α} {u : αX} {x : X} {v : αX} (h : u =ᶠ[F] v) :
theorem MapClusterPt.congrFun {X : Type u} [TopologicalSpace X] {α : Type u_1} {F : Filter α} {u : αX} {x : X} {v : αX} (h : u =ᶠ[F] v) :
MapClusterPt x F uMapClusterPt x F v

Alias of the forward direction of Filter.EventuallyEq.mapClusterPt_iff.

theorem MapClusterPt.mono {X : Type u} [TopologicalSpace X] {α : Type u_1} {F : Filter α} {u : αX} {x : X} {G : Filter α} (h : MapClusterPt x F u) (hle : F G) :
theorem MapClusterPt.tendsto_comp' {X : Type u} [TopologicalSpace X] {Y : Type v} {α : Type u_1} {F : Filter α} {u : αX} {x : X} [TopologicalSpace Y] {f : XY} {y : Y} (hf : Filter.Tendsto f (nhds x Filter.map u F) (nhds y)) (hu : MapClusterPt x F u) :
MapClusterPt y F (f u)
theorem MapClusterPt.tendsto_comp {X : Type u} [TopologicalSpace X] {Y : Type v} {α : Type u_1} {F : Filter α} {u : αX} {x : X} [TopologicalSpace Y] {f : XY} {y : Y} (hf : Filter.Tendsto f (nhds x) (nhds y)) (hu : MapClusterPt x F u) :
MapClusterPt y F (f u)
theorem MapClusterPt.continuousAt_comp {X : Type u} [TopologicalSpace X] {Y : Type v} {α : Type u_1} {F : Filter α} {u : αX} {x : X} [TopologicalSpace Y] {f : XY} (hf : ContinuousAt f x) (hu : MapClusterPt x F u) :
MapClusterPt (f x) F (f u)
theorem Filter.HasBasis.mapClusterPt_iff_frequently {X : Type u} [TopologicalSpace X] {α : Type u_1} {F : Filter α} {u : αX} {x : X} {ι : Sort u_3} {p : ιProp} {s : ιSet X} (hx : (nhds x).HasBasis p s) :
MapClusterPt x F u ∀ (i : ι), p i∃ᶠ (a : α) in F, u a s i
theorem mapClusterPt_iff_frequently {X : Type u} [TopologicalSpace X] {α : Type u_1} {F : Filter α} {u : αX} {x : X} :
MapClusterPt x F u snhds x, ∃ᶠ (a : α) in F, u a s
@[deprecated mapClusterPt_iff_frequently (since := "2025-03-16")]
theorem mapClusterPt_iff {X : Type u} [TopologicalSpace X] {α : Type u_1} {F : Filter α} {u : αX} {x : X} :
MapClusterPt x F u snhds x, ∃ᶠ (a : α) in F, u a s

Alias of mapClusterPt_iff_frequently.

theorem MapClusterPt.frequently {X : Type u} [TopologicalSpace X] {α : Type u_1} {F : Filter α} {u : αX} {x : X} (h : MapClusterPt x F u) {p : XProp} (hp : ∀ᶠ (y : X) in nhds x, p y) :
∃ᶠ (a : α) in F, p (u a)
theorem mapClusterPt_comp {X : Type u} [TopologicalSpace X] {α : Type u_1} {β : Type u_2} {F : Filter α} {x : X} {φ : αβ} {u : βX} :
theorem Filter.Tendsto.mapClusterPt {X : Type u} [TopologicalSpace X] {α : Type u_1} {F : Filter α} {u : αX} {x : X} [F.NeBot] (h : Tendsto u F (nhds x)) :
theorem MapClusterPt.of_comp {X : Type u} [TopologicalSpace X] {α : Type u_1} {β : Type u_2} {F : Filter α} {u : αX} {x : X} {φ : βα} {p : Filter β} (h : Filter.Tendsto φ p F) (H : MapClusterPt x p (u φ)) :
theorem accPt_sup {X : Type u} [TopologicalSpace X] (x : X) (F G : Filter X) :
AccPt x (F G) AccPt x F AccPt x G

x is an accumulation point of a set C iff it is a cluster point of C ∖ {x}.

theorem accPt_iff_nhds {X : Type u} [TopologicalSpace X] (x : X) (C : Set X) :
AccPt x (Filter.principal C) Unhds x, yU C, y x

x is an accumulation point of a set C iff every neighborhood of x contains a point of C other than x.

theorem accPt_iff_frequently {X : Type u} [TopologicalSpace X] (x : X) (C : Set X) :
AccPt x (Filter.principal C) ∃ᶠ (y : X) in nhds x, y x y C

x is an accumulation point of a set C iff there are points near x in C and different from x.

theorem AccPt.mono {X : Type u} [TopologicalSpace X] {x : X} {F G : Filter X} (h : AccPt x F) (hFG : F G) :
AccPt x G

If x is an accumulation point of F and F ≤ G, then x is an accumulation point of G.

theorem AccPt.clusterPt {X : Type u} [TopologicalSpace X] (x : X) (F : Filter X) (h : AccPt x F) :

The set of cluster points of a filter is closed. In particular, the set of limit points of a sequence is closed.

If x is not an isolated point of a topological space, then {x}ᶜ is dense in the whole space.

If x is not an isolated point of a topological space, then the closure of {x}ᶜ is the whole space.

@[simp]

If x is not an isolated point of a topological space, then the interior of {x} is empty.

theorem mem_closure_iff_nhds {X : Type u} [TopologicalSpace X] {x : X} {s : Set X} :
x closure s tnhds x, (t s).Nonempty
theorem mem_closure_iff_nhds' {X : Type u} [TopologicalSpace X] {x : X} {s : Set X} :
x closure s tnhds x, ∃ (y : s), y t
theorem mem_closure_iff_nhds_basis' {X : Type u} [TopologicalSpace X] {ι : Sort w} {x : X} {t : Set X} {p : ιProp} {s : ιSet X} (h : (nhds x).HasBasis p s) :
x closure t ∀ (i : ι), p i(s i t).Nonempty
theorem mem_closure_iff_nhds_basis {X : Type u} [TopologicalSpace X] {ι : Sort w} {x : X} {t : Set X} {p : ιProp} {s : ιSet X} (h : (nhds x).HasBasis p s) :
x closure t ∀ (i : ι), p iyt, y s i
@[simp]
theorem isClosed_iff_clusterPt {X : Type u} [TopologicalSpace X] {s : Set X} :
IsClosed s ∀ (a : X), ClusterPt a (Filter.principal s)a s
theorem isClosed_iff_nhds {X : Type u} [TopologicalSpace X] {s : Set X} :
IsClosed s ∀ (x : X), (∀ Unhds x, (U s).Nonempty)x s
theorem isClosed_iff_forall_filter {X : Type u} [TopologicalSpace X] {s : Set X} :
IsClosed s ∀ (x : X) (F : Filter X), F.NeBotF Filter.principal sF nhds xx s
theorem mem_closure_of_mem_closure_union {X : Type u} [TopologicalSpace X] {x : X} {s₁ s₂ : Set X} (h : x closure (s₁ s₂)) (h₁ : s₁ nhds x) :
x closure s₂