Documentation

Mathlib.Topology.EMetricSpace.Basic

Extended metric spaces #

This file is devoted to the definition and study of EMetricSpaces, i.e., metric spaces in which the distance is allowed to take the value ∞. This extended distance is called edist, and takes values in ℝ≥0∞.

Many definitions and theorems expected on emetric spaces are already introduced on uniform spaces and topological spaces. For example: open and closed sets, compactness, completeness, continuity and uniform continuity.

The class EMetricSpace therefore extends UniformSpace (and TopologicalSpace).

Since a lot of elementary properties don't require eq_of_edist_eq_zero we start setting up the theory of PseudoEMetricSpace, where we don't require edist x y = 0 → x = y and we specialize to EMetricSpace at the end.

theorem uniformity_dist_of_mem_uniformity {α : Type u} {β : Type v} [LinearOrder β] {U : Filter (α × α)} (z : β) (D : ααβ) (H : ∀ (s : Set (α × α)), s U ∃ ε > z, ∀ {a b : α}, D a b < ε(a, b) s) :
U = ⨅ (ε : β), ⨅ (_ : ε > z), Filter.principal {p : α × α | D p.1 p.2 < ε}

Characterizing uniformities associated to a (generalized) distance function D in terms of the elements of the uniformity.

theorem EDist.ext {α : Type u_2} (x : EDist α) (y : EDist α) (edist : EDist.edist = EDist.edist) :
x = y
theorem EDist.ext_iff {α : Type u_2} (x : EDist α) (y : EDist α) :
x = y edist = edist
class EDist (α : Type u_2) :
Type u_2

EDist α means that α is equipped with an extended distance.

