Ternary Cantor Set #
This file defines the Cantor ternary set and proves a few properties.
Main Definitions #
preCantorSet n: The ordernpre-Cantor set, defined inductively as the union of the images under the functions(· / 3)and((2 + ·) / 3), withpreCantorSet 0 := Set.Icc 0 1, i.e.preCantorSet 0is the unit interval [0,1].cantorSet: The ternary Cantor set, defined as the intersection of all pre-Cantor sets.cantorToTernary: given a numberxin the Cantor set, returns its ternary representation(d₀, d₁, ...)consisting only of digits0and2, such thatx = 0.d₀d₁...(seeofDigits_cantorToTernary).ofDigits_zero_two_sequence_mem_cantorSet: any such sequence corresponds to a number in the Cantor set.ofDigits_zero_two_sequence_unique: such a representation is unique.
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
- preCantorSet 0 = Set.Icc 0 1
- preCantorSet n.succ = (fun (x : ℝ) => x / 3) '' preCantorSet n ∪ (fun (x : ℝ) => (2 + x) / 3) '' preCantorSet n
Instances For
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
- cantorSet = ⋂ (n : ℕ), preCantorSet n
Instances For
Simple Properties #
The ternary Cantor set is a subset of [0,1].
The preCantor sets are closed.
The Cantor set as the set of 0–2 numbers in the ternary system. #
If x = 0.d₀d₁... in base-3 (ternary), and none of the digits dᵢ is 1,
then x belongs to the Cantor set.
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...).
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.
Instances For
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
- cantorToBinary x = Stream'.map (fun (x : ℝ) => if x ∈ Set.Icc 0 (1 / 3) then false else true) (cantorSequence x)
Instances For
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
- cantorToTernary x = Stream'.map (fun (x : Bool) => bif x then 2 else 0) (cantorToBinary x)