Documentation

Mathlib.Analysis.NormedSpace.ProdLp

L^p distance on products of two metric spaces #

Given two metric spaces, one can put the max distance on their product, but there is also a whole family of natural distances, indexed by a parameter p : ℝ≥0∞, that also induce the product topology. We define them in this file. For 0 < p < ∞, the distance on α × β is given by $$ d(x, y) = \left(d(x_1, y_1)^p + d(x_2, y_2)^p\right)^{1/p}. $$ For p = ∞ the distance is the supremum of the distances and p = 0 the distance is the cardinality of the elements that are not equal.

We give instances of this construction for emetric spaces, metric spaces, normed groups and normed spaces.

To avoid conflicting instances, all these are defined on a copy of the original Prod-type, named WithLp p (α × β). The assumption [Fact (1 ≤ p)] is required for the metric and normed space instances.

We ensure that the topology, bornology and uniform structure on WithLp p (α × β) are (defeq to) the product topology, product bornology and product uniformity, to be able to use freely continuity statements for the coordinate functions, for instance.

Implementation notes #

This files is a straight-forward adaption of Mathlib.Analysis.NormedSpace.PiLp.

@[simp]
theorem WithLp.zero_fst {p : ENNReal} {α : Type u_2} {β : Type u_3} [AddCommGroup α] [AddCommGroup β] :
0.1 = 0
@[simp]
theorem WithLp.zero_snd {p : ENNReal} {α : Type u_2} {β : Type u_3} [AddCommGroup α] [AddCommGroup β] :
0.2 = 0
@[simp]
theorem WithLp.add_fst {p : ENNReal} {α : Type u_2} {β : Type u_3} [AddCommGroup α] [AddCommGroup β] (x : WithLp p (α × β)) (y : WithLp p (α × β)) :
(x + y).1 = x.1 + y.1
@[simp]
theorem WithLp.add_snd {p : ENNReal} {α : Type u_2} {β : Type u_3} [AddCommGroup α] [AddCommGroup β] (x : WithLp p (α × β)) (y : WithLp p (α × β)) :
(x + y).2 = x.2 + y.2
@[simp]
theorem WithLp.sub_fst {p : ENNReal} {α : Type u_2} {β : Type u_3} [AddCommGroup α] [AddCommGroup β] (x : WithLp p (α × β)) (y : WithLp p (α × β)) :
(x - y).1 = x.1 - y.1
@[simp]
theorem WithLp.sub_snd {p : ENNReal} {α : Type u_2} {β : Type u_3} [AddCommGroup α] [AddCommGroup β] (x : WithLp p (α × β)) (y : WithLp p (α × β)) :
(x - y).2 = x.2 - y.2
@[simp]
theorem WithLp.neg_fst {p : ENNReal} {α : Type u_2} {β : Type u_3} [AddCommGroup α] [AddCommGroup β] (x : WithLp p (α × β)) :
(-x).1 = -x.1
@[simp]
theorem WithLp.neg_snd {p : ENNReal} {α : Type u_2} {β : Type u_3} [AddCommGroup α] [AddCommGroup β] (x : WithLp p (α × β)) :
(-x).2 = -x.2
@[simp]
theorem WithLp.smul_fst {p : ENNReal} {𝕜 : Type u_1} {α : Type u_2} {β : Type u_3} [Semiring 𝕜] [AddCommGroup α] [AddCommGroup β] (x : WithLp p (α × β)) (c : 𝕜) [Module 𝕜 α] [Module 𝕜 β] :
(c x).1 = c x.1
@[simp]
theorem WithLp.smul_snd {p : ENNReal} {𝕜 : Type u_1} {α : Type u_2} {β : Type u_3} [Semiring 𝕜] [AddCommGroup α] [AddCommGroup β] (x : WithLp p (α × β)) (c : 𝕜) [Module 𝕜 α] [Module 𝕜 β] :
(c x).2 = c x.2

Note that the unapplied versions of these lemmas are deliberately omitted, as they break the use of the type synonym.

