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) :
theorem Filter.HasBasis.clusterPt_iff {X : Type u} [TopologicalSpace X] {x : X} {ฮนX : Sort u_3} {ฮนF : Sort u_4} {pX : ฮนX โ†’ Prop} {sX : ฮนX โ†’ Set X} {pF : ฮนF โ†’ Prop} {sF : ฮนF โ†’ Set 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 โ†” โˆ€ s โˆˆ nhds x, โˆƒแถ  (y : X) in F, y โˆˆ s
theorem ClusterPt.frequently {X : Type u} [TopologicalSpace X] {x : X} {F : Filter X} {p : X โ†’ Prop} (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 i โ†’ x โˆˆ closure (s i)
theorem clusterPt_iff_forall_mem_closure {X : Type u} [TopologicalSpace X] {x : X} {F : Filter X} :
ClusterPt x F โ†” โˆ€ s โˆˆ F, x โˆˆ closure s
theorem clusterPt_principal_iff {X : Type u} [TopologicalSpace X] {x : X} {s : Set X} :
ClusterPt x (Filter.principal s) โ†” โˆ€ U โˆˆ nhds 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} :
MapClusterPt x F u โ†’ ClusterPt x (Filter.map u F)

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 u โ†’ MapClusterPt 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 : X โ†’ Y} {y : Y} (hf : Filter.Tendsto f (nhds x โŠ“ Filter.map u F) (nhds y)) (hu : MapClusterPt x 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 : X โ†’ Y} {y : Y} (hf : Filter.Tendsto f (nhds x) (nhds y)) (hu : MapClusterPt x 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 : X โ†’ Y} (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 โ†” โˆ€ s โˆˆ nhds 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 โ†” โˆ€ s โˆˆ nhds 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 : X โ†’ Prop} (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} :
MapClusterPt x F (u โˆ˜ ฯ†) โ†” MapClusterPt x (Filter.map ฯ† F) u
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) :

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) โ†” โˆ€ U โˆˆ nhds x, โˆƒ y โˆˆ U โˆฉ 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.

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 โ†” โˆ€ t โˆˆ nhds x, (t โˆฉ s).Nonempty
theorem mem_closure_iff_nhds' {X : Type u} [TopologicalSpace X] {x : X} {s : Set X} :
x โˆˆ closure s โ†” โˆ€ t โˆˆ nhds 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 i โ†’ โˆƒ y โˆˆ t, y โˆˆ s i
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), (โˆ€ U โˆˆ nhds 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.NeBot โ†’ F โ‰ค Filter.principal s โ†’ F โ‰ค nhds x โ†’ x โˆˆ 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โ‚‚