Documentation

Mathlib.Topology.DiscreteQuotient

Discrete quotients of a topological space. #

This file defines the type of discrete quotients of a topological space, denoted DiscreteQuotient X. To avoid quantifying over types, we model such quotients as setoids whose equivalence classes are clopen.

Definitions #

  1. DiscreteQuotient X is the type of discrete quotients of X. It is endowed with a coercion to Type, which is defined as the quotient associated to the setoid in question, and each such quotient is endowed with the discrete topology.
  2. Given S : DiscreteQuotient X, the projection X → S is denoted S.proj.
  3. When X is compact and S : DiscreteQuotient X, the space S is endowed with a Fintype instance.

Order structure #

The type DiscreteQuotient X is endowed with an instance of a SemilatticeInf with OrderTop. The partial ordering A ≤ B mathematically means that B.proj factors through A.proj. The top element is the trivial quotient, meaning that every element of X is collapsed to a point. Given h : A ≤ B, the map A → B is DiscreteQuotient.ofLE h.

Whenever X is a locally connected space, the type DiscreteQuotient X is also endowed with an instance of an OrderBot, where the bot element is given by the connectedComponentSetoid, i.e., x ~ y means that x and y belong to the same connected component. In particular, if X is a discrete topological space, then x ~ y is equivalent (propositionally, not definitionally) to x = y.

Given f : C(X, Y), we define a predicate DiscreteQuotient.LEComap f A B for A : DiscreteQuotient X and B : DiscreteQuotient Y, asserting that f descends to A → B. If cond : DiscreteQuotient.LEComap h A B, the function A → B is obtained by DiscreteQuotient.map f cond.

Theorems #

The two main results proved in this file are:

  1. DiscreteQuotient.eq_of_forall_proj_eq which states that when X is compact, T₂, and totally disconnected, any two elements of X are equal if their projections in Q agree for all Q : DiscreteQuotient X.

  2. DiscreteQuotient.exists_of_compat which states that when X is compact, then any system of elements of Q as Q : DiscreteQuotient X varies, which is compatible with respect to DiscreteQuotient.ofLE, must arise from some element of X.

Remarks #

The constructions in this file will be used to show that any profinite space is a limit of finite discrete spaces.

theorem DiscreteQuotient.ext_iff {X : Type u_5} :
∀ {inst : TopologicalSpace X} (x y : DiscreteQuotient X), x = y Setoid.r = Setoid.r
theorem DiscreteQuotient.ext {X : Type u_5} :
∀ {inst : TopologicalSpace X} (x y : DiscreteQuotient X), Setoid.r = Setoid.rx = y
structure DiscreteQuotient (X : Type u_5) [TopologicalSpace X] extends Setoid :
Type u_5

The type of discrete quotients of a topological space.

  • r : XXProp
  • iseqv : Equivalence Setoid.r
  • isOpen_setOf_rel : ∀ (x : X), IsOpen (setOf (self.Rel x))

    For every point x, the set { y | Rel x y } is an open set.