@[simp]
theorem WithLp.equiv_fst {p : ENNReal} {α : Type u_2} {β : Type u_3} (x : WithLp p (α × β)) :
((WithLp.equiv p (α × β)) x).1 = x.1
@[simp]
theorem WithLp.equiv_snd {p : ENNReal} {α : Type u_2} {β : Type u_3} (x : WithLp p (α × β)) :
((WithLp.equiv p (α × β)) x).2 = x.2
@[simp]
theorem WithLp.equiv_symm_fst {p : ENNReal} {α : Type u_2} {β : Type u_3} (x : α × β) :
((WithLp.equiv p (α × β)).symm x).1 = x.1
@[simp]
theorem WithLp.equiv_symm_snd {p : ENNReal} {α : Type u_2} {β : Type u_3} (x : α × β) :
((WithLp.equiv p (α × β)).symm x).2 = x.2

Definition of edist, dist and norm on WithLp p (α × β) #

In this section we define the edist, dist and norm functions on WithLp p (α × β) without assuming [Fact (1 ≤ p)] or metric properties of the spaces α and β. This allows us to provide the rewrite lemmas for each of three cases p = 0, p = ∞ and 0 < p.toReal.

instance WithLp.instProdEDist (p : ENNReal) (α : Type u_2) (β : Type u_3) [EDist α] [EDist β] :
EDist (WithLp p (α × β))

Endowing the space WithLp p (α × β) with the L^p edistance. We register this instance separate from WithLp.instProdPseudoEMetric since the latter requires the type class hypothesis [Fact (1 ≤ p)] in order to prove the triangle inequality.

Registering this separately allows for a future emetric-like structure on WithLp p (α × β) for p < 1 satisfying a relaxed triangle inequality. The terminology for this varies throughout the literature, but it is sometimes called a quasi-metric or semi-metric.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem WithLp.prod_edist_eq_card {α : Type u_2} {β : Type u_3} [EDist α] [EDist β] (f : WithLp 0 (α × β)) (g : WithLp 0 (α × β)) :
edist f g = (if edist f.1 g.1 = 0 then 0 else 1) + if edist f.2 g.2 = 0 then 0 else 1
theorem WithLp.prod_edist_eq_add {p : ENNReal} {α : Type u_2} {β : Type u_3} [EDist α] [EDist β] (hp : 0 < p.toReal) (f : WithLp p (α × β)) (g : WithLp p (α × β)) :
edist f g = (edist f.1 g.1 ^ p.toReal + edist f.2 g.2 ^ p.toReal) ^ (1 / p.toReal)
theorem WithLp.prod_edist_eq_sup {α : Type u_2} {β : Type u_3} [EDist α] [EDist β] (f : WithLp (α × β)) (g : WithLp (α × β)) :
edist f g = edist f.1 g.1 edist f.2 g.2
theorem WithLp.prod_edist_self (p : ENNReal) {α : Type u_2} {β : Type u_3} [PseudoEMetricSpace α] [PseudoEMetricSpace β] (f : WithLp p (α × β)) :
edist f f = 0

The distance from one point to itself is always zero.

This holds independent of p and does not require [Fact (1 ≤ p)]. We keep it separate from WithLp.instProdPseudoEMetricSpace so it can be used also for p < 1.

theorem WithLp.prod_edist_comm (p : ENNReal) {α : Type u_2} {β : Type u_3} [PseudoEMetricSpace α] [PseudoEMetricSpace β] (f : WithLp p (α × β)) (g : WithLp p (α × β)) :
edist f g = edist g f

The distance is symmetric.

This holds independent of p and does not require [Fact (1 ≤ p)]. We keep it separate from WithLp.instProdPseudoEMetricSpace so it can be used also for p < 1.

instance WithLp.instProdDist (p : ENNReal) (α : Type u_2) (β : Type u_3) [Dist α] [Dist β] :
Dist (WithLp p (α × β))

Endowing the space WithLp p (α × β) with the L^p distance. We register this instance separate from WithLp.instProdPseudoMetricSpace since the latter requires the type class hypothesis [Fact (1 ≤ p)] in order to prove the triangle inequality.

Registering this separately allows for a future metric-like structure on WithLp p (α × β) for p < 1 satisfying a relaxed triangle inequality. The terminology for this varies throughout the literature, but it is sometimes called a quasi-metric or semi-metric.

Equations
  • One or more equations did not get rendered due to their size.