Instances
    def uniformSpaceOfEDist {α : Type u} (edist : ααENNReal) (edist_self : ∀ (x : α), edist x x = 0) (edist_comm : ∀ (x y : α), edist x y = edist y x) (edist_triangle : ∀ (x y z : α), edist x z edist x y + edist y z) :

    Creating a uniform space from an extended distance.

    Equations
    Instances For
      class PseudoEMetricSpace (α : Type u) extends EDist :

      Extended (pseudo) metric spaces, with an extended distance edist possibly taking the value ∞

      Each pseudo_emetric space induces a canonical UniformSpace and hence a canonical TopologicalSpace. This is enforced in the type class definition, by extending the UniformSpace structure. When instantiating a PseudoEMetricSpace structure, the uniformity fields are not necessary, they will be filled in by default. There is a default value for the uniformity, that can be substituted in cases of interest, for instance when instantiating a PseudoEMetricSpace structure on a product.

      Continuity of edist is proved in Topology.Instances.ENNReal

      Instances
        theorem PseudoEMetricSpace.ext {α : Type u_2} {m : PseudoEMetricSpace α} {m' : PseudoEMetricSpace α} (h : PseudoEMetricSpace.toEDist = PseudoEMetricSpace.toEDist) :
        m = m'

        Two pseudo emetric space structures with the same edistance function coincide.

        theorem edist_triangle_left {α : Type u} [PseudoEMetricSpace α] (x : α) (y : α) (z : α) :
        edist x y edist z x + edist z y

        Triangle inequality for the extended distance

        theorem edist_triangle_right {α : Type u} [PseudoEMetricSpace α] (x : α) (y : α) (z : α) :
        edist x y edist x z + edist y z
        theorem edist_congr_right {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {z : α} (h : edist x y = 0) :
        edist x z = edist y z
        theorem edist_congr_left {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {z : α} (h : edist x y = 0) :
        edist z x = edist z y
        theorem edist_congr {α : Type u} [PseudoEMetricSpace α] {w : α} {x : α} {y : α} {z : α} (hl : edist w x = 0) (hr : edist y z = 0) :
        edist w y = edist x z
        theorem edist_triangle4 {α : Type u} [PseudoEMetricSpace α] (x : α) (y : α) (z : α) (t : α) :
        edist x t edist x y + edist y z + edist z t
        theorem edist_le_Ico_sum_edist {α : Type u} [PseudoEMetricSpace α] (f : α) {m : } {n : } (h : m n) :
        edist (f m) (f n) Finset.sum (Finset.Ico m n) fun (i : ) => edist (f i) (f (i + 1))

        The triangle (polygon) inequality for sequences of points; Finset.Ico version.

        theorem edist_le_range_sum_edist {α : Type u} [PseudoEMetricSpace α] (f : α) (n : ) :
        edist (f 0) (f n) Finset.sum (Finset.range n) fun (i : ) => edist (f i) (f (i + 1))

        The triangle (polygon) inequality for sequences of points; Finset.range version.

        theorem edist_le_Ico_sum_of_edist_le {α : Type u} [PseudoEMetricSpace α] {f : α} {m : } {n : } (hmn : m n) {d : ENNReal} (hd : ∀ {k : }, m kk < nedist (f k) (f (k + 1)) d k) :
        edist (f m) (f n) Finset.sum (Finset.Ico m n) fun (i : ) => d i

        A version of edist_le_Ico_sum_edist with each intermediate distance replaced with an upper estimate.

        theorem edist_le_range_sum_of_edist_le {α : Type u} [PseudoEMetricSpace α] {f : α} (n : ) {d : ENNReal} (hd : ∀ {k : }, k < nedist (f k) (f (k + 1)) d k) :
        edist (f 0) (f n) Finset.sum (Finset.range n) fun (i : ) => d i

        A version of edist_le_range_sum_edist with each intermediate distance replaced with an upper estimate.

        theorem uniformity_pseudoedist {α : Type u} [PseudoEMetricSpace α] :
        uniformity α = ⨅ (ε : ENNReal), ⨅ (_ : ε > 0), Filter.principal {p : α × α | edist p.1 p.2 < ε}

        Reformulation of the uniform structure in terms of the extended distance

        theorem uniformSpace_edist {α : Type u} [PseudoEMetricSpace α] :
        PseudoEMetricSpace.toUniformSpace = uniformSpaceOfEDist edist
        theorem uniformity_basis_edist {α : Type u} [PseudoEMetricSpace α] :
        Filter.HasBasis (uniformity α) (fun (ε : ENNReal) => 0 < ε) fun (ε : ENNReal) => {p : α × α | edist p.1 p.2 < ε}
        theorem mem_uniformity_edist {α : Type u} [PseudoEMetricSpace α] {s : Set (α × α)} :
        s uniformity α ∃ ε > 0, ∀ {a b : α}, edist a b < ε(a, b) s

        Characterization of the elements of the uniformity in terms of the extended distance

        theorem EMetric.mk_uniformity_basis {α : Type u} [PseudoEMetricSpace α] {β : Type u_2} {p : βProp} {f : βENNReal} (hf₀ : ∀ (x : β), p x0 < f x) (hf : ∀ (ε : ENNReal), 0 < ε∃ (x : β), p x f x ε) :
        Filter.HasBasis (uniformity α) p fun (x : β) => {p : α × α | edist p.1 p.2 < f x}

        Given f : β → ℝ≥0∞, if f sends {i | p i} to a set of positive numbers accumulating to zero, then f i-neighborhoods of the diagonal form a basis of 𝓤 α.

        For specific bases see uniformity_basis_edist, uniformity_basis_edist', uniformity_basis_edist_nnreal, and uniformity_basis_edist_inv_nat.

        theorem EMetric.mk_uniformity_basis_le {α : Type u} [PseudoEMetricSpace α] {β : Type u_2} {p : βProp} {f : βENNReal} (hf₀ : ∀ (x : β), p x0 < f x) (hf : ∀ (ε : ENNReal), 0 < ε∃ (x : β), p x f x ε) :
        Filter.HasBasis (uniformity α) p fun (x : β) => {p : α × α | edist p.1 p.2 f x}

        Given f : β → ℝ≥0∞, if f sends {i | p i} to a set of positive numbers accumulating to zero, then closed f i-neighborhoods of the diagonal form a basis of 𝓤 α.

        For specific bases see uniformity_basis_edist_le and uniformity_basis_edist_le'.

        theorem uniformity_basis_edist_le {α : Type u} [PseudoEMetricSpace α] :
        Filter.HasBasis (uniformity α) (fun (ε : ENNReal) => 0 < ε) fun (ε : ENNReal) => {p : α × α | edist p.1 p.2 ε}
        theorem uniformity_basis_edist' {α : Type u} [PseudoEMetricSpace α] (ε' : ENNReal) (hε' : 0 < ε') :
        Filter.HasBasis (uniformity α) (fun (ε : ENNReal) => ε Set.Ioo 0 ε') fun (ε : ENNReal) => {p : α × α | edist p.1 p.2 < ε}
        theorem uniformity_basis_edist_le' {α : Type u} [PseudoEMetricSpace α] (ε' : ENNReal) (hε' : 0 < ε') :
        Filter.HasBasis (uniformity α) (fun (ε : ENNReal) => ε Set.Ioo 0 ε') fun (ε : ENNReal) => {p : α × α | edist p.1 p.2 ε}
        theorem uniformity_basis_edist_nnreal {α : Type u} [PseudoEMetricSpace α] :
        Filter.HasBasis (uniformity α) (fun (ε : NNReal) => 0 < ε) fun (ε : NNReal) => {p : α × α | edist p.1 p.2 < ε}
        theorem uniformity_basis_edist_nnreal_le {α : Type u} [PseudoEMetricSpace α] :
        Filter.HasBasis (uniformity α) (fun (ε : NNReal) => 0 < ε) fun (ε : NNReal) => {p : α × α | edist p.1 p.2 ε}
        theorem uniformity_basis_edist_inv_nat {α : Type u} [PseudoEMetricSpace α] :
        Filter.HasBasis (uniformity α) (fun (x : ) => True) fun (n : ) => {p : α × α | edist p.1 p.2 < (n)⁻¹}
        theorem uniformity_basis_edist_inv_two_pow {α : Type u} [PseudoEMetricSpace α] :
        Filter.HasBasis (uniformity α) (fun (x : ) => True) fun (n : ) => {p : α × α | edist p.1 p.2 < 2⁻¹ ^ n}
        theorem edist_mem_uniformity {α : Type u} [PseudoEMetricSpace α] {ε : ENNReal} (ε0 : 0 < ε) :
        {p : α × α | edist p.1 p.2 < ε} uniformity α

        Fixed size neighborhoods of the diagonal belong to the uniform structure

        theorem EMetric.uniformContinuousOn_iff {α : Type u} {β : Type v} [PseudoEMetricSpace α] [PseudoEMetricSpace β] {f : αβ} {s : Set α} :
        UniformContinuousOn f s ε > 0, ∃ δ > 0, ∀ {a : α}, a s∀ {b : α}, b sedist a b < δedist (f a) (f b) < ε

        ε-δ characterization of uniform continuity on a set for pseudoemetric spaces

        theorem EMetric.uniformContinuous_iff {α : Type u} {β : Type v} [PseudoEMetricSpace α] [PseudoEMetricSpace β] {f : αβ} :
        UniformContinuous f ε > 0, ∃ δ > 0, ∀ {a b : α}, edist a b < δedist (f a) (f b) < ε

        ε-δ characterization of uniform continuity on pseudoemetric spaces

        theorem EMetric.uniformInducing_iff {α : Type u} {β : Type v} [PseudoEMetricSpace α] [PseudoEMetricSpace β] {f : αβ} :
        UniformInducing f UniformContinuous f δ > 0, ∃ ε > 0, ∀ {a b : α}, edist (f a) (f b) < εedist a b < δ
        theorem EMetric.uniformEmbedding_iff {α : Type u} {β : Type v} [PseudoEMetricSpace α] [PseudoEMetricSpace β] {f : αβ} :
        UniformEmbedding f Function.Injective f UniformContinuous f δ > 0, ∃ ε > 0, ∀ {a b : α}, edist (f a) (f b) < εedist a b < δ

        ε-δ characterization of uniform embeddings on pseudoemetric spaces

        theorem EMetric.controlled_of_uniformEmbedding {α : Type u} {β : Type v} [PseudoEMetricSpace α] [PseudoEMetricSpace β] {f : αβ} (h : UniformEmbedding f) :
        (ε > 0, ∃ δ > 0, ∀ {a b : α}, edist a b < δedist (f a) (f b) < ε) δ > 0, ∃ ε > 0, ∀ {a b : α}, edist (f a) (f b) < εedist a b < δ

        If a map between pseudoemetric spaces is a uniform embedding then the edistance between f x and f y is controlled in terms of the distance between x and y.

        In fact, this lemma holds for a UniformInducing map. TODO: generalize?

        theorem EMetric.cauchy_iff {α : Type u} [PseudoEMetricSpace α] {f : Filter α} :
        Cauchy f f ε > 0, ∃ t ∈ f, xt, yt, edist x y < ε

        ε-δ characterization of Cauchy sequences on pseudoemetric spaces

        theorem EMetric.complete_of_convergent_controlled_sequences {α : Type u} [PseudoEMetricSpace α] (B : ENNReal) (hB : ∀ (n : ), 0 < B n) (H : ∀ (u : α), (∀ (N n m : ), N nN medist (u n) (u m) < B N)∃ (x : α), Filter.Tendsto u Filter.atTop (nhds x)) :

        A very useful criterion to show that a space is complete is to show that all sequences which satisfy a bound of the form edist (u n) (u m) < B N for all n m ≥ N are converging. This is often applied for B N = 2^{-N}, i.e., with a very fast convergence to 0, which makes it possible to use arguments of converging series, while this is impossible to do in general for arbitrary Cauchy sequences.

        theorem EMetric.complete_of_cauchySeq_tendsto {α : Type u} [PseudoEMetricSpace α] :
        (∀ (u : α), CauchySeq u∃ (a : α), Filter.Tendsto u Filter.atTop (nhds a))CompleteSpace α

        A sequentially complete pseudoemetric space is complete.

        theorem EMetric.tendstoLocallyUniformlyOn_iff {α : Type u} {β : Type v} [PseudoEMetricSpace α] {ι : Type u_2} [TopologicalSpace β] {F : ιβα} {f : βα} {p : Filter ι} {s : Set β} :
        TendstoLocallyUniformlyOn F f p s ε > 0, xs, ∃ t ∈ nhdsWithin x s, ∀ᶠ (n : ι) in p, yt, edist (f y) (F n y) < ε

        Expressing locally uniform convergence on a set using edist.

        theorem EMetric.tendstoUniformlyOn_iff {α : Type u} {β : Type v} [PseudoEMetricSpace α] {ι : Type u_2} {F : ιβα} {f : βα} {p : Filter ι} {s : Set β} :
        TendstoUniformlyOn F f p s ε > 0, ∀ᶠ (n : ι) in p, xs, edist (f x) (F n x) < ε

        Expressing uniform convergence on a set using edist.

        theorem EMetric.tendstoLocallyUniformly_iff {α : Type u} {β : Type v} [PseudoEMetricSpace α] {ι : Type u_2} [TopologicalSpace β] {F : ιβα} {f : βα} {p : Filter ι} :
        TendstoLocallyUniformly F f p ε > 0, ∀ (x : β), ∃ t ∈ nhds x, ∀ᶠ (n : ι) in p, yt, edist (f y) (F n y) < ε

        Expressing locally uniform convergence using edist.

        theorem EMetric.tendstoUniformly_iff {α : Type u} {β : Type v} [PseudoEMetricSpace α] {ι : Type u_2} {F : ιβα} {f : βα} {p : Filter ι} :
        TendstoUniformly F f p ε > 0, ∀ᶠ (n : ι) in p, ∀ (x : β), edist (f x) (F n x) < ε

        Expressing uniform convergence using edist.

        Auxiliary function to replace the uniformity on a pseudoemetric space with a uniformity which is equal to the original one, but maybe not defeq. This is useful if one wants to construct a pseudoemetric space with a specified uniformity. See Note [forgetful inheritance] explaining why having definitionally the right uniformity is often important.

        Equations
        Instances For
          def PseudoEMetricSpace.induced {α : Type u_2} {β : Type u_3} (f : αβ) (m : PseudoEMetricSpace β) :

          The extended pseudometric induced by a function taking values in a pseudoemetric space.

          Equations
          Instances For

            Pseudoemetric space instance on subsets of pseudoemetric spaces

            Equations
            theorem Subtype.edist_eq {α : Type u} [PseudoEMetricSpace α] {p : αProp} (x : Subtype p) (y : Subtype p) :
            edist x y = edist x y

            The extended pseudodistance on a subset of a pseudoemetric space is the restriction of the original pseudodistance, by definition

            Pseudoemetric space instance on the additive opposite of a pseudoemetric space.

            Equations

            Pseudoemetric space instance on the multiplicative opposite of a pseudoemetric space.

            Equations
            theorem AddOpposite.edist_op {α : Type u} [PseudoEMetricSpace α] (x : α) (y : α) :
            theorem MulOpposite.edist_op {α : Type u} [PseudoEMetricSpace α] (x : α) (y : α) :
            Equations
            theorem ULift.edist_eq {α : Type u} [PseudoEMetricSpace α] (x : ULift.{u_2, u} α) (y : ULift.{u_2, u} α) :
            edist x y = edist x.down y.down
            @[simp]
            theorem ULift.edist_up_up {α : Type u} [PseudoEMetricSpace α] (x : α) (y : α) :
            edist { down := x } { down := y } = edist x y

            The product of two pseudoemetric spaces, with the max distance, is an extended pseudometric spaces. We make sure that the uniform structure thus constructed is the one corresponding to the product of uniform spaces, to avoid diamond problems.

            Equations
            theorem Prod.edist_eq {α : Type u} {β : Type v} [PseudoEMetricSpace α] [PseudoEMetricSpace β] (x : α × β) (y : α × β) :
            edist x y = max (edist x.1 y.1) (edist x.2 y.2)
            instance instEDistForAll {β : Type v} {π : βType u_2} [Fintype β] [(b : β) → EDist (π b)] :
            EDist ((b : β) → π b)
            Equations
            • instEDistForAll = { edist := fun (f g : (b : β) → π b) => Finset.sup Finset.univ fun (b : β) => edist (f b) (g b) }
            theorem edist_pi_def {β : Type v} {π : βType u_2} [Fintype β] [(b : β) → EDist (π b)] (f : (b : β) → π b) (g : (b : β) → π b) :
            edist f g = Finset.sup Finset.univ fun (b : β) => edist (f b) (g b)
            theorem edist_le_pi_edist {β : Type v} {π : βType u_2} [Fintype β] [(b : β) → EDist (π b)] (f : (b : β) → π b) (g : (b : β) → π b) (b : β) :
            edist (f b) (g b) edist f g
            theorem edist_pi_le_iff {β : Type v} {π : βType u_2} [Fintype β] [(b : β) → EDist (π b)] {f : (b : β) → π b} {g : (b : β) → π b} {d : ENNReal} :
            edist f g d ∀ (b : β), edist (f b) (g b) d
            theorem edist_pi_const_le {α : Type u} {β : Type v} [PseudoEMetricSpace α] [Fintype β] (a : α) (b : α) :
            (edist (fun (x : β) => a) fun (x : β) => b) edist a b
            @[simp]
            theorem edist_pi_const {α : Type u} {β : Type v} [PseudoEMetricSpace α] [Fintype β] [Nonempty β] (a : α) (b : α) :
            (edist (fun (x : β) => a) fun (x : β) => b) = edist a b
            instance pseudoEMetricSpacePi {β : Type v} {π : βType u_2} [Fintype β] [(b : β) → PseudoEMetricSpace (π b)] :
            PseudoEMetricSpace ((b : β) → π b)

            The product of a finite number of pseudoemetric spaces, with the max distance, is still a pseudoemetric space. This construction would also work for infinite products, but it would not give rise to the product topology. Hence, we only formalize it in the good situation of finitely many spaces.

            Equations
            def EMetric.ball {α : Type u} [PseudoEMetricSpace α] (x : α) (ε : ENNReal) :
            Set α

            EMetric.ball x ε is the set of all points y with edist y x < ε

            Equations
            Instances For
              @[simp]
              theorem EMetric.mem_ball {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {ε : ENNReal} :
              y EMetric.ball x ε edist y x < ε
              theorem EMetric.mem_ball' {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {ε : ENNReal} :
              y EMetric.ball x ε edist x y < ε
              def EMetric.closedBall {α : Type u} [PseudoEMetricSpace α] (x : α) (ε : ENNReal) :
              Set α

              EMetric.closedBall x ε is the set of all points y with edist y x ≤ ε

              Equations
              Instances For
                @[simp]
                theorem EMetric.mem_closedBall {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {ε : ENNReal} :
                theorem EMetric.mem_closedBall' {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {ε : ENNReal} :
                @[simp]
                theorem EMetric.closedBall_top {α : Type u} [PseudoEMetricSpace α] (x : α) :
                theorem EMetric.pos_of_mem_ball {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {ε : ENNReal} (hy : y EMetric.ball x ε) :
                0 < ε
                theorem EMetric.mem_ball_self {α : Type u} [PseudoEMetricSpace α] {x : α} {ε : ENNReal} (h : 0 < ε) :
                theorem EMetric.mem_ball_comm {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {ε : ENNReal} :
                theorem EMetric.ball_subset_ball {α : Type u} [PseudoEMetricSpace α] {x : α} {ε₁ : ENNReal} {ε₂ : ENNReal} (h : ε₁ ε₂) :
                EMetric.ball x ε₁ EMetric.ball x ε₂
                theorem EMetric.closedBall_subset_closedBall {α : Type u} [PseudoEMetricSpace α] {x : α} {ε₁ : ENNReal} {ε₂ : ENNReal} (h : ε₁ ε₂) :
                theorem EMetric.ball_disjoint {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {ε₁ : ENNReal} {ε₂ : ENNReal} (h : ε₁ + ε₂ edist x y) :
                Disjoint (EMetric.ball x ε₁) (EMetric.ball y ε₂)
                theorem EMetric.ball_subset {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {ε₁ : ENNReal} {ε₂ : ENNReal} (h : edist x y + ε₁ ε₂) (h' : edist x y ) :
                EMetric.ball x ε₁ EMetric.ball y ε₂
                theorem EMetric.exists_ball_subset_ball {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {ε : ENNReal} (h : y EMetric.ball x ε) :
                ∃ ε' > 0, EMetric.ball y ε' EMetric.ball x ε
                theorem EMetric.ball_eq_empty_iff {α : Type u} [PseudoEMetricSpace α] {x : α} {ε : ENNReal} :
                EMetric.ball x ε = ε = 0

                Relation “two points are at a finite edistance” is an equivalence relation.

                Equations
                • EMetric.edistLtTopSetoid = { r := fun (x y : α) => edist x y < , iseqv := }
                Instances For
                  @[simp]
                  theorem EMetric.ball_zero {α : Type u} [PseudoEMetricSpace α] {x : α} :
                  theorem EMetric.nhds_basis_eball {α : Type u} [PseudoEMetricSpace α] {x : α} :
                  Filter.HasBasis (nhds x) (fun (ε : ENNReal) => 0 < ε) (EMetric.ball x)
                  theorem EMetric.nhdsWithin_basis_eball {α : Type u} [PseudoEMetricSpace α] {x : α} {s : Set α} :
                  Filter.HasBasis (nhdsWithin x s) (fun (ε : ENNReal) => 0 < ε) fun (ε : ENNReal) => EMetric.ball x ε s
                  theorem EMetric.nhdsWithin_basis_closed_eball {α : Type u} [PseudoEMetricSpace α] {x : α} {s : Set α} :
                  Filter.HasBasis (nhdsWithin x s) (fun (ε : ENNReal) => 0 < ε) fun (ε : ENNReal) => EMetric.closedBall x ε s
                  theorem EMetric.nhds_eq {α : Type u} [PseudoEMetricSpace α] {x : α} :
                  nhds x = ⨅ (ε : ENNReal), ⨅ (_ : ε > 0), Filter.principal (EMetric.ball x ε)
                  theorem EMetric.mem_nhds_iff {α : Type u} [PseudoEMetricSpace α] {x : α} {s : Set α} :
                  s nhds x ∃ ε > 0, EMetric.ball x ε s
                  theorem EMetric.mem_nhdsWithin_iff {α : Type u} [PseudoEMetricSpace α] {x : α} {s : Set α} {t : Set α} :
                  s nhdsWithin x t ∃ ε > 0, EMetric.ball x ε t s
                  theorem EMetric.tendsto_nhdsWithin_nhdsWithin {α : Type u} {β : Type v} [PseudoEMetricSpace α] {s : Set α} [PseudoEMetricSpace β] {f : αβ} {t : Set β} {a : α} {b : β} :
                  Filter.Tendsto f (nhdsWithin a s) (nhdsWithin b t) ε > 0, ∃ δ > 0, ∀ ⦃x : α⦄, x sedist x a < δf x t edist (f x) b < ε
                  theorem EMetric.tendsto_nhdsWithin_nhds {α : Type u} {β : Type v} [PseudoEMetricSpace α] {s : Set α} [PseudoEMetricSpace β] {f : αβ} {a : α} {b : β} :
                  Filter.Tendsto f (nhdsWithin a s) (nhds b) ε > 0, ∃ δ > 0, ∀ {x : α}, x sedist x a < δedist (f x) b < ε
                  theorem EMetric.tendsto_nhds_nhds {α : Type u} {β : Type v} [PseudoEMetricSpace α] [PseudoEMetricSpace β] {f : αβ} {a : α} {b : β} :
                  Filter.Tendsto f (nhds a) (nhds b) ε > 0, ∃ δ > 0, ∀ ⦃x : α⦄, edist x a < δedist (f x) b < ε
                  theorem EMetric.isOpen_iff {α : Type u} [PseudoEMetricSpace α] {s : Set α} :
                  IsOpen s xs, ∃ ε > 0, EMetric.ball x ε s
                  theorem EMetric.isOpen_ball {α : Type u} [PseudoEMetricSpace α] {x : α} {ε : ENNReal} :
                  theorem EMetric.ball_mem_nhds {α : Type u} [PseudoEMetricSpace α] (x : α) {ε : ENNReal} (ε0 : 0 < ε) :
                  theorem EMetric.closedBall_mem_nhds {α : Type u} [PseudoEMetricSpace α] (x : α) {ε : ENNReal} (ε0 : 0 < ε) :
                  theorem EMetric.ball_prod_same {α : Type u} {β : Type v} [PseudoEMetricSpace α] [PseudoEMetricSpace β] (x : α) (y : β) (r : ENNReal) :
                  theorem EMetric.mem_closure_iff {α : Type u} [PseudoEMetricSpace α] {x : α} {s : Set α} :
                  x closure s ε > 0, ∃ y ∈ s, edist x y < ε

                  ε-characterization of the closure in pseudoemetric spaces

                  theorem EMetric.tendsto_nhds {α : Type u} {β : Type v} [PseudoEMetricSpace α] {f : Filter β} {u : βα} {a : α} :
                  Filter.Tendsto u f (nhds a) ε > 0, ∀ᶠ (x : β) in f, edist (u x) a < ε
                  theorem EMetric.tendsto_atTop {α : Type u} {β : Type v} [PseudoEMetricSpace α] [Nonempty β] [SemilatticeSup β] {u : βα} {a : α} :
                  Filter.Tendsto u Filter.atTop (nhds a) ε > 0, ∃ (N : β), nN, edist (u n) a < ε
                  theorem EMetric.inseparable_iff {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} :
                  theorem EMetric.cauchySeq_iff {α : Type u} {β : Type v} [PseudoEMetricSpace α] [Nonempty β] [SemilatticeSup β] {u : βα} :
                  CauchySeq u ε > 0, ∃ (N : β), ∀ (m : β), N m∀ (n : β), N nedist (u m) (u n) < ε

                  In a pseudoemetric space, Cauchy sequences are characterized by the fact that, eventually, the pseudoedistance between its elements is arbitrarily small

                  theorem EMetric.cauchySeq_iff' {α : Type u} {β : Type v} [PseudoEMetricSpace α] [Nonempty β] [SemilatticeSup β] {u : βα} :
                  CauchySeq u ε > 0, ∃ (N : β), nN, edist (u n) (u N) < ε

                  A variation around the emetric characterization of Cauchy sequences

                  theorem EMetric.cauchySeq_iff_NNReal {α : Type u} {β : Type v} [PseudoEMetricSpace α] [Nonempty β] [SemilatticeSup β] {u : βα} :
                  CauchySeq u ∀ (ε : NNReal), 0 < ε∃ (N : β), ∀ (n : β), N nedist (u n) (u N) < ε

                  A variation of the emetric characterization of Cauchy sequences that deals with ℝ≥0 upper bounds.

                  theorem EMetric.totallyBounded_iff {α : Type u} [PseudoEMetricSpace α] {s : Set α} :
                  TotallyBounded s ε > 0, ∃ (t : Set α), Set.Finite t s ⋃ y ∈ t, EMetric.ball y ε
                  theorem EMetric.totallyBounded_iff' {α : Type u} [PseudoEMetricSpace α] {s : Set α} :
                  TotallyBounded s ε > 0, ∃ t ⊆ s, Set.Finite t s ⋃ y ∈ t, EMetric.ball y ε
                  theorem EMetric.subset_countable_closure_of_almost_dense_set {α : Type u} [PseudoEMetricSpace α] (s : Set α) (hs : ε > 0, ∃ (t : Set α), Set.Countable t s ⋃ x ∈ t, EMetric.closedBall x ε) :
                  ∃ t ⊆ s, Set.Countable t s closure t

                  For a set s in a pseudo emetric space, if for every ε > 0 there exists a countable set that is ε-dense in s, then there exists a countable subset t ⊆ s that is dense in s.

                  If a set s is separable in a (pseudo extended) metric space, then it admits a countable dense subset. This is not obvious, as the countable set whose closure covers s given by the definition of separability does not need in general to be contained in s.

                  If a set s is separable, then the corresponding subtype is separable in a (pseudo extended) metric space. This is not obvious, as the countable set whose closure covers s does not need in general to be contained in s.

                  A compact set in a pseudo emetric space is separable, i.e., it is a subset of the closure of a countable set.

                  A sigma compact pseudo emetric space has second countable topology.

                  Equations
                  • =
                  theorem EMetric.secondCountable_of_almost_dense_set {α : Type u} [PseudoEMetricSpace α] (hs : ε > 0, ∃ (t : Set α), Set.Countable t ⋃ x ∈ t, EMetric.closedBall x ε = Set.univ) :
                  noncomputable def EMetric.diam {α : Type u} [PseudoEMetricSpace α] (s : Set α) :

                  The diameter of a set in a pseudoemetric space, named EMetric.diam

                  Equations
                  Instances For
                    theorem EMetric.diam_eq_sSup {α : Type u} [PseudoEMetricSpace α] (s : Set α) :
                    theorem EMetric.diam_le_iff {α : Type u} [PseudoEMetricSpace α] {s : Set α} {d : ENNReal} :
                    EMetric.diam s d xs, ys, edist x y d
                    theorem EMetric.diam_image_le_iff {α : Type u} {β : Type v} [PseudoEMetricSpace α] {d : ENNReal} {f : βα} {s : Set β} :
                    EMetric.diam (f '' s) d xs, ys, edist (f x) (f y) d
                    theorem EMetric.edist_le_of_diam_le {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {s : Set α} {d : ENNReal} (hx : x s) (hy : y s) (hd : EMetric.diam s d) :
                    edist x y d
                    theorem EMetric.edist_le_diam_of_mem {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {s : Set α} (hx : x s) (hy : y s) :

                    If two points belong to some set, their edistance is bounded by the diameter of the set

                    theorem EMetric.diam_le {α : Type u} [PseudoEMetricSpace α] {s : Set α} {d : ENNReal} (h : xs, ys, edist x y d) :

                    If the distance between any two points in a set is bounded by some constant, this constant bounds the diameter.

                    The diameter of a subsingleton vanishes.

                    @[simp]

                    The diameter of the empty set vanishes

                    @[simp]
                    theorem EMetric.diam_singleton {α : Type u} [PseudoEMetricSpace α] {x : α} :

                    The diameter of a singleton vanishes

                    @[simp]
                    @[simp]
                    theorem EMetric.diam_one {α : Type u} [PseudoEMetricSpace α] [One α] :
                    theorem EMetric.diam_iUnion_mem_option {α : Type u} [PseudoEMetricSpace α] {ι : Type u_2} (o : Option ι) (s : ιSet α) :
                    EMetric.diam (⋃ i ∈ o, s i) = ⨆ i ∈ o, EMetric.diam (s i)
                    theorem EMetric.diam_insert {α : Type u} [PseudoEMetricSpace α] {x : α} {s : Set α} :
                    EMetric.diam (insert x s) = max (⨆ y ∈ s, edist x y) (EMetric.diam s)
                    theorem EMetric.diam_pair {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} :
                    EMetric.diam {x, y} = edist x y
                    theorem EMetric.diam_triple {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {z : α} :
                    EMetric.diam {x, y, z} = max (max (edist x y) (edist x z)) (edist y z)
                    theorem EMetric.diam_mono {α : Type u} [PseudoEMetricSpace α] {s : Set α} {t : Set α} (h : s t) :

                    The diameter is monotonous with respect to inclusion

                    theorem EMetric.diam_union {α : Type u} [PseudoEMetricSpace α] {x : α} {y : α} {s : Set α} {t : Set α} (xs : x s) (yt : y t) :

                    The diameter of a union is controlled by the diameter of the sets, and the edistance between two points in the sets.

                    theorem EMetric.diam_union' {α : Type u} [PseudoEMetricSpace α] {s : Set α} {t : Set α} (h : Set.Nonempty (s t)) :
                    theorem EMetric.diam_ball {α : Type u} [PseudoEMetricSpace α] {x : α} {r : ENNReal} :
                    theorem EMetric.diam_pi_le_of_le {β : Type v} {π : βType u_2} [Fintype β] [(b : β) → PseudoEMetricSpace (π b)] {s : (b : β) → Set (π b)} {c : ENNReal} (h : ∀ (b : β), EMetric.diam (s b) c) :
                    EMetric.diam (Set.pi Set.univ s) c
                    class EMetricSpace (α : Type u) extends PseudoEMetricSpace :

                    We now define EMetricSpace, extending PseudoEMetricSpace.

                    Instances
                      theorem EMetricSpace.ext {α : Type u_2} {m : EMetricSpace α} {m' : EMetricSpace α} (h : PseudoEMetricSpace.toEDist = PseudoEMetricSpace.toEDist) :
                      m = m'
                      @[simp]
                      theorem edist_eq_zero {γ : Type w} [EMetricSpace γ] {x : γ} {y : γ} :
                      edist x y = 0 x = y

                      Characterize the equality of points by the vanishing of their extended distance

                      @[simp]
                      theorem zero_eq_edist {γ : Type w} [EMetricSpace γ] {x : γ} {y : γ} :
                      0 = edist x y x = y
                      theorem edist_le_zero {γ : Type w} [EMetricSpace γ] {x : γ} {y : γ} :
                      edist x y 0 x = y
                      @[simp]
                      theorem edist_pos {γ : Type w} [EMetricSpace γ] {x : γ} {y : γ} :
                      0 < edist x y x y
                      theorem eq_of_forall_edist_le {γ : Type w} [EMetricSpace γ] {x : γ} {y : γ} (h : ε > 0, edist x y ε) :
                      x = y

                      Two points coincide if their distance is < ε for all positive ε

                      An emetric space is separated

                      Equations
                      • =
                      theorem EMetric.uniformEmbedding_iff' {β : Type v} {γ : Type w} [EMetricSpace γ] [EMetricSpace β] {f : γβ} :
                      UniformEmbedding f (ε > 0, ∃ δ > 0, ∀ {a b : γ}, edist a b < δedist (f a) (f b) < ε) δ > 0, ∃ ε > 0, ∀ {a b : γ}, edist (f a) (f b) < εedist a b < δ

                      A map between emetric spaces is a uniform embedding if and only if the edistance between f x and f y is controlled in terms of the distance between x and y and conversely.

                      @[reducible]

                      If a PseudoEMetricSpace is a T₀ space, then it is an EMetricSpace.

                      Equations
                      Instances For

                        Auxiliary function to replace the uniformity on an emetric space with a uniformity which is equal to the original one, but maybe not defeq. This is useful if one wants to construct an emetric space with a specified uniformity. See Note [forgetful inheritance] explaining why having definitionally the right uniformity is often important.

                        Equations
                        Instances For
                          def EMetricSpace.induced {γ : Type u_2} {β : Type u_3} (f : γβ) (hf : Function.Injective f) (m : EMetricSpace β) :

                          The extended metric induced by an injective function taking values in an emetric space.

                          Equations
                          Instances For
                            instance instEMetricSpaceSubtype {α : Type u_2} {p : αProp} [EMetricSpace α] :

                            EMetric space instance on subsets of emetric spaces

                            Equations

                            EMetric space instance on the additive opposite of an emetric space.

                            Equations

                            EMetric space instance on the multiplicative opposite of an emetric space.

                            Equations
                            Equations
                            instance Prod.emetricSpaceMax {β : Type v} {γ : Type w} [EMetricSpace γ] [EMetricSpace β] :

                            The product of two emetric spaces, with the max distance, is an extended metric spaces. We make sure that the uniform structure thus constructed is the one corresponding to the product of uniform spaces, to avoid diamond problems.

                            Equations
                            theorem uniformity_edist {γ : Type w} [EMetricSpace γ] :
                            uniformity γ = ⨅ (ε : ENNReal), ⨅ (_ : ε > 0), Filter.principal {p : γ × γ | edist p.1 p.2 < ε}

                            Reformulation of the uniform structure in terms of the extended distance

                            instance emetricSpacePi {β : Type v} {π : βType u_2} [Fintype β] [(b : β) → EMetricSpace (π b)] :
                            EMetricSpace ((b : β) → π b)

                            The product of a finite number of emetric spaces, with the max distance, is still an emetric space. This construction would also work for infinite products, but it would not give rise to the product topology. Hence, we only formalize it in the good situation of finitely many spaces.

                            Equations
                            theorem EMetric.countable_closure_of_compact {γ : Type w} [EMetricSpace γ] {s : Set γ} (hs : IsCompact s) :
                            ∃ t ⊆ s, Set.Countable t s = closure t

                            A compact set in an emetric space is separable, i.e., it is the closure of a countable set.

                            theorem EMetric.diam_pos_iff' {γ : Type w} [EMetricSpace γ] {s : Set γ} :
                            0 < EMetric.diam s ∃ x ∈ s, ∃ y ∈ s, x y

                            Separation quotient #

                            Equations

                            Additive, Multiplicative #

                            The distance on those type synonyms is inherited without change.

                            instance instEDistAdditive {X : Type u_1} [EDist X] :
                            Equations
                            • instEDistAdditive = inst
                            Equations
                            • instEDistMultiplicative = inst
                            @[simp]
                            theorem edist_ofMul {X : Type u_1} [EDist X] (a : X) (b : X) :
                            edist (Additive.ofMul a) (Additive.ofMul b) = edist a b
                            @[simp]
                            theorem edist_ofAdd {X : Type u_1} [EDist X] (a : X) (b : X) :
                            edist (Multiplicative.ofAdd a) (Multiplicative.ofAdd b) = edist a b
                            @[simp]
                            theorem edist_toMul {X : Type u_1} [EDist X] (a : Additive X) (b : Additive X) :
                            edist (Additive.toMul a) (Additive.toMul b) = edist a b
                            @[simp]
                            theorem edist_toAdd {X : Type u_1} [EDist X] (a : Multiplicative X) (b : Multiplicative X) :
                            edist (Multiplicative.toAdd a) (Multiplicative.toAdd b) = edist a b
                            Equations
                            • instPseudoEMetricSpaceAdditive = inst
                            Equations
                            • instPseudoEMetricSpaceMultiplicative = inst
                            Equations
                            • instEMetricSpaceAdditive = inst
                            Equations
                            • instEMetricSpaceMultiplicative = inst

                            Order dual #

                            The distance on this type synonym is inherited without change.

                            instance instEDistOrderDual {X : Type u_1} [EDist X] :
                            Equations
                            • instEDistOrderDual = inst
                            @[simp]
                            theorem edist_toDual {X : Type u_1} [EDist X] (a : X) (b : X) :
                            edist (OrderDual.toDual a) (OrderDual.toDual b) = edist a b
                            @[simp]
                            theorem edist_ofDual {X : Type u_1} [EDist X] (a : Xᵒᵈ) (b : Xᵒᵈ) :
                            edist (OrderDual.ofDual a) (OrderDual.ofDual b) = edist a b
                            Equations
                            • instPseudoEMetricSpaceOrderDual = inst
                            Equations
                            • instEMetricSpaceOrderDual = inst