Instances For
    theorem DiscreteQuotient.isOpen_setOf_rel {X : Type u_5} [TopologicalSpace X] (self : DiscreteQuotient X) (x : X) :
    IsOpen (setOf (self.Rel x))

    For every point x, the set { y | Rel x y } is an open set.

    Construct a discrete quotient from a clopen set.

    Equations
    Instances For
      theorem DiscreteQuotient.refl {X : Type u_2} [TopologicalSpace X] (S : DiscreteQuotient X) (x : X) :
      S.Rel x x
      theorem DiscreteQuotient.symm {X : Type u_2} [TopologicalSpace X] (S : DiscreteQuotient X) (x : X) (y : X) :
      S.Rel x yS.Rel y x
      theorem DiscreteQuotient.trans {X : Type u_2} [TopologicalSpace X] (S : DiscreteQuotient X) (x : X) (y : X) (z : X) :
      S.Rel x yS.Rel y zS.Rel x z
      Equations
      Equations
      def DiscreteQuotient.proj {X : Type u_2} [TopologicalSpace X] (S : DiscreteQuotient X) :
      XQuotient S.toSetoid

      The projection from X to the given discrete quotient.

      Equations
      • S.proj = Quotient.mk''
      Instances For
        theorem DiscreteQuotient.fiber_eq {X : Type u_2} [TopologicalSpace X] (S : DiscreteQuotient X) (x : X) :
        S.proj ⁻¹' {S.proj x} = setOf (S.Rel x)
        theorem DiscreteQuotient.isClopen_preimage {X : Type u_2} [TopologicalSpace X] (S : DiscreteQuotient X) (A : Set (Quotient S.toSetoid)) :
        IsClopen (S.proj ⁻¹' A)
        theorem DiscreteQuotient.isOpen_preimage {X : Type u_2} [TopologicalSpace X] (S : DiscreteQuotient X) (A : Set (Quotient S.toSetoid)) :
        IsOpen (S.proj ⁻¹' A)
        theorem DiscreteQuotient.isClosed_preimage {X : Type u_2} [TopologicalSpace X] (S : DiscreteQuotient X) (A : Set (Quotient S.toSetoid)) :
        IsClosed (S.proj ⁻¹' A)
        Equations
        • DiscreteQuotient.instInf = { inf := fun (S₁ S₂ : DiscreteQuotient X) => { toSetoid := S₁.toSetoid S₂.toSetoid, isOpen_setOf_rel := } }
        Equations
        Equations
        Equations
        • DiscreteQuotient.instInhabited = { default := }
        Equations
        • S.inhabitedQuotient = { default := S.proj default }
        Equations
        • =

        The quotient by ⊤ : DiscreteQuotient X is a Subsingleton.

        Equations
        • =

        Comap a discrete quotient along a continuous map.

        Equations
        Instances For
          def DiscreteQuotient.ofLE {X : Type u_2} [TopologicalSpace X] {A : DiscreteQuotient X} {B : DiscreteQuotient X} (h : A B) :
          Quotient A.toSetoidQuotient B.toSetoid

          The map induced by a refinement of a discrete quotient.

          Equations
          Instances For
            @[simp]
            theorem DiscreteQuotient.ofLE_ofLE {X : Type u_2} [TopologicalSpace X] {A : DiscreteQuotient X} {B : DiscreteQuotient X} {C : DiscreteQuotient X} (h₁ : A B) (h₂ : B C) (x : Quotient A.toSetoid) :
            @[simp]
            theorem DiscreteQuotient.ofLE_proj {X : Type u_2} [TopologicalSpace X] {A : DiscreteQuotient X} {B : DiscreteQuotient X} (h : A B) (x : X) :
            DiscreteQuotient.ofLE h (A.proj x) = B.proj x
            @[simp]

            When X is a locally connected space, there is an OrderBot instance on DiscreteQuotient X. The bottom element is given by connectedComponentSetoid X

            Equations
            • DiscreteQuotient.instOrderBotOfLocallyConnectedSpace = OrderBot.mk
            theorem DiscreteQuotient.proj_bot_inj {X : Type u_2} [TopologicalSpace X] [DiscreteTopology X] {x : X} {y : X} :
            .proj x = .proj y x = y

            Given f : C(X, Y), DiscreteQuotient.LEComap f A B is defined as A ≤ B.comap f. Mathematically this means that f descends to a morphism A → B.

            Equations
            Instances For
              theorem DiscreteQuotient.LEComap.mono {X : Type u_2} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {f : C(X, Y)} {A : DiscreteQuotient X} {A' : DiscreteQuotient X} {B : DiscreteQuotient Y} {B' : DiscreteQuotient Y} (h : DiscreteQuotient.LEComap f A B) (hA : A' A) (hB : B B') :
              def DiscreteQuotient.map {X : Type u_2} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {A : DiscreteQuotient X} {B : DiscreteQuotient Y} (f : C(X, Y)) (cond : DiscreteQuotient.LEComap f A B) :
              Quotient A.toSetoidQuotient B.toSetoid

              Map a discrete quotient along a continuous map.

              Equations
              Instances For
                @[simp]
                theorem DiscreteQuotient.map_comp_proj {X : Type u_2} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {f : C(X, Y)} {A : DiscreteQuotient X} {B : DiscreteQuotient Y} (cond : DiscreteQuotient.LEComap f A B) :
                DiscreteQuotient.map f cond A.proj = B.proj f
                @[simp]
                theorem DiscreteQuotient.map_proj {X : Type u_2} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {f : C(X, Y)} {A : DiscreteQuotient X} {B : DiscreteQuotient Y} (cond : DiscreteQuotient.LEComap f A B) (x : X) :
                DiscreteQuotient.map f cond (A.proj x) = B.proj (f x)
                @[simp]
                theorem DiscreteQuotient.ofLE_map {X : Type u_2} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {f : C(X, Y)} {A : DiscreteQuotient X} {B : DiscreteQuotient Y} {B' : DiscreteQuotient Y} (cond : DiscreteQuotient.LEComap f A B) (h : B B') (a : Quotient A.toSetoid) :
                @[simp]
                theorem DiscreteQuotient.map_ofLE {X : Type u_2} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {f : C(X, Y)} {A : DiscreteQuotient X} {A' : DiscreteQuotient X} {B : DiscreteQuotient Y} (cond : DiscreteQuotient.LEComap f A B) (h : A' A) (c : Quotient A'.toSetoid) :
                theorem DiscreteQuotient.eq_of_forall_proj_eq {X : Type u_2} [TopologicalSpace X] [T2Space X] [CompactSpace X] [disc : TotallyDisconnectedSpace X] {x : X} {y : X} (h : ∀ (Q : DiscreteQuotient X), Q.proj x = Q.proj y) :
                x = y
                theorem DiscreteQuotient.fiber_subset_ofLE {X : Type u_2} [TopologicalSpace X] {A : DiscreteQuotient X} {B : DiscreteQuotient X} (h : A B) (a : Quotient A.toSetoid) :
                A.proj ⁻¹' {a} B.proj ⁻¹' {DiscreteQuotient.ofLE h a}
                theorem DiscreteQuotient.exists_of_compat {X : Type u_2} [TopologicalSpace X] [CompactSpace X] (Qs : (Q : DiscreteQuotient X) → Quotient Q.toSetoid) (compat : ∀ (A B : DiscreteQuotient X) (h : A B), DiscreteQuotient.ofLE h (Qs A) = Qs B) :
                ∃ (x : X), ∀ (Q : DiscreteQuotient X), Q.proj x = Qs Q

                If X is a compact space, then any discrete quotient of X is finite.

                Equations
                • =

                If X is a compact space, then we associate to any discrete quotient on X a finite set of clopen subsets of X, given by the fibers of proj.

                TODO: prove that these form a partition of X 

                Equations
                Instances For
                  theorem DiscreteQuotient.comp_finsetClopens (X : Type u_2) [TopologicalSpace X] [CompactSpace X] :
                  ((Set.image fun (t : TopologicalSpace.Clopens X) => t.carrier) Finset.toSet) DiscreteQuotient.finsetClopens X = fun (x : DiscreteQuotient X) => match x with | { toSetoid := f, isOpen_setOf_rel := isOpen_setOf_rel } => f.classes

                  A helper lemma to prove that finsetClopens X is injective, see finsetClopens_inj.

                  The discrete quotients of a compact space are in bijection with a subtype of the type of Finset (Clopens X).

                  TODO: show that this is precisely those finsets of clopens which form a partition of X.

                  Equations
                  Instances For

                    Any locally constant function induces a discrete quotient.

                    Equations
                    • f.discreteQuotient = { toSetoid := Setoid.comap f , isOpen_setOf_rel := }
                    Instances For
                      def LocallyConstant.lift {α : Type u_1} {X : Type u_2} [TopologicalSpace X] (f : LocallyConstant X α) :
                      LocallyConstant (Quotient f.discreteQuotient.toSetoid) α

                      The (locally constant) function from the discrete quotient associated to a locally constant function.

                      Equations
                      • f.lift = { toFun := fun (a : Quotient f.discreteQuotient.toSetoid) => a.liftOn' f , isLocallyConstant := }
                      Instances For
                        @[simp]
                        theorem LocallyConstant.lift_comp_proj {α : Type u_1} {X : Type u_2} [TopologicalSpace X] (f : LocallyConstant X α) :
                        f.lift f.discreteQuotient.proj = f