Documentation

Mathlib.Topology.MetricSpace.HausdorffDimension

Hausdorff dimension #

The Hausdorff dimension of a set X in an (extended) metric space is the unique number dimH s : ℝ≥0∞ such that for any d : ℝ≥0 we have

In this file we define dimH s to be the Hausdorff dimension of s, then prove some basic properties of Hausdorff dimension.

Main definitions #

Main results #

Basic properties of Hausdorff dimension #

(Pre)images under (anti)lipschitz and Hölder continuous maps #

Hausdorff measure in ℝⁿ #

Notations #

We use the following notation localized in MeasureTheory. It is defined in MeasureTheory.Measure.Hausdorff.

Implementation notes #

Tags #

Hausdorff measure, Hausdorff dimension, dimension

@[irreducible]
noncomputable def dimH {X : Type u_2} [EMetricSpace X] (s : Set X) :

Hausdorff dimension of a set in an (e)metric space.

Equations
    Instances For

      Basic properties #

      theorem dimH_def {X : Type u_2} [EMetricSpace X] [MeasurableSpace X] [BorelSpace X] (s : Set X) :
      dimH s = ⨆ (d : NNReal), ⨆ (_ : (MeasureTheory.Measure.hausdorffMeasure d) s = ), d

      Unfold the definition of dimH using [MeasurableSpace X] [BorelSpace X] from the environment.

      theorem dimH_le {X : Type u_2} [EMetricSpace X] [MeasurableSpace X] [BorelSpace X] {s : Set X} {d : ENNReal} (H : ∀ (d' : NNReal), (MeasureTheory.Measure.hausdorffMeasure d') s = d' d) :
      dimH s d
      theorem dimH_eq_iInf {X : Type u_2} [EMetricSpace X] [MeasurableSpace X] [BorelSpace X] (s : Set X) :
      dimH s = ⨅ (d : NNReal), ⨅ (_ : (MeasureTheory.Measure.hausdorffMeasure d) s = 0), d

      The Hausdorff dimension of a set s is the infimum of all d : ℝ≥0 such that the d-dimensional Hausdorff measure of s is zero. This infimum is taken in ℝ≥0∞. This gives an equivalent definition of the Hausdorff dimension.

      theorem dimH_mono {X : Type u_2} [EMetricSpace X] {s t : Set X} (h : s t) :
      theorem dimH_subsingleton {X : Type u_2} [EMetricSpace X] {s : Set X} (h : s.Subsingleton) :
      dimH s = 0
      theorem Set.Subsingleton.dimH_zero {X : Type u_2} [EMetricSpace X] {s : Set X} (h : s.Subsingleton) :
      dimH s = 0

      Alias of dimH_subsingleton.

      @[simp]
      theorem dimH_empty {X : Type u_2} [EMetricSpace X] :
      @[simp]
      theorem dimH_singleton {X : Type u_2} [EMetricSpace X] (x : X) :
      dimH {x} = 0
      @[simp]
      theorem dimH_iUnion {X : Type u_2} [EMetricSpace X] {ι : Sort u_4} [Countable ι] (s : ιSet X) :
      dimH (⋃ (i : ι), s i) = ⨆ (i : ι), dimH (s i)
      @[simp]
      theorem dimH_bUnion {ι : Type u_1} {X : Type u_2} [EMetricSpace X] {s : Set ι} (hs : s.Countable) (t : ιSet X) :
      dimH (⋃ is, t i) = is, dimH (t i)
      @[simp]
      theorem dimH_sUnion {X : Type u_2} [EMetricSpace X] {S : Set (Set X)} (hS : S.Countable) :
      dimH (⋃₀ S) = sS, dimH s
      @[simp]
      theorem dimH_union {X : Type u_2} [EMetricSpace X] (s t : Set X) :
      dimH (s t) = max (dimH s) (dimH t)
      theorem dimH_countable {X : Type u_2} [EMetricSpace X] {s : Set X} (hs : s.Countable) :
      dimH s = 0
      theorem Set.Countable.dimH_zero {X : Type u_2} [EMetricSpace X] {s : Set X} (hs : s.Countable) :
      dimH s = 0

      Alias of dimH_countable.

      theorem dimH_finite {X : Type u_2} [EMetricSpace X] {s : Set X} (hs : s.Finite) :
      dimH s = 0
      theorem Set.Finite.dimH_zero {X : Type u_2} [EMetricSpace X] {s : Set X} (hs : s.Finite) :
      dimH s = 0

      Alias of dimH_finite.

      @[simp]
      theorem dimH_coe_finset {X : Type u_2} [EMetricSpace X] (s : Finset X) :
      dimH s = 0
      theorem Finset.dimH_zero {X : Type u_2} [EMetricSpace X] (s : Finset X) :
      dimH s = 0

      Alias of dimH_coe_finset.

      Hausdorff dimension as the supremum of local Hausdorff dimensions #

      theorem exists_mem_nhdsWithin_lt_dimH_of_lt_dimH {X : Type u_2} [EMetricSpace X] [SecondCountableTopology X] {s : Set X} {r : ENNReal} (h : r < dimH s) :
      xs, tnhdsWithin x s, r < dimH t

      If r is less than the Hausdorff dimension of a set s in an (extended) metric space with second countable topology, then there exists a point x ∈ s such that every neighborhood t of x within s has Hausdorff dimension greater than r.

      In an (extended) metric space with second countable topology, the Hausdorff dimension of a set s is the supremum over x ∈ s of the limit superiors of dimH t along (𝓝[s] x).smallSets.

      theorem iSup_limsup_dimH {X : Type u_2} [EMetricSpace X] [SecondCountableTopology X] (s : Set X) :

      In an (extended) metric space with second countable topology, the Hausdorff dimension of a set s is the supremum over all x of the limit superiors of dimH t along (𝓝[s] x).smallSets.

      Hausdorff dimension and Hölder continuity #

      theorem HolderOnWith.dimH_image_le {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] {C r : NNReal} {f : XY} {s : Set X} (h : HolderOnWith C r f s) (hr : 0 < r) :
      dimH (f '' s) dimH s / r

      If f is a Hölder continuous map with exponent r > 0, then dimH (f '' s) ≤ dimH s / r.

      theorem HolderWith.dimH_image_le {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] {C r : NNReal} {f : XY} (h : HolderWith C r f) (hr : 0 < r) (s : Set X) :
      dimH (f '' s) dimH s / r

      If f : X → Y is Hölder continuous with a positive exponent r, then the Hausdorff dimension of the image of a set s is at most dimH s / r.

      theorem HolderWith.dimH_range_le {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] {C r : NNReal} {f : XY} (h : HolderWith C r f) (hr : 0 < r) :

      If f is a Hölder continuous map with exponent r > 0, then the Hausdorff dimension of its range is at most the Hausdorff dimension of its domain divided by r.

      theorem dimH_image_le_of_locally_holder_on {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] [SecondCountableTopology X] {r : NNReal} {f : XY} (hr : 0 < r) {s : Set X} (hf : xs, ∃ (C : NNReal), tnhdsWithin x s, HolderOnWith C r f t) :
      dimH (f '' s) dimH s / r

      If s is a set in a space X with second countable topology and f : X → Y is Hölder continuous in a neighborhood within s of every point x ∈ s with the same positive exponent r but possibly different coefficients, then the Hausdorff dimension of the image f '' s is at most the Hausdorff dimension of s divided by r.

      theorem dimH_range_le_of_locally_holder_on {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] [SecondCountableTopology X] {r : NNReal} {f : XY} (hr : 0 < r) (hf : ∀ (x : X), ∃ (C : NNReal), snhds x, HolderOnWith C r f s) :

      If f : X → Y is Hölder continuous in a neighborhood of every point x : X with the same positive exponent r but possibly different coefficients, then the Hausdorff dimension of the range of f is at most the Hausdorff dimension of X divided by r.

      Hausdorff dimension and Lipschitz continuity #

      theorem LipschitzOnWith.dimH_image_le {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] {K : NNReal} {f : XY} {s : Set X} (h : LipschitzOnWith K f s) :
      dimH (f '' s) dimH s

      If f : X → Y is Lipschitz continuous on s, then dimH (f '' s) ≤ dimH s.

      theorem LipschitzWith.dimH_image_le {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] {K : NNReal} {f : XY} (h : LipschitzWith K f) (s : Set X) :
      dimH (f '' s) dimH s

      If f is a Lipschitz continuous map, then dimH (f '' s) ≤ dimH s.

      theorem LipschitzWith.dimH_range_le {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] {K : NNReal} {f : XY} (h : LipschitzWith K f) :

      If f is a Lipschitz continuous map, then the Hausdorff dimension of its range is at most the Hausdorff dimension of its domain.

      theorem dimH_image_le_of_locally_lipschitzOn {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] [SecondCountableTopology X] {f : XY} {s : Set X} (hf : xs, ∃ (C : NNReal), tnhdsWithin x s, LipschitzOnWith C f t) :
      dimH (f '' s) dimH s

      If s is a set in an extended metric space X with second countable topology and f : X → Y is Lipschitz in a neighborhood within s of every point x ∈ s, then the Hausdorff dimension of the image f '' s is at most the Hausdorff dimension of s.

      theorem dimH_range_le_of_locally_lipschitzOn {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] [SecondCountableTopology X] {f : XY} (hf : ∀ (x : X), ∃ (C : NNReal), snhds x, LipschitzOnWith C f s) :

      If f : X → Y is Lipschitz in a neighborhood of each point x : X, then the Hausdorff dimension of range f is at most the Hausdorff dimension of X.

      theorem AntilipschitzWith.dimH_preimage_le {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] {K : NNReal} {f : XY} (hf : AntilipschitzWith K f) (s : Set Y) :
      theorem AntilipschitzWith.le_dimH_image {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] {K : NNReal} {f : XY} (hf : AntilipschitzWith K f) (s : Set X) :
      dimH s dimH (f '' s)

      Isometries preserve Hausdorff dimension #

      theorem Isometry.dimH_image {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] {f : XY} (hf : Isometry f) (s : Set X) :
      dimH (f '' s) = dimH s
      @[simp]
      theorem IsometryEquiv.dimH_image {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] (e : X ≃ᵢ Y) (s : Set X) :
      dimH (e '' s) = dimH s
      @[simp]
      theorem IsometryEquiv.dimH_preimage {X : Type u_2} {Y : Type u_3} [EMetricSpace X] [EMetricSpace Y] (e : X ≃ᵢ Y) (s : Set Y) :
      dimH (e ⁻¹' s) = dimH s
      @[simp]
      theorem ContinuousLinearEquiv.dimH_image {𝕜 : Type u_4} {E : Type u_5} {F : Type u_6} [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] (e : E ≃L[𝕜] F) (s : Set E) :
      dimH (e '' s) = dimH s
      @[simp]
      theorem ContinuousLinearEquiv.dimH_preimage {𝕜 : Type u_4} {E : Type u_5} {F : Type u_6} [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] (e : E ≃L[𝕜] F) (s : Set F) :
      dimH (e ⁻¹' s) = dimH s

      Hausdorff dimension in a real vector space #

      theorem Real.dimH_ball_pi {ι : Type u_1} [Fintype ι] (x : ι) {r : } (hr : 0 < r) :
      theorem Real.dimH_ball_pi_fin {n : } (x : Fin n) {r : } (hr : 0 < r) :
      dimH (Metric.ball x r) = n
      theorem Real.dimH_of_mem_nhds {E : Type u_4} [NormedAddCommGroup E] [NormedSpace E] [FiniteDimensional E] {x : E} {s : Set E} (h : s nhds x) :

      The Hausdorff dimension of any set in a finite-dimensional real normed space is finite.

      Hausdorff dimension and -smooth maps #

      -smooth maps are locally Lipschitz continuous, hence they do not increase the Hausdorff dimension of sets.

      theorem ContDiffOn.dimH_image_le {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [FiniteDimensional E] [NormedAddCommGroup F] [NormedSpace F] {f : EF} {s t : Set E} (hf : ContDiffOn 1 f s) (hc : Convex s) (ht : t s) :
      dimH (f '' t) dimH t

      Let f be a function defined on a finite dimensional real normed space. If f is -smooth on a convex set s, then the Hausdorff dimension of f '' s is less than or equal to the Hausdorff dimension of s.

      TODO: do we actually need Convex ℝ s?

      The Hausdorff dimension of the range of a -smooth function defined on a finite dimensional real normed space is at most the dimension of its domain as a vector space over .

      theorem ContDiffOn.dense_compl_image_of_dimH_lt_finrank {E : Type u_4} {F : Type u_5} [NormedAddCommGroup E] [NormedSpace E] [FiniteDimensional E] [NormedAddCommGroup F] [NormedSpace F] [FiniteDimensional F] {f : EF} {s t : Set E} (h : ContDiffOn 1 f s) (hc : Convex s) (ht : t s) (htF : dimH t < (Module.finrank F)) :
      Dense (f '' t)

      A particular case of Sard's Theorem. Let f : E → F be a map between finite dimensional real vector spaces. Suppose that f is smooth on a convex set s of Hausdorff dimension strictly less than the dimension of F. Then the complement of the image f '' s is dense in F.

      A particular case of Sard's Theorem. If f is a smooth map from a real vector space to a real vector space F of strictly larger dimension, then the complement of the range of f is dense in F.

      theorem dimH_orthogonalProjection_le {𝕜 : Type u_6} {E : Type u_7} [RCLike 𝕜] [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] (K : Submodule 𝕜 E) [K.HasOrthogonalProjection] (s : Set E) :

      The Hausdorff dimension of the orthogonal projection of a set s onto a subspace K is less than or equal to the Hausdorff dimension of s.