Documentation

Mathlib.Topology.EMetricSpace.Defs

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} [LT β] {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.

class EDist (α : Type u_2) :
Type u_2

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

  • edist : ααENNReal

    Extended distance between two points

Instances
    theorem EDist.ext {α : Type u_2} {x y : EDist α} (edist : edist = edist) :
    x = y
    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 α :

      A pseudo extended metric space is a type endowed with a ℝ≥0∞-valued distance edist satisfying reflexivity edist x x = 0, commutativity edist x y = edist y x, and the triangle inequality edist x z ≤ edist x y + edist y z.

      Note that we do not require edist x y = 0 → x = y. See extended metric spaces (EMetricSpace) for the similar class with that stronger assumption.

      Any pseudo extended metric space is a topological space and a uniform space (see TopologicalSpace, UniformSpace), where the topology and uniformity come from the metric. Note that a T1 pseudo extended metric space is just an extended metric space.

      We make the uniformity/topology part of the data instead of deriving it from the metric. This eg ensures that we do not get a diamond when doing [PseudoEMetricSpace α] [PseudoEMetricSpace β] : TopologicalSpace (α × β): The product metric and product topology agree, but not definitionally so. See Note [forgetful inheritance].

      Instances
        theorem PseudoEMetricSpace.ext {α : Type u_2} {m m' : PseudoEMetricSpace α} (h : toEDist = 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 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 uniformity_basis_edist {α : Type u} [PseudoEMetricSpace α] :
        (uniformity α).HasBasis (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 ε) :
        (uniformity α).HasBasis 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 ε) :
        (uniformity α).HasBasis 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 α] :
        (uniformity α).HasBasis (fun (ε : ENNReal) => 0 < ε) fun (ε : ENNReal) => {p : α × α | edist p.1 p.2 ε}
        theorem uniformity_basis_edist' {α : Type u} [PseudoEMetricSpace α] (ε' : ENNReal) (hε' : 0 < ε') :
        (uniformity α).HasBasis (fun (ε : ENNReal) => ε Set.Ioo 0 ε') fun (ε : ENNReal) => {p : α × α | edist p.1 p.2 < ε}
        theorem uniformity_basis_edist_le' {α : Type u} [PseudoEMetricSpace α] (ε' : ENNReal) (hε' : 0 < ε') :
        (uniformity α).HasBasis (fun (ε : ENNReal) => ε Set.Ioo 0 ε') fun (ε : ENNReal) => {p : α × α | edist p.1 p.2 ε}
        theorem uniformity_basis_edist_nnreal {α : Type u} [PseudoEMetricSpace α] :
        (uniformity α).HasBasis (fun (ε : NNReal) => 0 < ε) fun (ε : NNReal) => {p : α × α | edist p.1 p.2 < ε}
        theorem uniformity_basis_edist_nnreal_le {α : Type u} [PseudoEMetricSpace α] :
        (uniformity α).HasBasis (fun (ε : NNReal) => 0 < ε) fun (ε : NNReal) => {p : α × α | edist p.1 p.2 ε}
        theorem uniformity_basis_edist_inv_nat {α : Type u} [PseudoEMetricSpace α] :
        (uniformity α).HasBasis (fun (x : ) => True) fun (n : ) => {p : α × α | edist p.1 p.2 < (↑n)⁻¹}
        theorem uniformity_basis_edist_inv_two_pow {α : Type u} [PseudoEMetricSpace α] :
        (uniformity α).HasBasis (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

        @[reducible, inline]

        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. See note [reducible non-instances].

        Equations
        Instances For
          @[reducible, inline]
          abbrev PseudoEMetricSpace.induced {α : Type u_2} {β : Type u_3} (f : αβ) (m : PseudoEMetricSpace β) :

          The extended pseudometric induced by a function taking values in a pseudoemetric space. See note [reducible non-instances].

          Equations
          Instances For

            Pseudoemetric space instance on subsets of pseudoemetric spaces

            Equations
            theorem Subtype.edist_eq {α : Type u} [PseudoEMetricSpace α] {p : αProp} (x 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.

            @[simp]
            theorem Subtype.edist_mk_mk {α : Type u} [PseudoEMetricSpace α] {p : αProp} {x y : α} (hx : p x) (hy : p y) :
            edist x, hx y, hy = edist x y

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

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

            Equations

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

            Equations
            theorem MulOpposite.edist_unop {α : Type u} [PseudoEMetricSpace α] (x y : αᵐᵒᵖ) :
            edist (unop x) (unop y) = edist x y
            theorem AddOpposite.edist_unop {α : Type u} [PseudoEMetricSpace α] (x y : αᵃᵒᵖ) :
            edist (unop x) (unop y) = edist x y
            theorem MulOpposite.edist_op {α : Type u} [PseudoEMetricSpace α] (x y : α) :
            edist (op x) (op y) = edist x y
            theorem AddOpposite.edist_op {α : Type u} [PseudoEMetricSpace α] (x y : α) :
            edist (op x) (op y) = edist x y
            theorem ULift.edist_eq {α : Type u} [PseudoEMetricSpace α] (x y : ULift.{u_2, u} α) :
            @[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 = edist x.1 y.1 edist x.2 y.2
            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 ball x ε edist y x < ε
              theorem EMetric.mem_ball' {α : Type u} [PseudoEMetricSpace α] {x y : α} {ε : ENNReal} :
              y 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} :
                y closedBall x ε edist y x ε
                theorem EMetric.mem_closedBall' {α : Type u} [PseudoEMetricSpace α] {x y : α} {ε : ENNReal} :
                y closedBall x ε edist x y ε
                @[simp]
                theorem EMetric.ball_subset_closedBall {α : Type u} [PseudoEMetricSpace α] {x : α} {ε : ENNReal} :
                ball x ε closedBall x ε
                theorem EMetric.pos_of_mem_ball {α : Type u} [PseudoEMetricSpace α] {x y : α} {ε : ENNReal} (hy : y ball x ε) :
                0 < ε
                theorem EMetric.mem_ball_self {α : Type u} [PseudoEMetricSpace α] {x : α} {ε : ENNReal} (h : 0 < ε) :
                x ball x ε
                theorem EMetric.mem_closedBall_self {α : Type u} [PseudoEMetricSpace α] {x : α} {ε : ENNReal} :
                theorem EMetric.mem_ball_comm {α : Type u} [PseudoEMetricSpace α] {x y : α} {ε : ENNReal} :
                x ball y ε y ball x ε
                theorem EMetric.mem_closedBall_comm {α : Type u} [PseudoEMetricSpace α] {x y : α} {ε : ENNReal} :
                theorem EMetric.ball_subset_ball {α : Type u} [PseudoEMetricSpace α] {x : α} {ε₁ ε₂ : ENNReal} (h : ε₁ ε₂) :
                ball x ε₁ ball x ε₂
                theorem EMetric.closedBall_subset_closedBall {α : Type u} [PseudoEMetricSpace α] {x : α} {ε₁ ε₂ : ENNReal} (h : ε₁ ε₂) :
                closedBall x ε₁ closedBall x ε₂
                theorem EMetric.ball_disjoint {α : Type u} [PseudoEMetricSpace α] {x y : α} {ε₁ ε₂ : ENNReal} (h : ε₁ + ε₂ edist x y) :
                Disjoint (ball x ε₁) (ball y ε₂)
                theorem EMetric.ball_subset {α : Type u} [PseudoEMetricSpace α] {x y : α} {ε₁ ε₂ : ENNReal} (h : edist x y + ε₁ ε₂) (h' : edist x y ) :
                ball x ε₁ ball y ε₂
                theorem EMetric.exists_ball_subset_ball {α : Type u} [PseudoEMetricSpace α] {x y : α} {ε : ENNReal} (h : y ball x ε) :
                ε' > 0, ball y ε' ball x ε
                theorem EMetric.ball_eq_empty_iff {α : Type u} [PseudoEMetricSpace α] {x : α} {ε : ENNReal} :
                ball x ε = ε = 0

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

                Equations
                Instances For
                  @[simp]
                  theorem EMetric.ball_zero {α : Type u} [PseudoEMetricSpace α] {x : α} :
                  ball x 0 =
                  theorem EMetric.nhds_basis_eball {α : Type u} [PseudoEMetricSpace α] {x : α} :
                  (nhds x).HasBasis (fun (ε : ENNReal) => 0 < ε) (ball x)
                  theorem EMetric.nhdsWithin_basis_eball {α : Type u} [PseudoEMetricSpace α] {x : α} {s : Set α} :
                  (nhdsWithin x s).HasBasis (fun (ε : ENNReal) => 0 < ε) fun (ε : ENNReal) => ball x ε s
                  theorem EMetric.nhds_basis_closed_eball {α : Type u} [PseudoEMetricSpace α] {x : α} :
                  (nhds x).HasBasis (fun (ε : ENNReal) => 0 < ε) (closedBall x)
                  theorem EMetric.nhdsWithin_basis_closed_eball {α : Type u} [PseudoEMetricSpace α] {x : α} {s : Set α} :
                  (nhdsWithin x s).HasBasis (fun (ε : ENNReal) => 0 < ε) fun (ε : ENNReal) => closedBall x ε s
                  theorem EMetric.nhds_eq {α : Type u} [PseudoEMetricSpace α] {x : α} :
                  nhds x = ⨅ (ε : ENNReal), ⨅ (_ : ε > 0), Filter.principal (ball x ε)
                  theorem EMetric.mem_nhds_iff {α : Type u} [PseudoEMetricSpace α] {x : α} {s : Set α} :
                  s nhds x ε > 0, ball x ε s
                  theorem EMetric.mem_nhdsWithin_iff {α : Type u} [PseudoEMetricSpace α] {x : α} {s t : Set α} :
                  s nhdsWithin x t ε > 0, 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, ball x ε s
                  @[simp]
                  theorem EMetric.isOpen_ball {α : Type u} [PseudoEMetricSpace α] {x : α} {ε : ENNReal} :
                  IsOpen (ball x ε)
                  theorem EMetric.ball_mem_nhds {α : Type u} [PseudoEMetricSpace α] (x : α) {ε : ENNReal} (ε0 : 0 < ε) :
                  ball x ε nhds x
                  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) :
                  ball x r ×ˢ ball y r = ball (x, y) r
                  theorem EMetric.closedBall_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, ys, 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.subset_countable_closure_of_almost_dense_set {α : Type u} [PseudoEMetricSpace α] (s : Set α) (hs : ε > 0, ∃ (t : Set α), t.Countable s xt, closedBall x ε) :
                  ts, t.Countable 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.

                  class EMetricSpace (α : Type u) extends PseudoEMetricSpace α :

                  An extended metric space is a type endowed with a ℝ≥0∞-valued distance edist satisfying edist x y = 0 ↔ x = y, commutativity edist x y = edist y x, and the triangle inequality edist x z ≤ edist x y + edist y z.

                  See pseudo extended metric spaces (PseudoEMetricSpace) for the similar class with the edist x y = 0 ↔ x = y assumption weakened to edist x x = 0.

                  Any extended metric space is a T1 topological space and a uniform space (see TopologicalSpace, T1Space, UniformSpace), where the topology and uniformity come from the metric.

                  We make the uniformity/topology part of the data instead of deriving it from the metric. This eg ensures that we do not get a diamond when doing [EMetricSpace α] [EMetricSpace β] : TopologicalSpace (α × β): The product metric and product topology agree, but not definitionally so. See Note [forgetful inheritance].

                  Instances
                    @[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
                    @[simp]
                    theorem EMetric.closedBall_zero {γ : Type w} [EMetricSpace γ] (x : γ) :
                    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 ε

                    @[reducible, inline]

                    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. See note [reducible non-instances].

                    Equations
                    Instances For
                      @[reducible, inline]
                      abbrev 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. See Note [reducible non-instances].

                      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 multiplicative opposite of an emetric space.

                        Equations

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

                        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

                        Additive, Multiplicative #

                        The distance on those type synonyms is inherited without change.

                        instance instEDistAdditive {X : Type u_1} [EDist X] :
                        Equations
                        @[simp]
                        theorem edist_ofMul {X : Type u_1} [EDist X] (a b : X) :
                        @[simp]
                        theorem edist_ofAdd {X : Type u_1} [EDist X] (a b : X) :
                        @[simp]
                        theorem edist_toMul {X : Type u_1} [EDist X] (a b : Additive X) :
                        @[simp]

                        Order dual #

                        The distance on this type synonym is inherited without change.

                        instance instEDistOrderDual {X : Type u_1} [EDist X] :
                        Equations
                        @[simp]
                        theorem edist_toDual {X : Type u_1} [EDist X] (a b : X) :
                        @[simp]
                        theorem edist_ofDual {X : Type u_1} [EDist X] (a b : Xᵒᵈ) :