mathlib3 documentation

analysis.normed_space.pi_Lp

L^p distance on finite products of metric spaces #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

Given finitely many 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 Π i, α i is given by $$ d(x, y) = \left(\sum d(x_i, y_i)^p\right)^{1/p}. $$, whereas for p = 0 it is the cardinality of the set ${ i | x_i ≠ y_i}$. For p = ∞ the distance is the supremum of the distances.

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 Π-type, named pi_Lp p α. The assumpion [fact (1 ≤ p)] is required for the metric and normed space instances.

We ensure that the topology, bornology and uniform structure on pi_Lp 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 #

We only deal with the L^p distance on a product of finitely many metric spaces, which may be distinct. A closely related construction is lp, the L^p norm on a product of (possibly infinitely many) normed spaces, where the norm is $$ \left(\sum ‖f (x)‖^p \right)^{1/p}. $$ However, the topology induced by this construction is not the product topology, and some functions have infinite L^p norm. These subtleties are not present in the case of finitely many metric spaces, hence it is worth devoting a file to this specific case which is particularly well behaved.

Another related construction is measure_theory.Lp, the L^p norm on the space of functions from a measure space to a normed space, where the norm is $$ \left(\int ‖f (x)‖^p dμ\right)^{1/p}. $$ This has all the same subtleties as lp, and the further subtlety that this only defines a seminorm (as almost everywhere zero functions have zero L^p norm). The construction pi_Lp corresponds to the special case of measure_theory.Lp in which the basis is a finite space equipped with the counting measure.

To prove that the topology (and the uniform structure) on a finite product with the L^p distance are the same as those coming from the L^∞ distance, we could argue that the L^p and L^∞ norms are equivalent on ℝ^n for abstract (norm equivalence) reasons. Instead, we give a more explicit (easy) proof which provides a comparison between these two norms with explicit constants.

We also set up the theory for pseudo_emetric_space and pseudo_metric_space.

@[nolint]
def pi_Lp (p : ennreal) {ι : Type u_1} (α : ι Type u_2) :
Type (max u_1 u_2)

A copy of a Pi type, on which we will put the L^p distance. Since the Pi type itself is already endowed with the L^∞ distance, we need the type synonym to avoid confusing typeclass resolution. Also, we let it depend on p, to get a whole family of type on which we can put different distances.

Equations
Instances for pi_Lp
@[protected, instance]
def pi_Lp.inhabited (p : ennreal) {ι : Type u_1} (α : ι Type u_2) [Π (i : ι), inhabited (α i)] :
Equations
@[protected]
def pi_Lp.equiv (p : ennreal) {ι : Type u_3} (α : ι Type u_4) :
pi_Lp p α Π (i : ι), α i

Canonical bijection between pi_Lp p α and the original Pi type. We introduce it to be able to compare the L^p and L^∞ distances through it.

Equations

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

@[simp]
theorem pi_Lp.equiv_apply (p : ennreal) {ι : Type u_3} (α : ι Type u_4) (x : pi_Lp p α) (i : ι) :
(pi_Lp.equiv p α) x i = x i
@[simp]
theorem pi_Lp.equiv_symm_apply (p : ennreal) {ι : Type u_3} (α : ι Type u_4) (x : Π (i : ι), α i) (i : ι) :
((pi_Lp.equiv p α).symm) x i = x i

Definition of edist, dist and norm on pi_Lp #

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