theorem WithLp.prod_dist_eq_card {α : Type u_2} {β : Type u_3} [Dist α] [Dist β] (f : WithLp 0 (α × β)) (g : WithLp 0 (α × β)) :
dist f g = (if dist f.1 g.1 = 0 then 0 else 1) + if dist f.2 g.2 = 0 then 0 else 1
theorem WithLp.prod_dist_eq_add {p : ENNReal} {α : Type u_2} {β : Type u_3} [Dist α] [Dist β] (hp : 0 < p.toReal) (f : WithLp p (α × β)) (g : WithLp p (α × β)) :
dist f g = (dist f.1 g.1 ^ p.toReal + dist f.2 g.2 ^ p.toReal) ^ (1 / p.toReal)
theorem WithLp.prod_dist_eq_sup {α : Type u_2} {β : Type u_3} [Dist α] [Dist β] (f : WithLp (α × β)) (g : WithLp (α × β)) :
dist f g = dist f.1 g.1 dist f.2 g.2
instance WithLp.instProdNorm (p : ENNReal) (α : Type u_2) (β : Type u_3) [Norm α] [Norm β] :
Norm (WithLp p (α × β))

Endowing the space WithLp p (α × β) with the L^p norm. We register this instance separate from WithLp.instProdSeminormedAddCommGroup since the latter requires the type class hypothesis [Fact (1 ≤ p)] in order to prove the triangle inequality.

Registering this separately allows for a future norm-like structure on WithLp p (α × β) for p < 1 satisfying a relaxed triangle inequality. These are called quasi-norms.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem WithLp.prod_norm_eq_card {α : Type u_2} {β : Type u_3} [Norm α] [Norm β] (f : WithLp 0 (α × β)) :
f = (if f.1 = 0 then 0 else 1) + if f.2 = 0 then 0 else 1
theorem WithLp.prod_norm_eq_sup {α : Type u_2} {β : Type u_3} [Norm α] [Norm β] (f : WithLp (α × β)) :
theorem WithLp.prod_norm_eq_add {p : ENNReal} {α : Type u_2} {β : Type u_3} [Norm α] [Norm β] (hp : 0 < p.toReal) (f : WithLp p (α × β)) :
f = (f.1 ^ p.toReal + f.2 ^ p.toReal) ^ (1 / p.toReal)

The uniformity on finite L^p products is the product uniformity #

In this section, we put the L^p edistance on WithLp p (α × β), and we check that the uniformity coming from this edistance coincides with the product uniformity, by showing that the canonical map to the Prod type (with the L^∞ distance) is a uniform embedding, as it is both Lipschitz and antiLipschitz.

We only register this emetric space structure as a temporary instance, as the true instance (to be registered later) will have as uniformity exactly the product uniformity, instead of the one coming from the edistance (which is equal to it, but not defeq). See Note [forgetful inheritance] explaining why having definitionally the right uniformity is often important.

