Documentation

Mathlib.Topology.Connected.Clopen

Connected subsets and their relation to clopen sets #

In this file we show how connected subsets of a topological space are intimately connected to clopen sets.

Main declarations #

theorem IsPreconnected.subset_isClopen {α : Type u} [TopologicalSpace α] {s : Set α} {t : Set α} (hs : IsPreconnected s) (ht : IsClopen t) (hne : (s t).Nonempty) :
s t

Preconnected sets are either contained in or disjoint to any given clopen set.

theorem Sigma.isConnected_iff {ι : Type u_1} {π : ιType u_2} [(i : ι) → TopologicalSpace (π i)] {s : Set ((i : ι) × π i)} :
IsConnected s ∃ (i : ι) (t : Set (π i)), IsConnected t s = Sigma.mk i '' t
theorem Sigma.isPreconnected_iff {ι : Type u_1} {π : ιType u_2} [hι : Nonempty ι] [(i : ι) → TopologicalSpace (π i)] {s : Set ((i : ι) × π i)} :
IsPreconnected s ∃ (i : ι) (t : Set (π i)), IsPreconnected t s = Sigma.mk i '' t
theorem Sum.isConnected_iff {α : Type u} {β : Type v} [TopologicalSpace α] [TopologicalSpace β] {s : Set (α β)} :
IsConnected s (∃ (t : Set α), IsConnected t s = Sum.inl '' t) ∃ (t : Set β), IsConnected t s = Sum.inr '' t
theorem Sum.isPreconnected_iff {α : Type u} {β : Type v} [TopologicalSpace α] [TopologicalSpace β] {s : Set (α β)} :
IsPreconnected s (∃ (t : Set α), IsPreconnected t s = Sum.inl '' t) ∃ (t : Set β), IsPreconnected t s = Sum.inr '' t
theorem Continuous.exists_lift_sigma {α : Type u} {ι : Type u_1} {π : ιType u_2} [TopologicalSpace α] [ConnectedSpace α] [(i : ι) → TopologicalSpace (π i)] {f : α(i : ι) × π i} (hf : Continuous f) :
∃ (i : ι) (g : απ i), Continuous g f = Sigma.mk i g

A continuous map from a connected space to a disjoint union Σ i, π i can be lifted to one of the components π i. See also ContinuousMap.exists_lift_sigma for a version with bundled ContinuousMaps.

theorem nonempty_inter {α : Type u} [TopologicalSpace α] [PreconnectedSpace α] {s : Set α} {t : Set α} :
IsOpen sIsOpen ts t = Set.univs.Nonemptyt.Nonempty(s t).Nonempty
theorem isClopen_iff {α : Type u} [TopologicalSpace α] [PreconnectedSpace α] {s : Set α} :
IsClopen s s = s = Set.univ
theorem IsClopen.eq_univ {α : Type u} [TopologicalSpace α] [PreconnectedSpace α] {s : Set α} (h' : IsClopen s) (h : s.Nonempty) :
s = Set.univ
theorem subsingleton_of_disjoint_isClopen {α : Type u} {ι : Type u_1} [TopologicalSpace α] [PreconnectedSpace α] {s : ιSet α} (h_nonempty : ∀ (i : ι), (s i).Nonempty) (h_disj : Pairwise (Disjoint on s)) (h_clopen : ∀ (i : ι), IsClopen (s i)) :

In a preconnected space, any disjoint family of non-empty clopen subsets has at most one element.

theorem subsingleton_of_disjoint_isOpen_iUnion_eq_univ {α : Type u} {ι : Type u_1} [TopologicalSpace α] [PreconnectedSpace α] {s : ιSet α} (h_nonempty : ∀ (i : ι), (s i).Nonempty) (h_disj : Pairwise (Disjoint on s)) (h_open : ∀ (i : ι), IsOpen (s i)) (h_Union : ⋃ (i : ι), s i = Set.univ) :

In a preconnected space, any disjoint cover by non-empty open subsets has at most one element.

theorem subsingleton_of_disjoint_isClosed_iUnion_eq_univ {α : Type u} {ι : Type u_1} [TopologicalSpace α] [PreconnectedSpace α] {s : ιSet α} (h_nonempty : ∀ (i : ι), (s i).Nonempty) (h_disj : Pairwise (Disjoint on s)) [Finite ι] (h_closed : ∀ (i : ι), IsClosed (s i)) (h_Union : ⋃ (i : ι), s i = Set.univ) :

In a preconnected space, any finite disjoint cover by non-empty closed subsets has at most one element.

theorem frontier_eq_empty_iff {α : Type u} [TopologicalSpace α] [PreconnectedSpace α] {s : Set α} :
frontier s = s = s = Set.univ
theorem nonempty_frontier_iff {α : Type u} [TopologicalSpace α] [PreconnectedSpace α] {s : Set α} :
(frontier s).Nonempty s.Nonempty s Set.univ
theorem PreconnectedSpace.induction₂' {α : Type u} [TopologicalSpace α] [PreconnectedSpace α] (P : ααProp) (h : ∀ (x : α), ∀ᶠ (y : α) in nhds x, P x y P y x) (h' : Transitive P) (x : α) (y : α) :
P x y