@[protected, instance]
noncomputable def pi_Lp.has_edist (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [Π (i : ι), has_edist (β i)] :

Endowing the space pi_Lp p β with the L^p edistance. We register this instance separate from pi_Lp.pseudo_emetric 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 pi_Lp 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
theorem pi_Lp.edist_eq_card {ι : Type u_3} {β : ι Type u_5} [fintype ι] [Π (i : ι), has_edist (β i)] (f g : pi_Lp 0 β) :
theorem pi_Lp.edist_eq_sum {ι : Type u_3} {β : ι Type u_5} [fintype ι] [Π (i : ι), has_edist (β i)] {p : ennreal} (hp : 0 < p.to_real) (f g : pi_Lp p β) :
has_edist.edist f g = finset.univ.sum (λ (i : ι), has_edist.edist (f i) (g i) ^ p.to_real) ^ (1 / p.to_real)
theorem pi_Lp.edist_eq_supr {ι : Type u_3} {β : ι Type u_5} [fintype ι] [Π (i : ι), has_edist (β i)] (f g : pi_Lp β) :
has_edist.edist f g = (i : ι), has_edist.edist (f i) (g i)
@[protected]
theorem pi_Lp.edist_self (p : ennreal) {ι : Type u_3} {β : ι Type u_5} [fintype ι] [Π (i : ι), pseudo_emetric_space (β i)] (f : pi_Lp p β) :

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

@[protected]
theorem pi_Lp.edist_comm (p : ennreal) {ι : Type u_3} {β : ι Type u_5} [fintype ι] [Π (i : ι), pseudo_emetric_space (β i)] (f g : pi_Lp p β) :

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

@[protected, instance]
noncomputable def pi_Lp.has_dist (p : ennreal) {ι : Type u_3} (α : ι Type u_4) [fintype ι] [Π (i : ι), has_dist (α i)] :

Endowing the space pi_Lp p β with the L^p distance. We register this instance separate from pi_Lp.pseudo_metric 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 pi_Lp 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
theorem pi_Lp.dist_eq_card {ι : Type u_3} {α : ι Type u_4} [fintype ι] [Π (i : ι), has_dist (α i)] (f g : pi_Lp 0 α) :
theorem pi_Lp.dist_eq_sum {ι : Type u_3} {α : ι Type u_4} [fintype ι] [Π (i : ι), has_dist (α i)] {p : ennreal} (hp : 0 < p.to_real) (f g : pi_Lp p α) :
has_dist.dist f g = finset.univ.sum (λ (i : ι), has_dist.dist (f i) (g i) ^ p.to_real) ^ (1 / p.to_real)
theorem pi_Lp.dist_eq_csupr {ι : Type u_3} {α : ι Type u_4} [fintype ι] [Π (i : ι), has_dist (α i)] (f g : pi_Lp α) :
has_dist.dist f g = (i : ι), has_dist.dist (f i) (g i)
@[protected, instance]
noncomputable def pi_Lp.has_norm (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [Π (i : ι), has_norm (β i)] [Π (i : ι), has_zero (β i)] :

Endowing the space pi_Lp p β with the L^p norm. We register this instance separate from pi_Lp.seminormed_add_comm_group 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 pi_Lp p β for p < 1 satisfying a relaxed triangle inequality. These are called quasi-norms.

Equations
theorem pi_Lp.norm_eq_card {ι : Type u_3} {β : ι Type u_5} [fintype ι] [Π (i : ι), has_norm (β i)] [Π (i : ι), has_zero (β i)] (f : pi_Lp 0 β) :
theorem pi_Lp.norm_eq_csupr {ι : Type u_3} {β : ι Type u_5} [fintype ι] [Π (i : ι), has_norm (β i)] [Π (i : ι), has_zero (β i)] (f : pi_Lp β) :
f = (i : ι), f i
theorem pi_Lp.norm_eq_sum {p : ennreal} {ι : Type u_3} {β : ι Type u_5} [fintype ι] [Π (i : ι), has_norm (β i)] [Π (i : ι), has_zero (β i)] (hp : 0 < p.to_real) (f : pi_Lp p β) :
f = finset.univ.sum (λ (i : ι), f i ^ p.to_real) ^ (1 / p.to_real)

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

In this section, we put the L^p edistance on pi_Lp p α, and we check that the uniformity coming from this edistance coincides with the product uniformity, by showing that the canonical map to the Pi 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.

noncomputable def pi_Lp.pseudo_emetric_aux (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fact (1 p)] [Π (i : ι), pseudo_emetric_space (β i)] [fintype ι] :

Endowing the space pi_Lp 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 pseudo_emetric_space.replace_uniformity.

Equations
theorem pi_Lp.supr_edist_ne_top_aux {ι : Type u_1} [finite ι] {α : ι Type u_2} [Π (i : ι), pseudo_metric_space (α i)] (f g : pi_Lp α) :
( (i : ι), has_edist.edist (f i) (g i))

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

@[reducible]
noncomputable def pi_Lp.pseudo_metric_aux (p : ennreal) {ι : Type u_3} (α : ι Type u_4) [fact (1 p)] [Π (i : ι), pseudo_metric_space (α i)] [fintype ι] :

Endowing the space pi_Lp 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, pseudo_metric_space.replace_uniformity, and pseudo_metric_space.replace_bornology.

See note [reducible non-instances]

Equations
theorem pi_Lp.lipschitz_with_equiv_aux (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fact (1 p)] [Π (i : ι), pseudo_emetric_space (β i)] [fintype ι] :
theorem pi_Lp.antilipschitz_with_equiv_aux (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fact (1 p)] [Π (i : ι), pseudo_emetric_space (β i)] [fintype ι] :
theorem pi_Lp.aux_uniformity_eq (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fact (1 p)] [Π (i : ι), pseudo_emetric_space (β i)] [fintype ι] :
uniformity (pi_Lp p β) = uniformity (Π (i : ι), β i)
theorem pi_Lp.aux_cobounded_eq (p : ennreal) {ι : Type u_3} (α : ι Type u_4) [fact (1 p)] [Π (i : ι), pseudo_metric_space (α i)] [fintype ι] :

Instances on finite L^p products #

@[protected, instance]
def pi_Lp.uniform_space (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [Π (i : ι), uniform_space (β i)] :
Equations
theorem pi_Lp.uniform_continuous_equiv (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [Π (i : ι), uniform_space (β i)] :
theorem pi_Lp.uniform_continuous_equiv_symm (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [Π (i : ι), uniform_space (β i)] :
@[continuity]
theorem pi_Lp.continuous_equiv (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [Π (i : ι), uniform_space (β i)] :
@[continuity]
theorem pi_Lp.continuous_equiv_symm (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [Π (i : ι), uniform_space (β i)] :
@[protected, instance]
def pi_Lp.bornology (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [Π (i : ι), bornology (β i)] :
Equations
@[protected, instance]
noncomputable def pi_Lp.pseudo_emetric_space (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), pseudo_emetric_space (β i)] :

pseudoemetric space instance on the product of finitely many pseudoemetric spaces, using the L^p pseudoedistance, and having as uniformity the product uniformity.

Equations
@[protected, instance]
noncomputable def pi_Lp.emetric_space (p : ennreal) {ι : Type u_3} (α : ι Type u_4) [fintype ι] [fact (1 p)] [Π (i : ι), emetric_space (α i)] :

emetric space instance on the product of finitely many emetric spaces, using the L^p edistance, and having as uniformity the product uniformity.

Equations
@[protected, instance]
noncomputable def pi_Lp.pseudo_metric_space (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), pseudo_metric_space (β i)] :

pseudometric space instance on the product of finitely many psuedometric spaces, using the L^p distance, and having as uniformity the product uniformity.

Equations
@[protected, instance]
noncomputable def pi_Lp.metric_space (p : ennreal) {ι : Type u_3} (α : ι Type u_4) [fintype ι] [fact (1 p)] [Π (i : ι), metric_space (α i)] :

metric space instance on the product of finitely many metric spaces, using the L^p distance, and having as uniformity the product uniformity.

Equations
theorem pi_Lp.nndist_eq_sum {ι : Type u_3} [fintype ι] {p : ennreal} [fact (1 p)] {β : ι Type u_1} [Π (i : ι), pseudo_metric_space (β i)] (hp : p ) (x y : pi_Lp p β) :
has_nndist.nndist x y = finset.univ.sum (λ (i : ι), has_nndist.nndist (x i) (y i) ^ p.to_real) ^ (1 / p.to_real)
theorem pi_Lp.nndist_eq_supr {ι : Type u_3} [fintype ι] {β : ι Type u_1} [Π (i : ι), pseudo_metric_space (β i)] (x y : pi_Lp β) :
has_nndist.nndist x y = (i : ι), has_nndist.nndist (x i) (y i)
theorem pi_Lp.lipschitz_with_equiv (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), pseudo_emetric_space (β i)] :
theorem pi_Lp.antilipschitz_with_equiv (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), pseudo_emetric_space (β i)] :
theorem pi_Lp.infty_equiv_isometry {ι : Type u_3} (β : ι Type u_5) [fintype ι] [Π (i : ι), pseudo_emetric_space (β i)] :
@[protected, instance]
noncomputable def pi_Lp.normed_add_comm_group (p : ennreal) {ι : Type u_3} (α : ι Type u_4) [fintype ι] [fact (1 p)] [Π (i : ι), normed_add_comm_group (α i)] :

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

Equations
theorem pi_Lp.nnnorm_eq_sum {ι : Type u_3} [fintype ι] {p : ennreal} [fact (1 p)] {β : ι Type u_1} (hp : p ) [Π (i : ι), seminormed_add_comm_group (β i)] (f : pi_Lp p β) :
f‖₊ = finset.univ.sum (λ (i : ι), f i‖₊ ^ p.to_real) ^ (1 / p.to_real)
theorem pi_Lp.nnnorm_eq_csupr {ι : Type u_3} [fintype ι] {β : ι Type u_1} [Π (i : ι), seminormed_add_comm_group (β i)] (f : pi_Lp β) :
f‖₊ = (i : ι), f i‖₊
theorem pi_Lp.norm_eq_of_nat {ι : Type u_3} [fintype ι] {p : ennreal} [fact (1 p)] {β : ι Type u_1} [Π (i : ι), seminormed_add_comm_group (β i)] (n : ) (h : p = n) (f : pi_Lp p β) :
f = finset.univ.sum (λ (i : ι), f i ^ n) ^ (1 / n)
theorem pi_Lp.norm_eq_of_L2 {ι : Type u_3} [fintype ι] {β : ι Type u_1} [Π (i : ι), seminormed_add_comm_group (β i)] (x : pi_Lp 2 β) :
x = (finset.univ.sum (λ (i : ι), x i ^ 2))
theorem pi_Lp.nnnorm_eq_of_L2 {ι : Type u_3} [fintype ι] {β : ι Type u_1} [Π (i : ι), seminormed_add_comm_group (β i)] (x : pi_Lp 2 β) :
theorem pi_Lp.norm_sq_eq_of_L2 {ι : Type u_3} [fintype ι] (β : ι Type u_1) [Π (i : ι), seminormed_add_comm_group (β i)] (x : pi_Lp 2 β) :
x ^ 2 = finset.univ.sum (λ (i : ι), x i ^ 2)
theorem pi_Lp.dist_eq_of_L2 {ι : Type u_3} [fintype ι] {β : ι Type u_1} [Π (i : ι), seminormed_add_comm_group (β i)] (x y : pi_Lp 2 β) :
has_dist.dist x y = (finset.univ.sum (λ (i : ι), has_dist.dist (x i) (y i) ^ 2))
theorem pi_Lp.nndist_eq_of_L2 {ι : Type u_3} [fintype ι] {β : ι Type u_1} [Π (i : ι), seminormed_add_comm_group (β i)] (x y : pi_Lp 2 β) :
theorem pi_Lp.edist_eq_of_L2 {ι : Type u_3} [fintype ι] {β : ι Type u_1} [Π (i : ι), seminormed_add_comm_group (β i)] (x y : pi_Lp 2 β) :
has_edist.edist x y = finset.univ.sum (λ (i : ι), has_edist.edist (x i) (y i) ^ 2) ^ (1 / 2)
@[protected, instance]
def pi_Lp.normed_space (p : ennreal) (𝕜 : Type u_1) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] :
normed_space 𝕜 (pi_Lp p β)

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

Equations
@[protected, instance]
def pi_Lp.is_scalar_tower (p : ennreal) (𝕜 : Type u_1) (𝕜' : Type u_2) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [normed_field 𝕜'] [Π (i : ι), seminormed_add_comm_group (β i)] [has_smul 𝕜 𝕜'] [Π (i : ι), normed_space 𝕜 (β i)] [Π (i : ι), normed_space 𝕜' (β i)] [ (i : ι), is_scalar_tower 𝕜 𝕜' (β i)] :
is_scalar_tower 𝕜 𝕜' (pi_Lp p β)
@[protected, instance]
def pi_Lp.smul_comm_class (p : ennreal) (𝕜 : Type u_1) (𝕜' : Type u_2) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [normed_field 𝕜'] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] [Π (i : ι), normed_space 𝕜' (β i)] [ (i : ι), smul_comm_class 𝕜 𝕜' (β i)] :
smul_comm_class 𝕜 𝕜' (pi_Lp p β)
@[protected, instance]
def pi_Lp.finite_dimensional (p : ennreal) (𝕜 : Type u_1) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] [I : (i : ι), finite_dimensional 𝕜 (β i)] :
@[simp]
theorem pi_Lp.zero_apply {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] (i : ι) :
0 i = 0
@[simp]
theorem pi_Lp.add_apply {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] (x y : pi_Lp p β) (i : ι) :
(x + y) i = x i + y i
@[simp]
theorem pi_Lp.sub_apply {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] (x y : pi_Lp p β) (i : ι) :
(x - y) i = x i - y i
@[simp]
theorem pi_Lp.smul_apply {p : ennreal} {𝕜 : Type u_1} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] (c : 𝕜) (x : pi_Lp p β) (i : ι) :
(c x) i = c x i
@[simp]
theorem pi_Lp.neg_apply {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] (x : pi_Lp p β) (i : ι) :
(-x) i = -x i
noncomputable def pi_Lp.equivₗᵢ {𝕜 : Type u_1} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] :
pi_Lp β ≃ₗᵢ[𝕜] Π (i : ι), β i