def WithLp.prodPseudoEMetricAux (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [PseudoEMetricSpace α] [PseudoEMetricSpace β] :

Endowing the space WithLp p (α × β) with the L^p pseudoemetric structure. This definition is not satisfactory, as it does not register the fact that the topology and the uniform structure coincide with the product one. Therefore, we do not register it as an instance. Using this as a temporary pseudoemetric space instance, we will show that the uniform structure is equal (but not defeq) to the product one, and then register an instance in which we replace the uniform structure by the product one using this pseudoemetric space and PseudoEMetricSpace.replaceUniformity.

Equations
Instances For
    theorem WithLp.prod_sup_edist_ne_top_aux {α : Type u_2} {β : Type u_3} [PseudoMetricSpace α] [PseudoMetricSpace β] (f : WithLp (α × β)) (g : WithLp (α × β)) :
    edist f.1 g.1 edist f.2 g.2

    An auxiliary lemma used twice in the proof of WithLp.prodPseudoMetricAux below. Not intended for use outside this file.

    @[reducible]
    def WithLp.prodPseudoMetricAux (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [PseudoMetricSpace α] [PseudoMetricSpace β] :

    Endowing the space WithLp p (α × β) with the L^p pseudometric structure. This definition is not satisfactory, as it does not register the fact that the topology, the uniform structure, and the bornology coincide with the product ones. Therefore, we do not register it as an instance. Using this as a temporary pseudoemetric space instance, we will show that the uniform structure is equal (but not defeq) to the product one, and then register an instance in which we replace the uniform structure and the bornology by the product ones using this pseudometric space, PseudoMetricSpace.replaceUniformity, and PseudoMetricSpace.replaceBornology.

    See note [reducible non-instances]

    Equations
    Instances For
      theorem WithLp.prod_lipschitzWith_equiv_aux (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [PseudoEMetricSpace α] [PseudoEMetricSpace β] :
      LipschitzWith 1 (WithLp.equiv p (α × β))
      theorem WithLp.prod_antilipschitzWith_equiv_aux (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [PseudoEMetricSpace α] [PseudoEMetricSpace β] :
      AntilipschitzWith (2 ^ (1 / p).toReal) (WithLp.equiv p (α × β))
      theorem WithLp.prod_aux_uniformity_eq (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [PseudoEMetricSpace α] [PseudoEMetricSpace β] :
      uniformity (WithLp p (α × β)) = uniformity (α × β)
      theorem WithLp.prod_aux_cobounded_eq (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [PseudoMetricSpace α] [PseudoMetricSpace β] :

      Instances on L^p products #

      instance WithLp.instProdTopologicalSpace (p : ENNReal) (α : Type u_2) (β : Type u_3) [TopologicalSpace α] [TopologicalSpace β] :
      Equations
      theorem WithLp.prod_continuous_equiv (p : ENNReal) (α : Type u_2) (β : Type u_3) [TopologicalSpace α] [TopologicalSpace β] :
      Continuous (WithLp.equiv p (α × β))
      theorem WithLp.prod_continuous_equiv_symm (p : ENNReal) (α : Type u_2) (β : Type u_3) [TopologicalSpace α] [TopologicalSpace β] :
      Continuous (WithLp.equiv p (α × β)).symm
      instance WithLp.instProdT0Space (p : ENNReal) (α : Type u_2) (β : Type u_3) [TopologicalSpace α] [TopologicalSpace β] [T0Space α] [T0Space β] :
      T0Space (WithLp p (α × β))
      Equations
      • =
      instance WithLp.instProdUniformSpace (p : ENNReal) (α : Type u_2) (β : Type u_3) [UniformSpace α] [UniformSpace β] :
      UniformSpace (WithLp p (α × β))
      Equations
      instance WithLp.instProdCompleteSpace (p : ENNReal) (α : Type u_2) (β : Type u_3) [UniformSpace α] [UniformSpace β] [CompleteSpace α] [CompleteSpace β] :
      CompleteSpace (WithLp p (α × β))
      Equations
      • =
      instance WithLp.instProdBornology (p : ENNReal) (α : Type u_2) (β : Type u_3) [Bornology α] [Bornology β] :
      Bornology (WithLp p (α × β))
      Equations
      @[simp]
      theorem WithLp.prodContinuousLinearEquiv_apply (p : ENNReal) (𝕜 : Type u_1) (α : Type u_2) (β : Type u_3) [TopologicalSpace α] [TopologicalSpace β] [Semiring 𝕜] [AddCommGroup α] [AddCommGroup β] [Module 𝕜 α] [Module 𝕜 β] :
      (WithLp.prodContinuousLinearEquiv p 𝕜 α β) = (WithLp.equiv p (α × β))
      @[simp]
      theorem WithLp.prodContinuousLinearEquiv_symm_apply (p : ENNReal) (𝕜 : Type u_1) (α : Type u_2) (β : Type u_3) [TopologicalSpace α] [TopologicalSpace β] [Semiring 𝕜] [AddCommGroup α] [AddCommGroup β] [Module 𝕜 α] [Module 𝕜 β] :
      (WithLp.prodContinuousLinearEquiv p 𝕜 α β).symm = (WithLp.equiv p (α × β)).symm
      def WithLp.prodContinuousLinearEquiv (p : ENNReal) (𝕜 : Type u_1) (α : Type u_2) (β : Type u_3) [TopologicalSpace α] [TopologicalSpace β] [Semiring 𝕜] [AddCommGroup α] [AddCommGroup β] [Module 𝕜 α] [Module 𝕜 β] :
      WithLp p (α × β) ≃L[𝕜] α × β

      WithLp.equiv as a continuous linear equivalence.

      Equations
      Instances For

        Throughout the rest of the file, we assume 1 ≤ p

        instance WithLp.instProdPseudoEMetricSpace (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [PseudoEMetricSpace α] [PseudoEMetricSpace β] :

        PseudoEMetricSpace instance on the product of two pseudoemetric spaces, using the L^p pseudoedistance, and having as uniformity the product uniformity.

        Equations
        instance WithLp.instProdEMetricSpace (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [EMetricSpace α] [EMetricSpace β] :
        EMetricSpace (WithLp p (α × β))

        EMetricSpace instance on the product of two emetric spaces, using the L^p edistance, and having as uniformity the product uniformity.

        Equations
        instance WithLp.instProdPseudoMetricSpace (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [PseudoMetricSpace α] [PseudoMetricSpace β] :

        PseudoMetricSpace instance on the product of two pseudometric spaces, using the L^p distance, and having as uniformity the product uniformity.

        Equations
        instance WithLp.instProdMetricSpace (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [MetricSpace α] [MetricSpace β] :
        MetricSpace (WithLp p (α × β))

        MetricSpace instance on the product of two metric spaces, using the L^p distance, and having as uniformity the product uniformity.

        Equations
        theorem WithLp.prod_nndist_eq_add {p : ENNReal} {α : Type u_2} {β : Type u_3} [hp : Fact (1 p)] [PseudoMetricSpace α] [PseudoMetricSpace β] (hp : p ) (x : WithLp p (α × β)) (y : WithLp p (α × β)) :
        nndist x y = (nndist x.1 y.1 ^ p.toReal + nndist x.2 y.2 ^ p.toReal) ^ (1 / p.toReal)
        theorem WithLp.prod_nndist_eq_sup {α : Type u_2} {β : Type u_3} [PseudoMetricSpace α] [PseudoMetricSpace β] (x : WithLp (α × β)) (y : WithLp (α × β)) :
        nndist x y = nndist x.1 y.1 nndist x.2 y.2
        theorem WithLp.prod_lipschitzWith_equiv (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [PseudoEMetricSpace α] [PseudoEMetricSpace β] :
        LipschitzWith 1 (WithLp.equiv p (α × β))
        theorem WithLp.prod_antilipschitzWith_equiv (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [PseudoEMetricSpace α] [PseudoEMetricSpace β] :
        AntilipschitzWith (2 ^ (1 / p).toReal) (WithLp.equiv p (α × β))

        Seminormed group instance on the product of two normed groups, using the L^p norm.

        Equations
        instance WithLp.instProdNormedAddCommGroup (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [NormedAddCommGroup α] [NormedAddCommGroup β] :

        normed group instance on the product of two normed groups, using the L^p norm.

        Equations
        theorem WithLp.prod_norm_eq_of_nat {p : ENNReal} {α : Type u_2} {β : Type u_3} [hp : Fact (1 p)] [Norm α] [Norm β] (n : ) (h : p = n) (f : WithLp p (α × β)) :
        f = (f.1 ^ n + f.2 ^ n) ^ (1 / n)
        theorem WithLp.prod_nnnorm_eq_add {p : ENNReal} {α : Type u_2} {β : Type u_3} [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (hp : p ) (f : WithLp p (α × β)) :
        f‖₊ = (f.1‖₊ ^ p.toReal + f.2‖₊ ^ p.toReal) ^ (1 / p.toReal)
        @[simp]
        @[simp]
        theorem WithLp.prod_nnnorm_equiv_symm {α : Type u_2} {β : Type u_3} [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (f : α × β) :
        @[simp]
        theorem WithLp.prod_norm_equiv {α : Type u_2} {β : Type u_3} [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (f : WithLp (α × β)) :
        @[simp]
        theorem WithLp.prod_norm_equiv_symm {α : Type u_2} {β : Type u_3} [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (f : α × β) :
        (WithLp.equiv (α × β)).symm f = f
        theorem WithLp.prod_norm_eq_of_L2 {α : Type u_2} {β : Type u_3} [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (x : WithLp 2 (α × β)) :
        x = (x.1 ^ 2 + x.2 ^ 2).sqrt
        theorem WithLp.prod_nnnorm_eq_of_L2 {α : Type u_2} {β : Type u_3} [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (x : WithLp 2 (α × β)) :
        x‖₊ = NNReal.sqrt (x.1‖₊ ^ 2 + x.2‖₊ ^ 2)
        theorem WithLp.prod_norm_sq_eq_of_L2 {α : Type u_2} {β : Type u_3} [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (x : WithLp 2 (α × β)) :
        x ^ 2 = x.1 ^ 2 + x.2 ^ 2
        theorem WithLp.prod_dist_eq_of_L2 {α : Type u_2} {β : Type u_3} [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (x : WithLp 2 (α × β)) (y : WithLp 2 (α × β)) :
        dist x y = (dist x.1 y.1 ^ 2 + dist x.2 y.2 ^ 2).sqrt
        theorem WithLp.prod_nndist_eq_of_L2 {α : Type u_2} {β : Type u_3} [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (x : WithLp 2 (α × β)) (y : WithLp 2 (α × β)) :
        nndist x y = NNReal.sqrt (nndist x.1 y.1 ^ 2 + nndist x.2 y.2 ^ 2)
        theorem WithLp.prod_edist_eq_of_L2 {α : Type u_2} {β : Type u_3} [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (x : WithLp 2 (α × β)) (y : WithLp 2 (α × β)) :
        edist x y = (edist x.1 y.1 ^ 2 + edist x.2 y.2 ^ 2) ^ (1 / 2)
        @[simp]
        theorem WithLp.nnnorm_equiv_symm_fst (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (x : α) :
        (WithLp.equiv p (α × β)).symm (x, 0)‖₊ = x‖₊
        @[simp]
        theorem WithLp.nnnorm_equiv_symm_snd (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (y : β) :
        (WithLp.equiv p (α × β)).symm (0, y)‖₊ = y‖₊
        @[simp]
        theorem WithLp.norm_equiv_symm_fst (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (x : α) :
        (WithLp.equiv p (α × β)).symm (x, 0) = x
        @[simp]
        theorem WithLp.norm_equiv_symm_snd (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (y : β) :
        (WithLp.equiv p (α × β)).symm (0, y) = y
        @[simp]
        theorem WithLp.nndist_equiv_symm_fst (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (x₁ : α) (x₂ : α) :
        nndist ((WithLp.equiv p (α × β)).symm (x₁, 0)) ((WithLp.equiv p (α × β)).symm (x₂, 0)) = nndist x₁ x₂
        @[simp]
        theorem WithLp.nndist_equiv_symm_snd (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (y₁ : β) (y₂ : β) :
        nndist ((WithLp.equiv p (α × β)).symm (0, y₁)) ((WithLp.equiv p (α × β)).symm (0, y₂)) = nndist y₁ y₂
        @[simp]
        theorem WithLp.dist_equiv_symm_fst (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (x₁ : α) (x₂ : α) :
        dist ((WithLp.equiv p (α × β)).symm (x₁, 0)) ((WithLp.equiv p (α × β)).symm (x₂, 0)) = dist x₁ x₂
        @[simp]
        theorem WithLp.dist_equiv_symm_snd (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (y₁ : β) (y₂ : β) :
        dist ((WithLp.equiv p (α × β)).symm (0, y₁)) ((WithLp.equiv p (α × β)).symm (0, y₂)) = dist y₁ y₂
        @[simp]
        theorem WithLp.edist_equiv_symm_fst (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (x₁ : α) (x₂ : α) :
        edist ((WithLp.equiv p (α × β)).symm (x₁, 0)) ((WithLp.equiv p (α × β)).symm (x₂, 0)) = edist x₁ x₂
        @[simp]
        theorem WithLp.edist_equiv_symm_snd (p : ENNReal) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] (y₁ : β) (y₂ : β) :
        edist ((WithLp.equiv p (α × β)).symm (0, y₁)) ((WithLp.equiv p (α × β)).symm (0, y₂)) = edist y₁ y₂
        instance WithLp.instProdBoundedSMul (p : ENNReal) (𝕜 : Type u_1) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] [SeminormedRing 𝕜] [Module 𝕜 α] [Module 𝕜 β] [BoundedSMul 𝕜 α] [BoundedSMul 𝕜 β] :
        BoundedSMul 𝕜 (WithLp p (α × β))
        Equations
        • =
        def WithLp.prodEquivₗᵢ {𝕜 : Type u_1} {α : Type u_2} {β : Type u_3} [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] [SeminormedRing 𝕜] [Module 𝕜 α] [Module 𝕜 β] :
        WithLp (α × β) ≃ₗᵢ[𝕜] α × β

        The canonical map WithLp.equiv between WithLp ∞ (α × β) and α × β as a linear isometric equivalence.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          instance WithLp.instProdNormedSpace (p : ENNReal) (𝕜 : Type u_1) (α : Type u_2) (β : Type u_3) [hp : Fact (1 p)] [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] [NormedField 𝕜] [NormedSpace 𝕜 α] [NormedSpace 𝕜 β] :
          NormedSpace 𝕜 (WithLp p (α × β))

          The product of two normed spaces is a normed space, with the L^p norm.

          Equations