In a preconnected space, given a transitive relation P, if P x y and P y x are true for y close enough to x, then P x y holds for all x, y. This is a version of the fact that, if an equivalence relation has open classes, then it has a single equivalence class.

theorem PreconnectedSpace.induction₂ {α : Type u} [TopologicalSpace α] [PreconnectedSpace α] (P : ααProp) (h : ∀ (x : α), ∀ᶠ (y : α) in nhds x, P x y) (h' : Transitive P) (h'' : Symmetric P) (x : α) (y : α) :
P x y

In a preconnected space, if a symmetric transitive relation P x y is true for y close enough to x, then it holds for all x, y. This is a version of the fact that, if an equivalence relation has open classes, then it has a single equivalence class.

theorem IsPreconnected.induction₂' {α : Type u} [TopologicalSpace α] {s : Set α} (hs : IsPreconnected s) (P : ααProp) (h : xs, ∀ᶠ (y : α) in nhdsWithin x s, P x y P y x) (h' : ∀ (x y z : α), x sy sz sP x yP y zP x z) {x : α} {y : α} (hx : x s) (hy : y s) :
P x y

In a preconnected set, given a transitive relation P, if P x y and P y x are true for y close enough to x, then P x y holds for all x, y. This is a version of the fact that, if an equivalence relation has open classes, then it has a single equivalence class.

