Documentation

Mathlib.Topology.Instances.CantorSet

Ternary Cantor Set #

This file defines the Cantor ternary set and proves a few properties.

Main Definitions #

The order n pre-Cantor set, defined starting from [0, 1] and successively removing the middle third of each interval. Formally, the order n + 1 pre-Cantor set is the union of the images under the functions (· / 3) and ((2 + ·) / 3) of preCantorSet n.

Equations
Instances For
    @[simp]
    theorem preCantorSet_succ (n : ) :
    preCantorSet (n + 1) = (fun (x : ) => x / 3) '' preCantorSet n (fun (x : ) => (2 + x) / 3) '' preCantorSet n

    The Cantor set is the subset of the unit interval obtained as the intersection of all pre-Cantor sets. This means that the Cantor set is obtained by iteratively removing the open middle third of each subinterval, starting from the unit interval [0, 1].

    Equations
    Instances For

      Simple Properties #

      The ternary Cantor set is a subset of [0,1].

      theorem cantorSet_eq_union_halves :
      cantorSet = (fun (x : ) => x / 3) '' cantorSet (fun (x : ) => (2 + x) / 3) '' cantorSet

      The ternary Cantor set satisfies the equation C = C / 3 ∪ (2 / 3 + C / 3).

      The preCantor sets are closed.

      The ternary Cantor set is closed.

      The ternary Cantor set is compact.

      The Cantor set as the set of 0–2 numbers in the ternary system. #

      theorem ofDigits_zero_two_sequence_mem_cantorSet {a : Fin 3} (h : ∀ (n : ), a n 1) :

      If x = 0.d₀d₁... in base-3 (ternary), and none of the digits dᵢ is 1, then x belongs to the Cantor set.

      theorem ofDigits_zero_two_sequence_unique {a b : Fin 3} (ha : ∀ (n : ), a n 1) (hb : ∀ (n : ), b n 1) (h : Real.ofDigits a = Real.ofDigits b) :
      a = b

      If two base-3 representations using only digits 0 and 2 define the same number, then the sequences must be equal. This uniqueness fails for general base-3 representations (e.g. 0.1000... = 0.0222...).

      noncomputable def cantorStep (x : ) :

      Given x ∈ [0, 1/3] ∪ [2/3, 1] (i.e. a level of the Cantor set), this function rescales the interval containing x back to [0, 1]. Used to iteratively extract the ternary representation of x.

      Equations
      Instances For
        noncomputable def cantorSequence (x : ) :

        The infinite sequence obtained by repeatedly applying cantorStep to x.

        Equations
        Instances For
          noncomputable def cantorToBinary (x : ) :

          Points of the Cantor set correspond to infinite paths in the full binary tree. at each level n, the set preCantorSet (n + 1) splits each interval in preCantorSet n into two parts. Given x ∈ cantorSet, the point x lies in one of the intervals of preCantorSet n. This function tracks which of the two intervals in preCantorSet (n + 1) contains x at each step, producing the corresponding path as a stream of booleans.

          Equations
          Instances For
            noncomputable def cantorToTernary (x : ) :

            Given x in the Cantor set, return its ternary representation (d₀, d₁, …) using only digits 0 and 2, such that x = 0.d₀d₁... in base-3.

            Equations
            Instances For
              theorem cantorSet_eq_zero_two_ofDigits :
              cantorSet = {x : | ∃ (a : Fin 3), (∀ (i : ), a i 1) Real.ofDigits a = x}