The canonical map pi_Lp.equiv between pi_Lp ∞ β and Π i, β i as a linear isometric equivalence.

Equations
def linear_isometry_equiv.pi_Lp_congr_left (p : ennreal) (𝕜 : Type u_1) {ι : Type u_3} [fintype ι] [fact (1 p)] [normed_field 𝕜] {ι' : Type u_6} [fintype ι'] (E : Type u_7) [normed_add_comm_group E] [normed_space 𝕜 E] (e : ι ι') :
pi_Lp p (λ (i : ι), E) ≃ₗᵢ[𝕜] pi_Lp p (λ (i : ι'), E)

An equivalence of finite domains induces a linearly isometric equivalence of finitely supported functions

Equations
@[simp]
theorem linear_isometry_equiv.pi_Lp_congr_left_apply {p : ennreal} {𝕜 : Type u_1} {ι : Type u_3} [fintype ι] [fact (1 p)] [normed_field 𝕜] {ι' : Type u_6} [fintype ι'] {E : Type u_7} [normed_add_comm_group E] [normed_space 𝕜 E] (e : ι ι') (v : pi_Lp p (λ (i : ι), E)) :
@[simp]
theorem linear_isometry_equiv.pi_Lp_congr_left_symm {p : ennreal} {𝕜 : Type u_1} {ι : Type u_3} [fintype ι] [fact (1 p)] [normed_field 𝕜] {ι' : Type u_6} [fintype ι'] {E : Type u_7} [normed_add_comm_group E] [normed_space 𝕜 E] (e : ι ι') :
@[simp]
theorem linear_isometry_equiv.pi_Lp_congr_left_single {p : ennreal} {𝕜 : Type u_1} {ι : Type u_3} [fintype ι] [fact (1 p)] [normed_field 𝕜] {ι' : Type u_6} [fintype ι'] {E : Type u_7} [normed_add_comm_group E] [normed_space 𝕜 E] [decidable_eq ι] [decidable_eq ι'] (e : ι ι') (i : ι) (v : E) :
(linear_isometry_equiv.pi_Lp_congr_left p 𝕜 E e) (((pi_Lp.equiv p (λ (_x : ι), E)).symm) (pi.single i v)) = ((pi_Lp.equiv p (λ (_x : ι'), E)).symm) (pi.single (e i) v)
@[simp]
theorem pi_Lp.equiv_zero {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] :
(pi_Lp.equiv p β) 0 = 0
@[simp]
theorem pi_Lp.equiv_symm_zero {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] :
((pi_Lp.equiv p β).symm) 0 = 0
@[simp]
theorem pi_Lp.equiv_add {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] (x y : pi_Lp p β) :
(pi_Lp.equiv p β) (x + y) = (pi_Lp.equiv p β) x + (pi_Lp.equiv p β) y
@[simp]
theorem pi_Lp.equiv_symm_add {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] (x' y' : Π (i : ι), β i) :
((pi_Lp.equiv p β).symm) (x' + y') = ((pi_Lp.equiv p β).symm) x' + ((pi_Lp.equiv p β).symm) y'
@[simp]
theorem pi_Lp.equiv_sub {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] (x y : pi_Lp p β) :
(pi_Lp.equiv p β) (x - y) = (pi_Lp.equiv p β) x - (pi_Lp.equiv p β) y
@[simp]
theorem pi_Lp.equiv_symm_sub {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] (x' y' : Π (i : ι), β i) :
((pi_Lp.equiv p β).symm) (x' - y') = ((pi_Lp.equiv p β).symm) x' - ((pi_Lp.equiv p β).symm) y'
@[simp]
theorem pi_Lp.equiv_neg {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] (x : pi_Lp p β) :
(pi_Lp.equiv p β) (-x) = -(pi_Lp.equiv p β) x
@[simp]
theorem pi_Lp.equiv_symm_neg {p : ennreal} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] (x' : Π (i : ι), β i) :
((pi_Lp.equiv p β).symm) (-x') = -((pi_Lp.equiv p β).symm) x'
@[simp]
theorem pi_Lp.equiv_smul {p : ennreal} {𝕜 : Type u_1} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] (c : 𝕜) (x : pi_Lp p β) :
(pi_Lp.equiv p β) (c x) = c (pi_Lp.equiv p β) x
@[simp]
theorem pi_Lp.equiv_symm_smul {p : ennreal} {𝕜 : Type u_1} {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] (c : 𝕜) (x' : Π (i : ι), β i) :
((pi_Lp.equiv p β).symm) (c x') = c ((pi_Lp.equiv p β).symm) x'
@[simp]
theorem pi_Lp.nnnorm_equiv_symm_single (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] [decidable_eq ι] (i : ι) (b : β i) :
@[simp]
theorem pi_Lp.norm_equiv_symm_single (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] [decidable_eq ι] (i : ι) (b : β i) :
@[simp]
theorem pi_Lp.nndist_equiv_symm_single_same (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] [decidable_eq ι] (i : ι) (b₁ b₂ : β i) :
@[simp]
theorem pi_Lp.dist_equiv_symm_single_same (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] [decidable_eq ι] (i : ι) (b₁ b₂ : β i) :
has_dist.dist (((pi_Lp.equiv p β).symm) (pi.single i b₁)) (((pi_Lp.equiv p β).symm) (pi.single i b₂)) = has_dist.dist b₁ b₂
@[simp]
theorem pi_Lp.edist_equiv_symm_single_same (p : ennreal) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [Π (i : ι), seminormed_add_comm_group (β i)] [decidable_eq ι] (i : ι) (b₁ b₂ : β i) :
has_edist.edist (((pi_Lp.equiv p β).symm) (pi.single i b₁)) (((pi_Lp.equiv p β).symm) (pi.single i b₂)) = has_edist.edist b₁ b₂
theorem pi_Lp.nnnorm_equiv_symm_const {p : ennreal} {ι : Type u_3} [fintype ι] [fact (1 p)] {β : Type u_1} [seminormed_add_comm_group β] (hp : p ) (b : β) :
((pi_Lp.equiv p (λ (_x : ι), β)).symm) (function.const ι b)‖₊ = (fintype.card ι) ^ (1 / p).to_real * b‖₊

When p = ∞, this lemma does not hold without the additional assumption nonempty ι because the left-hand side simplifies to 0, while the right-hand side simplifies to ‖b‖₊. See pi_Lp.nnnorm_equiv_symm_const' for a version which exchanges the hypothesis p ≠ ∞ for nonempty ι.

theorem pi_Lp.nnnorm_equiv_symm_const' {p : ennreal} {ι : Type u_3} [fintype ι] [fact (1 p)] {β : Type u_1} [seminormed_add_comm_group β] [nonempty ι] (b : β) :
((pi_Lp.equiv p (λ (_x : ι), β)).symm) (function.const ι b)‖₊ = (fintype.card ι) ^ (1 / p).to_real * b‖₊

When is_empty ι, this lemma does not hold without the additional assumption p ≠ ∞ because the left-hand side simplifies to 0, while the right-hand side simplifies to ‖b‖₊. See pi_Lp.nnnorm_equiv_symm_const for a version which exchanges the hypothesis nonempty ι. for p ≠ ∞.

theorem pi_Lp.norm_equiv_symm_const {p : ennreal} {ι : Type u_3} [fintype ι] [fact (1 p)] {β : Type u_1} [seminormed_add_comm_group β] (hp : p ) (b : β) :
((pi_Lp.equiv p (λ (_x : ι), β)).symm) (function.const ι b) = (fintype.card ι) ^ (1 / p).to_real * b

When p = ∞, this lemma does not hold without the additional assumption nonempty ι because the left-hand side simplifies to 0, while the right-hand side simplifies to ‖b‖₊. See pi_Lp.norm_equiv_symm_const' for a version which exchanges the hypothesis p ≠ ∞ for nonempty ι.

theorem pi_Lp.norm_equiv_symm_const' {p : ennreal} {ι : Type u_3} [fintype ι] [fact (1 p)] {β : Type u_1} [seminormed_add_comm_group β] [nonempty ι] (b : β) :
((pi_Lp.equiv p (λ (_x : ι), β)).symm) (function.const ι b) = (fintype.card ι) ^ (1 / p).to_real * b

When is_empty ι, this lemma does not hold without the additional assumption p ≠ ∞ because the left-hand side simplifies to 0, while the right-hand side simplifies to ‖b‖₊. See pi_Lp.norm_equiv_symm_const for a version which exchanges the hypothesis nonempty ι. for p ≠ ∞.

theorem pi_Lp.nnnorm_equiv_symm_one {p : ennreal} {ι : Type u_3} [fintype ι] [fact (1 p)] {β : Type u_1} [seminormed_add_comm_group β] (hp : p ) [has_one β] :
((pi_Lp.equiv p (λ (_x : ι), β)).symm) 1‖₊ = (fintype.card ι) ^ (1 / p).to_real * 1‖₊
theorem pi_Lp.norm_equiv_symm_one {p : ennreal} {ι : Type u_3} [fintype ι] [fact (1 p)] {β : Type u_1} [seminormed_add_comm_group β] (hp : p ) [has_one β] :
((pi_Lp.equiv p (λ (_x : ι), β)).symm) 1 = (fintype.card ι) ^ (1 / p).to_real * 1
@[protected]
def pi_Lp.linear_equiv (p : ennreal) (𝕜 : Type u_1) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] :
pi_Lp p β ≃ₗ[𝕜] Π (i : ι), β i

pi_Lp.equiv as a linear equivalence.

Equations
@[simp]
theorem pi_Lp.linear_equiv_apply (p : ennreal) (𝕜 : Type u_1) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] :
@[simp]
theorem pi_Lp.linear_equiv_symm_apply (p : ennreal) (𝕜 : Type u_1) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] :
((pi_Lp.linear_equiv p 𝕜 β).symm) = ((pi_Lp.equiv p (λ (i : ι), β i)).symm)
@[protected]
def pi_Lp.continuous_linear_equiv (p : ennreal) (𝕜 : Type u_1) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] :
pi_Lp p β ≃L[𝕜] Π (i : ι), β i

pi_Lp.equiv as a continuous linear equivalence.

Equations
@[simp]
theorem pi_Lp.continuous_linear_equiv_apply (p : ennreal) (𝕜 : Type u_1) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] :
@[simp]
theorem pi_Lp.continuous_linear_equiv_symm_apply (p : ennreal) (𝕜 : Type u_1) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] :
@[simp]
theorem pi_Lp.continuous_linear_equiv_to_linear_equiv (p : ennreal) (𝕜 : Type u_1) {ι : Type u_3} (β : ι Type u_5) [fintype ι] [fact (1 p)] [normed_field 𝕜] [Π (i : ι), seminormed_add_comm_group (β i)] [Π (i : ι), normed_space 𝕜 (β i)] :
noncomputable def pi_Lp.basis_fun (p : ennreal) (𝕜 : Type u_1) (ι : Type u_3) [fintype ι] [fact (1 p)] [normed_field 𝕜] :
basis ι 𝕜 (pi_Lp p (λ (_x : ι), 𝕜))

A version of pi.basis_fun for pi_Lp.

Equations
@[simp]
theorem pi_Lp.basis_fun_apply (p : ennreal) (𝕜 : Type u_1) (ι : Type u_3) [fintype ι] [fact (1 p)] [normed_field 𝕜] [decidable_eq ι] (i : ι) :
(pi_Lp.basis_fun p 𝕜 ι) i = ((pi_Lp.equiv p (λ (i : ι), 𝕜)).symm) (pi.single i 1)
@[simp]
theorem pi_Lp.basis_fun_repr (p : ennreal) (𝕜 : Type u_1) (ι : Type u_3) [fintype ι] [fact (1 p)] [normed_field 𝕜] (x : pi_Lp p (λ (i : ι), 𝕜)) (i : ι) :
(((pi_Lp.basis_fun p 𝕜 ι).repr) x) i = x i
@[simp]
theorem pi_Lp.basis_fun_equiv_fun (p : ennreal) (𝕜 : Type u_1) (ι : Type u_3) [fintype ι] [fact (1 p)] [normed_field 𝕜] :
(pi_Lp.basis_fun p 𝕜 ι).equiv_fun = pi_Lp.linear_equiv p 𝕜 (λ (_x : ι), 𝕜)
theorem pi_Lp.basis_fun_eq_pi_basis_fun (p : ennreal) (𝕜 : Type u_1) (ι : Type u_3) [fintype ι] [fact (1 p)] [normed_field 𝕜] :
pi_Lp.basis_fun p 𝕜 ι = (pi.basis_fun 𝕜 ι).map (pi_Lp.linear_equiv p 𝕜 (λ (_x : ι), 𝕜)).symm
@[simp]
theorem pi_Lp.basis_fun_map (p : ennreal) (𝕜 : Type u_1) (ι : Type u_3) [fintype ι] [fact (1 p)] [normed_field 𝕜] :
(pi_Lp.basis_fun p 𝕜 ι).map (pi_Lp.linear_equiv p 𝕜 (λ (_x : ι), 𝕜)) = pi.basis_fun 𝕜 ι
theorem pi_Lp.basis_to_matrix_basis_fun_mul (p : ennreal) (𝕜 : Type u_1) (ι : Type u_3) [fintype ι] [fact (1 p)] [normed_field 𝕜] (b : basis ι 𝕜 (pi_Lp p (λ (i : ι), 𝕜))) (A : matrix ι ι 𝕜) :
(b.to_matrix (pi_Lp.basis_fun p 𝕜 ι)).mul A = matrix.of (λ (i j : ι), ((b.repr) (((pi_Lp.equiv p (λ (ᾰ : ι), 𝕜)).symm) (A.transpose j))) i)