theorem IsPreconnected.induction₂ {α : Type u} [TopologicalSpace α] {s : Set α} (hs : IsPreconnected s) (P : ααProp) (h : xs, ∀ᶠ (y : α) in nhdsWithin x s, P x y) (h' : ∀ (x y z : α), x sy sz sP x yP y zP x z) (h'' : ∀ (x y : α), x sy sP x yP y x) {x : α} {y : α} (hx : x s) (hy : y s) :
P x y

In a preconnected set, if a symmetric transitive relation P x y is true for y close enough to x, then it holds for all x, y. This is a version of the fact that, if an equivalence relation has open classes, then it has a single equivalence class.

theorem isPreconnected_iff_subset_of_disjoint {α : Type u} [TopologicalSpace α] {s : Set α} :
IsPreconnected s ∀ (u v : Set α), IsOpen uIsOpen vs u vs (u v) = s u s v

A set s is preconnected if and only if for every cover by two open sets that are disjoint on s, it is contained in one of the two covering sets.

theorem isConnected_iff_sUnion_disjoint_open {α : Type u} [TopologicalSpace α] {s : Set α} :
IsConnected s ∀ (U : Finset (Set α)), (∀ (u v : Set α), u Uv U(s (u v)).Nonemptyu = v)(uU, IsOpen u)s ⋃₀ UuU, s u

A set s is connected if and only if for every cover by a finite collection of open sets that are pairwise disjoint on s, it is contained in one of the members of the collection.

theorem disjoint_or_subset_of_isClopen {α : Type u} [TopologicalSpace α] {s : Set α} {t : Set α} (hs : IsPreconnected s) (ht : IsClopen t) :
Disjoint s t s t

Preconnected sets are either contained in or disjoint to any given clopen set.

theorem isPreconnected_iff_subset_of_disjoint_closed {α : Type u} [TopologicalSpace α] {s : Set α} :
IsPreconnected s ∀ (u v : Set α), IsClosed uIsClosed vs u vs (u v) = s u s v

A set s is preconnected if and only if for every cover by two closed sets that are disjoint on s, it is contained in one of the two covering sets.

theorem isPreconnected_iff_subset_of_fully_disjoint_closed {α : Type u} [TopologicalSpace α] {s : Set α} (hs : IsClosed s) :
IsPreconnected s ∀ (u v : Set α), IsClosed uIsClosed vs u vDisjoint u vs u s v

A closed set s is preconnected if and only if for every cover by two closed sets that are disjoint, it is contained in one of the two covering sets.

theorem IsClopen.connectedComponent_subset {α : Type u} [TopologicalSpace α] {s : Set α} {x : α} (hs : IsClopen s) (hx : x s) :
theorem connectedComponent_subset_iInter_isClopen {α : Type u} [TopologicalSpace α] {x : α} :
connectedComponent x ⋂ (Z : { Z : Set α // IsClopen Z x Z }), Z

The connected component of a point is always a subset of the intersection of all its clopen neighbourhoods.

theorem IsClopen.biUnion_connectedComponent_eq {α : Type u} [TopologicalSpace α] {Z : Set α} (h : IsClopen Z) :
xZ, connectedComponent x = Z

A clopen set is the union of its connected components.

theorem preimage_connectedComponent_connected {α : Type u} {β : Type v} [TopologicalSpace α] [TopologicalSpace β] {f : αβ} (connected_fibers : ∀ (t : β), IsConnected (f ⁻¹' {t})) (hcl : ∀ (T : Set β), IsClosed T IsClosed (f ⁻¹' T)) (t : β) :

The preimage of a connected component is preconnected if the function has connected fibers and a subset is closed iff the preimage is.

theorem QuotientMap.preimage_connectedComponent {α : Type u} {β : Type v} [TopologicalSpace α] [TopologicalSpace β] {f : αβ} (hf : QuotientMap f) (h_fibers : ∀ (y : β), IsConnected (f ⁻¹' {y})) (a : α) :
theorem QuotientMap.image_connectedComponent {α : Type u} {β : Type v} [TopologicalSpace α] [TopologicalSpace β] {f : αβ} (hf : QuotientMap f) (h_fibers : ∀ (y : β), IsConnected (f ⁻¹' {y})) (a : α) :

The setoid of connected components of a topological space

Equations
Instances For

    The quotient of a space by its connected components

    Equations
    Instances For

      Coercion from a topological space to the set of connected components of this space.

      Equations
      • ConnectedComponents.mk = Quotient.mk''
      Instances For
        Equations
        • ConnectedComponents.instCoeTC = { coe := ConnectedComponents.mk }
        Equations
        theorem ConnectedComponents.quotientMap_coe {α : Type u} [TopologicalSpace α] :
        QuotientMap ConnectedComponents.mk
        theorem ConnectedComponents.continuous_coe {α : Type u} [TopologicalSpace α] :
        Continuous ConnectedComponents.mk
        @[simp]
        theorem ConnectedComponents.range_coe {α : Type u} [TopologicalSpace α] :
        Set.range ConnectedComponents.mk = Set.univ

        The preimage of a singleton in connectedComponents is the connected component of an element in the equivalence class.

        theorem connectedComponents_preimage_image {α : Type u} [TopologicalSpace α] (U : Set α) :
        ConnectedComponents.mk ⁻¹' (ConnectedComponents.mk '' U) = xU, connectedComponent x

        The preimage of the image of a set under the quotient map to connectedComponents α is the union of the connected components of the elements in it.

        theorem isPreconnected_of_forall_constant {α : Type u} [TopologicalSpace α] {s : Set α} (hs : ∀ (f : αBool), ContinuousOn f sxs, ys, f x = f y) :

        If every map to Bool (a discrete two-element space), that is continuous on a set s, is constant on s, then s is preconnected

        theorem preconnectedSpace_of_forall_constant {α : Type u} [TopologicalSpace α] (hs : ∀ (f : αBool), Continuous f∀ (x y : α), f x = f y) :

        A PreconnectedSpace version of isPreconnected_of_forall_constant