Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 64abe48

Browse files
refactor(topology/metric_space/pi_Lp): generalize to pseudo_metric (#6978)
We generalize here the Lp distance to `pseudo_emetric` and similar concepts.
1 parent b7fbc17 commit 64abe48

File tree

1 file changed

+87
-40
lines changed

1 file changed

+87
-40
lines changed

src/topology/metric_space/pi_Lp.lean

Lines changed: 87 additions & 40 deletions
Original file line numberDiff line numberDiff line change
@@ -44,6 +44,8 @@ To prove that the topology (and the uniform structure) on a finite product with
4444
are the same as those coming from the `L^∞` distance, we could argue that the `L^p` and `L^∞` norms
4545
are equivalent on `ℝ^n` for abstract (norm equivalence) reasons. Instead, we give a more explicit
4646
(easy) proof which provides a comparison between these two norms with explicit constants.
47+
48+
We also set up the theory for `pseudo_emetric_space` and `pseudo_metric_space`.
4749
-/
4850

4951
open real set filter
@@ -67,7 +69,7 @@ instance {ι : Type*} (p : ℝ) (hp : 1 ≤ p) (α : ι → Type*) [∀ i, inhab
6769

6870
namespace pi_Lp
6971

70-
variables (p : ℝ) (hp : 1 ≤ p) (α : ι → Type*)
72+
variables (p : ℝ) (hp : 1 ≤ p) (α : ι → Type*) (β : ι → Type*)
7173

7274
/-- Canonical bijection between `pi_Lp p hp α` and the original Pi type. We introduce it to be able
7375
to compare the `L^p` and `L^∞` distances through it. -/
@@ -89,15 +91,15 @@ from the edistance (which is equal to it, but not defeq). See Note [forgetful in
8991
explaining why having definitionally the right uniformity is often important.
9092
-/
9193

92-
variables [∀ i, emetric_space (α i)] [fintype ι]
94+
variables [∀ i, emetric_space (α i)] [∀ i, pseudo_emetric_space (β i)] [fintype ι]
9395

94-
/-- Endowing the space `pi_Lp p hp α` with the `L^p` edistance. This definition is not satisfactory,
95-
as it does not register the fact that the topology and the uniform structure coincide with the
96-
product one. Therefore, we do not register it as an instance. Using this as a temporary emetric
97-
space instance, we will show that the uniform structure is equal (but not defeq) to the product one,
98-
and then register an instance in which we replace the uniform structure by the product one using
99-
this emetric space and `emetric_space.replace_uniformity`. -/
100-
def emetric_aux : emetric_space (pi_Lp p hp α) :=
96+
/-- Endowing the space `pi_Lp p hp β` with the `L^p` pseudoedistance. This definition is not
97+
satisfactory, as it does not register the fact that the topology and the uniform structure coincide
98+
with the product one. Therefore, we do not register it as an instance. Using this as a temporary
99+
pseudoemetric space instance, we will show that the uniform structure is equal (but not defeq) to
100+
the product one, and then register an instance in which we replace the uniform structure by the
101+
product one using this pseudoemetric space and `pseudo_emetric_space.replace_uniformity`. -/
102+
def pseudo_emetric_aux : pseudo_emetric_space (pi_Lp p hp β) :=
101103
have pos : 0 < p := lt_of_lt_of_le zero_lt_one hp,
102104
{ edist := λ f g, (∑ (i : ι), (edist (f i) (g i)) ^ p) ^ (1/p),
103105
edist_self := λ f, by simp [edist, ennreal.zero_rpow_of_pos pos,
@@ -113,16 +115,28 @@ have pos : 0 < p := lt_of_lt_of_le zero_lt_one hp,
113115
end
114116
... ≤
115117
(∑ (i : ι), edist (f i) (g i) ^ p) ^ (1 / p) + (∑ (i : ι), edist (g i) (h i) ^ p) ^ (1 / p) :
116-
ennreal.Lp_add_le _ _ _ hp,
117-
eq_of_edist_eq_zero := λ f g hfg,
118+
ennreal.Lp_add_le _ _ _ hp }
119+
120+
/-- Endowing the space `pi_Lp p hp α` with the `L^p` edistance. This definition is not satisfactory,
121+
as it does not register the fact that the topology and the uniform structure coincide with the
122+
product one. Therefore, we do not register it as an instance. Using this as a temporary emetric
123+
space instance, we will show that the uniform structure is equal (but not defeq) to the product one,
124+
and then register an instance in which we replace the uniform structure by the product one using
125+
this emetric space and `emetric_space.replace_uniformity`. -/
126+
def emetric_aux : emetric_space (pi_Lp p hp α) :=
127+
{ eq_of_edist_eq_zero := λ f g hfg,
118128
begin
119-
simp [edist, ennreal.rpow_eq_zero_iff, pos, asymm pos, finset.sum_eq_zero_iff_of_nonneg] at hfg,
129+
have pos : 0 < p := lt_of_lt_of_le zero_lt_one hp,
130+
letI h := pseudo_emetric_aux p hp α,
131+
have h : edist f g = (∑ (i : ι), (edist (f i) (g i)) ^ p) ^ (1/p) := rfl,
132+
simp [h, ennreal.rpow_eq_zero_iff, pos, asymm pos, finset.sum_eq_zero_iff_of_nonneg] at hfg,
120133
exact funext hfg
121-
end }
134+
end,
135+
..pseudo_emetric_aux p hp α }
122136

123-
local attribute [instance] pi_Lp.emetric_aux
137+
local attribute [instance] pi_Lp.emetric_aux pi_Lp.pseudo_emetric_aux
124138

125-
lemma lipschitz_with_equiv : lipschitz_with 1 (pi_Lp.equiv p hp α) :=
139+
lemma lipschitz_with_equiv : lipschitz_with 1 (pi_Lp.equiv p hp β) :=
126140
begin
127141
have pos : 0 < p := lt_of_lt_of_le zero_lt_one hp,
128142
have cancel : p * (1/p) = 1 := mul_div_cancel' 1 (ne_of_gt pos),
@@ -141,23 +155,23 @@ begin
141155
end
142156

143157
lemma antilipschitz_with_equiv :
144-
antilipschitz_with ((fintype.card ι : ℝ≥0) ^ (1/p)) (pi_Lp.equiv p hp α) :=
158+
antilipschitz_with ((fintype.card ι : ℝ≥0) ^ (1/p)) (pi_Lp.equiv p hp β) :=
145159
begin
146160
have pos : 0 < p := lt_of_lt_of_le zero_lt_one hp,
147161
have nonneg : 01 / p := one_div_nonneg.2 (le_of_lt pos),
148162
have cancel : p * (1/p) = 1 := mul_div_cancel' 1 (ne_of_gt pos),
149163
assume x y,
150164
simp [edist, -one_div],
151165
calc (∑ (i : ι), edist (x i) (y i) ^ p) ^ (1 / p) ≤
152-
(∑ (i : ι), edist (pi_Lp.equiv p hp α x) (pi_Lp.equiv p hp α y) ^ p) ^ (1 / p) :
166+
(∑ (i : ι), edist (pi_Lp.equiv p hp β x) (pi_Lp.equiv p hp β y) ^ p) ^ (1 / p) :
153167
begin
154168
apply ennreal.rpow_le_rpow _ nonneg,
155169
apply finset.sum_le_sum (λ i hi, _),
156170
apply ennreal.rpow_le_rpow _ (le_of_lt pos),
157171
exact finset.le_sup (finset.mem_univ i)
158172
end
159173
... = (((fintype.card ι : ℝ≥0)) ^ (1/p) : ℝ≥0) *
160-
edist (pi_Lp.equiv p hp α x) (pi_Lp.equiv p hp α y) :
174+
edist (pi_Lp.equiv p hp β x) (pi_Lp.equiv p hp β y) :
161175
begin
162176
simp only [nsmul_eq_mul, finset.card_univ, ennreal.rpow_one, finset.sum_const,
163177
ennreal.mul_rpow_of_nonneg _ _ nonneg, ←ennreal.rpow_mul, cancel],
@@ -168,13 +182,13 @@ begin
168182
end
169183

170184
lemma aux_uniformity_eq :
171-
𝓤 (pi_Lp p hp α) = @uniformity _ (Pi.uniform_space _) :=
185+
𝓤 (pi_Lp p hp β) = @uniformity _ (Pi.uniform_space _) :=
172186
begin
173-
have A : uniform_embedding (pi_Lp.equiv p hp α) :=
174-
(antilipschitz_with_equiv p hp α).uniform_embedding
175-
(lipschitz_with_equiv p hp α).uniform_continuous,
176-
have : (λ (x : pi_Lp p hp α × pi_Lp p hp α),
177-
((pi_Lp.equiv p hp α) x.fst, (pi_Lp.equiv p hp α) x.snd)) = id,
187+
have A : uniform_embedding (pi_Lp.equiv p hp β) :=
188+
(antilipschitz_with_equiv p hp β).uniform_embedding_of_injective (pi_Lp.equiv p hp β).injective
189+
(lipschitz_with_equiv p hp β).uniform_continuous,
190+
have : (λ (x : pi_Lp p hp β × pi_Lp p hp β),
191+
((pi_Lp.equiv p hp β) x.fst, (pi_Lp.equiv p hp β) x.snd)) = id,
178192
by ext i; refl,
179193
rw [← A.comap_uniformity, this, comap_id]
180194
end
@@ -183,20 +197,43 @@ end
183197

184198
/-! ### Instances on finite `L^p` products -/
185199

186-
instance uniform_space [∀ i, uniform_space (α i)] : uniform_space (pi_Lp p hp α) :=
200+
instance uniform_space [∀ i, uniform_space (β i)] : uniform_space (pi_Lp p hp β) :=
187201
Pi.uniform_space _
188202

189203
variable [fintype ι]
190204

205+
/-- pseudoemetric space instance on the product of finitely many pseudoemetric spaces, using the
206+
`L^p` pseudoedistance, and having as uniformity the product uniformity. -/
207+
instance [∀ i, pseudo_emetric_space (β i)] : pseudo_emetric_space (pi_Lp p hp β) :=
208+
(pseudo_emetric_aux p hp β).replace_uniformity (aux_uniformity_eq p hp β).symm
209+
191210
/-- emetric space instance on the product of finitely many emetric spaces, using the `L^p`
192211
edistance, and having as uniformity the product uniformity. -/
193212
instance [∀ i, emetric_space (α i)] : emetric_space (pi_Lp p hp α) :=
194213
(emetric_aux p hp α).replace_uniformity (aux_uniformity_eq p hp α).symm
195214

196-
protected lemma edist {p : ℝ} {hp : 1 ≤ p} {α : ι → Type*}
197-
[∀ i, emetric_space (α i)] (x y : pi_Lp p hp α) :
215+
protected lemma edist {p : ℝ} {hp : 1 ≤ p} {β : ι → Type*}
216+
[∀ i, pseudo_emetric_space (β i)] (x y : pi_Lp p hp β) :
198217
edist x y = (∑ (i : ι), (edist (x i) (y i)) ^ p) ^ (1/p) := rfl
199218

219+
/-- pseudometric space instance on the product of finitely many psuedometric spaces, using the
220+
`L^p` distance, and having as uniformity the product uniformity. -/
221+
instance [∀ i, pseudo_metric_space (β i)] : pseudo_metric_space (pi_Lp p hp β) :=
222+
begin
223+
/- we construct the instance from the pseudo emetric space instance to avoid checking again that
224+
the uniformity is the same as the product uniformity, but we register nevertheless a nice formula
225+
for the distance -/
226+
have pos : 0 < p := lt_of_lt_of_le zero_lt_one hp,
227+
refine pseudo_emetric_space.to_pseudo_metric_space_of_dist
228+
(λf g, (∑ (i : ι), (dist (f i) (g i)) ^ p) ^ (1/p)) (λ f g, _) (λ f g, _),
229+
{ simp [pi_Lp.edist, ennreal.rpow_eq_top_iff, asymm pos, pos,
230+
ennreal.sum_eq_top_iff, edist_ne_top] },
231+
{ have A : ∀ (i : ι), i ∈ (finset.univ : finset ι) → edist (f i) (g i) ^ p < ⊤ :=
232+
λ i hi, by simp [lt_top_iff_ne_top, edist_ne_top, le_of_lt pos],
233+
simp [dist, -one_div, pi_Lp.edist, ← ennreal.to_real_rpow,
234+
ennreal.to_real_sum A, dist_edist] }
235+
end
236+
200237
/-- metric space instance on the product of finitely many metric spaces, using the `L^p` distance,
201238
and having as uniformity the product uniformity. -/
202239
instance [∀ i, metric_space (α i)] : metric_space (pi_Lp p hp α) :=
@@ -215,30 +252,35 @@ begin
215252
ennreal.to_real_sum A, dist_edist] }
216253
end
217254

218-
protected lemma dist {p : ℝ} {hp : 1 ≤ p} {α : ι → Type*}
219-
[∀ i, metric_space (α i)] (x y : pi_Lp p hp α) :
255+
protected lemma dist {p : ℝ} {hp : 1 ≤ p} {β : ι → Type*}
256+
[∀ i, pseudo_metric_space (β i)] (x y : pi_Lp p hp β) :
220257
dist x y = (∑ (i : ι), (dist (x i) (y i)) ^ p) ^ (1/p) := rfl
221258

222-
/-- normed group instance on the product of finitely many normed groups, using the `L^p` norm. -/
223-
instance normed_group [∀i, normed_group (α i)] : normed_group (pi_Lp p hp α) :=
259+
/-- seminormed group instance on the product of finitely many normed groups, using the `L^p`
260+
norm. -/
261+
instance semi_normed_group [∀i, semi_normed_group (β i)] : semi_normed_group (pi_Lp p hp β) :=
224262
{ norm := λf, (∑ (i : ι), norm (f i) ^ p) ^ (1/p),
225263
dist_eq := λ x y, by { simp [pi_Lp.dist, dist_eq_norm, sub_eq_add_neg] },
226264
.. pi.add_comm_group }
227265

228-
lemma norm_eq {p : ℝ} {hp : 1 ≤ p} {α : ι → Type*}
229-
[∀i, normed_group (α i)] (f : pi_Lp p hp α) :
266+
/-- normed group instance on the product of finitely many normed groups, using the `L^p` norm. -/
267+
instance normed_group [∀i, normed_group (α i)] : normed_group (pi_Lp p hp α) :=
268+
{ ..pi_Lp.semi_normed_group p hp α }
269+
270+
lemma norm_eq {p : ℝ} {hp : 1 ≤ p} {β : ι → Type*}
271+
[∀i, semi_normed_group (β i)] (f : pi_Lp p hp β) :
230272
∥f∥ = (∑ (i : ι), ∥f i∥ ^ p) ^ (1/p) := rfl
231273

232-
lemma norm_eq_of_nat {p : ℝ} {hp : 1 ≤ p} {α : ι → Type*}
233-
[∀i, normed_group (α i)] (n : ℕ) (h : p = n) (f : pi_Lp p hp α) :
274+
lemma norm_eq_of_nat {p : ℝ} {hp : 1 ≤ p} {β : ι → Type*}
275+
[∀i, semi_normed_group (β i)] (n : ℕ) (h : p = n) (f : pi_Lp p hp β) :
234276
∥f∥ = (∑ (i : ι), ∥f i∥ ^ n) ^ (1/(n : ℝ)) :=
235277
by simp [norm_eq, h, real.sqrt_eq_rpow, ←real.rpow_nat_cast]
236278

237279
variables (𝕜 : Type*) [normed_field 𝕜]
238280

239-
/-- The product of finitely many normed spaces is a normed space, with the `L^p` norm. -/
240-
instance normed_space [∀i, normed_group (α i)] [∀i, normed_space 𝕜 (α i)] :
241-
normed_space 𝕜 (pi_Lp p hp α) :=
281+
/-- The product of finitely many seminormed spaces is a seminormed space, with the `L^p` norm. -/
282+
instance semi_normed_space [∀i, semi_normed_group (β i)] [∀i, semi_normed_space 𝕜 (β i)] :
283+
semi_normed_space 𝕜 (pi_Lp p hp β) :=
242284
{ norm_smul_le :=
243285
begin
244286
assume c f,
@@ -248,12 +290,17 @@ instance normed_space [∀i, normed_group (α i)] [∀i, normed_space 𝕜 (α i
248290
this, rpow_one],
249291
exact finset.sum_nonneg (λ i hi, rpow_nonneg_of_nonneg (norm_nonneg _) _)
250292
end,
251-
.. pi.semimodule ι α 𝕜 }
293+
.. pi.semimodule ι β 𝕜 }
294+
295+
/-- The product of finitely many normed spaces is a normed space, with the `L^p` norm. -/
296+
instance normed_space [∀i, normed_group (α i)] [∀i, normed_space 𝕜 (α i)] :
297+
normed_space 𝕜 (pi_Lp p hp α) :=
298+
{ ..pi_Lp.semi_normed_space p hp α 𝕜 }
252299

253300
/- Register simplification lemmas for the applications of `pi_Lp` elements, as the usual lemmas
254301
for Pi types will not trigger. -/
255302
variables {𝕜 p hp α}
256-
[∀i, normed_group (α i)] [∀i, normed_space 𝕜 (α i)] (c : 𝕜) (x y : pi_Lp p hp α) (i : ι)
303+
[∀i, semi_normed_group (β i)] [∀i, semi_normed_space 𝕜 (β i)] (c : 𝕜) (x y : pi_Lp p hp β) (i : ι)
257304

258305
@[simp] lemma add_apply : (x + y) i = x i + y i := rfl
259306
@[simp] lemma sub_apply : (x - y) i = x i - y i := rfl

0 commit comments

Comments
 (0)