From 793c15c63475fbc676cce4f726b4d3144bfa023b Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Tue, 25 Jul 2023 22:27:26 +0200 Subject: [PATCH 01/39] ProdLp, first quarter or so --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 242 +++++++++++++++++++++++ 1 file changed, 242 insertions(+) create mode 100644 Mathlib/Analysis/NormedSpace/ProdLp.lean diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean new file mode 100644 index 0000000000000..1326c7c548867 --- /dev/null +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -0,0 +1,242 @@ +/- +Copyright (c) 2023 Moritz Doll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Doll, Sébastien Gouëzel, Jireh Loreaux +-/ +import Mathlib.Analysis.MeanInequalities +import Mathlib.Data.Fintype.Order +import Mathlib.LinearAlgebra.Matrix.Basis + +local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue #2220 + +open Real Set Filter IsROrC Bornology BigOperators Uniformity Topology NNReal ENNReal + +noncomputable section + +/-- A copy of a Prod type, on which we will put the `L^p` distance. Since the Prod 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. -/ +@[nolint unusedArguments] +def ProdLp (_p : ℝ≥0∞) (α β : Type _) : Type _ := + α × β + +instance (p : ℝ≥0∞) (α β : Type _) [Inhabited α] [Inhabited β] : Inhabited (ProdLp p α β) := + ⟨default, default⟩ + +@[ext] +protected theorem ProdLp.ext {p : ℝ≥0∞} {α β : Type _} {x y : ProdLp p α β} + (h1 : x.fst = y.fst) (h2 : x.snd = y.snd) : x = y := Prod.ext h1 h2 + +namespace ProdLp + +variable (p : ℝ≥0∞) (𝕜 𝕜' : Type _) (α β : Type _) + +/-- Canonical bijection between `PiLp p α` and the original Pi type. We introduce it to be able +to compare the `L^p` and `L^∞` distances through it. -/ +protected def equiv : ProdLp p α β ≃ α × β := + Equiv.refl _ + +/-! Note that the unapplied versions of these lemmas are deliberately omitted, as they break +the use of the type synonym. -/ + +@[simp] +theorem equiv_apply (x : ProdLp p α β) : ProdLp.equiv p α β x = x := + rfl + +@[simp] +theorem equiv_symm_apply (x : α × β) : (ProdLp.equiv p α β).symm x = x := + rfl + +section DistNorm + +variable [Fintype ι] + +/-! +### Definition of `edist`, `dist` and `norm` on `PiLp` + +In this section we define the `edist`, `dist` and `norm` functions on `PiLp 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`. +-/ + + +section Edist + +variable [EDist α] [EDist β] + +/-- Endowing the space `ProdLp p α β` with the `L^p` edistance. We register this instance +separate from `ProdLp.instPseudoEMetric` 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 `ProdLp 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*. -/ +instance : EDist (ProdLp p α β) where + edist f g := + -- Porting note: can we drop the `_hp` entirely? + if _hp : p = 0 then 0 + else + if p = ∞ then edist f.fst g.fst ⊔ edist f.snd g.snd + else (edist f.fst g.fst ^ p.toReal + edist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) + +variable {α β} + +variable (x y : ProdLp p α β) (x' : α × β) + +theorem edist_eq_card (f g : ProdLp 0 α β) : edist f g = 0 := + if_pos rfl + +theorem edist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : ProdLp p α β) : + edist f g = (edist f.fst g.fst ^ p.toReal + edist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) := + let hp' := ENNReal.toReal_pos_iff.mp hp + (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) + +theorem edist_eq_sup (f g : ProdLp ∞ α β) : + edist f g = edist f.fst g.fst ⊔ edist f.snd g.snd := by + dsimp [edist] + exact if_neg ENNReal.top_ne_zero + +end Edist + +section EdistProp + +variable {α β} +variable [PseudoEMetricSpace α] [PseudoEMetricSpace β] + +/-- 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 edist_self (f : ProdLp p α β) : edist f f = 0 := by + rcases p.trichotomy with (rfl | rfl | h) + · simp [edist_eq_card] + · simp [edist_eq_sup] + · simp [edist_eq_add h, ENNReal.zero_rpow_of_pos h, ENNReal.zero_rpow_of_pos (inv_pos.2 <| h)] + +/-- 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 edist_comm (f g : ProdLp p α β) : edist f g = edist g f := by + rcases p.trichotomy with (rfl | rfl | h) + · simp only [edist_eq_card, eq_comm, Ne.def] + · simp only [edist_eq_sup, edist_comm] + · simp only [edist_eq_add h, edist_comm] + +end EdistProp + +section Dist + +variable [Dist α] [Dist β] + +/-- Endowing the space `PiLp 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 `PiLp 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*. -/ +instance : Dist (ProdLp p α β) where + dist f g := + if _hp : p = 0 then 0 --{ i | f i ≠ g i }.toFinite.toFinset.card + else + if p = ∞ then dist f.fst g.fst ⊔ dist f.snd g.snd + else (dist f.fst g.fst ^ p.toReal + dist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) + +variable {α β} + +theorem dist_eq_card (f g : ProdLp 0 α β) : dist f g = 0 := + if_pos rfl + +theorem dist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : ProdLp p α β) : + dist f g = (dist f.fst g.fst ^ p.toReal + dist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) := + let hp' := ENNReal.toReal_pos_iff.mp hp + (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) + +theorem dist_eq_sup (f g : ProdLp ∞ α β) : + dist f g = dist f.fst g.fst ⊔ dist f.snd g.snd := by + dsimp [dist] + exact if_neg ENNReal.top_ne_zero + +end Dist + +section Norm + +variable [Norm α] [Zero α] [Norm β] [Zero β] + +/-- Endowing the space `PiLp p β` with the `L^p` norm. We register this instance +separate from `PiLp.seminormedAddCommGroup` 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 `PiLp p β` for `p < 1` +satisfying a relaxed triangle inequality. These are called *quasi-norms*. -/ +instance instNorm : Norm (ProdLp p α β) where + norm f := + if _hp : p = 0 then 0 -- { i | f i ≠ 0 }.toFinite.toFinset.card + else if p = ∞ then ‖f.fst‖ ⊔ ‖f.snd‖ + else (‖f.fst‖ ^ p.toReal + ‖f.snd‖ ^ p.toReal) ^ (1 / p.toReal) + +variable {p β} + +theorem norm_eq_card (f : ProdLp 0 α β) : ‖f‖ = 0 := + if_pos rfl + +theorem norm_eq_sup (f : ProdLp ∞ α β) : ‖f‖ = ‖f.fst‖ ⊔ ‖f.snd‖ := by + dsimp [Norm.norm] + exact if_neg ENNReal.top_ne_zero + +theorem norm_eq_add (hp : 0 < p.toReal) (f : ProdLp p α β) : + ‖f‖ = (‖f.fst‖ ^ p.toReal + ‖f.snd‖ ^ p.toReal) ^ (1 / p.toReal) := + let hp' := ENNReal.toReal_pos_iff.mp hp + (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) + +end Norm + +end DistNorm + +/-! ### Instances on finite `L^p` products -/ + + +instance instUniformSpace [UniformSpace α] [UniformSpace β] : UniformSpace (ProdLp p α β) := + instUniformSpaceProd + +theorem uniformContinuous_equiv [UniformSpace α] [UniformSpace β] : + UniformContinuous (ProdLp.equiv p α β) := + uniformContinuous_id + +theorem uniformContinuous_equiv_symm [UniformSpace α] [UniformSpace β] : + UniformContinuous (ProdLp.equiv p α β).symm := + uniformContinuous_id + +@[continuity] +theorem continuous_equiv [UniformSpace α] [UniformSpace β] : Continuous (ProdLp.equiv p α β) := + continuous_id + +@[continuity] +theorem continuous_equiv_symm [UniformSpace α] [UniformSpace β] : + Continuous (ProdLp.equiv p α β).symm := + continuous_id + +instance instBornology [Bornology α] [Bornology β] : Bornology (ProdLp p α β) := + Prod.instBornology + +-- throughout the rest of the file, we assume `1 ≤ p` +variable [Fact (1 ≤ p)] + +/-- pseudoemetric space instance on the product of finitely many pseudoemetric spaces, using the +`L^p` pseudoedistance, and having as uniformity the product uniformity. -/ +instance [PseudoEMetricSpace α] [PseudoEMetricSpace β] : PseudoEMetricSpace (ProdLp p α β) := + (pseudoEmetricAux p α β).replaceUniformity (aux_uniformity_eq p β).symm + +/-- emetric space instance on the product of finitely many emetric spaces, using the `L^p` +edistance, and having as uniformity the product uniformity. -/ +instance [EMetricSpace α] [EMetricSpace β] : EMetricSpace (ProdLp p α β) := + @EMetricSpace.ofT0PseudoEMetricSpace (ProdLp p α β) _ Prod.instT0Space + +/-- pseudometric space instance on the product of finitely many pseudometric spaces, using the +`L^p` distance, and having as uniformity the product uniformity. -/ +instance [∀ i, PseudoMetricSpace (β i)] : PseudoMetricSpace (PiLp p β) := + ((pseudoMetricAux p β).replaceUniformity (aux_uniformity_eq p β).symm).replaceBornology fun s => + Filter.ext_iff.1 (aux_cobounded_eq p β).symm sᶜ + +/-- metric space instance on the product of finitely many metric spaces, using the `L^p` distance, +and having as uniformity the product uniformity. -/ +instance [∀ i, MetricSpace (α i)] : MetricSpace (PiLp p α) := + MetricSpace.ofT0PseudoMetricSpace _ From fb9744ca84ee21ffa7e26f47624f888c1b3bcfc0 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 26 Jul 2023 16:32:37 +0200 Subject: [PATCH 02/39] two proofs --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 175 ++++++++++++++++++++++- 1 file changed, 173 insertions(+), 2 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 1326c7c548867..c507efdf9cdcd 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -50,8 +50,6 @@ theorem equiv_symm_apply (x : α × β) : (ProdLp.equiv p α β).symm x = x := section DistNorm -variable [Fintype ι] - /-! ### Definition of `edist`, `dist` and `norm` on `PiLp` @@ -191,6 +189,179 @@ end Norm end DistNorm +section Aux + +/-! +### The uniformity on finite `L^p` products is the product uniformity + +In this section, we put the `L^p` edistance on `PiLp 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. +-/ + + +variable [Fact (1 ≤ p)] --[∀ i, PseudoMetricSpace (α i)] [∀ i, PseudoEMetricSpace (β i)] + +/-- Endowing the space `PiLp 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`. -/ +def pseudoEmetricAux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : + PseudoEMetricSpace (ProdLp p α β) where + edist_self := ProdLp.edist_self p + edist_comm := ProdLp.edist_comm p + edist_triangle f g h := by + rcases p.dichotomy with (rfl | hp) + · simp only [edist_eq_sup] + exact sup_le ((edist_triangle _ g.fst _).trans <| add_le_add le_sup_left le_sup_left) + ((edist_triangle _ g.snd _).trans <| add_le_add le_sup_right le_sup_right) + · simp only [edist_eq_add (zero_lt_one.trans_le hp)] + calc + (edist f.fst h.fst ^ p.toReal + edist f.snd h.snd ^ p.toReal) ^ (1 / p.toReal) ≤ + ((edist f.fst g.fst + edist g.fst h.fst) ^ p.toReal + + (edist f.snd g.snd + edist g.snd h.snd) ^ p.toReal) ^ (1 / p.toReal) := by + apply ENNReal.rpow_le_rpow _ (one_div_nonneg.2 <| zero_le_one.trans hp) + exact add_le_add (ENNReal.rpow_le_rpow (edist_triangle _ _ _) (zero_le_one.trans hp)) + (ENNReal.rpow_le_rpow (edist_triangle _ _ _) (zero_le_one.trans hp)) + _ ≤ + (edist f.fst g.fst ^ p.toReal + edist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) + + (edist g.fst h.fst ^ p.toReal + edist g.snd h.snd ^ p.toReal) ^ (1 / p.toReal) := by + have := ENNReal.Lp_add_le {0, 1} + (if · = 0 then edist f.fst g.fst else edist f.snd g.snd) + (if · = 0 then edist g.fst h.fst else edist g.snd h.snd) hp + simp only [Finset.mem_singleton, not_false_eq_true, Finset.sum_insert, + Finset.sum_singleton] at this + exact this + +attribute [local instance] ProdLp.pseudoEmetricAux + +/-- An auxiliary lemma used twice in the proof of `PiLp.pseudoMetricAux` below. Not intended for +use outside this file. -/ +theorem sup_edist_ne_top_aux {α β : Type _} + [PseudoMetricSpace α] [PseudoMetricSpace β] (f g : ProdLp ∞ α β) : + edist f.fst g.fst ⊔ edist f.snd g.snd ≠ ⊤ := by + refine ne_of_lt ?_ + simp [edist, PseudoMetricSpace.edist_dist] + +/-- Endowing the space `PiLp 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] -/ +@[reducible] +def pseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : + PseudoMetricSpace (ProdLp p α β) := + PseudoEMetricSpace.toPseudoMetricSpaceOfDist dist + (fun f g => by + rcases p.dichotomy with (rfl | h) + · exact sup_edist_ne_top_aux f g + · rw [edist_eq_add (zero_lt_one.trans_le h)] + refine + ENNReal.rpow_ne_top_of_nonneg (one_div_nonneg.2 (zero_le_one.trans h)) (ne_of_lt ?_) + simp [ENNReal.add_lt_top, ENNReal.rpow_lt_top_of_nonneg, edist_ne_top] ) + fun f g => by + rcases p.dichotomy with (rfl | h) + · rw [edist_eq_sup, dist_eq_sup] + refine' le_antisymm (sup_le _ _) _ + · rw [← ENNReal.ofReal_le_iff_le_toReal (sup_edist_ne_top_aux f g), ← + PseudoMetricSpace.edist_dist] + exact le_sup_left + · rw [← ENNReal.ofReal_le_iff_le_toReal (sup_edist_ne_top_aux f g), ← + PseudoMetricSpace.edist_dist] + exact le_sup_right + · refine ENNReal.toReal_le_of_le_ofReal ?_ ?_ + · simp only [ge_iff_le, le_sup_iff, dist_nonneg] + · change PseudoMetricSpace.edist _ _ ⊔ PseudoMetricSpace.edist _ _ ≤ _ + simp [PseudoMetricSpace.edist_dist, ENNReal.ofReal_le_ofReal] + · have h1 : edist f.fst g.fst ^ p.toReal ≠ ⊤ := + ENNReal.rpow_ne_top_of_nonneg (zero_le_one.trans h) (edist_ne_top _ _) + have h2 : edist f.snd g.snd ^ p.toReal ≠ ⊤ := + ENNReal.rpow_ne_top_of_nonneg (zero_le_one.trans h) (edist_ne_top _ _) + simp only [edist_eq_add (zero_lt_one.trans_le h), dist_edist, ENNReal.toReal_rpow, + dist_eq_add (zero_lt_one.trans_le h), ← ENNReal.toReal_add h1 h2] + +attribute [local instance] ProdLp.pseudoMetricAux + +theorem lipschitzWith_equiv_aux : LipschitzWith 1 (PiLp.equiv p β) := by + intro x y + rcases p.dichotomy with (rfl | h) + · simpa only [ENNReal.coe_one, one_mul, edist_eq_iSup, edist, Finset.sup_le_iff, Finset.mem_univ, + forall_true_left] using le_iSup fun i => edist (x i) (y i) + · have cancel : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le h).ne' + rw [edist_eq_sum (zero_lt_one.trans_le h)] + simp only [edist, forall_prop_of_true, one_mul, Finset.mem_univ, Finset.sup_le_iff, + ENNReal.coe_one] + intro i + calc + edist (x i) (y i) = (edist (x i) (y i) ^ p.toReal) ^ (1 / p.toReal) := by + simp [← ENNReal.rpow_mul, cancel, -one_div] + _ ≤ (∑ i, edist (x i) (y i) ^ p.toReal) ^ (1 / p.toReal) := by + apply ENNReal.rpow_le_rpow _ (one_div_nonneg.2 <| zero_le_one.trans h) + exact Finset.single_le_sum (fun i _ => (bot_le : (0 : ℝ≥0∞) ≤ _)) (Finset.mem_univ i) +#align pi_Lp.lipschitz_with_equiv_aux PiLp.lipschitzWith_equiv_aux + +theorem antilipschitzWith_equiv_aux : + AntilipschitzWith ((Fintype.card ι : ℝ≥0) ^ (1 / p).toReal) (PiLp.equiv p β) := by + intro x y + rcases p.dichotomy with (rfl | h) + · simp only [edist_eq_iSup, ENNReal.div_top, ENNReal.zero_toReal, NNReal.rpow_zero, + ENNReal.coe_one, one_mul, iSup_le_iff] + -- Porting note: `Finset.le_sup` needed some help + exact fun i => Finset.le_sup (f := fun i => edist (x i) (y i)) (Finset.mem_univ i) + · have pos : 0 < p.toReal := zero_lt_one.trans_le h + have nonneg : 0 ≤ 1 / p.toReal := one_div_nonneg.2 (le_of_lt pos) + have cancel : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (ne_of_gt pos) + rw [edist_eq_sum pos, ENNReal.toReal_div 1 p] + simp only [edist, ← one_div, ENNReal.one_toReal] + calc + (∑ i, edist (x i) (y i) ^ p.toReal) ^ (1 / p.toReal) ≤ + (∑ _i, edist (PiLp.equiv p β x) (PiLp.equiv p β y) ^ p.toReal) ^ (1 / p.toReal) := by + refine ENNReal.rpow_le_rpow ?_ nonneg + swap + refine Finset.sum_le_sum fun i _ => ?_ + apply ENNReal.rpow_le_rpow _ (le_of_lt pos) + exact Finset.le_sup (f := fun i => edist (x i) (y i)) (Finset.mem_univ i) + _ = + ((Fintype.card ι : ℝ≥0) ^ (1 / p.toReal) : ℝ≥0) * + edist (PiLp.equiv p β x) (PiLp.equiv p β y) := by + simp only [nsmul_eq_mul, Finset.card_univ, ENNReal.rpow_one, Finset.sum_const, + ENNReal.mul_rpow_of_nonneg _ _ nonneg, ← ENNReal.rpow_mul, cancel] + have : (Fintype.card ι : ℝ≥0∞) = (Fintype.card ι : ℝ≥0) := + (ENNReal.coe_nat (Fintype.card ι)).symm + rw [this, ENNReal.coe_rpow_of_nonneg _ nonneg] +#align pi_Lp.antilipschitz_with_equiv_aux PiLp.antilipschitzWith_equiv_aux + +theorem aux_uniformity_eq : 𝓤 (PiLp p β) = 𝓤[Pi.uniformSpace _] := by + have A : UniformInducing (PiLp.equiv p β) := + (antilipschitzWith_equiv_aux p β).uniformInducing + (lipschitzWith_equiv_aux p β).uniformContinuous + have : (fun x : PiLp p β × PiLp p β => ((PiLp.equiv p β) x.fst, (PiLp.equiv p β) x.snd)) = id := + by ext i <;> rfl + rw [← A.comap_uniformity, this, comap_id] +#align pi_Lp.aux_uniformity_eq PiLp.aux_uniformity_eq + +theorem aux_cobounded_eq : cobounded (PiLp p α) = @cobounded _ Pi.instBornology := + calc + cobounded (PiLp p α) = comap (PiLp.equiv p α) (cobounded _) := + le_antisymm (antilipschitzWith_equiv_aux p α).tendsto_cobounded.le_comap + (lipschitzWith_equiv_aux p α).comap_cobounded_le + _ = _ := comap_id +#align pi_Lp.aux_cobounded_eq PiLp.aux_cobounded_eq + +end Aux + /-! ### Instances on finite `L^p` products -/ From 3288d0329c13ea4747bc1f81cac79165d1525847 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 26 Jul 2023 16:45:59 +0200 Subject: [PATCH 03/39] more fixes --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 44 ++++++++++++------------ 1 file changed, 22 insertions(+), 22 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index c507efdf9cdcd..df95fce06a4e8 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -294,7 +294,8 @@ def pseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : attribute [local instance] ProdLp.pseudoMetricAux -theorem lipschitzWith_equiv_aux : LipschitzWith 1 (PiLp.equiv p β) := by +theorem lipschitzWith_equiv_aux [PseudoMetricSpace α] [PseudoMetricSpace β] : + LipschitzWith 1 (ProdLp.equiv p α β) := by intro x y rcases p.dichotomy with (rfl | h) · simpa only [ENNReal.coe_one, one_mul, edist_eq_iSup, edist, Finset.sup_le_iff, Finset.mem_univ, @@ -310,10 +311,9 @@ theorem lipschitzWith_equiv_aux : LipschitzWith 1 (PiLp.equiv p β) := by _ ≤ (∑ i, edist (x i) (y i) ^ p.toReal) ^ (1 / p.toReal) := by apply ENNReal.rpow_le_rpow _ (one_div_nonneg.2 <| zero_le_one.trans h) exact Finset.single_le_sum (fun i _ => (bot_le : (0 : ℝ≥0∞) ≤ _)) (Finset.mem_univ i) -#align pi_Lp.lipschitz_with_equiv_aux PiLp.lipschitzWith_equiv_aux -theorem antilipschitzWith_equiv_aux : - AntilipschitzWith ((Fintype.card ι : ℝ≥0) ^ (1 / p).toReal) (PiLp.equiv p β) := by +theorem antilipschitzWith_equiv_aux [PseudoMetricSpace α] [PseudoMetricSpace β] : + AntilipschitzWith ((2 : ℝ≥0) ^ (1 / p).toReal) (ProdLp.equiv p α β) := by intro x y rcases p.dichotomy with (rfl | h) · simp only [edist_eq_iSup, ENNReal.div_top, ENNReal.zero_toReal, NNReal.rpow_zero, @@ -341,24 +341,24 @@ theorem antilipschitzWith_equiv_aux : have : (Fintype.card ι : ℝ≥0∞) = (Fintype.card ι : ℝ≥0) := (ENNReal.coe_nat (Fintype.card ι)).symm rw [this, ENNReal.coe_rpow_of_nonneg _ nonneg] -#align pi_Lp.antilipschitz_with_equiv_aux PiLp.antilipschitzWith_equiv_aux -theorem aux_uniformity_eq : 𝓤 (PiLp p β) = 𝓤[Pi.uniformSpace _] := by - have A : UniformInducing (PiLp.equiv p β) := - (antilipschitzWith_equiv_aux p β).uniformInducing - (lipschitzWith_equiv_aux p β).uniformContinuous - have : (fun x : PiLp p β × PiLp p β => ((PiLp.equiv p β) x.fst, (PiLp.equiv p β) x.snd)) = id := +theorem aux_uniformity_eq [PseudoMetricSpace α] [PseudoMetricSpace β] : + 𝓤 (ProdLp p α β) = 𝓤[instUniformSpaceProd] := by + have A : UniformInducing (ProdLp.equiv p α β) := + (antilipschitzWith_equiv_aux p α β).uniformInducing + (lipschitzWith_equiv_aux p α β).uniformContinuous + have : (fun x : ProdLp p α β × ProdLp p α β => + ((ProdLp.equiv p α β) x.fst, (ProdLp.equiv p α β) x.snd)) = id := by ext i <;> rfl rw [← A.comap_uniformity, this, comap_id] -#align pi_Lp.aux_uniformity_eq PiLp.aux_uniformity_eq -theorem aux_cobounded_eq : cobounded (PiLp p α) = @cobounded _ Pi.instBornology := +theorem aux_cobounded_eq [PseudoMetricSpace α] [PseudoMetricSpace β] : + cobounded (ProdLp p α β) = @cobounded _ Prod.instBornology := calc - cobounded (PiLp p α) = comap (PiLp.equiv p α) (cobounded _) := - le_antisymm (antilipschitzWith_equiv_aux p α).tendsto_cobounded.le_comap - (lipschitzWith_equiv_aux p α).comap_cobounded_le + cobounded (ProdLp p α β) = comap (ProdLp.equiv p α β) (cobounded _) := + le_antisymm (antilipschitzWith_equiv_aux p α β).tendsto_cobounded.le_comap + (lipschitzWith_equiv_aux p α β).comap_cobounded_le _ = _ := comap_id -#align pi_Lp.aux_cobounded_eq PiLp.aux_cobounded_eq end Aux @@ -394,20 +394,20 @@ variable [Fact (1 ≤ p)] /-- pseudoemetric space instance on the product of finitely many pseudoemetric spaces, using the `L^p` pseudoedistance, and having as uniformity the product uniformity. -/ instance [PseudoEMetricSpace α] [PseudoEMetricSpace β] : PseudoEMetricSpace (ProdLp p α β) := - (pseudoEmetricAux p α β).replaceUniformity (aux_uniformity_eq p β).symm + (pseudoEmetricAux p α β).replaceUniformity (aux_uniformity_eq p α β).symm /-- emetric space instance on the product of finitely many emetric spaces, using the `L^p` edistance, and having as uniformity the product uniformity. -/ instance [EMetricSpace α] [EMetricSpace β] : EMetricSpace (ProdLp p α β) := - @EMetricSpace.ofT0PseudoEMetricSpace (ProdLp p α β) _ Prod.instT0Space + @EMetricSpace.ofT0PseudoEMetricSpace (ProdLp p α β) _ instT0SpaceProdInstTopologicalSpaceProd /-- pseudometric space instance on the product of finitely many pseudometric spaces, using the `L^p` distance, and having as uniformity the product uniformity. -/ -instance [∀ i, PseudoMetricSpace (β i)] : PseudoMetricSpace (PiLp p β) := - ((pseudoMetricAux p β).replaceUniformity (aux_uniformity_eq p β).symm).replaceBornology fun s => - Filter.ext_iff.1 (aux_cobounded_eq p β).symm sᶜ +instance [PseudoMetricSpace α] [PseudoMetricSpace β] : PseudoMetricSpace (ProdLp p α β) := + ((pseudoMetricAux p α β).replaceUniformity (aux_uniformity_eq p α β).symm).replaceBornology + fun s => Filter.ext_iff.1 (aux_cobounded_eq p α β).symm sᶜ /-- metric space instance on the product of finitely many metric spaces, using the `L^p` distance, and having as uniformity the product uniformity. -/ -instance [∀ i, MetricSpace (α i)] : MetricSpace (PiLp p α) := +instance [MetricSpace α] [MetricSpace β] : MetricSpace (ProdLp p α β) := MetricSpace.ofT0PseudoMetricSpace _ From b5e84244558e50266f86fa5e9bbea7be6a3cc801 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 26 Jul 2023 20:18:49 +0200 Subject: [PATCH 04/39] half way point --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 126 +++++++++++++++-------- 1 file changed, 82 insertions(+), 44 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index df95fce06a4e8..ebef8ff23959e 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -21,7 +21,8 @@ different distances. -/ def ProdLp (_p : ℝ≥0∞) (α β : Type _) : Type _ := α × β -instance (p : ℝ≥0∞) (α β : Type _) [Inhabited α] [Inhabited β] : Inhabited (ProdLp p α β) := +instance ProdLp.instInhabited (p : ℝ≥0∞) (α β : Type _) [Inhabited α] [Inhabited β] : + Inhabited (ProdLp p α β) := ⟨default, default⟩ @[ext] @@ -70,7 +71,7 @@ separate from `ProdLp.instPseudoEMetric` since the latter requires the type clas Registering this separately allows for a future emetric-like structure on `ProdLp 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*. -/ -instance : EDist (ProdLp p α β) where +instance instEDist : EDist (ProdLp p α β) where edist f g := -- Porting note: can we drop the `_hp` entirely? if _hp : p = 0 then 0 @@ -131,7 +132,7 @@ separate from `pi_Lp.pseudo_metric` since the latter requires the type class hyp Registering this separately allows for a future metric-like structure on `PiLp 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*. -/ -instance : Dist (ProdLp p α β) where +instance instDist : Dist (ProdLp p α β) where dist f g := if _hp : p = 0 then 0 --{ i | f i ≠ g i }.toFinite.toFinset.card else @@ -269,7 +270,7 @@ def pseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : · exact sup_edist_ne_top_aux f g · rw [edist_eq_add (zero_lt_one.trans_le h)] refine - ENNReal.rpow_ne_top_of_nonneg (one_div_nonneg.2 (zero_le_one.trans h)) (ne_of_lt ?_) + ENNReal.rpow_ne_top_of_nonneg (by positivity) (ne_of_lt ?_) simp [ENNReal.add_lt_top, ENNReal.rpow_lt_top_of_nonneg, edist_ne_top] ) fun f g => by rcases p.dichotomy with (rfl | h) @@ -283,8 +284,7 @@ def pseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : exact le_sup_right · refine ENNReal.toReal_le_of_le_ofReal ?_ ?_ · simp only [ge_iff_le, le_sup_iff, dist_nonneg] - · change PseudoMetricSpace.edist _ _ ⊔ PseudoMetricSpace.edist _ _ ≤ _ - simp [PseudoMetricSpace.edist_dist, ENNReal.ofReal_le_ofReal] + · simp [edist, PseudoMetricSpace.edist_dist, ENNReal.ofReal_le_ofReal] · have h1 : edist f.fst g.fst ^ p.toReal ≠ ⊤ := ENNReal.rpow_ne_top_of_nonneg (zero_le_one.trans h) (edist_ne_top _ _) have h2 : edist f.snd g.snd ^ p.toReal ≠ ⊤ := @@ -294,55 +294,58 @@ def pseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : attribute [local instance] ProdLp.pseudoMetricAux -theorem lipschitzWith_equiv_aux [PseudoMetricSpace α] [PseudoMetricSpace β] : +theorem lipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : LipschitzWith 1 (ProdLp.equiv p α β) := by intro x y rcases p.dichotomy with (rfl | h) - · simpa only [ENNReal.coe_one, one_mul, edist_eq_iSup, edist, Finset.sup_le_iff, Finset.mem_univ, - forall_true_left] using le_iSup fun i => edist (x i) (y i) + · simp only [equiv_apply, coe_one, one_mul, le_refl] · have cancel : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le h).ne' - rw [edist_eq_sum (zero_lt_one.trans_le h)] - simp only [edist, forall_prop_of_true, one_mul, Finset.mem_univ, Finset.sup_le_iff, - ENNReal.coe_one] - intro i - calc - edist (x i) (y i) = (edist (x i) (y i) ^ p.toReal) ^ (1 / p.toReal) := by - simp [← ENNReal.rpow_mul, cancel, -one_div] - _ ≤ (∑ i, edist (x i) (y i) ^ p.toReal) ^ (1 / p.toReal) := by - apply ENNReal.rpow_le_rpow _ (one_div_nonneg.2 <| zero_le_one.trans h) - exact Finset.single_le_sum (fun i _ => (bot_le : (0 : ℝ≥0∞) ≤ _)) (Finset.mem_univ i) - -theorem antilipschitzWith_equiv_aux [PseudoMetricSpace α] [PseudoMetricSpace β] : + rw [edist_eq_add (zero_lt_one.trans_le h)] + simp only [edist, forall_prop_of_true, one_mul, ENNReal.coe_one, equiv_apply, ge_iff_le, + sup_le_iff] + constructor + · calc + edist x.fst y.fst ≤ (edist x.fst y.fst ^ p.toReal) ^ (1 / p.toReal) := by + simp only [← ENNReal.rpow_mul, cancel, ENNReal.rpow_one, le_refl] + _ ≤ (edist x.fst y.fst ^ p.toReal + edist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) := by + apply ENNReal.rpow_le_rpow _ (by positivity) + simp only [self_le_add_right] + · calc + edist x.snd y.snd ≤ (edist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) := by + simp only [← ENNReal.rpow_mul, cancel, ENNReal.rpow_one, le_refl] + _ ≤ (edist x.fst y.fst ^ p.toReal + edist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) := by + apply ENNReal.rpow_le_rpow _ (by positivity) + simp only [self_le_add_left] + +example (a : ENNReal) : a + a = 2*a := by exact Eq.symm (two_mul a) + +theorem antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : AntilipschitzWith ((2 : ℝ≥0) ^ (1 / p).toReal) (ProdLp.equiv p α β) := by intro x y rcases p.dichotomy with (rfl | h) - · simp only [edist_eq_iSup, ENNReal.div_top, ENNReal.zero_toReal, NNReal.rpow_zero, - ENNReal.coe_one, one_mul, iSup_le_iff] - -- Porting note: `Finset.le_sup` needed some help - exact fun i => Finset.le_sup (f := fun i => edist (x i) (y i)) (Finset.mem_univ i) - · have pos : 0 < p.toReal := zero_lt_one.trans_le h - have nonneg : 0 ≤ 1 / p.toReal := one_div_nonneg.2 (le_of_lt pos) + · simp [edist] + · have pos : 0 < p.toReal := by positivity + have nonneg : 0 ≤ 1 / p.toReal := by positivity have cancel : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (ne_of_gt pos) - rw [edist_eq_sum pos, ENNReal.toReal_div 1 p] + rw [edist_eq_add pos, ENNReal.toReal_div 1 p] simp only [edist, ← one_div, ENNReal.one_toReal] calc - (∑ i, edist (x i) (y i) ^ p.toReal) ^ (1 / p.toReal) ≤ - (∑ _i, edist (PiLp.equiv p β x) (PiLp.equiv p β y) ^ p.toReal) ^ (1 / p.toReal) := by - refine ENNReal.rpow_le_rpow ?_ nonneg - swap - refine Finset.sum_le_sum fun i _ => ?_ - apply ENNReal.rpow_le_rpow _ (le_of_lt pos) - exact Finset.le_sup (f := fun i => edist (x i) (y i)) (Finset.mem_univ i) + (edist x.fst y.fst ^ p.toReal + edist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) ≤ + (edist (ProdLp.equiv p α β x) (ProdLp.equiv p α β y) ^ p.toReal + + edist (ProdLp.equiv p α β x) (ProdLp.equiv p α β y) ^ p.toReal) ^ (1 / p.toReal) := by + refine ENNReal.rpow_le_rpow (add_le_add ?_ ?_) nonneg + · refine ENNReal.rpow_le_rpow ?_ (le_of_lt pos) + simp [edist] + · refine ENNReal.rpow_le_rpow ?_ (le_of_lt pos) + simp [edist] _ = - ((Fintype.card ι : ℝ≥0) ^ (1 / p.toReal) : ℝ≥0) * - edist (PiLp.equiv p β x) (PiLp.equiv p β y) := by - simp only [nsmul_eq_mul, Finset.card_univ, ENNReal.rpow_one, Finset.sum_const, - ENNReal.mul_rpow_of_nonneg _ _ nonneg, ← ENNReal.rpow_mul, cancel] - have : (Fintype.card ι : ℝ≥0∞) = (Fintype.card ι : ℝ≥0) := - (ENNReal.coe_nat (Fintype.card ι)).symm - rw [this, ENNReal.coe_rpow_of_nonneg _ nonneg] - -theorem aux_uniformity_eq [PseudoMetricSpace α] [PseudoMetricSpace β] : + ((2 : ℝ≥0) ^ (1 / p.toReal) : ℝ≥0) * + edist (ProdLp.equiv p α β x) (ProdLp.equiv p α β y) := by + simp only [equiv_apply, ← two_mul, ENNReal.mul_rpow_of_nonneg _ _ nonneg, + ← ENNReal.rpow_mul, cancel, ENNReal.rpow_one, ← ENNReal.coe_rpow_of_nonneg _ nonneg, + coe_ofNat] + +theorem aux_uniformity_eq [PseudoEMetricSpace α] [PseudoEMetricSpace β] : 𝓤 (ProdLp p α β) = 𝓤[instUniformSpaceProd] := by have A : UniformInducing (ProdLp.equiv p α β) := (antilipschitzWith_equiv_aux p α β).uniformInducing @@ -411,3 +414,38 @@ instance [PseudoMetricSpace α] [PseudoMetricSpace β] : PseudoMetricSpace (Prod and having as uniformity the product uniformity. -/ instance [MetricSpace α] [MetricSpace β] : MetricSpace (ProdLp p α β) := MetricSpace.ofT0PseudoMetricSpace _ + +variable {p α β} + +theorem nndist_eq_sum [PseudoMetricSpace α] [PseudoMetricSpace β] + (hp : p ≠ ∞) (x y : ProdLp p α β) : + nndist x y = (nndist x.fst y.fst ^ p.toReal + nndist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) := + -- Porting note: was `Subtype.ext` + NNReal.eq <| by + push_cast + exact dist_eq_add (p.toReal_pos_iff_ne_top.mpr hp) _ _ + +theorem nndist_eq_iSup [PseudoMetricSpace α] [PseudoMetricSpace β] (x y : ProdLp ∞ α β) : + nndist x y = nndist x.fst y.fst ⊔ nndist x.snd y.snd := + -- Porting note: was `Subtype.ext` + NNReal.eq <| by + push_cast + exact dist_eq_sup _ _ + +variable (p α β) + +theorem lipschitzWith_equiv [PseudoEMetricSpace α] [PseudoEMetricSpace β] : + LipschitzWith 1 (ProdLp.equiv p α β) := + lipschitzWith_equiv_aux p α β + +theorem antilipschitzWith_equiv [PseudoEMetricSpace α] [PseudoEMetricSpace β] : + AntilipschitzWith ((2 : ℝ≥0) ^ (1 / p).toReal) (ProdLp.equiv p α β) := + antilipschitzWith_equiv_aux p α β + +theorem infty_equiv_isometry [PseudoEMetricSpace α] [PseudoEMetricSpace β] : + Isometry (ProdLp.equiv ∞ α β) := + fun x y => + le_antisymm (by simpa only [ENNReal.coe_one, one_mul] using lipschitzWith_equiv ∞ α β x y) + (by + simpa only [ENNReal.div_top, ENNReal.zero_toReal, NNReal.rpow_zero, ENNReal.coe_one, + one_mul] using antilipschitzWith_equiv ∞ α β x y) From bc719aa7602f167cd9f1eb8ab8c1d0af93829f43 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 2 Aug 2023 23:28:02 +1000 Subject: [PATCH 05/39] more stuff --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 136 ++++++++++++++++++++++- 1 file changed, 134 insertions(+), 2 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index ebef8ff23959e..9b5e6302d1224 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -172,7 +172,7 @@ instance instNorm : Norm (ProdLp p α β) where else if p = ∞ then ‖f.fst‖ ⊔ ‖f.snd‖ else (‖f.fst‖ ^ p.toReal + ‖f.snd‖ ^ p.toReal) ^ (1 / p.toReal) -variable {p β} +variable {p α β} theorem norm_eq_card (f : ProdLp 0 α β) : ‖f‖ = 0 := if_pos rfl @@ -207,7 +207,7 @@ explaining why having definitionally the right uniformity is often important. -/ -variable [Fact (1 ≤ p)] --[∀ i, PseudoMetricSpace (α i)] [∀ i, PseudoEMetricSpace (β i)] +variable [Fact (1 ≤ p)] /-- Endowing the space `PiLp 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 @@ -449,3 +449,135 @@ theorem infty_equiv_isometry [PseudoEMetricSpace α] [PseudoEMetricSpace β] : (by simpa only [ENNReal.div_top, ENNReal.zero_toReal, NNReal.rpow_zero, ENNReal.coe_one, one_mul] using antilipschitzWith_equiv ∞ α β x y) + +/-- seminormed group instance on the product of finitely many normed groups, using the `L^p` +norm. -/ +instance instSeminormedAddCommGroup [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] : + SeminormedAddCommGroup (ProdLp p α β) := + { Prod.instAddCommGroupSum with + dist_eq := fun x y => by + rcases p.dichotomy with (rfl | h) + · simp only [dist_eq_sup, norm_eq_sup, dist_eq_norm] + rfl + · have : p ≠ ∞ := by + intro hp + rw [hp, ENNReal.top_toReal] at h + linarith + simp only [dist_eq_add (zero_lt_one.trans_le h), norm_eq_add (zero_lt_one.trans_le h), + dist_eq_norm] + rfl } + +/-- normed group instance on the product of finitely many normed groups, using the `L^p` norm. -/ +instance normedAddCommGroup [NormedAddCommGroup α] [NormedAddCommGroup β] : + NormedAddCommGroup (ProdLp p α β) := + { ProdLp.instSeminormedAddCommGroup p α β with + eq_of_dist_eq_zero := eq_of_dist_eq_zero } + +section norm_of + +variable {p α β} +variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] + +theorem nnnorm_eq_add (hp : p ≠ ∞) (f : ProdLp p α β) : + ‖f‖₊ = (‖f.fst‖₊ ^ p.toReal + ‖f.snd‖₊ ^ p.toReal) ^ (1 / p.toReal) := by + ext + simp [norm_eq_add (p.toReal_pos_iff_ne_top.mpr hp)] + +theorem nnnorm_eq_sup (f : ProdLp ∞ α β) : ‖f‖₊ = ‖f.fst‖₊ ⊔ ‖f.snd‖₊ := by + ext + norm_cast + +theorem norm_eq_of_nat (n : ℕ) (h : p = n) (f : ProdLp p α β) : + ‖f‖ = (‖f.fst‖ ^ n + ‖f.snd‖ ^ n) ^ (1 / (n : ℝ)) := by + have := p.toReal_pos_iff_ne_top.mpr (ne_of_eq_of_ne h <| ENNReal.nat_ne_top n) + simp only [one_div, h, Real.rpow_nat_cast, ENNReal.toReal_nat, eq_self_iff_true, Finset.sum_congr, + norm_eq_add this] + +theorem norm_eq_of_L2 (x : ProdLp 2 α β) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by + rw [norm_eq_of_nat 2 (by norm_cast) _] -- Porting note: was `convert` + rw [Real.sqrt_eq_rpow] + norm_cast + +theorem nnnorm_eq_of_L2 (x : ProdLp 2 α β) : ‖x‖₊ = NNReal.sqrt (‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2) := + -- Porting note: was `Subtype.ext` + NNReal.eq <| by + push_cast + exact norm_eq_of_L2 x + +variable (α β) + +theorem norm_sq_eq_of_L2 (x : ProdLp 2 α β) : ‖x‖ ^ 2 = ‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2 := by + suffices ‖x‖₊ ^ 2 = ‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2 by + simpa only [NNReal.coe_sum] using congr_arg ((↑) : ℝ≥0 → ℝ) this + rw [nnnorm_eq_of_L2, NNReal.sq_sqrt] + +variable {α β} + +theorem dist_eq_of_L2 (x y : ProdLp 2 α β) : + dist x y = (dist x.fst y.fst ^ 2 + dist x.snd y.snd ^ 2).sqrt := by + simp_rw [dist_eq_norm, norm_eq_of_L2, Pi.sub_apply] + rfl -- Porting note: `Pi.sub_apply` doesn't work + +theorem nndist_eq_of_L2 (x y : ProdLp 2 α β) : + nndist x y = NNReal.sqrt (nndist x.fst y.fst ^ 2 + nndist x.snd y.snd ^ 2) := + -- Porting note: was `Subtype.ext` + NNReal.eq <| by + push_cast + exact dist_eq_of_L2 _ _ + +theorem edist_eq_of_L2 (x y : ProdLp 2 α β) : + edist x y = (edist x.fst y.fst ^ 2 + edist x.snd y.snd ^ 2) ^ (1 / 2 : ℝ) := by + simp [ProdLp.edist_eq_add] + +end norm_of + +variable [NormedField 𝕜] [NormedField 𝕜'] + +section normed_space_inst + +variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] + [SeminormedAddCommGroup β] [NormedSpace 𝕜 β] + +-- Porting note: added +instance instModule : Module 𝕜 (ProdLp p α β) := + { Prod.module with } + +/-- The product of finitely many normed spaces is a normed space, with the `L^p` norm. -/ +instance instNormedSpace : + NormedSpace 𝕜 (ProdLp p α β) := + { instModule p 𝕜 α β with + norm_smul_le := fun c f => by + rcases p.dichotomy with (rfl | hp) + · letI : Module 𝕜 (ProdLp ∞ α β) := Prod.module + suffices ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊ by exact_mod_cast NNReal.coe_mono this.le + simp only [nnnorm_eq_sup, NNReal.mul_sup, ← nnnorm_smul] + -- Porting note: added + congr + · have : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le hp).ne' + -- Porting note: added to replace Pi.smul_apply + --have smul_apply : ∀ i : ι, (c • f) i = c • (f i) := fun i => rfl + simp only [norm_eq_add (zero_lt_one.trans_le hp), norm_smul, Real.mul_rpow, norm_nonneg] + rw [mul_rpow (rpow_nonneg_of_nonneg (norm_nonneg _) _), ← rpow_mul (norm_nonneg _), this, + Real.rpow_one] + exact Finset.sum_nonneg fun i _ => rpow_nonneg_of_nonneg (norm_nonneg _) _ } + +section towers + +variable [NormedSpace 𝕜' α] [NormedSpace 𝕜' β] + +instance isScalarTower [SMul 𝕜 𝕜'] [IsScalarTower 𝕜 𝕜' α] [IsScalarTower 𝕜 𝕜' β] : + IsScalarTower 𝕜 𝕜' (ProdLp p α β) := + Prod.isScalarTower + +instance smulCommClass [SMulCommClass 𝕜 𝕜' α] [SMulCommClass 𝕜 𝕜' β] : + SMulCommClass 𝕜 𝕜' (ProdLp p α β) := + Prod.smulCommClass + +end towers + +instance finiteDimensional [FiniteDimensional 𝕜 α] [FiniteDimensional 𝕜 α] : + FiniteDimensional 𝕜 (ProdLp p α β) := + FiniteDimensional.finiteDimensional_prod' _ _ + + +end normed_space_inst From ec7429ca1260a1a0c9318ba50222263e19cdb3ee Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Fri, 4 Aug 2023 19:33:20 +1000 Subject: [PATCH 06/39] fixes --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 20 +++++++++++--------- 1 file changed, 11 insertions(+), 9 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 9b5e6302d1224..340f0c89cf342 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -555,29 +555,31 @@ instance instNormedSpace : congr · have : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le hp).ne' -- Porting note: added to replace Pi.smul_apply - --have smul_apply : ∀ i : ι, (c • f) i = c • (f i) := fun i => rfl - simp only [norm_eq_add (zero_lt_one.trans_le hp), norm_smul, Real.mul_rpow, norm_nonneg] - rw [mul_rpow (rpow_nonneg_of_nonneg (norm_nonneg _) _), ← rpow_mul (norm_nonneg _), this, - Real.rpow_one] - exact Finset.sum_nonneg fun i _ => rpow_nonneg_of_nonneg (norm_nonneg _) _ } + have smul_fst : (c • f).fst = c • f.fst := rfl + have smul_snd : (c • f).snd = c • f.snd := rfl + simp only [norm_eq_add (zero_lt_one.trans_le hp), norm_smul, Real.mul_rpow, norm_nonneg, + smul_fst, smul_snd] + rw [← mul_add, mul_rpow (rpow_nonneg_of_nonneg (norm_nonneg _) _), + ← rpow_mul (norm_nonneg _), this, Real.rpow_one] + positivity } section towers variable [NormedSpace 𝕜' α] [NormedSpace 𝕜' β] -instance isScalarTower [SMul 𝕜 𝕜'] [IsScalarTower 𝕜 𝕜' α] [IsScalarTower 𝕜 𝕜' β] : +instance instIsScalarTower [SMul 𝕜 𝕜'] [IsScalarTower 𝕜 𝕜' α] [IsScalarTower 𝕜 𝕜' β] : IsScalarTower 𝕜 𝕜' (ProdLp p α β) := Prod.isScalarTower -instance smulCommClass [SMulCommClass 𝕜 𝕜' α] [SMulCommClass 𝕜 𝕜' β] : +instance instSMulCommClass [SMulCommClass 𝕜 𝕜' α] [SMulCommClass 𝕜 𝕜' β] : SMulCommClass 𝕜 𝕜' (ProdLp p α β) := Prod.smulCommClass end towers -instance finiteDimensional [FiniteDimensional 𝕜 α] [FiniteDimensional 𝕜 α] : +instance instFiniteDimensional [FiniteDimensional 𝕜 α] [FiniteDimensional 𝕜 β] : FiniteDimensional 𝕜 (ProdLp p α β) := - FiniteDimensional.finiteDimensional_prod' _ _ + Module.Finite.prod end normed_space_inst From b70be84c6153cd57c25de74f1fa824d6a91abf51 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Fri, 4 Aug 2023 22:07:15 +1000 Subject: [PATCH 07/39] stuff --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 311 +++++++++++++++++++++-- 1 file changed, 296 insertions(+), 15 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 340f0c89cf342..8b1f88d3abe0e 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -33,7 +33,7 @@ namespace ProdLp variable (p : ℝ≥0∞) (𝕜 𝕜' : Type _) (α β : Type _) -/-- Canonical bijection between `PiLp p α` and the original Pi type. We introduce it to be able +/-- Canonical bijection between `ProdLp p α β` and the original Pi type. We introduce it to be able to compare the `L^p` and `L^∞` distances through it. -/ protected def equiv : ProdLp p α β ≃ α × β := Equiv.refl _ @@ -52,9 +52,9 @@ theorem equiv_symm_apply (x : α × β) : (ProdLp.equiv p α β).symm x = x := section DistNorm /-! -### Definition of `edist`, `dist` and `norm` on `PiLp` +### Definition of `edist`, `dist` and `norm` on `ProdLp` -In this section we define the `edist`, `dist` and `norm` functions on `PiLp p α` without assuming +In this section we define the `edist`, `dist` and `norm` functions on `ProdLp 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`. -/ @@ -125,11 +125,11 @@ section Dist variable [Dist α] [Dist β] -/-- Endowing the space `PiLp p β` with the `L^p` distance. We register this instance -separate from `pi_Lp.pseudo_metric` since the latter requires the type class hypothesis +/-- Endowing the space `ProdLp p α β` with the `L^p` distance. We register this instance +separate from `ProdLp.instPseudoMetricSpace` 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 `PiLp p β` for `p < 1` +Registering this separately allows for a future metric-like structure on `ProdLp 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*. -/ instance instDist : Dist (ProdLp p α β) where @@ -160,11 +160,11 @@ section Norm variable [Norm α] [Zero α] [Norm β] [Zero β] -/-- Endowing the space `PiLp p β` with the `L^p` norm. We register this instance -separate from `PiLp.seminormedAddCommGroup` since the latter requires the type class hypothesis -`[Fact (1 ≤ p)]` in order to prove the triangle inequality. +/-- Endowing the space `ProdLp p α β` with the `L^p` norm. We register this instance +separate from `ProdLp.instSeminormedAddCommGroup` 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 `PiLp p β` for `p < 1` +Registering this separately allows for a future norm-like structure on `ProdLp p α β` for `p < 1` satisfying a relaxed triangle inequality. These are called *quasi-norms*. -/ instance instNorm : Norm (ProdLp p α β) where norm f := @@ -195,7 +195,7 @@ section Aux /-! ### The uniformity on finite `L^p` products is the product uniformity -In this section, we put the `L^p` edistance on `PiLp p α`, and we check that the uniformity +In this section, we put the `L^p` edistance on `ProdLp 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. @@ -209,7 +209,7 @@ explaining why having definitionally the right uniformity is often important. variable [Fact (1 ≤ p)] -/-- Endowing the space `PiLp p β` with the `L^p` pseudoemetric structure. This definition is not +/-- Endowing the space `ProdLp 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 @@ -244,7 +244,7 @@ def pseudoEmetricAux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : attribute [local instance] ProdLp.pseudoEmetricAux -/-- An auxiliary lemma used twice in the proof of `PiLp.pseudoMetricAux` below. Not intended for +/-- An auxiliary lemma used twice in the proof of `ProdLp.pseudoMetricAux` below. Not intended for use outside this file. -/ theorem sup_edist_ne_top_aux {α β : Type _} [PseudoMetricSpace α] [PseudoMetricSpace β] (f g : ProdLp ∞ α β) : @@ -252,7 +252,7 @@ theorem sup_edist_ne_top_aux {α β : Type _} refine ne_of_lt ?_ simp [edist, PseudoMetricSpace.edist_dist] -/-- Endowing the space `PiLp p α` with the `L^p` pseudometric structure. This definition is not +/-- Endowing the space `ProdLp 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 @@ -581,5 +581,286 @@ instance instFiniteDimensional [FiniteDimensional 𝕜 α] [FiniteDimensional FiniteDimensional 𝕜 (ProdLp p α β) := Module.Finite.prod - end normed_space_inst + +/- Register simplification lemmas for the applications of `ProdLp` elements, as the usual lemmas +for Pi types will not trigger. -/ +variable {𝕜 𝕜' p α β} +variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] + [SeminormedAddCommGroup β] [NormedSpace 𝕜 β] + +section algebra + +variable (x y : ProdLp p α β) (c : 𝕜) + +@[simp] +theorem zero_fst : (0 : ProdLp p α β).fst = 0 := + rfl + +@[simp] +theorem zero_snd : (0 : ProdLp p α β).snd = 0 := + rfl + +@[simp] +theorem add_fst : (x + y).fst = x.fst + y.fst := + rfl + +@[simp] +theorem add_snd : (x + y).snd = x.snd + y.snd := + rfl + +@[simp] +theorem sub_fst : (x - y).fst = x.fst - y.fst := + rfl + +@[simp] +theorem sub_snd : (x - y).snd = x.snd - y.snd := + rfl + +@[simp] +theorem smul_fst : (c • x).fst = c • x.fst := + rfl + +@[simp] +theorem smul_snd : (c • x).snd = c • x.snd := + rfl + +@[simp] +theorem neg_fst : (-x).fst = -x.fst := + rfl + +@[simp] +theorem neg_snd : (-x).snd = -x.snd := + rfl + +end algebra + +section Equiv + +/-- The canonical map `ProdLp.equiv` between `ProdLp ∞ β` and `Π i, β i` as a linear isometric +equivalence. -/ +def equivₗᵢ : ProdLp ∞ α β ≃ₗᵢ[𝕜] α × β := + { ProdLp.equiv ∞ α β with + map_add' := fun f g => rfl + map_smul' := fun c f => rfl + norm_map' := fun f => by simp } + +variable (x y : ProdLp p α β) (x' y' : α × β) (c : 𝕜) + +@[simp] +theorem equiv_zero : ProdLp.equiv p α β 0 = 0 := + rfl + +@[simp] +theorem equiv_symm_zero : (ProdLp.equiv p α β).symm 0 = 0 := + rfl + +@[simp] +theorem equiv_add : ProdLp.equiv p α β (x + y) = ProdLp.equiv p α β x + ProdLp.equiv p α β y := + rfl + +@[simp] +theorem equiv_symm_add : (ProdLp.equiv p α β).symm (x' + y') = + (ProdLp.equiv p α β).symm x' + (ProdLp.equiv p α β).symm y' := + rfl + +@[simp] +theorem equiv_sub : ProdLp.equiv p α β (x - y) = ProdLp.equiv p α β x - ProdLp.equiv p α β y := + rfl + +@[simp] +theorem equiv_symm_sub : (ProdLp.equiv p α β).symm (x' - y') = + (ProdLp.equiv p α β).symm x' - (ProdLp.equiv p α β).symm y' := + rfl + +@[simp] +theorem equiv_neg : ProdLp.equiv p α β (-x) = -ProdLp.equiv p α β x := + rfl + +@[simp] +theorem equiv_symm_neg : (ProdLp.equiv p α β).symm (-x') = -(ProdLp.equiv p α β).symm x' := + rfl + +@[simp] +theorem equiv_smul : ProdLp.equiv p α β (c • x) = c • ProdLp.equiv p α β x := + rfl + +@[simp] +theorem equiv_symm_smul : (ProdLp.equiv p α β).symm (c • x') = c • (ProdLp.equiv p α β).symm x' := + rfl + +end Equiv + +#exit + +section Single + +variable (p) + +variable [DecidableEq ι] + +-- Porting note: added `hp` +@[simp] +theorem nnnorm_equiv_symm_single [hp : Fact (1 ≤ p)] (i : ι) (b : β i) : + ‖(ProdLp.equiv p β).symm (Pi.single i b)‖₊ = ‖b‖₊ := by + clear x y -- Porting note: added + haveI : Nonempty ι := ⟨i⟩ + induction p using ENNReal.recTopCoe generalizing hp with + | top => + simp_rw [nnnorm_eq_ciSup, equiv_symm_apply] + refine' ciSup_eq_of_forall_le_of_forall_lt_exists_gt (fun j => _) fun n hn => ⟨i, hn.trans_eq _⟩ + · obtain rfl | hij := Decidable.eq_or_ne i j + · rw [Pi.single_eq_same] + · rw [Pi.single_eq_of_ne' hij, nnnorm_zero] + exact zero_le _ + · rw [Pi.single_eq_same] + | coe p => + have hp0 : (p : ℝ) ≠ 0 := by + exact_mod_cast (zero_lt_one.trans_le <| Fact.out (p := 1 ≤ (p : ℝ≥0∞))).ne' + rw [nnnorm_eq_sum ENNReal.coe_ne_top, ENNReal.coe_toReal, Fintype.sum_eq_single i, + equiv_symm_apply, Pi.single_eq_same, ← NNReal.rpow_mul, one_div, mul_inv_cancel hp0, + NNReal.rpow_one] + intro j hij + rw [equiv_symm_apply, Pi.single_eq_of_ne hij, nnnorm_zero, NNReal.zero_rpow hp0] + +@[simp] +theorem norm_equiv_symm_single (i : ι) (b : β i) : ‖(ProdLp.equiv p β).symm (Pi.single i b)‖ = ‖b‖ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_single p β i b + +@[simp] +theorem nndist_equiv_symm_single_same (i : ι) (b₁ b₂ : β i) : + nndist ((ProdLp.equiv p β).symm (Pi.single i b₁)) ((ProdLp.equiv p β).symm (Pi.single i b₂)) = + nndist b₁ b₂ := by + rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← equiv_symm_sub, ← Pi.single_sub, + nnnorm_equiv_symm_single] + +@[simp] +theorem dist_equiv_symm_single_same (i : ι) (b₁ b₂ : β i) : + dist ((ProdLp.equiv p β).symm (Pi.single i b₁)) ((ProdLp.equiv p β).symm (Pi.single i b₂)) = + dist b₁ b₂ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_same p β i b₁ b₂ + +@[simp] +theorem edist_equiv_symm_single_same (i : ι) (b₁ b₂ : β i) : + edist ((ProdLp.equiv p β).symm (Pi.single i b₁)) ((ProdLp.equiv p β).symm (Pi.single i b₂)) = + edist b₁ b₂ := by + -- Porting note: was `simpa using` + simp only [edist_nndist, nndist_equiv_symm_single_same p β i b₁ b₂] + +end Single + +/-- 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 +`ProdLp.nnnorm_equiv_symm_const'` for a version which exchanges the hypothesis `p ≠ ∞` for +`Nonempty ι`. -/ +theorem nnnorm_equiv_symm_const {β} [SeminormedAddCommGroup β] (hp : p ≠ ∞) (b : β) : + ‖(ProdLp.equiv p fun _ : ι => β).symm (Function.const _ b)‖₊ = + (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖b‖₊ := by + rcases p.dichotomy with (h | h) + · exact False.elim (hp h) + · have ne_zero : p.toReal ≠ 0 := (zero_lt_one.trans_le h).ne' + simp_rw [nnnorm_eq_sum hp, equiv_symm_apply, Function.const_apply, Finset.sum_const, + Finset.card_univ, nsmul_eq_mul, NNReal.mul_rpow, ← NNReal.rpow_mul, + mul_one_div_cancel ne_zero, NNReal.rpow_one, ENNReal.toReal_div, ENNReal.one_toReal] + +/-- When `IsEmpty ι`, 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 +`ProdLp.nnnorm_equiv_symm_const` for a version which exchanges the hypothesis `Nonempty ι`. +for `p ≠ ∞`. -/ +theorem nnnorm_equiv_symm_const' {β} [SeminormedAddCommGroup β] [Nonempty ι] (b : β) : + ‖(ProdLp.equiv p fun _ : ι => β).symm (Function.const _ b)‖₊ = + (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖b‖₊ := by + clear x y -- Porting note: added to avoid spurious arguments + rcases em <| p = ∞ with (rfl | hp) + · simp only [equiv_symm_apply, ENNReal.div_top, ENNReal.zero_toReal, NNReal.rpow_zero, one_mul, + nnnorm_eq_ciSup, Function.const_apply, ciSup_const] + · exact nnnorm_equiv_symm_const hp 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 +`ProdLp.norm_equiv_symm_const'` for a version which exchanges the hypothesis `p ≠ ∞` for +`Nonempty ι`. -/ +theorem norm_equiv_symm_const {β} [SeminormedAddCommGroup β] (hp : p ≠ ∞) (b : β) : + ‖(ProdLp.equiv p fun _ : ι => β).symm (Function.const _ b)‖ = + (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖b‖ := + (congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_const hp b).trans <| by simp + +/-- When `IsEmpty ι`, 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 +`ProdLp.norm_equiv_symm_const` for a version which exchanges the hypothesis `Nonempty ι`. +for `p ≠ ∞`. -/ +theorem norm_equiv_symm_const' {β} [SeminormedAddCommGroup β] [Nonempty ι] (b : β) : + ‖(ProdLp.equiv p fun _ : ι => β).symm (Function.const _ b)‖ = + (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖b‖ := + (congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_const' b).trans <| by simp + +theorem nnnorm_equiv_symm_one {β} [SeminormedAddCommGroup β] (hp : p ≠ ∞) [One β] : + ‖(ProdLp.equiv p fun _ : ι => β).symm 1‖₊ = + (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖(1 : β)‖₊ := + (nnnorm_equiv_symm_const hp (1 : β)).trans rfl + +theorem norm_equiv_symm_one {β} [SeminormedAddCommGroup β] (hp : p ≠ ∞) [One β] : + ‖(ProdLp.equiv p fun _ : ι => β).symm 1‖ = (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖(1 : β)‖ := + (norm_equiv_symm_const hp (1 : β)).trans rfl + +variable (𝕜 p) + +/-- `ProdLp.equiv` as a linear equivalence. -/ +@[simps (config := { fullyApplied := false })] +protected def linearEquiv : ProdLp p β ≃ₗ[𝕜] ∀ i, β i := + { LinearEquiv.refl _ _ with + toFun := ProdLp.equiv _ _ + invFun := (ProdLp.equiv _ _).symm } + +/-- `ProdLp.equiv` as a continuous linear equivalence. -/ +@[simps! (config := { fullyApplied := false }) apply symm_apply] +protected def continuousLinearEquiv : ProdLp p β ≃L[𝕜] ∀ i, β i where + toLinearEquiv := ProdLp.linearEquiv _ _ _ + continuous_toFun := continuous_equiv _ _ + continuous_invFun := continuous_equiv_symm _ _ + +section Basis + +variable (ι) + +/-- A version of `Pi.basisFun` for `ProdLp`. -/ +def basisFun : Basis ι 𝕜 (ProdLp p fun _ : ι => 𝕜) := + Basis.ofEquivFun (ProdLp.linearEquiv p 𝕜 fun _ : ι => 𝕜) + +@[simp] +theorem basisFun_apply [DecidableEq ι] (i) : + basisFun p 𝕜 ι i = (ProdLp.equiv p _).symm (Pi.single i 1) := by + simp_rw [basisFun, Basis.coe_ofEquivFun, ProdLp.linearEquiv_symm_apply, Pi.single] + +@[simp] +theorem basisFun_repr (x : ProdLp p fun _ : ι => 𝕜) (i : ι) : (basisFun p 𝕜 ι).repr x i = x i := + rfl + +@[simp] +theorem basisFun_equivFun : (basisFun p 𝕜 ι).equivFun = ProdLp.linearEquiv p 𝕜 fun _ : ι => 𝕜 := + Basis.equivFun_ofEquivFun _ + +theorem basisFun_eq_pi_basisFun : + basisFun p 𝕜 ι = (Pi.basisFun 𝕜 ι).map (ProdLp.linearEquiv p 𝕜 fun _ : ι => 𝕜).symm := + rfl + +@[simp] +theorem basisFun_map : + (basisFun p 𝕜 ι).map (ProdLp.linearEquiv p 𝕜 fun _ : ι => 𝕜) = Pi.basisFun 𝕜 ι := + rfl + +open Matrix + +nonrec theorem basis_toMatrix_basisFun_mul (b : Basis ι 𝕜 (ProdLp p fun _ : ι => 𝕜)) + (A : Matrix ι ι 𝕜) : + b.toMatrix (ProdLp.basisFun _ _ _) ⬝ A = + Matrix.of fun i j => b.repr ((ProdLp.equiv _ _).symm (Aᵀ j)) i := by + have := basis_toMatrix_basisFun_mul (b.map (ProdLp.linearEquiv _ 𝕜 _)) A + simp_rw [← ProdLp.basisFun_map p, Basis.map_repr, LinearEquiv.trans_apply, + ProdLp.linearEquiv_symm_apply, Basis.toMatrix_map, Function.comp, Basis.map_apply, + LinearEquiv.symm_apply_apply] at this + exact this + +end Basis + +end ProdLp From 44021a8631f576a0e4e305897a3ae8a550c3e153 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sat, 5 Aug 2023 02:40:37 +1000 Subject: [PATCH 08/39] last bit for this file? --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 203 +++++++---------------- 1 file changed, 60 insertions(+), 143 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 8b1f88d3abe0e..e1ed809b48ec3 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -73,7 +73,6 @@ satisfying a relaxed triangle inequality. The terminology for this varies throug literature, but it is sometimes called a *quasi-metric* or *semi-metric*. -/ instance instEDist : EDist (ProdLp p α β) where edist f g := - -- Porting note: can we drop the `_hp` entirely? if _hp : p = 0 then 0 else if p = ∞ then edist f.fst g.fst ⊔ edist f.snd g.snd @@ -691,176 +690,94 @@ theorem equiv_symm_smul : (ProdLp.equiv p α β).symm (c • x') = c • (ProdLp end Equiv -#exit - section Single -variable (p) - -variable [DecidableEq ι] +variable (p α β) --- Porting note: added `hp` @[simp] -theorem nnnorm_equiv_symm_single [hp : Fact (1 ≤ p)] (i : ι) (b : β i) : - ‖(ProdLp.equiv p β).symm (Pi.single i b)‖₊ = ‖b‖₊ := by - clear x y -- Porting note: added - haveI : Nonempty ι := ⟨i⟩ +theorem nnnorm_equiv_symm_fst [hp : Fact (1 ≤ p)] (x : α) : + ‖(ProdLp.equiv p α β).symm (x, 0)‖₊ = ‖x‖₊ := by induction p using ENNReal.recTopCoe generalizing hp with | top => - simp_rw [nnnorm_eq_ciSup, equiv_symm_apply] - refine' ciSup_eq_of_forall_le_of_forall_lt_exists_gt (fun j => _) fun n hn => ⟨i, hn.trans_eq _⟩ - · obtain rfl | hij := Decidable.eq_or_ne i j - · rw [Pi.single_eq_same] - · rw [Pi.single_eq_of_ne' hij, nnnorm_zero] - exact zero_le _ - · rw [Pi.single_eq_same] + simp [nnnorm_eq_sup] | coe p => have hp0 : (p : ℝ) ≠ 0 := by exact_mod_cast (zero_lt_one.trans_le <| Fact.out (p := 1 ≤ (p : ℝ≥0∞))).ne' - rw [nnnorm_eq_sum ENNReal.coe_ne_top, ENNReal.coe_toReal, Fintype.sum_eq_single i, - equiv_symm_apply, Pi.single_eq_same, ← NNReal.rpow_mul, one_div, mul_inv_cancel hp0, - NNReal.rpow_one] - intro j hij - rw [equiv_symm_apply, Pi.single_eq_of_ne hij, nnnorm_zero, NNReal.zero_rpow hp0] + simp [nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] @[simp] -theorem norm_equiv_symm_single (i : ι) (b : β i) : ‖(ProdLp.equiv p β).symm (Pi.single i b)‖ = ‖b‖ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_single p β i b +theorem nnnorm_equiv_symm_snd [hp : Fact (1 ≤ p)] (y : β) : + ‖(ProdLp.equiv p α β).symm (0, y)‖₊ = ‖y‖₊ := by + induction p using ENNReal.recTopCoe generalizing hp with + | top => + simp [nnnorm_eq_sup] + | coe p => + have hp0 : (p : ℝ) ≠ 0 := by + exact_mod_cast (zero_lt_one.trans_le <| Fact.out (p := 1 ≤ (p : ℝ≥0∞))).ne' + simp [nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] @[simp] -theorem nndist_equiv_symm_single_same (i : ι) (b₁ b₂ : β i) : - nndist ((ProdLp.equiv p β).symm (Pi.single i b₁)) ((ProdLp.equiv p β).symm (Pi.single i b₂)) = - nndist b₁ b₂ := by - rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← equiv_symm_sub, ← Pi.single_sub, - nnnorm_equiv_symm_single] +theorem norm_equiv_symm_fst (x : α) : ‖(ProdLp.equiv p α β).symm (x, 0)‖ = ‖x‖ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_fst p α β x @[simp] -theorem dist_equiv_symm_single_same (i : ι) (b₁ b₂ : β i) : - dist ((ProdLp.equiv p β).symm (Pi.single i b₁)) ((ProdLp.equiv p β).symm (Pi.single i b₂)) = - dist b₁ b₂ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_same p β i b₁ b₂ +theorem norm_equiv_symm_snd (y : β) : ‖(ProdLp.equiv p α β).symm (0, y)‖ = ‖y‖ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_snd p α β y @[simp] -theorem edist_equiv_symm_single_same (i : ι) (b₁ b₂ : β i) : - edist ((ProdLp.equiv p β).symm (Pi.single i b₁)) ((ProdLp.equiv p β).symm (Pi.single i b₂)) = - edist b₁ b₂ := by - -- Porting note: was `simpa using` - simp only [edist_nndist, nndist_equiv_symm_single_same p β i b₁ b₂] - -end Single - -/-- 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 -`ProdLp.nnnorm_equiv_symm_const'` for a version which exchanges the hypothesis `p ≠ ∞` for -`Nonempty ι`. -/ -theorem nnnorm_equiv_symm_const {β} [SeminormedAddCommGroup β] (hp : p ≠ ∞) (b : β) : - ‖(ProdLp.equiv p fun _ : ι => β).symm (Function.const _ b)‖₊ = - (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖b‖₊ := by - rcases p.dichotomy with (h | h) - · exact False.elim (hp h) - · have ne_zero : p.toReal ≠ 0 := (zero_lt_one.trans_le h).ne' - simp_rw [nnnorm_eq_sum hp, equiv_symm_apply, Function.const_apply, Finset.sum_const, - Finset.card_univ, nsmul_eq_mul, NNReal.mul_rpow, ← NNReal.rpow_mul, - mul_one_div_cancel ne_zero, NNReal.rpow_one, ENNReal.toReal_div, ENNReal.one_toReal] - -/-- When `IsEmpty ι`, 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 -`ProdLp.nnnorm_equiv_symm_const` for a version which exchanges the hypothesis `Nonempty ι`. -for `p ≠ ∞`. -/ -theorem nnnorm_equiv_symm_const' {β} [SeminormedAddCommGroup β] [Nonempty ι] (b : β) : - ‖(ProdLp.equiv p fun _ : ι => β).symm (Function.const _ b)‖₊ = - (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖b‖₊ := by - clear x y -- Porting note: added to avoid spurious arguments - rcases em <| p = ∞ with (rfl | hp) - · simp only [equiv_symm_apply, ENNReal.div_top, ENNReal.zero_toReal, NNReal.rpow_zero, one_mul, - nnnorm_eq_ciSup, Function.const_apply, ciSup_const] - · exact nnnorm_equiv_symm_const hp 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 -`ProdLp.norm_equiv_symm_const'` for a version which exchanges the hypothesis `p ≠ ∞` for -`Nonempty ι`. -/ -theorem norm_equiv_symm_const {β} [SeminormedAddCommGroup β] (hp : p ≠ ∞) (b : β) : - ‖(ProdLp.equiv p fun _ : ι => β).symm (Function.const _ b)‖ = - (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖b‖ := - (congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_const hp b).trans <| by simp - -/-- When `IsEmpty ι`, 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 -`ProdLp.norm_equiv_symm_const` for a version which exchanges the hypothesis `Nonempty ι`. -for `p ≠ ∞`. -/ -theorem norm_equiv_symm_const' {β} [SeminormedAddCommGroup β] [Nonempty ι] (b : β) : - ‖(ProdLp.equiv p fun _ : ι => β).symm (Function.const _ b)‖ = - (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖b‖ := - (congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_const' b).trans <| by simp - -theorem nnnorm_equiv_symm_one {β} [SeminormedAddCommGroup β] (hp : p ≠ ∞) [One β] : - ‖(ProdLp.equiv p fun _ : ι => β).symm 1‖₊ = - (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖(1 : β)‖₊ := - (nnnorm_equiv_symm_const hp (1 : β)).trans rfl - -theorem norm_equiv_symm_one {β} [SeminormedAddCommGroup β] (hp : p ≠ ∞) [One β] : - ‖(ProdLp.equiv p fun _ : ι => β).symm 1‖ = (Fintype.card ι : ℝ≥0) ^ (1 / p).toReal * ‖(1 : β)‖ := - (norm_equiv_symm_const hp (1 : β)).trans rfl - -variable (𝕜 p) - -/-- `ProdLp.equiv` as a linear equivalence. -/ -@[simps (config := { fullyApplied := false })] -protected def linearEquiv : ProdLp p β ≃ₗ[𝕜] ∀ i, β i := - { LinearEquiv.refl _ _ with - toFun := ProdLp.equiv _ _ - invFun := (ProdLp.equiv _ _).symm } - -/-- `ProdLp.equiv` as a continuous linear equivalence. -/ -@[simps! (config := { fullyApplied := false }) apply symm_apply] -protected def continuousLinearEquiv : ProdLp p β ≃L[𝕜] ∀ i, β i where - toLinearEquiv := ProdLp.linearEquiv _ _ _ - continuous_toFun := continuous_equiv _ _ - continuous_invFun := continuous_equiv_symm _ _ - -section Basis - -variable (ι) - -/-- A version of `Pi.basisFun` for `ProdLp`. -/ -def basisFun : Basis ι 𝕜 (ProdLp p fun _ : ι => 𝕜) := - Basis.ofEquivFun (ProdLp.linearEquiv p 𝕜 fun _ : ι => 𝕜) +theorem nndist_equiv_symm_single_fst (x₁ x₂ : α) : + nndist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = + nndist x₁ x₂ := by + rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← equiv_symm_sub, Prod.mk_sub_mk, sub_zero, + nnnorm_equiv_symm_fst] @[simp] -theorem basisFun_apply [DecidableEq ι] (i) : - basisFun p 𝕜 ι i = (ProdLp.equiv p _).symm (Pi.single i 1) := by - simp_rw [basisFun, Basis.coe_ofEquivFun, ProdLp.linearEquiv_symm_apply, Pi.single] +theorem nndist_equiv_symm_single_snd (y₁ y₂ : β) : + nndist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = + nndist y₁ y₂ := by + rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← equiv_symm_sub, Prod.mk_sub_mk, sub_zero, + nnnorm_equiv_symm_snd] @[simp] -theorem basisFun_repr (x : ProdLp p fun _ : ι => 𝕜) (i : ι) : (basisFun p 𝕜 ι).repr x i = x i := - rfl +theorem dist_equiv_symm_single_fst (x₁ x₂ : α) : + dist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = + dist x₁ x₂ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_fst p α β x₁ x₂ @[simp] -theorem basisFun_equivFun : (basisFun p 𝕜 ι).equivFun = ProdLp.linearEquiv p 𝕜 fun _ : ι => 𝕜 := - Basis.equivFun_ofEquivFun _ +theorem dist_equiv_symm_single_snd (y₁ y₂ : β) : + dist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = + dist y₁ y₂ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_snd p α β y₁ y₂ -theorem basisFun_eq_pi_basisFun : - basisFun p 𝕜 ι = (Pi.basisFun 𝕜 ι).map (ProdLp.linearEquiv p 𝕜 fun _ : ι => 𝕜).symm := - rfl +@[simp] +theorem edist_equiv_symm_single_fst (x₁ x₂ : α) : + edist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = + edist x₁ x₂ := by + simp only [edist_nndist, nndist_equiv_symm_single_fst p α β x₁ x₂] @[simp] -theorem basisFun_map : - (basisFun p 𝕜 ι).map (ProdLp.linearEquiv p 𝕜 fun _ : ι => 𝕜) = Pi.basisFun 𝕜 ι := - rfl +theorem edist_equiv_symm_single_snd (y₁ y₂ : β) : + edist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = + edist y₁ y₂ := by + simp only [edist_nndist, nndist_equiv_symm_single_snd p α β y₁ y₂] + +end Single -open Matrix +variable (𝕜 p α β) -nonrec theorem basis_toMatrix_basisFun_mul (b : Basis ι 𝕜 (ProdLp p fun _ : ι => 𝕜)) - (A : Matrix ι ι 𝕜) : - b.toMatrix (ProdLp.basisFun _ _ _) ⬝ A = - Matrix.of fun i j => b.repr ((ProdLp.equiv _ _).symm (Aᵀ j)) i := by - have := basis_toMatrix_basisFun_mul (b.map (ProdLp.linearEquiv _ 𝕜 _)) A - simp_rw [← ProdLp.basisFun_map p, Basis.map_repr, LinearEquiv.trans_apply, - ProdLp.linearEquiv_symm_apply, Basis.toMatrix_map, Function.comp, Basis.map_apply, - LinearEquiv.symm_apply_apply] at this - exact this +/-- `ProdLp.equiv` as a linear equivalence. -/ +@[simps (config := { fullyApplied := false })] +protected def linearEquiv : ProdLp p α β ≃ₗ[𝕜] α × β := + { LinearEquiv.refl _ _ with + toFun := ProdLp.equiv _ _ _ + invFun := (ProdLp.equiv _ _ _).symm } -end Basis +/-- `ProdLp.equiv` as a continuous linear equivalence. -/ +@[simps! (config := { fullyApplied := false }) apply symm_apply] +protected def continuousLinearEquiv : ProdLp p α β ≃L[𝕜] α × β where + toLinearEquiv := ProdLp.linearEquiv _ _ _ _ + continuous_toFun := continuous_equiv _ _ _ + continuous_invFun := continuous_equiv_symm _ _ _ end ProdLp From 86fda9c995230dd1847fee8cb60c8c4a89ba8305 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sat, 5 Aug 2023 14:06:10 +1000 Subject: [PATCH 09/39] L2 inner product instance --- .../Analysis/InnerProductSpace/ProdL2.lean | 45 +++++++++++++++++++ 1 file changed, 45 insertions(+) create mode 100644 Mathlib/Analysis/InnerProductSpace/ProdL2.lean diff --git a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean new file mode 100644 index 0000000000000..24d7279f0f995 --- /dev/null +++ b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean @@ -0,0 +1,45 @@ +/- +Copyright (c) 2023 Moritz Doll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Doll +-/ + +import Mathlib.Analysis.InnerProductSpace.Projection +import Mathlib.Analysis.NormedSpace.ProdLp + +/-! +# `L²` inner product space structure on products of inner product spaces + +The `L²` norm on product of two inner product spaces is compatible with an inner product +$$ +\langle x, y\rangle = \langle x_1, y_1 \rangle + \langle x_2, y_2 \rangle. +$$ +This is recorded in this file as an inner product space instance on `ProdLp 2`. +-/ + +noncomputable section + +variable {𝕜 : Type _} [IsROrC 𝕜] + +instance ProdLp.instInnerProductSpace (E F : Type _) [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] + [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] : + InnerProductSpace 𝕜 (ProdLp 2 E F) where + inner x y := inner x.fst y.fst + inner x.snd y.snd + norm_sq_eq_inner x := by + simp [ProdLp.norm_sq_eq_of_L2, ← norm_sq_eq_inner] + conj_symm x y := by + simp + add_left x y z := by + simp only [add_fst, add_snd, inner_add_left] + ring + smul_left x y r := by + simp only [smul_fst, inner_smul_left, smul_snd] + ring + +variable {E : Type _} [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +variable {F : Type _} [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] + +@[simp] +theorem ProdLp.inner_apply (x y : ProdLp 2 E F) : + @inner 𝕜 _ _ x y = inner x.fst y.fst + inner x.snd y.snd := rfl From 2c9f00878674aea671d49fef1d299040120954fe Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sat, 5 Aug 2023 14:27:54 +1000 Subject: [PATCH 10/39] doc-strings --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 92 ++++++++++++++++-------- 1 file changed, 61 insertions(+), 31 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index e1ed809b48ec3..6ef53eb287c3e 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -3,9 +3,37 @@ Copyright (c) 2023 Moritz Doll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll, Sébastien Gouëzel, Jireh Loreaux -/ + import Mathlib.Analysis.MeanInequalities -import Mathlib.Data.Fintype.Order -import Mathlib.LinearAlgebra.Matrix.Basis + +/-! +# `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. + +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 +`ProdLp 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 `ProdLp 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`. We deviate from +`PiLp` in that we use for `p = 0` the junk value `d(x, y) = 0`. + +-/ local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue #2220 @@ -33,8 +61,8 @@ namespace ProdLp variable (p : ℝ≥0∞) (𝕜 𝕜' : Type _) (α β : Type _) -/-- Canonical bijection between `ProdLp p α β` and the original Pi type. We introduce it to be able -to compare the `L^p` and `L^∞` distances through it. -/ +/-- Canonical bijection between `ProdLp p α β` and the original Prod type. We introduce it to be +able to compare the `L^p` and `L^∞` distances through it. -/ protected def equiv : ProdLp p α β ≃ α × β := Equiv.refl _ @@ -54,9 +82,9 @@ section DistNorm /-! ### Definition of `edist`, `dist` and `norm` on `ProdLp` -In this section we define the `edist`, `dist` and `norm` functions on `ProdLp 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`. +In this section we define the `edist`, `dist` and `norm` functions on `ProdLp 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.to_real`. -/ @@ -82,7 +110,8 @@ variable {α β} variable (x y : ProdLp p α β) (x' : α × β) -theorem edist_eq_card (f g : ProdLp 0 α β) : edist f g = 0 := +@[simp] +protected theorem edist_eq_zero (f g : ProdLp 0 α β) : edist f g = 0 := if_pos rfl theorem edist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : ProdLp p α β) : @@ -103,18 +132,18 @@ variable {α β} variable [PseudoEMetricSpace α] [PseudoEMetricSpace β] /-- 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`. -/ +from `ProdLp.instPseudoEMetricSpace` so it can be used also for `p < 1`. -/ protected theorem edist_self (f : ProdLp p α β) : edist f f = 0 := by rcases p.trichotomy with (rfl | rfl | h) - · simp [edist_eq_card] + · simp · simp [edist_eq_sup] · simp [edist_eq_add h, ENNReal.zero_rpow_of_pos h, ENNReal.zero_rpow_of_pos (inv_pos.2 <| h)] /-- 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`. -/ +from `ProdLp.instPseudoEMetricSpace` so it can be used also for `p < 1`. -/ protected theorem edist_comm (f g : ProdLp p α β) : edist f g = edist g f := by rcases p.trichotomy with (rfl | rfl | h) - · simp only [edist_eq_card, eq_comm, Ne.def] + · simp only [ProdLp.edist_eq_zero, eq_comm, Ne.def] · simp only [edist_eq_sup, edist_comm] · simp only [edist_eq_add h, edist_comm] @@ -167,13 +196,14 @@ Registering this separately allows for a future norm-like structure on `ProdLp p satisfying a relaxed triangle inequality. These are called *quasi-norms*. -/ instance instNorm : Norm (ProdLp p α β) where norm f := - if _hp : p = 0 then 0 -- { i | f i ≠ 0 }.toFinite.toFinset.card + if _hp : p = 0 then 0 else if p = ∞ then ‖f.fst‖ ⊔ ‖f.snd‖ else (‖f.fst‖ ^ p.toReal + ‖f.snd‖ ^ p.toReal) ^ (1 / p.toReal) variable {p α β} -theorem norm_eq_card (f : ProdLp 0 α β) : ‖f‖ = 0 := +@[simp] +protected theorem norm_eq_zero (f : ProdLp 0 α β) : ‖f‖ = 0 := if_pos rfl theorem norm_eq_sup (f : ProdLp ∞ α β) : ‖f‖ = ‖f.fst‖ ⊔ ‖f.snd‖ := by @@ -364,7 +394,7 @@ theorem aux_cobounded_eq [PseudoMetricSpace α] [PseudoMetricSpace β] : end Aux -/-! ### Instances on finite `L^p` products -/ +/-! ### Instances on `L^p` products -/ instance instUniformSpace [UniformSpace α] [UniformSpace β] : UniformSpace (ProdLp p α β) := @@ -393,25 +423,27 @@ instance instBornology [Bornology α] [Bornology β] : Bornology (ProdLp p α β -- throughout the rest of the file, we assume `1 ≤ p` variable [Fact (1 ≤ p)] -/-- pseudoemetric space instance on the product of finitely many pseudoemetric spaces, using the +/-- `PseudoEMetricSpace` instance on the product of two pseudoemetric spaces, using the `L^p` pseudoedistance, and having as uniformity the product uniformity. -/ -instance [PseudoEMetricSpace α] [PseudoEMetricSpace β] : PseudoEMetricSpace (ProdLp p α β) := +instance instPseudoEMetricSpace [PseudoEMetricSpace α] [PseudoEMetricSpace β] : + PseudoEMetricSpace (ProdLp p α β) := (pseudoEmetricAux p α β).replaceUniformity (aux_uniformity_eq p α β).symm -/-- emetric space instance on the product of finitely many emetric spaces, using the `L^p` +/-- `EMetricSpace` instance on the product of two emetric spaces, using the `L^p` edistance, and having as uniformity the product uniformity. -/ -instance [EMetricSpace α] [EMetricSpace β] : EMetricSpace (ProdLp p α β) := +instance instEMetricSpace [EMetricSpace α] [EMetricSpace β] : EMetricSpace (ProdLp p α β) := @EMetricSpace.ofT0PseudoEMetricSpace (ProdLp p α β) _ instT0SpaceProdInstTopologicalSpaceProd -/-- pseudometric space instance on the product of finitely many pseudometric spaces, using the +/-- `PseudoMetricSpace` instance on the product of two pseudometric spaces, using the `L^p` distance, and having as uniformity the product uniformity. -/ -instance [PseudoMetricSpace α] [PseudoMetricSpace β] : PseudoMetricSpace (ProdLp p α β) := +instance instPseudoMetricSpace [PseudoMetricSpace α] [PseudoMetricSpace β] : + PseudoMetricSpace (ProdLp p α β) := ((pseudoMetricAux p α β).replaceUniformity (aux_uniformity_eq p α β).symm).replaceBornology fun s => Filter.ext_iff.1 (aux_cobounded_eq p α β).symm sᶜ -/-- metric space instance on the product of finitely many metric spaces, using the `L^p` distance, +/-- `MetricSpace` instance on the product of two metric spaces, using the `L^p` distance, and having as uniformity the product uniformity. -/ -instance [MetricSpace α] [MetricSpace β] : MetricSpace (ProdLp p α β) := +instance instMetricSpace [MetricSpace α] [MetricSpace β] : MetricSpace (ProdLp p α β) := MetricSpace.ofT0PseudoMetricSpace _ variable {p α β} @@ -449,7 +481,7 @@ theorem infty_equiv_isometry [PseudoEMetricSpace α] [PseudoEMetricSpace β] : simpa only [ENNReal.div_top, ENNReal.zero_toReal, NNReal.rpow_zero, ENNReal.coe_one, one_mul] using antilipschitzWith_equiv ∞ α β x y) -/-- seminormed group instance on the product of finitely many normed groups, using the `L^p` +/-- Seminormed group instance on the product of two normed groups, using the `L^p` norm. -/ instance instSeminormedAddCommGroup [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] : SeminormedAddCommGroup (ProdLp p α β) := @@ -466,7 +498,7 @@ instance instSeminormedAddCommGroup [SeminormedAddCommGroup α] [SeminormedAddCo dist_eq_norm] rfl } -/-- normed group instance on the product of finitely many normed groups, using the `L^p` norm. -/ +/-- normed group instance on the product of two normed groups, using the `L^p` norm. -/ instance normedAddCommGroup [NormedAddCommGroup α] [NormedAddCommGroup β] : NormedAddCommGroup (ProdLp p α β) := { ProdLp.instSeminormedAddCommGroup p α β with @@ -515,7 +547,7 @@ variable {α β} theorem dist_eq_of_L2 (x y : ProdLp 2 α β) : dist x y = (dist x.fst y.fst ^ 2 + dist x.snd y.snd ^ 2).sqrt := by simp_rw [dist_eq_norm, norm_eq_of_L2, Pi.sub_apply] - rfl -- Porting note: `Pi.sub_apply` doesn't work + rfl theorem nndist_eq_of_L2 (x y : ProdLp 2 α β) : nndist x y = NNReal.sqrt (nndist x.fst y.fst ^ 2 + nndist x.snd y.snd ^ 2) := @@ -541,7 +573,7 @@ variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] instance instModule : Module 𝕜 (ProdLp p α β) := { Prod.module with } -/-- The product of finitely many normed spaces is a normed space, with the `L^p` norm. -/ +/-- The product of two normed spaces is a normed space, with the `L^p` norm. -/ instance instNormedSpace : NormedSpace 𝕜 (ProdLp p α β) := { instModule p 𝕜 α β with @@ -550,10 +582,8 @@ instance instNormedSpace : · letI : Module 𝕜 (ProdLp ∞ α β) := Prod.module suffices ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊ by exact_mod_cast NNReal.coe_mono this.le simp only [nnnorm_eq_sup, NNReal.mul_sup, ← nnnorm_smul] - -- Porting note: added - congr + rfl · have : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le hp).ne' - -- Porting note: added to replace Pi.smul_apply have smul_fst : (c • f).fst = c • f.fst := rfl have smul_snd : (c • f).snd = c • f.snd := rfl simp only [norm_eq_add (zero_lt_one.trans_le hp), norm_smul, Real.mul_rpow, norm_nonneg, @@ -583,7 +613,7 @@ instance instFiniteDimensional [FiniteDimensional 𝕜 α] [FiniteDimensional end normed_space_inst /- Register simplification lemmas for the applications of `ProdLp` elements, as the usual lemmas -for Pi types will not trigger. -/ +for Prod types will not trigger. -/ variable {𝕜 𝕜' p α β} variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] [SeminormedAddCommGroup β] [NormedSpace 𝕜 β] From 9160bf3d5969d17df95b624e24700b0e5ecef841 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sat, 5 Aug 2023 14:30:14 +1000 Subject: [PATCH 11/39] import all files --- Mathlib.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib.lean b/Mathlib.lean index 25c6266458e5f..8a1ef8b010ded 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -667,6 +667,7 @@ import Mathlib.Analysis.InnerProductSpace.Orientation import Mathlib.Analysis.InnerProductSpace.Orthogonal import Mathlib.Analysis.InnerProductSpace.PiL2 import Mathlib.Analysis.InnerProductSpace.Positive +import Mathlib.Analysis.InnerProductSpace.ProdL2 import Mathlib.Analysis.InnerProductSpace.Projection import Mathlib.Analysis.InnerProductSpace.Rayleigh import Mathlib.Analysis.InnerProductSpace.Spectrum @@ -746,6 +747,7 @@ import Mathlib.Analysis.NormedSpace.Multilinear import Mathlib.Analysis.NormedSpace.OperatorNorm import Mathlib.Analysis.NormedSpace.PiLp import Mathlib.Analysis.NormedSpace.Pointwise +import Mathlib.Analysis.NormedSpace.ProdLp import Mathlib.Analysis.NormedSpace.QuaternionExponential import Mathlib.Analysis.NormedSpace.Ray import Mathlib.Analysis.NormedSpace.RieszLemma From 58c0ff7b992dabee2cb04fd9ee2bed54a716a155 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sat, 5 Aug 2023 17:10:54 +1000 Subject: [PATCH 12/39] linter --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 54 +++++++++--------------- 1 file changed, 20 insertions(+), 34 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 6ef53eb287c3e..93bac2a054462 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -725,8 +725,8 @@ section Single variable (p α β) @[simp] -theorem nnnorm_equiv_symm_fst [hp : Fact (1 ≤ p)] (x : α) : - ‖(ProdLp.equiv p α β).symm (x, 0)‖₊ = ‖x‖₊ := by +theorem nnnorm_fst [hp : Fact (1 ≤ p)] (x : α) : + nnnorm (E := ProdLp p α β) (x, 0) = ‖x‖₊ := by induction p using ENNReal.recTopCoe generalizing hp with | top => simp [nnnorm_eq_sup] @@ -736,8 +736,8 @@ theorem nnnorm_equiv_symm_fst [hp : Fact (1 ≤ p)] (x : α) : simp [nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] @[simp] -theorem nnnorm_equiv_symm_snd [hp : Fact (1 ≤ p)] (y : β) : - ‖(ProdLp.equiv p α β).symm (0, y)‖₊ = ‖y‖₊ := by +theorem nnnorm_snd [hp : Fact (1 ≤ p)] (y : β) : + nnnorm (E := ProdLp p α β) (0, y) = ‖y‖₊ := by induction p using ENNReal.recTopCoe generalizing hp with | top => simp [nnnorm_eq_sup] @@ -747,50 +747,36 @@ theorem nnnorm_equiv_symm_snd [hp : Fact (1 ≤ p)] (y : β) : simp [nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] @[simp] -theorem norm_equiv_symm_fst (x : α) : ‖(ProdLp.equiv p α β).symm (x, 0)‖ = ‖x‖ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_fst p α β x +theorem norm_fst (x : α) : norm (E := ProdLp p α β) (x, 0) = ‖x‖ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_fst p α β x @[simp] -theorem norm_equiv_symm_snd (y : β) : ‖(ProdLp.equiv p α β).symm (0, y)‖ = ‖y‖ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_snd p α β y +theorem norm_snd (y : β) : norm (E := ProdLp p α β) (0, y) = ‖y‖ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_snd p α β y @[simp] -theorem nndist_equiv_symm_single_fst (x₁ x₂ : α) : - nndist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = - nndist x₁ x₂ := by - rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← equiv_symm_sub, Prod.mk_sub_mk, sub_zero, - nnnorm_equiv_symm_fst] +theorem nndist_fst (x₁ x₂ : α) : nndist (α := ProdLp p α β) (x₁, 0) (x₂, 0) = nndist x₁ x₂ := by + rw [nndist_eq_nnnorm, nndist_eq_nnnorm, Prod.mk_sub_mk, sub_zero, nnnorm_fst] @[simp] -theorem nndist_equiv_symm_single_snd (y₁ y₂ : β) : - nndist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = - nndist y₁ y₂ := by - rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← equiv_symm_sub, Prod.mk_sub_mk, sub_zero, - nnnorm_equiv_symm_snd] +theorem nndist_snd (y₁ y₂ : β) : nndist (α := ProdLp p α β) (0, y₁) (0, y₂) = nndist y₁ y₂ := by + rw [nndist_eq_nnnorm, nndist_eq_nnnorm, Prod.mk_sub_mk, sub_zero, nnnorm_snd] @[simp] -theorem dist_equiv_symm_single_fst (x₁ x₂ : α) : - dist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = - dist x₁ x₂ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_fst p α β x₁ x₂ +theorem dist_fst (x₁ x₂ : α) : dist (α := ProdLp p α β) (x₁, 0) (x₂, 0) = dist x₁ x₂ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_fst p α β x₁ x₂ @[simp] -theorem dist_equiv_symm_single_snd (y₁ y₂ : β) : - dist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = - dist y₁ y₂ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_snd p α β y₁ y₂ +theorem dist_snd (y₁ y₂ : β) : dist (α := ProdLp p α β) (0, y₁) (0, y₂) = dist y₁ y₂ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_snd p α β y₁ y₂ @[simp] -theorem edist_equiv_symm_single_fst (x₁ x₂ : α) : - edist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = - edist x₁ x₂ := by - simp only [edist_nndist, nndist_equiv_symm_single_fst p α β x₁ x₂] +theorem edist_fst (x₁ x₂ : α) : edist (α := ProdLp p α β) (x₁, 0) (x₂, 0) = edist x₁ x₂ := by + simp only [edist_nndist, nndist_fst p α β x₁ x₂] @[simp] -theorem edist_equiv_symm_single_snd (y₁ y₂ : β) : - edist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = - edist y₁ y₂ := by - simp only [edist_nndist, nndist_equiv_symm_single_snd p α β y₁ y₂] +theorem edist_snd (y₁ y₂ : β) : edist (α := ProdLp p α β) (0, y₁) (0, y₂) = edist y₁ y₂ := by + simp only [edist_nndist, nndist_snd p α β y₁ y₂] end Single From ab93a9e6e1438c66462b9ebc5a49977d8d7a2483 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sat, 5 Aug 2023 17:16:01 +1000 Subject: [PATCH 13/39] old names --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 93bac2a054462..4a007332a5c39 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -448,7 +448,7 @@ instance instMetricSpace [MetricSpace α] [MetricSpace β] : MetricSpace (ProdLp variable {p α β} -theorem nndist_eq_sum [PseudoMetricSpace α] [PseudoMetricSpace β] +theorem nndist_eq_add [PseudoMetricSpace α] [PseudoMetricSpace β] (hp : p ≠ ∞) (x y : ProdLp p α β) : nndist x y = (nndist x.fst y.fst ^ p.toReal + nndist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) := -- Porting note: was `Subtype.ext` @@ -456,7 +456,7 @@ theorem nndist_eq_sum [PseudoMetricSpace α] [PseudoMetricSpace β] push_cast exact dist_eq_add (p.toReal_pos_iff_ne_top.mpr hp) _ _ -theorem nndist_eq_iSup [PseudoMetricSpace α] [PseudoMetricSpace β] (x y : ProdLp ∞ α β) : +theorem nndist_eq_sup [PseudoMetricSpace α] [PseudoMetricSpace β] (x y : ProdLp ∞ α β) : nndist x y = nndist x.fst y.fst ⊔ nndist x.snd y.snd := -- Porting note: was `Subtype.ext` NNReal.eq <| by From 5401fc184d7334282a338ee33ebc0f85fa8db029 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sun, 6 Aug 2023 23:58:48 +1000 Subject: [PATCH 14/39] Eric's suggestion --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 27 ++++++++++++++++++------ 1 file changed, 20 insertions(+), 7 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 4a007332a5c39..04c55c056743b 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -92,6 +92,7 @@ section Edist variable [EDist α] [EDist β] +open Classical in /-- Endowing the space `ProdLp p α β` with the `L^p` edistance. We register this instance separate from `ProdLp.instPseudoEMetric` since the latter requires the type class hypothesis `[Fact (1 ≤ p)]` in order to prove the triangle inequality. @@ -101,7 +102,8 @@ satisfying a relaxed triangle inequality. The terminology for this varies throug literature, but it is sometimes called a *quasi-metric* or *semi-metric*. -/ instance instEDist : EDist (ProdLp p α β) where edist f g := - if _hp : p = 0 then 0 + if _hp : p = 0 then + (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) else if p = ∞ then edist f.fst g.fst ⊔ edist f.snd g.snd else (edist f.fst g.fst ^ p.toReal + edist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) @@ -110,8 +112,10 @@ variable {α β} variable (x y : ProdLp p α β) (x' : α × β) +open Classical in @[simp] -protected theorem edist_eq_zero (f g : ProdLp 0 α β) : edist f g = 0 := +protected theorem edist_eq_card (f g : ProdLp 0 α β) : + edist f g = (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := if_pos rfl theorem edist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : ProdLp p α β) : @@ -139,11 +143,12 @@ protected theorem edist_self (f : ProdLp p α β) : edist f f = 0 := by · simp [edist_eq_sup] · simp [edist_eq_add h, ENNReal.zero_rpow_of_pos h, ENNReal.zero_rpow_of_pos (inv_pos.2 <| h)] +open Classical in /-- This holds independent of `p` and does not require `[Fact (1 ≤ p)]`. We keep it separate from `ProdLp.instPseudoEMetricSpace` so it can be used also for `p < 1`. -/ protected theorem edist_comm (f g : ProdLp p α β) : edist f g = edist g f := by rcases p.trichotomy with (rfl | rfl | h) - · simp only [ProdLp.edist_eq_zero, eq_comm, Ne.def] + · simp [ProdLp.edist_eq_card, eq_comm] · simp only [edist_eq_sup, edist_comm] · simp only [edist_eq_add h, edist_comm] @@ -153,6 +158,7 @@ section Dist variable [Dist α] [Dist β] +open Classical in /-- Endowing the space `ProdLp p α β` with the `L^p` distance. We register this instance separate from `ProdLp.instPseudoMetricSpace` since the latter requires the type class hypothesis `[Fact (1 ≤ p)]` in order to prove the triangle inequality. @@ -162,14 +168,17 @@ satisfying a relaxed triangle inequality. The terminology for this varies throug literature, but it is sometimes called a *quasi-metric* or *semi-metric*. -/ instance instDist : Dist (ProdLp p α β) where dist f g := - if _hp : p = 0 then 0 --{ i | f i ≠ g i }.toFinite.toFinset.card + if _hp : p = 0 then + (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) else if p = ∞ then dist f.fst g.fst ⊔ dist f.snd g.snd else (dist f.fst g.fst ^ p.toReal + dist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) variable {α β} -theorem dist_eq_card (f g : ProdLp 0 α β) : dist f g = 0 := +open Classical in +theorem dist_eq_card (f g : ProdLp 0 α β) : dist f g = + (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := if_pos rfl theorem dist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : ProdLp p α β) : @@ -188,6 +197,7 @@ section Norm variable [Norm α] [Zero α] [Norm β] [Zero β] +open Classical in /-- Endowing the space `ProdLp p α β` with the `L^p` norm. We register this instance separate from `ProdLp.instSeminormedAddCommGroup` since the latter requires the type class hypothesis `[Fact (1 ≤ p)]` in order to prove the triangle inequality. @@ -196,14 +206,17 @@ Registering this separately allows for a future norm-like structure on `ProdLp p satisfying a relaxed triangle inequality. These are called *quasi-norms*. -/ instance instNorm : Norm (ProdLp p α β) where norm f := - if _hp : p = 0 then 0 + if _hp : p = 0 then + (if f.fst = 0 then 0 else 1) + (if f.snd = 0 then 0 else 1) else if p = ∞ then ‖f.fst‖ ⊔ ‖f.snd‖ else (‖f.fst‖ ^ p.toReal + ‖f.snd‖ ^ p.toReal) ^ (1 / p.toReal) variable {p α β} +open Classical in @[simp] -protected theorem norm_eq_zero (f : ProdLp 0 α β) : ‖f‖ = 0 := +protected theorem norm_eq_card (f : ProdLp 0 α β) : + ‖f‖ = (if f.fst = 0 then 0 else 1) + (if f.snd = 0 then 0 else 1) := if_pos rfl theorem norm_eq_sup (f : ProdLp ∞ α β) : ‖f‖ = ‖f.fst‖ ⊔ ‖f.snd‖ := by From eca9b3f2d1d9be9c729cb4a3b8933689696661cf Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Mon, 7 Aug 2023 15:29:27 +1000 Subject: [PATCH 15/39] Revert "linter" This reverts commit 58c0ff7b992dabee2cb04fd9ee2bed54a716a155. --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 54 +++++++++++++++--------- 1 file changed, 34 insertions(+), 20 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 04c55c056743b..209ceafdde47c 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -738,8 +738,8 @@ section Single variable (p α β) @[simp] -theorem nnnorm_fst [hp : Fact (1 ≤ p)] (x : α) : - nnnorm (E := ProdLp p α β) (x, 0) = ‖x‖₊ := by +theorem nnnorm_equiv_symm_fst [hp : Fact (1 ≤ p)] (x : α) : + ‖(ProdLp.equiv p α β).symm (x, 0)‖₊ = ‖x‖₊ := by induction p using ENNReal.recTopCoe generalizing hp with | top => simp [nnnorm_eq_sup] @@ -749,8 +749,8 @@ theorem nnnorm_fst [hp : Fact (1 ≤ p)] (x : α) : simp [nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] @[simp] -theorem nnnorm_snd [hp : Fact (1 ≤ p)] (y : β) : - nnnorm (E := ProdLp p α β) (0, y) = ‖y‖₊ := by +theorem nnnorm_equiv_symm_snd [hp : Fact (1 ≤ p)] (y : β) : + ‖(ProdLp.equiv p α β).symm (0, y)‖₊ = ‖y‖₊ := by induction p using ENNReal.recTopCoe generalizing hp with | top => simp [nnnorm_eq_sup] @@ -760,36 +760,50 @@ theorem nnnorm_snd [hp : Fact (1 ≤ p)] (y : β) : simp [nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] @[simp] -theorem norm_fst (x : α) : norm (E := ProdLp p α β) (x, 0) = ‖x‖ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_fst p α β x +theorem norm_equiv_symm_fst (x : α) : ‖(ProdLp.equiv p α β).symm (x, 0)‖ = ‖x‖ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_fst p α β x @[simp] -theorem norm_snd (y : β) : norm (E := ProdLp p α β) (0, y) = ‖y‖ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_snd p α β y +theorem norm_equiv_symm_snd (y : β) : ‖(ProdLp.equiv p α β).symm (0, y)‖ = ‖y‖ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_snd p α β y @[simp] -theorem nndist_fst (x₁ x₂ : α) : nndist (α := ProdLp p α β) (x₁, 0) (x₂, 0) = nndist x₁ x₂ := by - rw [nndist_eq_nnnorm, nndist_eq_nnnorm, Prod.mk_sub_mk, sub_zero, nnnorm_fst] +theorem nndist_equiv_symm_single_fst (x₁ x₂ : α) : + nndist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = + nndist x₁ x₂ := by + rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← equiv_symm_sub, Prod.mk_sub_mk, sub_zero, + nnnorm_equiv_symm_fst] @[simp] -theorem nndist_snd (y₁ y₂ : β) : nndist (α := ProdLp p α β) (0, y₁) (0, y₂) = nndist y₁ y₂ := by - rw [nndist_eq_nnnorm, nndist_eq_nnnorm, Prod.mk_sub_mk, sub_zero, nnnorm_snd] +theorem nndist_equiv_symm_single_snd (y₁ y₂ : β) : + nndist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = + nndist y₁ y₂ := by + rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← equiv_symm_sub, Prod.mk_sub_mk, sub_zero, + nnnorm_equiv_symm_snd] @[simp] -theorem dist_fst (x₁ x₂ : α) : dist (α := ProdLp p α β) (x₁, 0) (x₂, 0) = dist x₁ x₂ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_fst p α β x₁ x₂ +theorem dist_equiv_symm_single_fst (x₁ x₂ : α) : + dist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = + dist x₁ x₂ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_fst p α β x₁ x₂ @[simp] -theorem dist_snd (y₁ y₂ : β) : dist (α := ProdLp p α β) (0, y₁) (0, y₂) = dist y₁ y₂ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_snd p α β y₁ y₂ +theorem dist_equiv_symm_single_snd (y₁ y₂ : β) : + dist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = + dist y₁ y₂ := + congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_snd p α β y₁ y₂ @[simp] -theorem edist_fst (x₁ x₂ : α) : edist (α := ProdLp p α β) (x₁, 0) (x₂, 0) = edist x₁ x₂ := by - simp only [edist_nndist, nndist_fst p α β x₁ x₂] +theorem edist_equiv_symm_single_fst (x₁ x₂ : α) : + edist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = + edist x₁ x₂ := by + simp only [edist_nndist, nndist_equiv_symm_single_fst p α β x₁ x₂] @[simp] -theorem edist_snd (y₁ y₂ : β) : edist (α := ProdLp p α β) (0, y₁) (0, y₂) = edist y₁ y₂ := by - simp only [edist_nndist, nndist_snd p α β y₁ y₂] +theorem edist_equiv_symm_single_snd (y₁ y₂ : β) : + edist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = + edist y₁ y₂ := by + simp only [edist_nndist, nndist_equiv_symm_single_snd p α β y₁ y₂] end Single From 6a9789bf63e9d5f1f2afb30962bd707ffc05fa14 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Mon, 7 Aug 2023 15:41:55 +1000 Subject: [PATCH 16/39] fixes --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 20 ++++++++++++++------ 1 file changed, 14 insertions(+), 6 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 209ceafdde47c..89a187b344202 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -70,11 +70,19 @@ protected def equiv : ProdLp p α β ≃ α × β := the use of the type synonym. -/ @[simp] -theorem equiv_apply (x : ProdLp p α β) : ProdLp.equiv p α β x = x := +theorem equiv_fst (x : ProdLp p α β) : (ProdLp.equiv p α β x).fst = x.fst := rfl @[simp] -theorem equiv_symm_apply (x : α × β) : (ProdLp.equiv p α β).symm x = x := +theorem equiv_snd (x : ProdLp p α β) : (ProdLp.equiv p α β x).snd = x.snd := + rfl + +@[simp] +theorem equiv_symm_fst (x : α × β) : ((ProdLp.equiv p α β).symm x).fst = x.fst := + rfl + +@[simp] +theorem equiv_symm_snd (x : α × β) : ((ProdLp.equiv p α β).symm x).snd = x.snd := rfl section DistNorm @@ -340,10 +348,10 @@ theorem lipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] LipschitzWith 1 (ProdLp.equiv p α β) := by intro x y rcases p.dichotomy with (rfl | h) - · simp only [equiv_apply, coe_one, one_mul, le_refl] + · simp [edist] · have cancel : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le h).ne' rw [edist_eq_add (zero_lt_one.trans_le h)] - simp only [edist, forall_prop_of_true, one_mul, ENNReal.coe_one, equiv_apply, ge_iff_le, + simp only [edist, forall_prop_of_true, one_mul, ENNReal.coe_one, ge_iff_le, sup_le_iff] constructor · calc @@ -383,7 +391,7 @@ theorem antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace _ = ((2 : ℝ≥0) ^ (1 / p.toReal) : ℝ≥0) * edist (ProdLp.equiv p α β x) (ProdLp.equiv p α β y) := by - simp only [equiv_apply, ← two_mul, ENNReal.mul_rpow_of_nonneg _ _ nonneg, + simp only [← two_mul, ENNReal.mul_rpow_of_nonneg _ _ nonneg, ← ENNReal.rpow_mul, cancel, ENNReal.rpow_one, ← ENNReal.coe_rpow_of_nonneg _ nonneg, coe_ofNat] @@ -685,7 +693,7 @@ def equivₗᵢ : ProdLp ∞ α β ≃ₗᵢ[𝕜] α × β := { ProdLp.equiv ∞ α β with map_add' := fun f g => rfl map_smul' := fun c f => rfl - norm_map' := fun f => by simp } + norm_map' := fun f => by simp [Norm.norm] } variable (x y : ProdLp p α β) (x' y' : α × β) (c : 𝕜) From c1a753b8d50fe92b042a6e84593c055040357b40 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sat, 12 Aug 2023 19:31:58 +1000 Subject: [PATCH 17/39] fixes --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 89a187b344202..c0e4c4c53ff56 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -30,8 +30,7 @@ statements for the coordinate functions, for instance. # Implementation notes -This files is a straight-forward adaption of `Mathlib.Analysis.NormedSpace.PiLp`. We deviate from -`PiLp` in that we use for `p = 0` the junk value `d(x, y) = 0`. +This files is a straight-forward adaption of `Mathlib.Analysis.NormedSpace.PiLp`. -/ @@ -506,7 +505,7 @@ theorem infty_equiv_isometry [PseudoEMetricSpace α] [PseudoEMetricSpace β] : norm. -/ instance instSeminormedAddCommGroup [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] : SeminormedAddCommGroup (ProdLp p α β) := - { Prod.instAddCommGroupSum with + { Prod.instAddCommGroup with dist_eq := fun x y => by rcases p.dichotomy with (rfl | h) · simp only [dist_eq_sup, norm_eq_sup, dist_eq_norm] @@ -592,7 +591,7 @@ variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] -- Porting note: added instance instModule : Module 𝕜 (ProdLp p α β) := - { Prod.module with } + { Prod.instModule with } /-- The product of two normed spaces is a normed space, with the `L^p` norm. -/ instance instNormedSpace : @@ -600,7 +599,7 @@ instance instNormedSpace : { instModule p 𝕜 α β with norm_smul_le := fun c f => by rcases p.dichotomy with (rfl | hp) - · letI : Module 𝕜 (ProdLp ∞ α β) := Prod.module + · letI : Module 𝕜 (ProdLp ∞ α β) := Prod.instModule suffices ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊ by exact_mod_cast NNReal.coe_mono this.le simp only [nnnorm_eq_sup, NNReal.mul_sup, ← nnnorm_smul] rfl From 63a1aff8c1ecbe21971b94fd0aefec24a3cc08f6 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Tue, 15 Aug 2023 12:20:17 +1000 Subject: [PATCH 18/39] search and replace --- .../Analysis/InnerProductSpace/ProdL2.lean | 4 +- Mathlib/Analysis/NormedSpace/ProdLp.lean | 305 ++++++------------ 2 files changed, 109 insertions(+), 200 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean index 24d7279f0f995..50fb946746ce3 100644 --- a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean @@ -23,7 +23,7 @@ variable {𝕜 : Type _} [IsROrC 𝕜] instance ProdLp.instInnerProductSpace (E F : Type _) [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] : - InnerProductSpace 𝕜 (ProdLp 2 E F) where + InnerProductSpace 𝕜 (WithLp 2 (E × F)) where inner x y := inner x.fst y.fst + inner x.snd y.snd norm_sq_eq_inner x := by simp [ProdLp.norm_sq_eq_of_L2, ← norm_sq_eq_inner] @@ -41,5 +41,5 @@ variable {E : Type _} [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] variable {F : Type _} [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] @[simp] -theorem ProdLp.inner_apply (x y : ProdLp 2 E F) : +theorem ProdLp.inner_apply (x y : WithLp 2 (E × F)) : @inner 𝕜 _ _ x y = inner x.fst y.fst + inner x.snd y.snd := rfl diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index c0e4c4c53ff56..1c18de9f50d7a 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -5,6 +5,7 @@ Authors: Moritz Doll, Sébastien Gouëzel, Jireh Loreaux -/ import Mathlib.Analysis.MeanInequalities +import Mathlib.Analysis.NormedSpace.WithLp /-! # `L^p` distance on products of two metric spaces @@ -21,11 +22,11 @@ We give instances of this construction for emetric spaces, metric spaces, normed spaces. To avoid conflicting instances, all these are defined on a copy of the original Prod-type, named -`ProdLp p α β`. The assumption `[Fact (1 ≤ p)]` is required for the metric and normed space +`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 `ProdLp p α β` are (defeq to) the -product topology, product bornology and product uniformity, to be able to use freely continuity +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 @@ -40,48 +41,27 @@ open Real Set Filter IsROrC Bornology BigOperators Uniformity Topology NNReal EN noncomputable section -/-- A copy of a Prod type, on which we will put the `L^p` distance. Since the Prod 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. -/ -@[nolint unusedArguments] -def ProdLp (_p : ℝ≥0∞) (α β : Type _) : Type _ := - α × β - -instance ProdLp.instInhabited (p : ℝ≥0∞) (α β : Type _) [Inhabited α] [Inhabited β] : - Inhabited (ProdLp p α β) := - ⟨default, default⟩ - -@[ext] -protected theorem ProdLp.ext {p : ℝ≥0∞} {α β : Type _} {x y : ProdLp p α β} - (h1 : x.fst = y.fst) (h2 : x.snd = y.snd) : x = y := Prod.ext h1 h2 - namespace ProdLp variable (p : ℝ≥0∞) (𝕜 𝕜' : Type _) (α β : Type _) -/-- Canonical bijection between `ProdLp p α β` and the original Prod type. We introduce it to be -able to compare the `L^p` and `L^∞` distances through it. -/ -protected def equiv : ProdLp p α β ≃ α × β := - Equiv.refl _ - /-! Note that the unapplied versions of these lemmas are deliberately omitted, as they break the use of the type synonym. -/ @[simp] -theorem equiv_fst (x : ProdLp p α β) : (ProdLp.equiv p α β x).fst = x.fst := +theorem equiv_fst (x : WithLp p (α × β)) : (WithLp.equiv p (α × β) x).fst = x.fst := rfl @[simp] -theorem equiv_snd (x : ProdLp p α β) : (ProdLp.equiv p α β x).snd = x.snd := +theorem equiv_snd (x : WithLp p (α × β)) : (WithLp.equiv p (α × β) x).snd = x.snd := rfl @[simp] -theorem equiv_symm_fst (x : α × β) : ((ProdLp.equiv p α β).symm x).fst = x.fst := +theorem equiv_symm_fst (x : α × β) : ((WithLp.equiv p (α × β)).symm x).fst = x.fst := rfl @[simp] -theorem equiv_symm_snd (x : α × β) : ((ProdLp.equiv p α β).symm x).snd = x.snd := +theorem equiv_symm_snd (x : α × β) : ((WithLp.equiv p (α × β)).symm x).snd = x.snd := rfl section DistNorm @@ -89,7 +69,7 @@ section DistNorm /-! ### Definition of `edist`, `dist` and `norm` on `ProdLp` -In this section we define the `edist`, `dist` and `norm` functions on `ProdLp p α β` without +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.to_real`. -/ @@ -100,14 +80,14 @@ section Edist variable [EDist α] [EDist β] open Classical in -/-- Endowing the space `ProdLp p α β` with the `L^p` edistance. We register this instance +/-- Endowing the space `WithLp p (α × β)` with the `L^p` edistance. We register this instance separate from `ProdLp.instPseudoEMetric` 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 `ProdLp p α β` for `p < 1` -satisfying a relaxed triangle inequality. The terminology for this varies throughout the +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*. -/ -instance instEDist : EDist (ProdLp p α β) where +instance instEDist : EDist (WithLp p (α × β)) where edist f g := if _hp : p = 0 then (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) @@ -117,20 +97,20 @@ instance instEDist : EDist (ProdLp p α β) where variable {α β} -variable (x y : ProdLp p α β) (x' : α × β) +variable (x y : WithLp p (α × β)) (x' : α × β) open Classical in @[simp] -protected theorem edist_eq_card (f g : ProdLp 0 α β) : +protected theorem edist_eq_card (f g : WithLp 0 (α × β)) : edist f g = (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := if_pos rfl -theorem edist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : ProdLp p α β) : +theorem edist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : edist f g = (edist f.fst g.fst ^ p.toReal + edist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) := let hp' := ENNReal.toReal_pos_iff.mp hp (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) -theorem edist_eq_sup (f g : ProdLp ∞ α β) : +theorem edist_eq_sup (f g : WithLp ∞ (α × β)) : edist f g = edist f.fst g.fst ⊔ edist f.snd g.snd := by dsimp [edist] exact if_neg ENNReal.top_ne_zero @@ -144,7 +124,7 @@ variable [PseudoEMetricSpace α] [PseudoEMetricSpace β] /-- This holds independent of `p` and does not require `[Fact (1 ≤ p)]`. We keep it separate from `ProdLp.instPseudoEMetricSpace` so it can be used also for `p < 1`. -/ -protected theorem edist_self (f : ProdLp p α β) : edist f f = 0 := by +protected theorem edist_self (f : WithLp p (α × β)) : edist f f = 0 := by rcases p.trichotomy with (rfl | rfl | h) · simp · simp [edist_eq_sup] @@ -153,7 +133,7 @@ protected theorem edist_self (f : ProdLp p α β) : edist f f = 0 := by open Classical in /-- This holds independent of `p` and does not require `[Fact (1 ≤ p)]`. We keep it separate from `ProdLp.instPseudoEMetricSpace` so it can be used also for `p < 1`. -/ -protected theorem edist_comm (f g : ProdLp p α β) : edist f g = edist g f := by +protected theorem edist_comm (f g : WithLp p (α × β)) : edist f g = edist g f := by rcases p.trichotomy with (rfl | rfl | h) · simp [ProdLp.edist_eq_card, eq_comm] · simp only [edist_eq_sup, edist_comm] @@ -166,14 +146,14 @@ section Dist variable [Dist α] [Dist β] open Classical in -/-- Endowing the space `ProdLp p α β` with the `L^p` distance. We register this instance +/-- Endowing the space `WithLp p (α × β)` with the `L^p` distance. We register this instance separate from `ProdLp.instPseudoMetricSpace` 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 `ProdLp p α β` for `p < 1` -satisfying a relaxed triangle inequality. The terminology for this varies throughout the +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*. -/ -instance instDist : Dist (ProdLp p α β) where +instance instDist : Dist (WithLp p (α × β)) where dist f g := if _hp : p = 0 then (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) @@ -184,16 +164,16 @@ instance instDist : Dist (ProdLp p α β) where variable {α β} open Classical in -theorem dist_eq_card (f g : ProdLp 0 α β) : dist f g = +theorem dist_eq_card (f g : WithLp 0 (α × β)) : dist f g = (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := if_pos rfl -theorem dist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : ProdLp p α β) : +theorem dist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : dist f g = (dist f.fst g.fst ^ p.toReal + dist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) := let hp' := ENNReal.toReal_pos_iff.mp hp (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) -theorem dist_eq_sup (f g : ProdLp ∞ α β) : +theorem dist_eq_sup (f g : WithLp ∞ (α × β)) : dist f g = dist f.fst g.fst ⊔ dist f.snd g.snd := by dsimp [dist] exact if_neg ENNReal.top_ne_zero @@ -205,13 +185,13 @@ section Norm variable [Norm α] [Zero α] [Norm β] [Zero β] open Classical in -/-- Endowing the space `ProdLp p α β` with the `L^p` norm. We register this instance +/-- Endowing the space `WithLp p (α × β)` with the `L^p` norm. We register this instance separate from `ProdLp.instSeminormedAddCommGroup` 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 `ProdLp p α β` for `p < 1` -satisfying a relaxed triangle inequality. These are called *quasi-norms*. -/ -instance instNorm : Norm (ProdLp p α β) where +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*. -/ +instance instNorm : Norm (WithLp p (α × β)) where norm f := if _hp : p = 0 then (if f.fst = 0 then 0 else 1) + (if f.snd = 0 then 0 else 1) @@ -222,15 +202,15 @@ variable {p α β} open Classical in @[simp] -protected theorem norm_eq_card (f : ProdLp 0 α β) : +protected theorem norm_eq_card (f : WithLp 0 (α × β)) : ‖f‖ = (if f.fst = 0 then 0 else 1) + (if f.snd = 0 then 0 else 1) := if_pos rfl -theorem norm_eq_sup (f : ProdLp ∞ α β) : ‖f‖ = ‖f.fst‖ ⊔ ‖f.snd‖ := by +theorem norm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖ = ‖f.fst‖ ⊔ ‖f.snd‖ := by dsimp [Norm.norm] exact if_neg ENNReal.top_ne_zero -theorem norm_eq_add (hp : 0 < p.toReal) (f : ProdLp p α β) : +theorem norm_eq_add (hp : 0 < p.toReal) (f : WithLp p (α × β)) : ‖f‖ = (‖f.fst‖ ^ p.toReal + ‖f.snd‖ ^ p.toReal) ^ (1 / p.toReal) := let hp' := ENNReal.toReal_pos_iff.mp hp (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) @@ -244,7 +224,7 @@ section Aux /-! ### The uniformity on finite `L^p` products is the product uniformity -In this section, we put the `L^p` edistance on `ProdLp p α β`, and we check that the 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 Pi type (with the `L^∞` distance) is a uniform embedding, as it is both Lipschitz and antiLipschitz. @@ -258,14 +238,14 @@ explaining why having definitionally the right uniformity is often important. variable [Fact (1 ≤ p)] -/-- Endowing the space `ProdLp 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`. -/ +/-- 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`. -/ def pseudoEmetricAux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : - PseudoEMetricSpace (ProdLp p α β) where + PseudoEMetricSpace (WithLp p (α × β)) where edist_self := ProdLp.edist_self p edist_comm := ProdLp.edist_comm p edist_triangle f g h := by @@ -296,12 +276,12 @@ attribute [local instance] ProdLp.pseudoEmetricAux /-- An auxiliary lemma used twice in the proof of `ProdLp.pseudoMetricAux` below. Not intended for use outside this file. -/ theorem sup_edist_ne_top_aux {α β : Type _} - [PseudoMetricSpace α] [PseudoMetricSpace β] (f g : ProdLp ∞ α β) : + [PseudoMetricSpace α] [PseudoMetricSpace β] (f g : WithLp ∞ (α × β)) : edist f.fst g.fst ⊔ edist f.snd g.snd ≠ ⊤ := by refine ne_of_lt ?_ simp [edist, PseudoMetricSpace.edist_dist] -/-- Endowing the space `ProdLp p α β` with the `L^p` pseudometric structure. This definition is not +/-- 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 @@ -312,7 +292,7 @@ structure and the bornology by the product ones using this pseudometric space, See note [reducible non-instances] -/ @[reducible] def pseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : - PseudoMetricSpace (ProdLp p α β) := + PseudoMetricSpace (WithLp p (α × β)) := PseudoEMetricSpace.toPseudoMetricSpaceOfDist dist (fun f g => by rcases p.dichotomy with (rfl | h) @@ -344,7 +324,7 @@ def pseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : attribute [local instance] ProdLp.pseudoMetricAux theorem lipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : - LipschitzWith 1 (ProdLp.equiv p α β) := by + LipschitzWith 1 (WithLp.equiv p (α × β)) := by intro x y rcases p.dichotomy with (rfl | h) · simp [edist] @@ -369,7 +349,7 @@ theorem lipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] example (a : ENNReal) : a + a = 2*a := by exact Eq.symm (two_mul a) theorem antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : - AntilipschitzWith ((2 : ℝ≥0) ^ (1 / p).toReal) (ProdLp.equiv p α β) := by + AntilipschitzWith ((2 : ℝ≥0) ^ (1 / p).toReal) (WithLp.equiv p (α × β)) := by intro x y rcases p.dichotomy with (rfl | h) · simp [edist] @@ -380,8 +360,8 @@ theorem antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace simp only [edist, ← one_div, ENNReal.one_toReal] calc (edist x.fst y.fst ^ p.toReal + edist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) ≤ - (edist (ProdLp.equiv p α β x) (ProdLp.equiv p α β y) ^ p.toReal + - edist (ProdLp.equiv p α β x) (ProdLp.equiv p α β y) ^ p.toReal) ^ (1 / p.toReal) := by + (edist (WithLp.equiv p (α × β) x) (WithLp.equiv p (α × β) y) ^ p.toReal + + edist (WithLp.equiv p (α × β) x) (WithLp.equiv p (α × β) y) ^ p.toReal) ^ (1 / p.toReal) := by refine ENNReal.rpow_le_rpow (add_le_add ?_ ?_) nonneg · refine ENNReal.rpow_le_rpow ?_ (le_of_lt pos) simp [edist] @@ -389,25 +369,25 @@ theorem antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace simp [edist] _ = ((2 : ℝ≥0) ^ (1 / p.toReal) : ℝ≥0) * - edist (ProdLp.equiv p α β x) (ProdLp.equiv p α β y) := by + edist (WithLp.equiv p (α × β) x) (WithLp.equiv p (α × β) y) := by simp only [← two_mul, ENNReal.mul_rpow_of_nonneg _ _ nonneg, ← ENNReal.rpow_mul, cancel, ENNReal.rpow_one, ← ENNReal.coe_rpow_of_nonneg _ nonneg, coe_ofNat] theorem aux_uniformity_eq [PseudoEMetricSpace α] [PseudoEMetricSpace β] : - 𝓤 (ProdLp p α β) = 𝓤[instUniformSpaceProd] := by - have A : UniformInducing (ProdLp.equiv p α β) := + 𝓤 (WithLp p (α × β)) = 𝓤[instUniformSpaceProd] := by + have A : UniformInducing (WithLp.equiv p (α × β)) := (antilipschitzWith_equiv_aux p α β).uniformInducing (lipschitzWith_equiv_aux p α β).uniformContinuous - have : (fun x : ProdLp p α β × ProdLp p α β => - ((ProdLp.equiv p α β) x.fst, (ProdLp.equiv p α β) x.snd)) = id := + have : (fun x : WithLp p (α × β) × WithLp p (α × β) => + ((WithLp.equiv p (α × β)) x.fst, (WithLp.equiv p (α × β)) x.snd)) = id := by ext i <;> rfl rw [← A.comap_uniformity, this, comap_id] theorem aux_cobounded_eq [PseudoMetricSpace α] [PseudoMetricSpace β] : - cobounded (ProdLp p α β) = @cobounded _ Prod.instBornology := + cobounded (WithLp p (α × β)) = @cobounded _ Prod.instBornology := calc - cobounded (ProdLp p α β) = comap (ProdLp.equiv p α β) (cobounded _) := + cobounded (WithLp p (α × β)) = comap (WithLp.equiv p (α × β)) (cobounded _) := le_antisymm (antilipschitzWith_equiv_aux p α β).tendsto_cobounded.le_comap (lipschitzWith_equiv_aux p α β).comap_cobounded_le _ = _ := comap_id @@ -417,27 +397,27 @@ end Aux /-! ### Instances on `L^p` products -/ -instance instUniformSpace [UniformSpace α] [UniformSpace β] : UniformSpace (ProdLp p α β) := +instance instUniformSpace [UniformSpace α] [UniformSpace β] : UniformSpace (WithLp p (α × β)) := instUniformSpaceProd theorem uniformContinuous_equiv [UniformSpace α] [UniformSpace β] : - UniformContinuous (ProdLp.equiv p α β) := + UniformContinuous (WithLp.equiv p (α × β)) := uniformContinuous_id theorem uniformContinuous_equiv_symm [UniformSpace α] [UniformSpace β] : - UniformContinuous (ProdLp.equiv p α β).symm := + UniformContinuous (WithLp.equiv p (α × β)).symm := uniformContinuous_id @[continuity] -theorem continuous_equiv [UniformSpace α] [UniformSpace β] : Continuous (ProdLp.equiv p α β) := +theorem continuous_equiv [UniformSpace α] [UniformSpace β] : Continuous (WithLp.equiv p (α × β)) := continuous_id @[continuity] theorem continuous_equiv_symm [UniformSpace α] [UniformSpace β] : - Continuous (ProdLp.equiv p α β).symm := + Continuous (WithLp.equiv p (α × β)).symm := continuous_id -instance instBornology [Bornology α] [Bornology β] : Bornology (ProdLp p α β) := +instance instBornology [Bornology α] [Bornology β] : Bornology (WithLp p (α × β)) := Prod.instBornology -- throughout the rest of the file, we assume `1 ≤ p` @@ -446,37 +426,37 @@ variable [Fact (1 ≤ p)] /-- `PseudoEMetricSpace` instance on the product of two pseudoemetric spaces, using the `L^p` pseudoedistance, and having as uniformity the product uniformity. -/ instance instPseudoEMetricSpace [PseudoEMetricSpace α] [PseudoEMetricSpace β] : - PseudoEMetricSpace (ProdLp p α β) := + PseudoEMetricSpace (WithLp p (α × β)) := (pseudoEmetricAux p α β).replaceUniformity (aux_uniformity_eq p α β).symm /-- `EMetricSpace` instance on the product of two emetric spaces, using the `L^p` edistance, and having as uniformity the product uniformity. -/ -instance instEMetricSpace [EMetricSpace α] [EMetricSpace β] : EMetricSpace (ProdLp p α β) := - @EMetricSpace.ofT0PseudoEMetricSpace (ProdLp p α β) _ instT0SpaceProdInstTopologicalSpaceProd +instance instEMetricSpace [EMetricSpace α] [EMetricSpace β] : EMetricSpace (WithLp p (α × β)) := + @EMetricSpace.ofT0PseudoEMetricSpace (WithLp p (α × β)) _ instT0SpaceProdInstTopologicalSpaceProd /-- `PseudoMetricSpace` instance on the product of two pseudometric spaces, using the `L^p` distance, and having as uniformity the product uniformity. -/ instance instPseudoMetricSpace [PseudoMetricSpace α] [PseudoMetricSpace β] : - PseudoMetricSpace (ProdLp p α β) := + PseudoMetricSpace (WithLp p (α × β)) := ((pseudoMetricAux p α β).replaceUniformity (aux_uniformity_eq p α β).symm).replaceBornology fun s => Filter.ext_iff.1 (aux_cobounded_eq p α β).symm sᶜ /-- `MetricSpace` instance on the product of two metric spaces, using the `L^p` distance, and having as uniformity the product uniformity. -/ -instance instMetricSpace [MetricSpace α] [MetricSpace β] : MetricSpace (ProdLp p α β) := +instance instMetricSpace [MetricSpace α] [MetricSpace β] : MetricSpace (WithLp p (α × β)) := MetricSpace.ofT0PseudoMetricSpace _ variable {p α β} theorem nndist_eq_add [PseudoMetricSpace α] [PseudoMetricSpace β] - (hp : p ≠ ∞) (x y : ProdLp p α β) : + (hp : p ≠ ∞) (x y : WithLp p (α × β)) : nndist x y = (nndist x.fst y.fst ^ p.toReal + nndist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) := -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast exact dist_eq_add (p.toReal_pos_iff_ne_top.mpr hp) _ _ -theorem nndist_eq_sup [PseudoMetricSpace α] [PseudoMetricSpace β] (x y : ProdLp ∞ α β) : +theorem nndist_eq_sup [PseudoMetricSpace α] [PseudoMetricSpace β] (x y : WithLp ∞ (α × β)) : nndist x y = nndist x.fst y.fst ⊔ nndist x.snd y.snd := -- Porting note: was `Subtype.ext` NNReal.eq <| by @@ -486,15 +466,15 @@ theorem nndist_eq_sup [PseudoMetricSpace α] [PseudoMetricSpace β] (x y : ProdL variable (p α β) theorem lipschitzWith_equiv [PseudoEMetricSpace α] [PseudoEMetricSpace β] : - LipschitzWith 1 (ProdLp.equiv p α β) := + LipschitzWith 1 (WithLp.equiv p (α × β)) := lipschitzWith_equiv_aux p α β theorem antilipschitzWith_equiv [PseudoEMetricSpace α] [PseudoEMetricSpace β] : - AntilipschitzWith ((2 : ℝ≥0) ^ (1 / p).toReal) (ProdLp.equiv p α β) := + AntilipschitzWith ((2 : ℝ≥0) ^ (1 / p).toReal) (WithLp.equiv p (α × β)) := antilipschitzWith_equiv_aux p α β theorem infty_equiv_isometry [PseudoEMetricSpace α] [PseudoEMetricSpace β] : - Isometry (ProdLp.equiv ∞ α β) := + Isometry (WithLp.equiv ∞ (α × β)) := fun x y => le_antisymm (by simpa only [ENNReal.coe_one, one_mul] using lipschitzWith_equiv ∞ α β x y) (by @@ -504,8 +484,8 @@ theorem infty_equiv_isometry [PseudoEMetricSpace α] [PseudoEMetricSpace β] : /-- Seminormed group instance on the product of two normed groups, using the `L^p` norm. -/ instance instSeminormedAddCommGroup [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] : - SeminormedAddCommGroup (ProdLp p α β) := - { Prod.instAddCommGroup with + SeminormedAddCommGroup (WithLp p (α × β)) := + { WithLp.instAddCommGroup _ _ with dist_eq := fun x y => by rcases p.dichotomy with (rfl | h) · simp only [dist_eq_sup, norm_eq_sup, dist_eq_norm] @@ -520,7 +500,7 @@ instance instSeminormedAddCommGroup [SeminormedAddCommGroup α] [SeminormedAddCo /-- normed group instance on the product of two normed groups, using the `L^p` norm. -/ instance normedAddCommGroup [NormedAddCommGroup α] [NormedAddCommGroup β] : - NormedAddCommGroup (ProdLp p α β) := + NormedAddCommGroup (WithLp p (α × β)) := { ProdLp.instSeminormedAddCommGroup p α β with eq_of_dist_eq_zero := eq_of_dist_eq_zero } @@ -529,27 +509,27 @@ section norm_of variable {p α β} variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] -theorem nnnorm_eq_add (hp : p ≠ ∞) (f : ProdLp p α β) : +theorem nnnorm_eq_add (hp : p ≠ ∞) (f : WithLp p (α × β)) : ‖f‖₊ = (‖f.fst‖₊ ^ p.toReal + ‖f.snd‖₊ ^ p.toReal) ^ (1 / p.toReal) := by ext simp [norm_eq_add (p.toReal_pos_iff_ne_top.mpr hp)] -theorem nnnorm_eq_sup (f : ProdLp ∞ α β) : ‖f‖₊ = ‖f.fst‖₊ ⊔ ‖f.snd‖₊ := by +theorem nnnorm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖₊ = ‖f.fst‖₊ ⊔ ‖f.snd‖₊ := by ext norm_cast -theorem norm_eq_of_nat (n : ℕ) (h : p = n) (f : ProdLp p α β) : +theorem norm_eq_of_nat (n : ℕ) (h : p = n) (f : WithLp p (α × β)) : ‖f‖ = (‖f.fst‖ ^ n + ‖f.snd‖ ^ n) ^ (1 / (n : ℝ)) := by have := p.toReal_pos_iff_ne_top.mpr (ne_of_eq_of_ne h <| ENNReal.nat_ne_top n) simp only [one_div, h, Real.rpow_nat_cast, ENNReal.toReal_nat, eq_self_iff_true, Finset.sum_congr, norm_eq_add this] -theorem norm_eq_of_L2 (x : ProdLp 2 α β) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by +theorem norm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by rw [norm_eq_of_nat 2 (by norm_cast) _] -- Porting note: was `convert` rw [Real.sqrt_eq_rpow] norm_cast -theorem nnnorm_eq_of_L2 (x : ProdLp 2 α β) : ‖x‖₊ = NNReal.sqrt (‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2) := +theorem nnnorm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖₊ = NNReal.sqrt (‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2) := -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast @@ -557,26 +537,26 @@ theorem nnnorm_eq_of_L2 (x : ProdLp 2 α β) : ‖x‖₊ = NNReal.sqrt (‖x.fs variable (α β) -theorem norm_sq_eq_of_L2 (x : ProdLp 2 α β) : ‖x‖ ^ 2 = ‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2 := by +theorem norm_sq_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ ^ 2 = ‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2 := by suffices ‖x‖₊ ^ 2 = ‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2 by simpa only [NNReal.coe_sum] using congr_arg ((↑) : ℝ≥0 → ℝ) this rw [nnnorm_eq_of_L2, NNReal.sq_sqrt] variable {α β} -theorem dist_eq_of_L2 (x y : ProdLp 2 α β) : +theorem dist_eq_of_L2 (x y : WithLp 2 (α × β)) : dist x y = (dist x.fst y.fst ^ 2 + dist x.snd y.snd ^ 2).sqrt := by simp_rw [dist_eq_norm, norm_eq_of_L2, Pi.sub_apply] rfl -theorem nndist_eq_of_L2 (x y : ProdLp 2 α β) : +theorem nndist_eq_of_L2 (x y : WithLp 2 (α × β)) : nndist x y = NNReal.sqrt (nndist x.fst y.fst ^ 2 + nndist x.snd y.snd ^ 2) := -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast exact dist_eq_of_L2 _ _ -theorem edist_eq_of_L2 (x y : ProdLp 2 α β) : +theorem edist_eq_of_L2 (x y : WithLp 2 (α × β)) : edist x y = (edist x.fst y.fst ^ 2 + edist x.snd y.snd ^ 2) ^ (1 / 2 : ℝ) := by simp [ProdLp.edist_eq_add] @@ -589,17 +569,13 @@ section normed_space_inst variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] [SeminormedAddCommGroup β] [NormedSpace 𝕜 β] --- Porting note: added -instance instModule : Module 𝕜 (ProdLp p α β) := - { Prod.instModule with } - /-- The product of two normed spaces is a normed space, with the `L^p` norm. -/ instance instNormedSpace : - NormedSpace 𝕜 (ProdLp p α β) := - { instModule p 𝕜 α β with + NormedSpace 𝕜 (WithLp p (α × β)) := + { WithLp.instModule p 𝕜 (α × β) with norm_smul_le := fun c f => by rcases p.dichotomy with (rfl | hp) - · letI : Module 𝕜 (ProdLp ∞ α β) := Prod.instModule + · letI : Module 𝕜 (WithLp ∞ (α × β)) := Prod.instModule suffices ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊ by exact_mod_cast NNReal.coe_mono this.le simp only [nnnorm_eq_sup, NNReal.mul_sup, ← nnnorm_smul] rfl @@ -612,24 +588,6 @@ instance instNormedSpace : ← rpow_mul (norm_nonneg _), this, Real.rpow_one] positivity } -section towers - -variable [NormedSpace 𝕜' α] [NormedSpace 𝕜' β] - -instance instIsScalarTower [SMul 𝕜 𝕜'] [IsScalarTower 𝕜 𝕜' α] [IsScalarTower 𝕜 𝕜' β] : - IsScalarTower 𝕜 𝕜' (ProdLp p α β) := - Prod.isScalarTower - -instance instSMulCommClass [SMulCommClass 𝕜 𝕜' α] [SMulCommClass 𝕜 𝕜' β] : - SMulCommClass 𝕜 𝕜' (ProdLp p α β) := - Prod.smulCommClass - -end towers - -instance instFiniteDimensional [FiniteDimensional 𝕜 α] [FiniteDimensional 𝕜 β] : - FiniteDimensional 𝕜 (ProdLp p α β) := - Module.Finite.prod - end normed_space_inst /- Register simplification lemmas for the applications of `ProdLp` elements, as the usual lemmas @@ -640,14 +598,14 @@ variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] section algebra -variable (x y : ProdLp p α β) (c : 𝕜) +variable (x y : WithLp p (α × β)) (c : 𝕜) @[simp] -theorem zero_fst : (0 : ProdLp p α β).fst = 0 := +theorem zero_fst : (0 : WithLp p (α × β)).fst = 0 := rfl @[simp] -theorem zero_snd : (0 : ProdLp p α β).snd = 0 := +theorem zero_snd : (0 : WithLp p (α × β)).snd = 0 := rfl @[simp] @@ -688,55 +646,13 @@ section Equiv /-- The canonical map `ProdLp.equiv` between `ProdLp ∞ β` and `Π i, β i` as a linear isometric equivalence. -/ -def equivₗᵢ : ProdLp ∞ α β ≃ₗᵢ[𝕜] α × β := - { ProdLp.equiv ∞ α β with +def equivₗᵢ : WithLp ∞ (α × β) ≃ₗᵢ[𝕜] α × β := + { WithLp.equiv ∞ (α × β) with map_add' := fun f g => rfl map_smul' := fun c f => rfl norm_map' := fun f => by simp [Norm.norm] } -variable (x y : ProdLp p α β) (x' y' : α × β) (c : 𝕜) - -@[simp] -theorem equiv_zero : ProdLp.equiv p α β 0 = 0 := - rfl - -@[simp] -theorem equiv_symm_zero : (ProdLp.equiv p α β).symm 0 = 0 := - rfl - -@[simp] -theorem equiv_add : ProdLp.equiv p α β (x + y) = ProdLp.equiv p α β x + ProdLp.equiv p α β y := - rfl - -@[simp] -theorem equiv_symm_add : (ProdLp.equiv p α β).symm (x' + y') = - (ProdLp.equiv p α β).symm x' + (ProdLp.equiv p α β).symm y' := - rfl - -@[simp] -theorem equiv_sub : ProdLp.equiv p α β (x - y) = ProdLp.equiv p α β x - ProdLp.equiv p α β y := - rfl - -@[simp] -theorem equiv_symm_sub : (ProdLp.equiv p α β).symm (x' - y') = - (ProdLp.equiv p α β).symm x' - (ProdLp.equiv p α β).symm y' := - rfl - -@[simp] -theorem equiv_neg : ProdLp.equiv p α β (-x) = -ProdLp.equiv p α β x := - rfl - -@[simp] -theorem equiv_symm_neg : (ProdLp.equiv p α β).symm (-x') = -(ProdLp.equiv p α β).symm x' := - rfl - -@[simp] -theorem equiv_smul : ProdLp.equiv p α β (c • x) = c • ProdLp.equiv p α β x := - rfl - -@[simp] -theorem equiv_symm_smul : (ProdLp.equiv p α β).symm (c • x') = c • (ProdLp.equiv p α β).symm x' := - rfl +variable (x y : WithLp p (α × β)) (x' y' : α × β) (c : 𝕜) end Equiv @@ -746,7 +662,7 @@ variable (p α β) @[simp] theorem nnnorm_equiv_symm_fst [hp : Fact (1 ≤ p)] (x : α) : - ‖(ProdLp.equiv p α β).symm (x, 0)‖₊ = ‖x‖₊ := by + ‖(WithLp.equiv p (α × β)).symm (x, 0)‖₊ = ‖x‖₊ := by induction p using ENNReal.recTopCoe generalizing hp with | top => simp [nnnorm_eq_sup] @@ -757,7 +673,7 @@ theorem nnnorm_equiv_symm_fst [hp : Fact (1 ≤ p)] (x : α) : @[simp] theorem nnnorm_equiv_symm_snd [hp : Fact (1 ≤ p)] (y : β) : - ‖(ProdLp.equiv p α β).symm (0, y)‖₊ = ‖y‖₊ := by + ‖(WithLp.equiv p (α × β)).symm (0, y)‖₊ = ‖y‖₊ := by induction p using ENNReal.recTopCoe generalizing hp with | top => simp [nnnorm_eq_sup] @@ -767,48 +683,48 @@ theorem nnnorm_equiv_symm_snd [hp : Fact (1 ≤ p)] (y : β) : simp [nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] @[simp] -theorem norm_equiv_symm_fst (x : α) : ‖(ProdLp.equiv p α β).symm (x, 0)‖ = ‖x‖ := +theorem norm_equiv_symm_fst (x : α) : ‖(WithLp.equiv p (α × β)).symm (x, 0)‖ = ‖x‖ := congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_fst p α β x @[simp] -theorem norm_equiv_symm_snd (y : β) : ‖(ProdLp.equiv p α β).symm (0, y)‖ = ‖y‖ := +theorem norm_equiv_symm_snd (y : β) : ‖(WithLp.equiv p (α × β)).symm (0, y)‖ = ‖y‖ := congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_snd p α β y @[simp] theorem nndist_equiv_symm_single_fst (x₁ x₂ : α) : - nndist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = + nndist ((WithLp.equiv p (α × β)).symm (x₁, 0)) ((WithLp.equiv p (α × β)).symm (x₂, 0)) = nndist x₁ x₂ := by - rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← equiv_symm_sub, Prod.mk_sub_mk, sub_zero, + rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← WithLp.equiv_symm_sub, Prod.mk_sub_mk, sub_zero, nnnorm_equiv_symm_fst] @[simp] theorem nndist_equiv_symm_single_snd (y₁ y₂ : β) : - nndist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = + nndist ((WithLp.equiv p (α × β)).symm (0, y₁)) ((WithLp.equiv p (α × β)).symm (0, y₂)) = nndist y₁ y₂ := by - rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← equiv_symm_sub, Prod.mk_sub_mk, sub_zero, + rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← WithLp.equiv_symm_sub, Prod.mk_sub_mk, sub_zero, nnnorm_equiv_symm_snd] @[simp] theorem dist_equiv_symm_single_fst (x₁ x₂ : α) : - dist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = + dist ((WithLp.equiv p (α × β)).symm (x₁, 0)) ((WithLp.equiv p (α × β)).symm (x₂, 0)) = dist x₁ x₂ := congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_fst p α β x₁ x₂ @[simp] theorem dist_equiv_symm_single_snd (y₁ y₂ : β) : - dist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = + dist ((WithLp.equiv p (α × β)).symm (0, y₁)) ((WithLp.equiv p (α × β)).symm (0, y₂)) = dist y₁ y₂ := congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_snd p α β y₁ y₂ @[simp] theorem edist_equiv_symm_single_fst (x₁ x₂ : α) : - edist ((ProdLp.equiv p α β).symm (x₁, 0)) ((ProdLp.equiv p α β).symm (x₂, 0)) = + edist ((WithLp.equiv p (α × β)).symm (x₁, 0)) ((WithLp.equiv p (α × β)).symm (x₂, 0)) = edist x₁ x₂ := by simp only [edist_nndist, nndist_equiv_symm_single_fst p α β x₁ x₂] @[simp] theorem edist_equiv_symm_single_snd (y₁ y₂ : β) : - edist ((ProdLp.equiv p α β).symm (0, y₁)) ((ProdLp.equiv p α β).symm (0, y₂)) = + edist ((WithLp.equiv p (α × β)).symm (0, y₁)) ((WithLp.equiv p (α × β)).symm (0, y₂)) = edist y₁ y₂ := by simp only [edist_nndist, nndist_equiv_symm_single_snd p α β y₁ y₂] @@ -816,17 +732,10 @@ end Single variable (𝕜 p α β) -/-- `ProdLp.equiv` as a linear equivalence. -/ -@[simps (config := { fullyApplied := false })] -protected def linearEquiv : ProdLp p α β ≃ₗ[𝕜] α × β := - { LinearEquiv.refl _ _ with - toFun := ProdLp.equiv _ _ _ - invFun := (ProdLp.equiv _ _ _).symm } - -/-- `ProdLp.equiv` as a continuous linear equivalence. -/ +/-- `WithLp.equiv` as a continuous linear equivalence. -/ @[simps! (config := { fullyApplied := false }) apply symm_apply] -protected def continuousLinearEquiv : ProdLp p α β ≃L[𝕜] α × β where - toLinearEquiv := ProdLp.linearEquiv _ _ _ _ +protected def continuousLinearEquiv : WithLp p (α × β) ≃L[𝕜] α × β where + toLinearEquiv := WithLp.linearEquiv _ _ _ continuous_toFun := continuous_equiv _ _ _ continuous_invFun := continuous_equiv_symm _ _ _ From 49df75a4678ca113c92ed308fbf15d1917ea4188 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Tue, 15 Aug 2023 12:24:15 +1000 Subject: [PATCH 19/39] fixed Type* --- Mathlib/Analysis/InnerProductSpace/ProdL2.lean | 8 ++++---- Mathlib/Analysis/NormedSpace/ProdLp.lean | 2 +- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean index 50fb946746ce3..41d225261d073 100644 --- a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean @@ -19,7 +19,7 @@ This is recorded in this file as an inner product space instance on `ProdLp 2`. noncomputable section -variable {𝕜 : Type _} [IsROrC 𝕜] +variable {𝕜 : Type*} [IsROrC 𝕜] instance ProdLp.instInnerProductSpace (E F : Type _) [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] : @@ -36,10 +36,10 @@ instance ProdLp.instInnerProductSpace (E F : Type _) [NormedAddCommGroup E] [Inn simp only [smul_fst, inner_smul_left, smul_snd] ring -variable {E : Type _} [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] +variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] -variable {F : Type _} [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] +variable {F : Type*} [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] @[simp] theorem ProdLp.inner_apply (x y : WithLp 2 (E × F)) : - @inner 𝕜 _ _ x y = inner x.fst y.fst + inner x.snd y.snd := rfl + inner (𝕜 := 𝕜) x y = inner x.fst y.fst + inner x.snd y.snd := rfl diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 1c18de9f50d7a..0a8fa14db08a4 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -43,7 +43,7 @@ noncomputable section namespace ProdLp -variable (p : ℝ≥0∞) (𝕜 𝕜' : Type _) (α β : Type _) +variable (p : ℝ≥0∞) (𝕜 𝕜' : Type*) (α β : Type*) /-! Note that the unapplied versions of these lemmas are deliberately omitted, as they break the use of the type synonym. -/ From 7703091c6af7a5c172c42696b6f7cc4067e20e8d Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Tue, 15 Aug 2023 13:07:09 +1000 Subject: [PATCH 20/39] linter --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 12 +++++------- 1 file changed, 5 insertions(+), 7 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 0a8fa14db08a4..ad12b2c819f54 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -281,8 +281,8 @@ theorem sup_edist_ne_top_aux {α β : Type _} refine ne_of_lt ?_ simp [edist, PseudoMetricSpace.edist_dist] -/-- 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 +/-- 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 @@ -360,16 +360,14 @@ theorem antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace simp only [edist, ← one_div, ENNReal.one_toReal] calc (edist x.fst y.fst ^ p.toReal + edist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) ≤ - (edist (WithLp.equiv p (α × β) x) (WithLp.equiv p (α × β) y) ^ p.toReal + - edist (WithLp.equiv p (α × β) x) (WithLp.equiv p (α × β) y) ^ p.toReal) ^ (1 / p.toReal) := by + (edist (WithLp.equiv p _ x) (WithLp.equiv p _ y) ^ p.toReal + + edist (WithLp.equiv p _ x) (WithLp.equiv p _ y) ^ p.toReal) ^ (1 / p.toReal) := by refine ENNReal.rpow_le_rpow (add_le_add ?_ ?_) nonneg · refine ENNReal.rpow_le_rpow ?_ (le_of_lt pos) simp [edist] · refine ENNReal.rpow_le_rpow ?_ (le_of_lt pos) simp [edist] - _ = - ((2 : ℝ≥0) ^ (1 / p.toReal) : ℝ≥0) * - edist (WithLp.equiv p (α × β) x) (WithLp.equiv p (α × β) y) := by + _ = (2 : ℝ≥0) ^ (1 / p.toReal) * edist (WithLp.equiv p _ x) (WithLp.equiv p _ y) := by simp only [← two_mul, ENNReal.mul_rpow_of_nonneg _ _ nonneg, ← ENNReal.rpow_mul, cancel, ENNReal.rpow_one, ← ENNReal.coe_rpow_of_nonneg _ nonneg, coe_ofNat] From 28ac6a5d57965b20fd01dbd61977a27062dff36e Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Tue, 15 Aug 2023 20:30:39 +1000 Subject: [PATCH 21/39] naming --- .../Analysis/InnerProductSpace/ProdL2.lean | 17 +- Mathlib/Analysis/NormedSpace/ProdLp.lean | 201 +++++++++--------- 2 files changed, 110 insertions(+), 108 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean index 41d225261d073..00c5a5c967b24 100644 --- a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean @@ -17,16 +17,17 @@ $$ This is recorded in this file as an inner product space instance on `ProdLp 2`. -/ -noncomputable section +namespace WithLp -variable {𝕜 : Type*} [IsROrC 𝕜] +variable {𝕜 : Type*} (E F : Type*) +variable [IsROrC 𝕜] [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] [NormedAddCommGroup F] + [InnerProductSpace 𝕜 F] -instance ProdLp.instInnerProductSpace (E F : Type _) [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] - [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] : +noncomputable instance instProdInnerProductSpace : InnerProductSpace 𝕜 (WithLp 2 (E × F)) where inner x y := inner x.fst y.fst + inner x.snd y.snd norm_sq_eq_inner x := by - simp [ProdLp.norm_sq_eq_of_L2, ← norm_sq_eq_inner] + simp [prod_norm_sq_eq_of_L2, ← norm_sq_eq_inner] conj_symm x y := by simp add_left x y z := by @@ -36,10 +37,8 @@ instance ProdLp.instInnerProductSpace (E F : Type _) [NormedAddCommGroup E] [Inn simp only [smul_fst, inner_smul_left, smul_snd] ring -variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] - -variable {F : Type*} [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] +variable {E F} @[simp] -theorem ProdLp.inner_apply (x y : WithLp 2 (E × F)) : +theorem prod_inner_apply (x y : WithLp 2 (E × F)) : inner (𝕜 := 𝕜) x y = inner x.fst y.fst + inner x.snd y.snd := rfl diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index ad12b2c819f54..b3ab7c25b2853 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -41,7 +41,7 @@ open Real Set Filter IsROrC Bornology BigOperators Uniformity Topology NNReal EN noncomputable section -namespace ProdLp +namespace WithLp variable (p : ℝ≥0∞) (𝕜 𝕜' : Type*) (α β : Type*) @@ -87,7 +87,7 @@ separate from `ProdLp.instPseudoEMetric` since the latter requires the type clas 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*. -/ -instance instEDist : EDist (WithLp p (α × β)) where +instance instProdEDist : EDist (WithLp p (α × β)) where edist f g := if _hp : p = 0 then (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) @@ -101,16 +101,16 @@ variable (x y : WithLp p (α × β)) (x' : α × β) open Classical in @[simp] -protected theorem edist_eq_card (f g : WithLp 0 (α × β)) : +theorem prod_edist_eq_card (f g : WithLp 0 (α × β)) : edist f g = (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := if_pos rfl -theorem edist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : +theorem prod_edist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : edist f g = (edist f.fst g.fst ^ p.toReal + edist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) := let hp' := ENNReal.toReal_pos_iff.mp hp (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) -theorem edist_eq_sup (f g : WithLp ∞ (α × β)) : +theorem prod_edist_eq_sup (f g : WithLp ∞ (α × β)) : edist f g = edist f.fst g.fst ⊔ edist f.snd g.snd := by dsimp [edist] exact if_neg ENNReal.top_ne_zero @@ -124,20 +124,21 @@ variable [PseudoEMetricSpace α] [PseudoEMetricSpace β] /-- This holds independent of `p` and does not require `[Fact (1 ≤ p)]`. We keep it separate from `ProdLp.instPseudoEMetricSpace` so it can be used also for `p < 1`. -/ -protected theorem edist_self (f : WithLp p (α × β)) : edist f f = 0 := by +theorem prod_edist_self (f : WithLp p (α × β)) : edist f f = 0 := by rcases p.trichotomy with (rfl | rfl | h) · simp - · simp [edist_eq_sup] - · simp [edist_eq_add h, ENNReal.zero_rpow_of_pos h, ENNReal.zero_rpow_of_pos (inv_pos.2 <| h)] + · simp [prod_edist_eq_sup] + · simp [prod_edist_eq_add h, ENNReal.zero_rpow_of_pos h, + ENNReal.zero_rpow_of_pos (inv_pos.2 <| h)] open Classical in /-- This holds independent of `p` and does not require `[Fact (1 ≤ p)]`. We keep it separate from `ProdLp.instPseudoEMetricSpace` so it can be used also for `p < 1`. -/ -protected theorem edist_comm (f g : WithLp p (α × β)) : edist f g = edist g f := by +theorem prod_edist_comm (f g : WithLp p (α × β)) : edist f g = edist g f := by rcases p.trichotomy with (rfl | rfl | h) - · simp [ProdLp.edist_eq_card, eq_comm] - · simp only [edist_eq_sup, edist_comm] - · simp only [edist_eq_add h, edist_comm] + · simp [prod_edist_eq_card, eq_comm] + · simp only [prod_edist_eq_sup, edist_comm] + · simp only [prod_edist_eq_add h, edist_comm] end EdistProp @@ -153,7 +154,7 @@ separate from `ProdLp.instPseudoMetricSpace` since the latter requires the type 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*. -/ -instance instDist : Dist (WithLp p (α × β)) where +instance instProdDist : Dist (WithLp p (α × β)) where dist f g := if _hp : p = 0 then (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) @@ -164,16 +165,16 @@ instance instDist : Dist (WithLp p (α × β)) where variable {α β} open Classical in -theorem dist_eq_card (f g : WithLp 0 (α × β)) : dist f g = +theorem prod_dist_eq_card (f g : WithLp 0 (α × β)) : dist f g = (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := if_pos rfl -theorem dist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : +theorem prod_dist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : dist f g = (dist f.fst g.fst ^ p.toReal + dist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) := let hp' := ENNReal.toReal_pos_iff.mp hp (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) -theorem dist_eq_sup (f g : WithLp ∞ (α × β)) : +theorem prod_dist_eq_sup (f g : WithLp ∞ (α × β)) : dist f g = dist f.fst g.fst ⊔ dist f.snd g.snd := by dsimp [dist] exact if_neg ENNReal.top_ne_zero @@ -191,7 +192,7 @@ 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*. -/ -instance instNorm : Norm (WithLp p (α × β)) where +instance instProdNorm : Norm (WithLp p (α × β)) where norm f := if _hp : p = 0 then (if f.fst = 0 then 0 else 1) + (if f.snd = 0 then 0 else 1) @@ -202,15 +203,15 @@ variable {p α β} open Classical in @[simp] -protected theorem norm_eq_card (f : WithLp 0 (α × β)) : +theorem prod_norm_eq_card (f : WithLp 0 (α × β)) : ‖f‖ = (if f.fst = 0 then 0 else 1) + (if f.snd = 0 then 0 else 1) := if_pos rfl -theorem norm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖ = ‖f.fst‖ ⊔ ‖f.snd‖ := by +theorem prod_norm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖ = ‖f.fst‖ ⊔ ‖f.snd‖ := by dsimp [Norm.norm] exact if_neg ENNReal.top_ne_zero -theorem norm_eq_add (hp : 0 < p.toReal) (f : WithLp p (α × β)) : +theorem prod_norm_eq_add (hp : 0 < p.toReal) (f : WithLp p (α × β)) : ‖f‖ = (‖f.fst‖ ^ p.toReal + ‖f.snd‖ ^ p.toReal) ^ (1 / p.toReal) := let hp' := ENNReal.toReal_pos_iff.mp hp (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) @@ -244,16 +245,16 @@ coincide with the product one. Therefore, we do not register it as an instance. 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`. -/ -def pseudoEmetricAux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : +def prodPseudoEmetricAux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : PseudoEMetricSpace (WithLp p (α × β)) where - edist_self := ProdLp.edist_self p - edist_comm := ProdLp.edist_comm p + edist_self := prod_edist_self p + edist_comm := prod_edist_comm p edist_triangle f g h := by rcases p.dichotomy with (rfl | hp) - · simp only [edist_eq_sup] + · simp only [prod_edist_eq_sup] exact sup_le ((edist_triangle _ g.fst _).trans <| add_le_add le_sup_left le_sup_left) ((edist_triangle _ g.snd _).trans <| add_le_add le_sup_right le_sup_right) - · simp only [edist_eq_add (zero_lt_one.trans_le hp)] + · simp only [prod_edist_eq_add (zero_lt_one.trans_le hp)] calc (edist f.fst h.fst ^ p.toReal + edist f.snd h.snd ^ p.toReal) ^ (1 / p.toReal) ≤ ((edist f.fst g.fst + edist g.fst h.fst) ^ p.toReal + @@ -271,11 +272,11 @@ def pseudoEmetricAux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : Finset.sum_singleton] at this exact this -attribute [local instance] ProdLp.pseudoEmetricAux +attribute [local instance] WithLp.prodPseudoEmetricAux /-- An auxiliary lemma used twice in the proof of `ProdLp.pseudoMetricAux` below. Not intended for use outside this file. -/ -theorem sup_edist_ne_top_aux {α β : Type _} +theorem prod_sup_edist_ne_top_aux {α β : Type _} [PseudoMetricSpace α] [PseudoMetricSpace β] (f g : WithLp ∞ (α × β)) : edist f.fst g.fst ⊔ edist f.snd g.snd ≠ ⊤ := by refine ne_of_lt ?_ @@ -291,25 +292,25 @@ structure and the bornology by the product ones using this pseudometric space, See note [reducible non-instances] -/ @[reducible] -def pseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : +def prodPseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : PseudoMetricSpace (WithLp p (α × β)) := PseudoEMetricSpace.toPseudoMetricSpaceOfDist dist (fun f g => by rcases p.dichotomy with (rfl | h) - · exact sup_edist_ne_top_aux f g - · rw [edist_eq_add (zero_lt_one.trans_le h)] + · exact prod_sup_edist_ne_top_aux f g + · rw [prod_edist_eq_add (zero_lt_one.trans_le h)] refine ENNReal.rpow_ne_top_of_nonneg (by positivity) (ne_of_lt ?_) simp [ENNReal.add_lt_top, ENNReal.rpow_lt_top_of_nonneg, edist_ne_top] ) fun f g => by rcases p.dichotomy with (rfl | h) - · rw [edist_eq_sup, dist_eq_sup] + · rw [prod_edist_eq_sup, prod_dist_eq_sup] refine' le_antisymm (sup_le _ _) _ - · rw [← ENNReal.ofReal_le_iff_le_toReal (sup_edist_ne_top_aux f g), ← - PseudoMetricSpace.edist_dist] + · rw [← ENNReal.ofReal_le_iff_le_toReal (prod_sup_edist_ne_top_aux f g), + ← PseudoMetricSpace.edist_dist] exact le_sup_left - · rw [← ENNReal.ofReal_le_iff_le_toReal (sup_edist_ne_top_aux f g), ← - PseudoMetricSpace.edist_dist] + · rw [← ENNReal.ofReal_le_iff_le_toReal (prod_sup_edist_ne_top_aux f g), + ← PseudoMetricSpace.edist_dist] exact le_sup_right · refine ENNReal.toReal_le_of_le_ofReal ?_ ?_ · simp only [ge_iff_le, le_sup_iff, dist_nonneg] @@ -318,18 +319,18 @@ def pseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : ENNReal.rpow_ne_top_of_nonneg (zero_le_one.trans h) (edist_ne_top _ _) have h2 : edist f.snd g.snd ^ p.toReal ≠ ⊤ := ENNReal.rpow_ne_top_of_nonneg (zero_le_one.trans h) (edist_ne_top _ _) - simp only [edist_eq_add (zero_lt_one.trans_le h), dist_edist, ENNReal.toReal_rpow, - dist_eq_add (zero_lt_one.trans_le h), ← ENNReal.toReal_add h1 h2] + simp only [prod_edist_eq_add (zero_lt_one.trans_le h), dist_edist, ENNReal.toReal_rpow, + prod_dist_eq_add (zero_lt_one.trans_le h), ← ENNReal.toReal_add h1 h2] -attribute [local instance] ProdLp.pseudoMetricAux +attribute [local instance] WithLp.prodPseudoMetricAux -theorem lipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : +theorem prod_lipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : LipschitzWith 1 (WithLp.equiv p (α × β)) := by intro x y rcases p.dichotomy with (rfl | h) · simp [edist] · have cancel : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le h).ne' - rw [edist_eq_add (zero_lt_one.trans_le h)] + rw [prod_edist_eq_add (zero_lt_one.trans_le h)] simp only [edist, forall_prop_of_true, one_mul, ENNReal.coe_one, ge_iff_le, sup_le_iff] constructor @@ -348,7 +349,7 @@ theorem lipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] example (a : ENNReal) : a + a = 2*a := by exact Eq.symm (two_mul a) -theorem antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : +theorem prod_antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : AntilipschitzWith ((2 : ℝ≥0) ^ (1 / p).toReal) (WithLp.equiv p (α × β)) := by intro x y rcases p.dichotomy with (rfl | h) @@ -356,7 +357,7 @@ theorem antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace · have pos : 0 < p.toReal := by positivity have nonneg : 0 ≤ 1 / p.toReal := by positivity have cancel : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (ne_of_gt pos) - rw [edist_eq_add pos, ENNReal.toReal_div 1 p] + rw [prod_edist_eq_add pos, ENNReal.toReal_div 1 p] simp only [edist, ← one_div, ENNReal.one_toReal] calc (edist x.fst y.fst ^ p.toReal + edist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) ≤ @@ -372,22 +373,22 @@ theorem antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace ← ENNReal.rpow_mul, cancel, ENNReal.rpow_one, ← ENNReal.coe_rpow_of_nonneg _ nonneg, coe_ofNat] -theorem aux_uniformity_eq [PseudoEMetricSpace α] [PseudoEMetricSpace β] : +theorem prod_aux_uniformity_eq [PseudoEMetricSpace α] [PseudoEMetricSpace β] : 𝓤 (WithLp p (α × β)) = 𝓤[instUniformSpaceProd] := by have A : UniformInducing (WithLp.equiv p (α × β)) := - (antilipschitzWith_equiv_aux p α β).uniformInducing - (lipschitzWith_equiv_aux p α β).uniformContinuous + (prod_antilipschitzWith_equiv_aux p α β).uniformInducing + (prod_lipschitzWith_equiv_aux p α β).uniformContinuous have : (fun x : WithLp p (α × β) × WithLp p (α × β) => ((WithLp.equiv p (α × β)) x.fst, (WithLp.equiv p (α × β)) x.snd)) = id := by ext i <;> rfl rw [← A.comap_uniformity, this, comap_id] -theorem aux_cobounded_eq [PseudoMetricSpace α] [PseudoMetricSpace β] : +theorem prod_aux_cobounded_eq [PseudoMetricSpace α] [PseudoMetricSpace β] : cobounded (WithLp p (α × β)) = @cobounded _ Prod.instBornology := calc cobounded (WithLp p (α × β)) = comap (WithLp.equiv p (α × β)) (cobounded _) := - le_antisymm (antilipschitzWith_equiv_aux p α β).tendsto_cobounded.le_comap - (lipschitzWith_equiv_aux p α β).comap_cobounded_le + le_antisymm (prod_antilipschitzWith_equiv_aux p α β).tendsto_cobounded.le_comap + (prod_lipschitzWith_equiv_aux p α β).comap_cobounded_le _ = _ := comap_id end Aux @@ -423,83 +424,84 @@ variable [Fact (1 ≤ p)] /-- `PseudoEMetricSpace` instance on the product of two pseudoemetric spaces, using the `L^p` pseudoedistance, and having as uniformity the product uniformity. -/ -instance instPseudoEMetricSpace [PseudoEMetricSpace α] [PseudoEMetricSpace β] : +instance instProdPseudoEMetricSpace [PseudoEMetricSpace α] [PseudoEMetricSpace β] : PseudoEMetricSpace (WithLp p (α × β)) := - (pseudoEmetricAux p α β).replaceUniformity (aux_uniformity_eq p α β).symm + (prodPseudoEmetricAux p α β).replaceUniformity (prod_aux_uniformity_eq p α β).symm /-- `EMetricSpace` instance on the product of two emetric spaces, using the `L^p` edistance, and having as uniformity the product uniformity. -/ -instance instEMetricSpace [EMetricSpace α] [EMetricSpace β] : EMetricSpace (WithLp p (α × β)) := +instance instProdEMetricSpace [EMetricSpace α] [EMetricSpace β] : EMetricSpace (WithLp p (α × β)) := @EMetricSpace.ofT0PseudoEMetricSpace (WithLp p (α × β)) _ instT0SpaceProdInstTopologicalSpaceProd /-- `PseudoMetricSpace` instance on the product of two pseudometric spaces, using the `L^p` distance, and having as uniformity the product uniformity. -/ -instance instPseudoMetricSpace [PseudoMetricSpace α] [PseudoMetricSpace β] : +instance instProdPseudoMetricSpace [PseudoMetricSpace α] [PseudoMetricSpace β] : PseudoMetricSpace (WithLp p (α × β)) := - ((pseudoMetricAux p α β).replaceUniformity (aux_uniformity_eq p α β).symm).replaceBornology - fun s => Filter.ext_iff.1 (aux_cobounded_eq p α β).symm sᶜ + ((prodPseudoMetricAux p α β).replaceUniformity + (prod_aux_uniformity_eq p α β).symm).replaceBornology + fun s => Filter.ext_iff.1 (prod_aux_cobounded_eq p α β).symm sᶜ /-- `MetricSpace` instance on the product of two metric spaces, using the `L^p` distance, and having as uniformity the product uniformity. -/ -instance instMetricSpace [MetricSpace α] [MetricSpace β] : MetricSpace (WithLp p (α × β)) := +instance instProdMetricSpace [MetricSpace α] [MetricSpace β] : MetricSpace (WithLp p (α × β)) := MetricSpace.ofT0PseudoMetricSpace _ variable {p α β} -theorem nndist_eq_add [PseudoMetricSpace α] [PseudoMetricSpace β] +theorem prod_nndist_eq_add [PseudoMetricSpace α] [PseudoMetricSpace β] (hp : p ≠ ∞) (x y : WithLp p (α × β)) : nndist x y = (nndist x.fst y.fst ^ p.toReal + nndist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) := -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast - exact dist_eq_add (p.toReal_pos_iff_ne_top.mpr hp) _ _ + exact prod_dist_eq_add (p.toReal_pos_iff_ne_top.mpr hp) _ _ -theorem nndist_eq_sup [PseudoMetricSpace α] [PseudoMetricSpace β] (x y : WithLp ∞ (α × β)) : +theorem prod_nndist_eq_sup [PseudoMetricSpace α] [PseudoMetricSpace β] (x y : WithLp ∞ (α × β)) : nndist x y = nndist x.fst y.fst ⊔ nndist x.snd y.snd := -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast - exact dist_eq_sup _ _ + exact prod_dist_eq_sup _ _ variable (p α β) -theorem lipschitzWith_equiv [PseudoEMetricSpace α] [PseudoEMetricSpace β] : +theorem prod_lipschitzWith_equiv [PseudoEMetricSpace α] [PseudoEMetricSpace β] : LipschitzWith 1 (WithLp.equiv p (α × β)) := - lipschitzWith_equiv_aux p α β + prod_lipschitzWith_equiv_aux p α β -theorem antilipschitzWith_equiv [PseudoEMetricSpace α] [PseudoEMetricSpace β] : +theorem prod_antilipschitzWith_equiv [PseudoEMetricSpace α] [PseudoEMetricSpace β] : AntilipschitzWith ((2 : ℝ≥0) ^ (1 / p).toReal) (WithLp.equiv p (α × β)) := - antilipschitzWith_equiv_aux p α β + prod_antilipschitzWith_equiv_aux p α β theorem infty_equiv_isometry [PseudoEMetricSpace α] [PseudoEMetricSpace β] : Isometry (WithLp.equiv ∞ (α × β)) := fun x y => - le_antisymm (by simpa only [ENNReal.coe_one, one_mul] using lipschitzWith_equiv ∞ α β x y) + le_antisymm (by simpa only [ENNReal.coe_one, one_mul] using prod_lipschitzWith_equiv ∞ α β x y) (by simpa only [ENNReal.div_top, ENNReal.zero_toReal, NNReal.rpow_zero, ENNReal.coe_one, - one_mul] using antilipschitzWith_equiv ∞ α β x y) + one_mul] using prod_antilipschitzWith_equiv ∞ α β x y) /-- Seminormed group instance on the product of two normed groups, using the `L^p` norm. -/ -instance instSeminormedAddCommGroup [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] : +instance instProdSeminormedAddCommGroup [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] : SeminormedAddCommGroup (WithLp p (α × β)) := { WithLp.instAddCommGroup _ _ with dist_eq := fun x y => by rcases p.dichotomy with (rfl | h) - · simp only [dist_eq_sup, norm_eq_sup, dist_eq_norm] + · simp only [prod_dist_eq_sup, prod_norm_eq_sup, dist_eq_norm] rfl · have : p ≠ ∞ := by intro hp rw [hp, ENNReal.top_toReal] at h linarith - simp only [dist_eq_add (zero_lt_one.trans_le h), norm_eq_add (zero_lt_one.trans_le h), - dist_eq_norm] + simp only [prod_dist_eq_add (zero_lt_one.trans_le h), + prod_norm_eq_add (zero_lt_one.trans_le h), dist_eq_norm] rfl } /-- normed group instance on the product of two normed groups, using the `L^p` norm. -/ -instance normedAddCommGroup [NormedAddCommGroup α] [NormedAddCommGroup β] : +instance instProdNormedAddCommGroup [NormedAddCommGroup α] [NormedAddCommGroup β] : NormedAddCommGroup (WithLp p (α × β)) := - { ProdLp.instSeminormedAddCommGroup p α β with + { instProdSeminormedAddCommGroup p α β with eq_of_dist_eq_zero := eq_of_dist_eq_zero } section norm_of @@ -507,56 +509,57 @@ section norm_of variable {p α β} variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] -theorem nnnorm_eq_add (hp : p ≠ ∞) (f : WithLp p (α × β)) : +theorem prod_nnnorm_eq_add (hp : p ≠ ∞) (f : WithLp p (α × β)) : ‖f‖₊ = (‖f.fst‖₊ ^ p.toReal + ‖f.snd‖₊ ^ p.toReal) ^ (1 / p.toReal) := by ext - simp [norm_eq_add (p.toReal_pos_iff_ne_top.mpr hp)] + simp [prod_norm_eq_add (p.toReal_pos_iff_ne_top.mpr hp)] -theorem nnnorm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖₊ = ‖f.fst‖₊ ⊔ ‖f.snd‖₊ := by +theorem prod_nnnorm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖₊ = ‖f.fst‖₊ ⊔ ‖f.snd‖₊ := by ext norm_cast -theorem norm_eq_of_nat (n : ℕ) (h : p = n) (f : WithLp p (α × β)) : +theorem prod_norm_eq_of_nat (n : ℕ) (h : p = n) (f : WithLp p (α × β)) : ‖f‖ = (‖f.fst‖ ^ n + ‖f.snd‖ ^ n) ^ (1 / (n : ℝ)) := by have := p.toReal_pos_iff_ne_top.mpr (ne_of_eq_of_ne h <| ENNReal.nat_ne_top n) simp only [one_div, h, Real.rpow_nat_cast, ENNReal.toReal_nat, eq_self_iff_true, Finset.sum_congr, - norm_eq_add this] + prod_norm_eq_add this] -theorem norm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by - rw [norm_eq_of_nat 2 (by norm_cast) _] -- Porting note: was `convert` +theorem prod_norm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by + rw [prod_norm_eq_of_nat 2 (by norm_cast) _] -- Porting note: was `convert` rw [Real.sqrt_eq_rpow] norm_cast -theorem nnnorm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖₊ = NNReal.sqrt (‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2) := +theorem prod_nnnorm_eq_of_L2 (x : WithLp 2 (α × β)) : + ‖x‖₊ = NNReal.sqrt (‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2) := -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast - exact norm_eq_of_L2 x + exact prod_norm_eq_of_L2 x variable (α β) -theorem norm_sq_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ ^ 2 = ‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2 := by +theorem prod_norm_sq_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ ^ 2 = ‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2 := by suffices ‖x‖₊ ^ 2 = ‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2 by simpa only [NNReal.coe_sum] using congr_arg ((↑) : ℝ≥0 → ℝ) this - rw [nnnorm_eq_of_L2, NNReal.sq_sqrt] + rw [prod_nnnorm_eq_of_L2, NNReal.sq_sqrt] variable {α β} -theorem dist_eq_of_L2 (x y : WithLp 2 (α × β)) : +theorem prod_dist_eq_of_L2 (x y : WithLp 2 (α × β)) : dist x y = (dist x.fst y.fst ^ 2 + dist x.snd y.snd ^ 2).sqrt := by - simp_rw [dist_eq_norm, norm_eq_of_L2, Pi.sub_apply] + simp_rw [dist_eq_norm, prod_norm_eq_of_L2, Pi.sub_apply] rfl -theorem nndist_eq_of_L2 (x y : WithLp 2 (α × β)) : +theorem prod_nndist_eq_of_L2 (x y : WithLp 2 (α × β)) : nndist x y = NNReal.sqrt (nndist x.fst y.fst ^ 2 + nndist x.snd y.snd ^ 2) := -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast - exact dist_eq_of_L2 _ _ + exact prod_dist_eq_of_L2 _ _ -theorem edist_eq_of_L2 (x y : WithLp 2 (α × β)) : +theorem prod_edist_eq_of_L2 (x y : WithLp 2 (α × β)) : edist x y = (edist x.fst y.fst ^ 2 + edist x.snd y.snd ^ 2) ^ (1 / 2 : ℝ) := by - simp [ProdLp.edist_eq_add] + simp [prod_edist_eq_add] end norm_of @@ -568,20 +571,20 @@ variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] [SeminormedAddCommGroup β] [NormedSpace 𝕜 β] /-- The product of two normed spaces is a normed space, with the `L^p` norm. -/ -instance instNormedSpace : +instance instProdNormedSpace : NormedSpace 𝕜 (WithLp p (α × β)) := { WithLp.instModule p 𝕜 (α × β) with norm_smul_le := fun c f => by rcases p.dichotomy with (rfl | hp) · letI : Module 𝕜 (WithLp ∞ (α × β)) := Prod.instModule suffices ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊ by exact_mod_cast NNReal.coe_mono this.le - simp only [nnnorm_eq_sup, NNReal.mul_sup, ← nnnorm_smul] + simp only [prod_nnnorm_eq_sup, NNReal.mul_sup, ← nnnorm_smul] rfl · have : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le hp).ne' have smul_fst : (c • f).fst = c • f.fst := rfl have smul_snd : (c • f).snd = c • f.snd := rfl - simp only [norm_eq_add (zero_lt_one.trans_le hp), norm_smul, Real.mul_rpow, norm_nonneg, - smul_fst, smul_snd] + simp only [prod_norm_eq_add (zero_lt_one.trans_le hp), norm_smul, Real.mul_rpow, + norm_nonneg, smul_fst, smul_snd] rw [← mul_add, mul_rpow (rpow_nonneg_of_nonneg (norm_nonneg _) _), ← rpow_mul (norm_nonneg _), this, Real.rpow_one] positivity } @@ -644,7 +647,7 @@ section Equiv /-- The canonical map `ProdLp.equiv` between `ProdLp ∞ β` and `Π i, β i` as a linear isometric equivalence. -/ -def equivₗᵢ : WithLp ∞ (α × β) ≃ₗᵢ[𝕜] α × β := +def prodEquivₗᵢ : WithLp ∞ (α × β) ≃ₗᵢ[𝕜] α × β := { WithLp.equiv ∞ (α × β) with map_add' := fun f g => rfl map_smul' := fun c f => rfl @@ -663,22 +666,22 @@ theorem nnnorm_equiv_symm_fst [hp : Fact (1 ≤ p)] (x : α) : ‖(WithLp.equiv p (α × β)).symm (x, 0)‖₊ = ‖x‖₊ := by induction p using ENNReal.recTopCoe generalizing hp with | top => - simp [nnnorm_eq_sup] + simp [prod_nnnorm_eq_sup] | coe p => have hp0 : (p : ℝ) ≠ 0 := by exact_mod_cast (zero_lt_one.trans_le <| Fact.out (p := 1 ≤ (p : ℝ≥0∞))).ne' - simp [nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] + simp [prod_nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] @[simp] theorem nnnorm_equiv_symm_snd [hp : Fact (1 ≤ p)] (y : β) : ‖(WithLp.equiv p (α × β)).symm (0, y)‖₊ = ‖y‖₊ := by induction p using ENNReal.recTopCoe generalizing hp with | top => - simp [nnnorm_eq_sup] + simp [prod_nnnorm_eq_sup] | coe p => have hp0 : (p : ℝ) ≠ 0 := by exact_mod_cast (zero_lt_one.trans_le <| Fact.out (p := 1 ≤ (p : ℝ≥0∞))).ne' - simp [nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] + simp [prod_nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] @[simp] theorem norm_equiv_symm_fst (x : α) : ‖(WithLp.equiv p (α × β)).symm (x, 0)‖ = ‖x‖ := @@ -732,9 +735,9 @@ variable (𝕜 p α β) /-- `WithLp.equiv` as a continuous linear equivalence. -/ @[simps! (config := { fullyApplied := false }) apply symm_apply] -protected def continuousLinearEquiv : WithLp p (α × β) ≃L[𝕜] α × β where +protected def prodContinuousLinearEquiv : WithLp p (α × β) ≃L[𝕜] α × β where toLinearEquiv := WithLp.linearEquiv _ _ _ continuous_toFun := continuous_equiv _ _ _ continuous_invFun := continuous_equiv_symm _ _ _ -end ProdLp +end WithLp From beace9e4f3d6763c8f3707d14f7b0d826fdbe59a Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 16 Aug 2023 11:58:13 +1000 Subject: [PATCH 22/39] moved algebra lemmas up --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 156 ++++++++++++----------- 1 file changed, 83 insertions(+), 73 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index b3ab7c25b2853..e9e46c5f61529 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -41,13 +41,68 @@ open Real Set Filter IsROrC Bornology BigOperators Uniformity Topology NNReal EN noncomputable section +variable (p : ℝ≥0∞) (𝕜 𝕜' α β : Type*) + namespace WithLp -variable (p : ℝ≥0∞) (𝕜 𝕜' : Type*) (α β : Type*) +section algebra + +variable {p 𝕜 α β} + +variable [Semiring 𝕜] [AddCommGroup α] [AddCommGroup β] +variable (x y : WithLp p (α × β)) (c : 𝕜) + +@[simp] +theorem zero_fst : (0 : WithLp p (α × β)).fst = 0 := + rfl + +@[simp] +theorem zero_snd : (0 : WithLp p (α × β)).snd = 0 := + rfl + +@[simp] +theorem add_fst : (x + y).fst = x.fst + y.fst := + rfl + +@[simp] +theorem add_snd : (x + y).snd = x.snd + y.snd := + rfl + +@[simp] +theorem sub_fst : (x - y).fst = x.fst - y.fst := + rfl + +@[simp] +theorem sub_snd : (x - y).snd = x.snd - y.snd := + rfl + +@[simp] +theorem neg_fst : (-x).fst = -x.fst := + rfl + +@[simp] +theorem neg_snd : (-x).snd = -x.snd := + rfl + +variable [Module 𝕜 α] [Module 𝕜 β] + +@[simp] +theorem smul_fst : (c • x).fst = c • x.fst := + rfl + +@[simp] +theorem smul_snd : (c • x).snd = c • x.snd := + rfl + +end algebra /-! Note that the unapplied versions of these lemmas are deliberately omitted, as they break the use of the type synonym. -/ +section equiv + +variable {p α β} + @[simp] theorem equiv_fst (x : WithLp p (α × β)) : (WithLp.equiv p (α × β) x).fst = x.fst := rfl @@ -64,6 +119,8 @@ theorem equiv_symm_fst (x : α × β) : ((WithLp.equiv p (α × β)).symm x).fst theorem equiv_symm_snd (x : α × β) : ((WithLp.equiv p (α × β)).symm x).snd = x.snd := rfl +end equiv + section DistNorm /-! @@ -71,11 +128,11 @@ section DistNorm 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.to_real`. +the rewrite lemmas for each of three cases `p = 0`, `p = ∞` and `0 < p.toReal`. -/ -section Edist +section EDist variable [EDist α] [EDist β] @@ -95,8 +152,7 @@ instance instProdEDist : EDist (WithLp p (α × β)) where if p = ∞ then edist f.fst g.fst ⊔ edist f.snd g.snd else (edist f.fst g.fst ^ p.toReal + edist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) -variable {α β} - +variable {p α β} variable (x y : WithLp p (α × β)) (x' : α × β) open Classical in @@ -105,7 +161,7 @@ theorem prod_edist_eq_card (f g : WithLp 0 (α × β)) : edist f g = (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := if_pos rfl -theorem prod_edist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : +theorem prod_edist_eq_add (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : edist f g = (edist f.fst g.fst ^ p.toReal + edist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) := let hp' := ENNReal.toReal_pos_iff.mp hp (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) @@ -115,15 +171,15 @@ theorem prod_edist_eq_sup (f g : WithLp ∞ (α × β)) : dsimp [edist] exact if_neg ENNReal.top_ne_zero -end Edist +end EDist -section EdistProp +section EDistProp variable {α β} variable [PseudoEMetricSpace α] [PseudoEMetricSpace β] /-- This holds independent of `p` and does not require `[Fact (1 ≤ p)]`. We keep it separate -from `ProdLp.instPseudoEMetricSpace` so it can be used also for `p < 1`. -/ +from `WithLp.instProdPseudoEMetricSpace` so it can be used also for `p < 1`. -/ theorem prod_edist_self (f : WithLp p (α × β)) : edist f f = 0 := by rcases p.trichotomy with (rfl | rfl | h) · simp @@ -133,14 +189,14 @@ theorem prod_edist_self (f : WithLp p (α × β)) : edist f f = 0 := by open Classical in /-- This holds independent of `p` and does not require `[Fact (1 ≤ p)]`. We keep it separate -from `ProdLp.instPseudoEMetricSpace` so it can be used also for `p < 1`. -/ +from `WithLp.instProdPseudoEMetricSpace` so it can be used also for `p < 1`. -/ theorem prod_edist_comm (f g : WithLp p (α × β)) : edist f g = edist g f := by rcases p.trichotomy with (rfl | rfl | h) · simp [prod_edist_eq_card, eq_comm] · simp only [prod_edist_eq_sup, edist_comm] · simp only [prod_edist_eq_add h, edist_comm] -end EdistProp +end EDistProp section Dist @@ -148,7 +204,7 @@ variable [Dist α] [Dist β] open Classical in /-- Endowing the space `WithLp p (α × β)` with the `L^p` distance. We register this instance -separate from `ProdLp.instPseudoMetricSpace` since the latter requires the type class hypothesis +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 @@ -162,14 +218,14 @@ instance instProdDist : Dist (WithLp p (α × β)) where if p = ∞ then dist f.fst g.fst ⊔ dist f.snd g.snd else (dist f.fst g.fst ^ p.toReal + dist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) -variable {α β} +variable {p α β} open Classical in theorem prod_dist_eq_card (f g : WithLp 0 (α × β)) : dist f g = (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := if_pos rfl -theorem prod_dist_eq_add {p : ℝ≥0∞} (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : +theorem prod_dist_eq_add (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : dist f g = (dist f.fst g.fst ^ p.toReal + dist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) := let hp' := ENNReal.toReal_pos_iff.mp hp (if_neg hp'.1.ne').trans (if_neg hp'.2.ne) @@ -187,7 +243,7 @@ variable [Norm α] [Zero α] [Norm β] [Zero β] open Classical in /-- Endowing the space `WithLp p (α × β)` with the `L^p` norm. We register this instance -separate from `ProdLp.instSeminormedAddCommGroup` since the latter requires the type class +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 @@ -274,14 +330,18 @@ def prodPseudoEmetricAux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : attribute [local instance] WithLp.prodPseudoEmetricAux -/-- An auxiliary lemma used twice in the proof of `ProdLp.pseudoMetricAux` below. Not intended for -use outside this file. -/ -theorem prod_sup_edist_ne_top_aux {α β : Type _} - [PseudoMetricSpace α] [PseudoMetricSpace β] (f g : WithLp ∞ (α × β)) : +variable {α β} + +/-- An auxiliary lemma used twice in the proof of `WithLp.prodPseudoMetricAux` below. Not intended +for use outside this file. -/ +theorem prod_sup_edist_ne_top_aux [PseudoMetricSpace α] [PseudoMetricSpace β] + (f g : WithLp ∞ (α × β)) : edist f.fst g.fst ⊔ edist f.snd g.snd ≠ ⊤ := by refine ne_of_lt ?_ simp [edist, PseudoMetricSpace.edist_dist] +variable (α β) + /-- 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 @@ -347,8 +407,6 @@ theorem prod_lipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace apply ENNReal.rpow_le_rpow _ (by positivity) simp only [self_le_add_left] -example (a : ENNReal) : a + a = 2*a := by exact Eq.symm (two_mul a) - theorem prod_antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : AntilipschitzWith ((2 : ℝ≥0) ^ (1 / p).toReal) (WithLp.equiv p (α × β)) := by intro x y @@ -565,7 +623,7 @@ end norm_of variable [NormedField 𝕜] [NormedField 𝕜'] -section normed_space_inst +section InstNormedSpace variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] [SeminormedAddCommGroup β] [NormedSpace 𝕜 β] @@ -589,60 +647,14 @@ instance instProdNormedSpace : ← rpow_mul (norm_nonneg _), this, Real.rpow_one] positivity } -end normed_space_inst +end InstNormedSpace -/- Register simplification lemmas for the applications of `ProdLp` elements, as the usual lemmas -for Prod types will not trigger. -/ +/- Register simplification lemmas for the applications of `WithLp p (α × β)` elements, as the usual +lemmas for `Prod` will not trigger. -/ variable {𝕜 𝕜' p α β} variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] [SeminormedAddCommGroup β] [NormedSpace 𝕜 β] -section algebra - -variable (x y : WithLp p (α × β)) (c : 𝕜) - -@[simp] -theorem zero_fst : (0 : WithLp p (α × β)).fst = 0 := - rfl - -@[simp] -theorem zero_snd : (0 : WithLp p (α × β)).snd = 0 := - rfl - -@[simp] -theorem add_fst : (x + y).fst = x.fst + y.fst := - rfl - -@[simp] -theorem add_snd : (x + y).snd = x.snd + y.snd := - rfl - -@[simp] -theorem sub_fst : (x - y).fst = x.fst - y.fst := - rfl - -@[simp] -theorem sub_snd : (x - y).snd = x.snd - y.snd := - rfl - -@[simp] -theorem smul_fst : (c • x).fst = c • x.fst := - rfl - -@[simp] -theorem smul_snd : (c • x).snd = c • x.snd := - rfl - -@[simp] -theorem neg_fst : (-x).fst = -x.fst := - rfl - -@[simp] -theorem neg_snd : (-x).snd = -x.snd := - rfl - -end algebra - section Equiv /-- The canonical map `ProdLp.equiv` between `ProdLp ∞ β` and `Π i, β i` as a linear isometric @@ -653,8 +665,6 @@ def prodEquivₗᵢ : WithLp ∞ (α × β) ≃ₗᵢ[𝕜] α × β := map_smul' := fun c f => rfl norm_map' := fun f => by simp [Norm.norm] } -variable (x y : WithLp p (α × β)) (x' y' : α × β) (c : 𝕜) - end Equiv section Single From e15ece9d3dc16840175c154b2a66d19910227e11 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 16 Aug 2023 12:04:40 +1000 Subject: [PATCH 23/39] simple corrections --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index e9e46c5f61529..7b427b05d0550 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -16,7 +16,8 @@ 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. +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. @@ -185,7 +186,7 @@ theorem prod_edist_self (f : WithLp p (α × β)) : edist f f = 0 := by · simp · simp [prod_edist_eq_sup] · simp [prod_edist_eq_add h, ENNReal.zero_rpow_of_pos h, - ENNReal.zero_rpow_of_pos (inv_pos.2 <| h)] + ENNReal.zero_rpow_of_pos (inv_pos.2 <| h)] open Classical in /-- This holds independent of `p` and does not require `[Fact (1 ≤ p)]`. We keep it separate @@ -336,9 +337,8 @@ variable {α β} for use outside this file. -/ theorem prod_sup_edist_ne_top_aux [PseudoMetricSpace α] [PseudoMetricSpace β] (f g : WithLp ∞ (α × β)) : - edist f.fst g.fst ⊔ edist f.snd g.snd ≠ ⊤ := by - refine ne_of_lt ?_ - simp [edist, PseudoMetricSpace.edist_dist] + edist f.fst g.fst ⊔ edist f.snd g.snd ≠ ⊤ := + ne_of_lt <| by simp [edist, PseudoMetricSpace.edist_dist] variable (α β) @@ -657,7 +657,7 @@ variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] section Equiv -/-- The canonical map `ProdLp.equiv` between `ProdLp ∞ β` and `Π i, β i` as a linear isometric +/-- The canonical map `WithLp.equiv` between `WithLp ∞ (α × β)` and `α × β` as a linear isometric equivalence. -/ def prodEquivₗᵢ : WithLp ∞ (α × β) ≃ₗᵢ[𝕜] α × β := { WithLp.equiv ∞ (α × β) with From 1a864f6494d6ef6539111360aaad78eb85a40d1e Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 16 Aug 2023 12:24:05 +1000 Subject: [PATCH 24/39] more suggestions --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 71 +++++++++++------------- 1 file changed, 32 insertions(+), 39 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 7b427b05d0550..1552c87848239 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -359,8 +359,7 @@ def prodPseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : rcases p.dichotomy with (rfl | h) · exact prod_sup_edist_ne_top_aux f g · rw [prod_edist_eq_add (zero_lt_one.trans_le h)] - refine - ENNReal.rpow_ne_top_of_nonneg (by positivity) (ne_of_lt ?_) + refine ENNReal.rpow_ne_top_of_nonneg (by positivity) (ne_of_lt ?_) simp [ENNReal.add_lt_top, ENNReal.rpow_lt_top_of_nonneg, edist_ne_top] ) fun f g => by rcases p.dichotomy with (rfl | h) @@ -391,8 +390,7 @@ theorem prod_lipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricSpace · simp [edist] · have cancel : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le h).ne' rw [prod_edist_eq_add (zero_lt_one.trans_le h)] - simp only [edist, forall_prop_of_true, one_mul, ENNReal.coe_one, ge_iff_le, - sup_le_iff] + simp only [edist, forall_prop_of_true, one_mul, ENNReal.coe_one, ge_iff_le, sup_le_iff] constructor · calc edist x.fst y.fst ≤ (edist x.fst y.fst ^ p.toReal) ^ (1 / p.toReal) := by @@ -427,9 +425,8 @@ theorem prod_antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricS · refine ENNReal.rpow_le_rpow ?_ (le_of_lt pos) simp [edist] _ = (2 : ℝ≥0) ^ (1 / p.toReal) * edist (WithLp.equiv p _ x) (WithLp.equiv p _ y) := by - simp only [← two_mul, ENNReal.mul_rpow_of_nonneg _ _ nonneg, - ← ENNReal.rpow_mul, cancel, ENNReal.rpow_one, ← ENNReal.coe_rpow_of_nonneg _ nonneg, - coe_ofNat] + simp only [← two_mul, ENNReal.mul_rpow_of_nonneg _ _ nonneg, ← ENNReal.rpow_mul, cancel, + ENNReal.rpow_one, ← ENNReal.coe_rpow_of_nonneg _ nonneg, coe_ofNat] theorem prod_aux_uniformity_eq [PseudoEMetricSpace α] [PseudoEMetricSpace β] : 𝓤 (WithLp p (α × β)) = 𝓤[instUniformSpaceProd] := by @@ -542,19 +539,18 @@ theorem infty_equiv_isometry [PseudoEMetricSpace α] [PseudoEMetricSpace β] : /-- Seminormed group instance on the product of two normed groups, using the `L^p` norm. -/ instance instProdSeminormedAddCommGroup [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] : - SeminormedAddCommGroup (WithLp p (α × β)) := - { WithLp.instAddCommGroup _ _ with - dist_eq := fun x y => by - rcases p.dichotomy with (rfl | h) - · simp only [prod_dist_eq_sup, prod_norm_eq_sup, dist_eq_norm] - rfl - · have : p ≠ ∞ := by - intro hp - rw [hp, ENNReal.top_toReal] at h - linarith - simp only [prod_dist_eq_add (zero_lt_one.trans_le h), - prod_norm_eq_add (zero_lt_one.trans_le h), dist_eq_norm] - rfl } + SeminormedAddCommGroup (WithLp p (α × β)) where + dist_eq := fun x y => by + rcases p.dichotomy with (rfl | h) + · simp only [prod_dist_eq_sup, prod_norm_eq_sup, dist_eq_norm] + rfl + · have : p ≠ ∞ := by + intro hp + rw [hp, ENNReal.top_toReal] at h + linarith + simp only [prod_dist_eq_add (zero_lt_one.trans_le h), + prod_norm_eq_add (zero_lt_one.trans_le h), dist_eq_norm] + rfl /-- normed group instance on the product of two normed groups, using the `L^p` norm. -/ instance instProdNormedAddCommGroup [NormedAddCommGroup α] [NormedAddCommGroup β] : @@ -583,13 +579,12 @@ theorem prod_norm_eq_of_nat (n : ℕ) (h : p = n) (f : WithLp p (α × β)) : prod_norm_eq_add this] theorem prod_norm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by - rw [prod_norm_eq_of_nat 2 (by norm_cast) _] -- Porting note: was `convert` + rw [prod_norm_eq_of_nat 2 (by norm_cast) _] rw [Real.sqrt_eq_rpow] norm_cast theorem prod_nnnorm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖₊ = NNReal.sqrt (‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2) := - -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast exact prod_norm_eq_of_L2 x @@ -610,7 +605,6 @@ theorem prod_dist_eq_of_L2 (x y : WithLp 2 (α × β)) : theorem prod_nndist_eq_of_L2 (x y : WithLp 2 (α × β)) : nndist x y = NNReal.sqrt (nndist x.fst y.fst ^ 2 + nndist x.snd y.snd ^ 2) := - -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast exact prod_dist_eq_of_L2 _ _ @@ -630,22 +624,21 @@ variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] /-- The product of two normed spaces is a normed space, with the `L^p` norm. -/ instance instProdNormedSpace : - NormedSpace 𝕜 (WithLp p (α × β)) := - { WithLp.instModule p 𝕜 (α × β) with - norm_smul_le := fun c f => by - rcases p.dichotomy with (rfl | hp) - · letI : Module 𝕜 (WithLp ∞ (α × β)) := Prod.instModule - suffices ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊ by exact_mod_cast NNReal.coe_mono this.le - simp only [prod_nnnorm_eq_sup, NNReal.mul_sup, ← nnnorm_smul] - rfl - · have : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le hp).ne' - have smul_fst : (c • f).fst = c • f.fst := rfl - have smul_snd : (c • f).snd = c • f.snd := rfl - simp only [prod_norm_eq_add (zero_lt_one.trans_le hp), norm_smul, Real.mul_rpow, - norm_nonneg, smul_fst, smul_snd] - rw [← mul_add, mul_rpow (rpow_nonneg_of_nonneg (norm_nonneg _) _), - ← rpow_mul (norm_nonneg _), this, Real.rpow_one] - positivity } + NormedSpace 𝕜 (WithLp p (α × β)) where + norm_smul_le := fun c f => by + rcases p.dichotomy with (rfl | hp) + · letI : Module 𝕜 (WithLp ∞ (α × β)) := Prod.instModule + suffices ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊ by exact_mod_cast NNReal.coe_mono this.le + simp only [prod_nnnorm_eq_sup, NNReal.mul_sup, ← nnnorm_smul] + rfl + · have : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le hp).ne' + have smul_fst : (c • f).fst = c • f.fst := rfl + have smul_snd : (c • f).snd = c • f.snd := rfl + simp only [prod_norm_eq_add (zero_lt_one.trans_le hp), norm_smul, Real.mul_rpow, + norm_nonneg, smul_fst, smul_snd] + rw [← mul_add, mul_rpow (rpow_nonneg_of_nonneg (norm_nonneg _) _), + ← rpow_mul (norm_nonneg _), this, Real.rpow_one] + positivity end InstNormedSpace From e49da985f21f541ca15e61b226513e82202357b7 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 16 Aug 2023 13:16:28 +1000 Subject: [PATCH 25/39] removed porting notes --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 1552c87848239..81b75700d7c90 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -506,14 +506,12 @@ variable {p α β} theorem prod_nndist_eq_add [PseudoMetricSpace α] [PseudoMetricSpace β] (hp : p ≠ ∞) (x y : WithLp p (α × β)) : nndist x y = (nndist x.fst y.fst ^ p.toReal + nndist x.snd y.snd ^ p.toReal) ^ (1 / p.toReal) := - -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast exact prod_dist_eq_add (p.toReal_pos_iff_ne_top.mpr hp) _ _ theorem prod_nndist_eq_sup [PseudoMetricSpace α] [PseudoMetricSpace β] (x y : WithLp ∞ (α × β)) : nndist x y = nndist x.fst y.fst ⊔ nndist x.snd y.snd := - -- Porting note: was `Subtype.ext` NNReal.eq <| by push_cast exact prod_dist_eq_sup _ _ From e0048b9a3948d4632e80c5af95ee089bef138852 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 16 Aug 2023 13:32:04 +1000 Subject: [PATCH 26/39] more corrections --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 72 +++++++++++++++--------- 1 file changed, 44 insertions(+), 28 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 81b75700d7c90..3db6ce9a2493b 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -42,7 +42,7 @@ open Real Set Filter IsROrC Bornology BigOperators Uniformity Topology NNReal EN noncomputable section -variable (p : ℝ≥0∞) (𝕜 𝕜' α β : Type*) +variable (p : ℝ≥0∞) (𝕜 α β : Type*) namespace WithLp @@ -302,7 +302,7 @@ coincide with the product one. Therefore, we do not register it as an instance. 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`. -/ -def prodPseudoEmetricAux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : +def prodPseudoEMetricAux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : PseudoEMetricSpace (WithLp p (α × β)) where edist_self := prod_edist_self p edist_comm := prod_edist_comm p @@ -329,7 +329,7 @@ def prodPseudoEmetricAux [PseudoEMetricSpace α] [PseudoEMetricSpace β] : Finset.sum_singleton] at this exact this -attribute [local instance] WithLp.prodPseudoEmetricAux +attribute [local instance] WithLp.prodPseudoEMetricAux variable {α β} @@ -450,30 +450,56 @@ end Aux /-! ### Instances on `L^p` products -/ +section TopologicalSpace -instance instUniformSpace [UniformSpace α] [UniformSpace β] : UniformSpace (WithLp p (α × β)) := - instUniformSpaceProd +variable [TopologicalSpace α] [TopologicalSpace β] -theorem uniformContinuous_equiv [UniformSpace α] [UniformSpace β] : - UniformContinuous (WithLp.equiv p (α × β)) := - uniformContinuous_id - -theorem uniformContinuous_equiv_symm [UniformSpace α] [UniformSpace β] : - UniformContinuous (WithLp.equiv p (α × β)).symm := - uniformContinuous_id +instance instTopologicalSpace : TopologicalSpace (WithLp p (α × β)) := + instTopologicalSpaceProd @[continuity] -theorem continuous_equiv [UniformSpace α] [UniformSpace β] : Continuous (WithLp.equiv p (α × β)) := +theorem continuous_equiv : Continuous (WithLp.equiv p (α × β)) := continuous_id @[continuity] -theorem continuous_equiv_symm [UniformSpace α] [UniformSpace β] : - Continuous (WithLp.equiv p (α × β)).symm := +theorem continuous_equiv_symm : Continuous (WithLp.equiv p (α × β)).symm := continuous_id +end TopologicalSpace + +section UniformSpace + +variable [UniformSpace α] [UniformSpace β] + +instance instUniformSpace : UniformSpace (WithLp p (α × β)) := + instUniformSpaceProd + +theorem uniformContinuous_equiv : UniformContinuous (WithLp.equiv p (α × β)) := + uniformContinuous_id + +theorem uniformContinuous_equiv_symm : UniformContinuous (WithLp.equiv p (α × β)).symm := + uniformContinuous_id + +end UniformSpace + instance instBornology [Bornology α] [Bornology β] : Bornology (WithLp p (α × β)) := Prod.instBornology +section ContinuousLinearEquiv + +variable [TopologicalSpace α] [TopologicalSpace β] +variable [Semiring 𝕜] [AddCommGroup α] [AddCommGroup β] +variable [Module 𝕜 α] [Module 𝕜 β] + +/-- `WithLp.equiv` as a continuous linear equivalence. -/ +@[simps! (config := { fullyApplied := false }) apply symm_apply] +protected def prodContinuousLinearEquiv : WithLp p (α × β) ≃L[𝕜] α × β where + toLinearEquiv := WithLp.linearEquiv _ _ _ + continuous_toFun := continuous_equiv _ _ _ + continuous_invFun := continuous_equiv_symm _ _ _ + +end ContinuousLinearEquiv + -- throughout the rest of the file, we assume `1 ≤ p` variable [Fact (1 ≤ p)] @@ -481,7 +507,7 @@ variable [Fact (1 ≤ p)] `L^p` pseudoedistance, and having as uniformity the product uniformity. -/ instance instProdPseudoEMetricSpace [PseudoEMetricSpace α] [PseudoEMetricSpace β] : PseudoEMetricSpace (WithLp p (α × β)) := - (prodPseudoEmetricAux p α β).replaceUniformity (prod_aux_uniformity_eq p α β).symm + (prodPseudoEMetricAux p α β).replaceUniformity (prod_aux_uniformity_eq p α β).symm /-- `EMetricSpace` instance on the product of two emetric spaces, using the `L^p` edistance, and having as uniformity the product uniformity. -/ @@ -613,7 +639,7 @@ theorem prod_edist_eq_of_L2 (x y : WithLp 2 (α × β)) : end norm_of -variable [NormedField 𝕜] [NormedField 𝕜'] +variable [NormedField 𝕜] section InstNormedSpace @@ -621,8 +647,7 @@ variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] [SeminormedAddCommGroup β] [NormedSpace 𝕜 β] /-- The product of two normed spaces is a normed space, with the `L^p` norm. -/ -instance instProdNormedSpace : - NormedSpace 𝕜 (WithLp p (α × β)) where +instance instProdNormedSpace : NormedSpace 𝕜 (WithLp p (α × β)) where norm_smul_le := fun c f => by rcases p.dichotomy with (rfl | hp) · letI : Module 𝕜 (WithLp ∞ (α × β)) := Prod.instModule @@ -732,13 +757,4 @@ theorem edist_equiv_symm_single_snd (y₁ y₂ : β) : end Single -variable (𝕜 p α β) - -/-- `WithLp.equiv` as a continuous linear equivalence. -/ -@[simps! (config := { fullyApplied := false }) apply symm_apply] -protected def prodContinuousLinearEquiv : WithLp p (α × β) ≃L[𝕜] α × β where - toLinearEquiv := WithLp.linearEquiv _ _ _ - continuous_toFun := continuous_equiv _ _ _ - continuous_invFun := continuous_equiv_symm _ _ _ - end WithLp From daabeab349f3a19181c150be330fc54d6997fcdb Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 16 Aug 2023 13:45:56 +1000 Subject: [PATCH 27/39] more corrections --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 19 +++++++++---------- 1 file changed, 9 insertions(+), 10 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 3db6ce9a2493b..d89c77e3f0c75 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -294,7 +294,7 @@ explaining why having definitionally the right uniformity is often important. -/ -variable [Fact (1 ≤ p)] +variable [hp : Fact (1 ≤ p)] /-- 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 @@ -501,7 +501,7 @@ protected def prodContinuousLinearEquiv : WithLp p (α × β) ≃L[𝕜] α × end ContinuousLinearEquiv -- throughout the rest of the file, we assume `1 ≤ p` -variable [Fact (1 ≤ p)] +variable [hp : Fact (1 ≤ p)] /-- `PseudoEMetricSpace` instance on the product of two pseudoemetric spaces, using the `L^p` pseudoedistance, and having as uniformity the product uniformity. -/ @@ -639,12 +639,11 @@ theorem prod_edist_eq_of_L2 (x y : WithLp 2 (α × β)) : end norm_of -variable [NormedField 𝕜] +variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] section InstNormedSpace -variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] - [SeminormedAddCommGroup β] [NormedSpace 𝕜 β] +variable [NormedField 𝕜] [NormedSpace 𝕜 α] [NormedSpace 𝕜 β] /-- The product of two normed spaces is a normed space, with the `L^p` norm. -/ instance instProdNormedSpace : NormedSpace 𝕜 (WithLp p (α × β)) where @@ -667,12 +666,12 @@ end InstNormedSpace /- Register simplification lemmas for the applications of `WithLp p (α × β)` elements, as the usual lemmas for `Prod` will not trigger. -/ -variable {𝕜 𝕜' p α β} -variable [SeminormedAddCommGroup α] [NormedSpace 𝕜 α] - [SeminormedAddCommGroup β] [NormedSpace 𝕜 β] +variable {𝕜 p α β} section Equiv +variable [NormedField 𝕜] [NormedSpace 𝕜 α] [NormedSpace 𝕜 β] + /-- The canonical map `WithLp.equiv` between `WithLp ∞ (α × β)` and `α × β` as a linear isometric equivalence. -/ def prodEquivₗᵢ : WithLp ∞ (α × β) ≃ₗᵢ[𝕜] α × β := @@ -688,7 +687,7 @@ section Single variable (p α β) @[simp] -theorem nnnorm_equiv_symm_fst [hp : Fact (1 ≤ p)] (x : α) : +theorem nnnorm_equiv_symm_fst (x : α) : ‖(WithLp.equiv p (α × β)).symm (x, 0)‖₊ = ‖x‖₊ := by induction p using ENNReal.recTopCoe generalizing hp with | top => @@ -699,7 +698,7 @@ theorem nnnorm_equiv_symm_fst [hp : Fact (1 ≤ p)] (x : α) : simp [prod_nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] @[simp] -theorem nnnorm_equiv_symm_snd [hp : Fact (1 ≤ p)] (y : β) : +theorem nnnorm_equiv_symm_snd (y : β) : ‖(WithLp.equiv p (α × β)).symm (0, y)‖₊ = ‖y‖₊ := by induction p using ENNReal.recTopCoe generalizing hp with | top => From 17d9d2b778cbde60f13ccaf37086d98d86353cff Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 16 Aug 2023 16:15:05 +1000 Subject: [PATCH 28/39] more instances and rearranging --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 91 ++++++++++++------------ 1 file changed, 47 insertions(+), 44 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index d89c77e3f0c75..b3dab5948eca1 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -48,6 +48,9 @@ namespace WithLp section algebra +/- Register simplification lemmas for the applications of `WithLp p (α × β)` elements, as the usual +lemmas for `Prod` will not trigger. -/ + variable {p 𝕜 α β} variable [Semiring 𝕜] [AddCommGroup α] [AddCommGroup β] @@ -465,6 +468,11 @@ theorem continuous_equiv : Continuous (WithLp.equiv p (α × β)) := theorem continuous_equiv_symm : Continuous (WithLp.equiv p (α × β)).symm := continuous_id +variable [T0Space α] [T0Space β] + +instance instT0Space : T0Space (WithLp p (α × β)) := + instT0SpaceProdInstTopologicalSpaceProd + end TopologicalSpace section UniformSpace @@ -480,6 +488,11 @@ theorem uniformContinuous_equiv : UniformContinuous (WithLp.equiv p (α × β)) theorem uniformContinuous_equiv_symm : UniformContinuous (WithLp.equiv p (α × β)).symm := uniformContinuous_id +variable [CompleteSpace α] [CompleteSpace β] + +instance instCompleteSpace : CompleteSpace (WithLp p (α × β)) := + CompleteSpace.prod + end UniformSpace instance instBornology [Bornology α] [Bornology β] : Bornology (WithLp p (α × β)) := @@ -512,7 +525,7 @@ instance instProdPseudoEMetricSpace [PseudoEMetricSpace α] [PseudoEMetricSpace /-- `EMetricSpace` instance on the product of two emetric spaces, using the `L^p` edistance, and having as uniformity the product uniformity. -/ instance instProdEMetricSpace [EMetricSpace α] [EMetricSpace β] : EMetricSpace (WithLp p (α × β)) := - @EMetricSpace.ofT0PseudoEMetricSpace (WithLp p (α × β)) _ instT0SpaceProdInstTopologicalSpaceProd + EMetricSpace.ofT0PseudoEMetricSpace (WithLp p (α × β)) /-- `PseudoMetricSpace` instance on the product of two pseudometric spaces, using the `L^p` distance, and having as uniformity the product uniformity. -/ @@ -641,51 +654,8 @@ end norm_of variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] -section InstNormedSpace - -variable [NormedField 𝕜] [NormedSpace 𝕜 α] [NormedSpace 𝕜 β] - -/-- The product of two normed spaces is a normed space, with the `L^p` norm. -/ -instance instProdNormedSpace : NormedSpace 𝕜 (WithLp p (α × β)) where - norm_smul_le := fun c f => by - rcases p.dichotomy with (rfl | hp) - · letI : Module 𝕜 (WithLp ∞ (α × β)) := Prod.instModule - suffices ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊ by exact_mod_cast NNReal.coe_mono this.le - simp only [prod_nnnorm_eq_sup, NNReal.mul_sup, ← nnnorm_smul] - rfl - · have : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le hp).ne' - have smul_fst : (c • f).fst = c • f.fst := rfl - have smul_snd : (c • f).snd = c • f.snd := rfl - simp only [prod_norm_eq_add (zero_lt_one.trans_le hp), norm_smul, Real.mul_rpow, - norm_nonneg, smul_fst, smul_snd] - rw [← mul_add, mul_rpow (rpow_nonneg_of_nonneg (norm_nonneg _) _), - ← rpow_mul (norm_nonneg _), this, Real.rpow_one] - positivity - -end InstNormedSpace - -/- Register simplification lemmas for the applications of `WithLp p (α × β)` elements, as the usual -lemmas for `Prod` will not trigger. -/ -variable {𝕜 p α β} - -section Equiv - -variable [NormedField 𝕜] [NormedSpace 𝕜 α] [NormedSpace 𝕜 β] - -/-- The canonical map `WithLp.equiv` between `WithLp ∞ (α × β)` and `α × β` as a linear isometric -equivalence. -/ -def prodEquivₗᵢ : WithLp ∞ (α × β) ≃ₗᵢ[𝕜] α × β := - { WithLp.equiv ∞ (α × β) with - map_add' := fun f g => rfl - map_smul' := fun c f => rfl - norm_map' := fun f => by simp [Norm.norm] } - -end Equiv - section Single -variable (p α β) - @[simp] theorem nnnorm_equiv_symm_fst (x : α) : ‖(WithLp.equiv p (α × β)).symm (x, 0)‖₊ = ‖x‖₊ := by @@ -756,4 +726,37 @@ theorem edist_equiv_symm_single_snd (y₁ y₂ : β) : end Single +section NormedSpace + +variable [NormedField 𝕜] [NormedSpace 𝕜 α] [NormedSpace 𝕜 β] + +/-- The product of two normed spaces is a normed space, with the `L^p` norm. -/ +instance instProdNormedSpace : NormedSpace 𝕜 (WithLp p (α × β)) where + norm_smul_le := fun c f => by + rcases p.dichotomy with (rfl | hp) + · letI : Module 𝕜 (WithLp ∞ (α × β)) := Prod.instModule + suffices ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊ by exact_mod_cast NNReal.coe_mono this.le + simp only [prod_nnnorm_eq_sup, NNReal.mul_sup, ← nnnorm_smul] + rfl + · have : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le hp).ne' + have smul_fst : (c • f).fst = c • f.fst := rfl + have smul_snd : (c • f).snd = c • f.snd := rfl + simp only [prod_norm_eq_add (zero_lt_one.trans_le hp), norm_smul, Real.mul_rpow, + norm_nonneg, smul_fst, smul_snd] + rw [← mul_add, mul_rpow (rpow_nonneg_of_nonneg (norm_nonneg _) _), + ← rpow_mul (norm_nonneg _), this, Real.rpow_one] + positivity + +variable {𝕜 p α β} + +/-- The canonical map `WithLp.equiv` between `WithLp ∞ (α × β)` and `α × β` as a linear isometric +equivalence. -/ +def prodEquivₗᵢ : WithLp ∞ (α × β) ≃ₗᵢ[𝕜] α × β := + { WithLp.equiv ∞ (α × β) with + map_add' := fun f g => rfl + map_smul' := fun c f => rfl + norm_map' := fun f => by simp [Norm.norm] } + +end NormedSpace + end WithLp From 8d85e4aa8eb5956253fe9e84c07f7808d0a5274a Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Thu, 17 Aug 2023 12:07:49 +1000 Subject: [PATCH 29/39] naming --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 38 +++++++++++++----------- 1 file changed, 21 insertions(+), 17 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index b3dab5948eca1..8e8e8a92755a6 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -152,9 +152,10 @@ instance instProdEDist : EDist (WithLp p (α × β)) where edist f g := if _hp : p = 0 then (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) + else if p = ∞ then + edist f.fst g.fst ⊔ edist f.snd g.snd else - if p = ∞ then edist f.fst g.fst ⊔ edist f.snd g.snd - else (edist f.fst g.fst ^ p.toReal + edist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) + (edist f.fst g.fst ^ p.toReal + edist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) variable {p α β} variable (x y : WithLp p (α × β)) (x' : α × β) @@ -218,9 +219,10 @@ instance instProdDist : Dist (WithLp p (α × β)) where dist f g := if _hp : p = 0 then (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) + else if p = ∞ then + dist f.fst g.fst ⊔ dist f.snd g.snd else - if p = ∞ then dist f.fst g.fst ⊔ dist f.snd g.snd - else (dist f.fst g.fst ^ p.toReal + dist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) + (dist f.fst g.fst ^ p.toReal + dist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) variable {p α β} @@ -256,8 +258,10 @@ instance instProdNorm : Norm (WithLp p (α × β)) where norm f := if _hp : p = 0 then (if f.fst = 0 then 0 else 1) + (if f.snd = 0 then 0 else 1) - else if p = ∞ then ‖f.fst‖ ⊔ ‖f.snd‖ - else (‖f.fst‖ ^ p.toReal + ‖f.snd‖ ^ p.toReal) ^ (1 / p.toReal) + else if p = ∞ then + ‖f.fst‖ ⊔ ‖f.snd‖ + else + (‖f.fst‖ ^ p.toReal + ‖f.snd‖ ^ p.toReal) ^ (1 / p.toReal) variable {p α β} @@ -457,20 +461,20 @@ section TopologicalSpace variable [TopologicalSpace α] [TopologicalSpace β] -instance instTopologicalSpace : TopologicalSpace (WithLp p (α × β)) := +instance instProdTopologicalSpace : TopologicalSpace (WithLp p (α × β)) := instTopologicalSpaceProd @[continuity] -theorem continuous_equiv : Continuous (WithLp.equiv p (α × β)) := +theorem prod_continuous_equiv : Continuous (WithLp.equiv p (α × β)) := continuous_id @[continuity] -theorem continuous_equiv_symm : Continuous (WithLp.equiv p (α × β)).symm := +theorem prod_continuous_equiv_symm : Continuous (WithLp.equiv p (α × β)).symm := continuous_id variable [T0Space α] [T0Space β] -instance instT0Space : T0Space (WithLp p (α × β)) := +instance instProdT0Space : T0Space (WithLp p (α × β)) := instT0SpaceProdInstTopologicalSpaceProd end TopologicalSpace @@ -479,23 +483,23 @@ section UniformSpace variable [UniformSpace α] [UniformSpace β] -instance instUniformSpace : UniformSpace (WithLp p (α × β)) := +instance instProdUniformSpace : UniformSpace (WithLp p (α × β)) := instUniformSpaceProd -theorem uniformContinuous_equiv : UniformContinuous (WithLp.equiv p (α × β)) := +theorem prod_uniformContinuous_equiv : UniformContinuous (WithLp.equiv p (α × β)) := uniformContinuous_id -theorem uniformContinuous_equiv_symm : UniformContinuous (WithLp.equiv p (α × β)).symm := +theorem prod_uniformContinuous_equiv_symm : UniformContinuous (WithLp.equiv p (α × β)).symm := uniformContinuous_id variable [CompleteSpace α] [CompleteSpace β] -instance instCompleteSpace : CompleteSpace (WithLp p (α × β)) := +instance instProdCompleteSpace : CompleteSpace (WithLp p (α × β)) := CompleteSpace.prod end UniformSpace -instance instBornology [Bornology α] [Bornology β] : Bornology (WithLp p (α × β)) := +instance instProdBornology [Bornology α] [Bornology β] : Bornology (WithLp p (α × β)) := Prod.instBornology section ContinuousLinearEquiv @@ -508,8 +512,8 @@ variable [Module 𝕜 α] [Module 𝕜 β] @[simps! (config := { fullyApplied := false }) apply symm_apply] protected def prodContinuousLinearEquiv : WithLp p (α × β) ≃L[𝕜] α × β where toLinearEquiv := WithLp.linearEquiv _ _ _ - continuous_toFun := continuous_equiv _ _ _ - continuous_invFun := continuous_equiv_symm _ _ _ + continuous_toFun := prod_continuous_equiv _ _ _ + continuous_invFun := prod_continuous_equiv_symm _ _ _ end ContinuousLinearEquiv From 1c28a303f8b0aedadd72481eb54407433f8f0995 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Thu, 17 Aug 2023 12:19:32 +1000 Subject: [PATCH 30/39] doc-string and examples --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 18 +++++++++++++++--- 1 file changed, 15 insertions(+), 3 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 8e8e8a92755a6..e07941776c394 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -517,7 +517,7 @@ protected def prodContinuousLinearEquiv : WithLp p (α × β) ≃L[𝕜] α × end ContinuousLinearEquiv --- throughout the rest of the file, we assume `1 ≤ p` +/-! Throughout the rest of the file, we assume `1 ≤ p` -/ variable [hp : Fact (1 ≤ p)] /-- `PseudoEMetricSpace` instance on the product of two pseudoemetric spaces, using the @@ -599,6 +599,19 @@ instance instProdNormedAddCommGroup [NormedAddCommGroup α] [NormedAddCommGroup { instProdSeminormedAddCommGroup p α β with eq_of_dist_eq_zero := eq_of_dist_eq_zero } +example [NormedAddCommGroup α] [NormedAddCommGroup β] : + (instProdNormedAddCommGroup p α β).toMetricSpace.toUniformSpace.toTopologicalSpace = + instProdTopologicalSpace p α β := + rfl + +example [NormedAddCommGroup α] [NormedAddCommGroup β] : + (instProdNormedAddCommGroup p α β).toMetricSpace.toUniformSpace = instProdUniformSpace p α β := + rfl + +example [NormedAddCommGroup α] [NormedAddCommGroup β] : + (instProdNormedAddCommGroup p α β).toMetricSpace.toBornology = instProdBornology p α β := + rfl + section norm_of variable {p α β} @@ -620,8 +633,7 @@ theorem prod_norm_eq_of_nat (n : ℕ) (h : p = n) (f : WithLp p (α × β)) : prod_norm_eq_add this] theorem prod_norm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by - rw [prod_norm_eq_of_nat 2 (by norm_cast) _] - rw [Real.sqrt_eq_rpow] + rw [prod_norm_eq_of_nat 2 (by norm_cast) _, Real.sqrt_eq_rpow] norm_cast theorem prod_nnnorm_eq_of_L2 (x : WithLp 2 (α × β)) : From 1d9a2698ed4ce0d10c24d4165bb9c43da0bc702c Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sat, 19 Aug 2023 12:40:29 +1000 Subject: [PATCH 31/39] corrections --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 19 +++++++++---------- 1 file changed, 9 insertions(+), 10 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index e07941776c394..73d8fb32cade4 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -160,11 +160,10 @@ instance instProdEDist : EDist (WithLp p (α × β)) where variable {p α β} variable (x y : WithLp p (α × β)) (x' : α × β) -open Classical in @[simp] -theorem prod_edist_eq_card (f g : WithLp 0 (α × β)) : - edist f g = (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := - if_pos rfl +theorem prod_edist_eq_card [DecidableEq α] [DecidableEq β] (f g : WithLp 0 (α × β)) : + edist f g = (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := by + convert if_pos rfl theorem prod_edist_eq_add (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : edist f g = (edist f.fst g.fst ^ p.toReal + edist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) := @@ -187,7 +186,8 @@ variable [PseudoEMetricSpace α] [PseudoEMetricSpace β] from `WithLp.instProdPseudoEMetricSpace` so it can be used also for `p < 1`. -/ theorem prod_edist_self (f : WithLp p (α × β)) : edist f f = 0 := by rcases p.trichotomy with (rfl | rfl | h) - · simp + · classical + simp · simp [prod_edist_eq_sup] · simp [prod_edist_eq_add h, ENNReal.zero_rpow_of_pos h, ENNReal.zero_rpow_of_pos (inv_pos.2 <| h)] @@ -226,10 +226,9 @@ instance instProdDist : Dist (WithLp p (α × β)) where variable {p α β} -open Classical in -theorem prod_dist_eq_card (f g : WithLp 0 (α × β)) : dist f g = - (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := - if_pos rfl +theorem prod_dist_eq_card [DecidableEq α] [DecidableEq β] (f g : WithLp 0 (α × β)) : dist f g = + (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := by + convert if_pos rfl theorem prod_dist_eq_add (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : dist f g = (dist f.fst g.fst ^ p.toReal + dist f.snd g.snd ^ p.toReal) ^ (1 / p.toReal) := @@ -653,7 +652,7 @@ variable {α β} theorem prod_dist_eq_of_L2 (x y : WithLp 2 (α × β)) : dist x y = (dist x.fst y.fst ^ 2 + dist x.snd y.snd ^ 2).sqrt := by - simp_rw [dist_eq_norm, prod_norm_eq_of_L2, Pi.sub_apply] + simp_rw [dist_eq_norm, prod_norm_eq_of_L2] rfl theorem prod_nndist_eq_of_L2 (x y : WithLp 2 (α × β)) : From a06efda2867fa11e647d03523d4b3576c2f3c653 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sat, 19 Aug 2023 20:41:43 +1000 Subject: [PATCH 32/39] Update Mathlib/Analysis/NormedSpace/ProdLp.lean Co-authored-by: Eric Wieser --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 73d8fb32cade4..b98931a2d9cec 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -264,11 +264,10 @@ instance instProdNorm : Norm (WithLp p (α × β)) where variable {p α β} -open Classical in @[simp] -theorem prod_norm_eq_card (f : WithLp 0 (α × β)) : - ‖f‖ = (if f.fst = 0 then 0 else 1) + (if f.snd = 0 then 0 else 1) := - if_pos rfl +theorem prod_norm_eq_card (f : WithLp 0 (α × β)) [DecidableEq α] [DecidableEq β] : + ‖f‖ = (if f.fst = 0 then 0 else 1) + (if f.snd = 0 then 0 else 1) := by + convert if_pos rfl theorem prod_norm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖ = ‖f.fst‖ ⊔ ‖f.snd‖ := by dsimp [Norm.norm] From 874ef83dd61a2f4fd7888dc6aa130e12712c2d4f Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sat, 2 Sep 2023 12:00:20 +0800 Subject: [PATCH 33/39] corrections --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 21 ++++++++++----------- 1 file changed, 10 insertions(+), 11 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index b98931a2d9cec..07608cba0d012 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -36,7 +36,7 @@ This files is a straight-forward adaption of `Mathlib.Analysis.NormedSpace.PiLp` -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue #2220 +local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 open Real Set Filter IsROrC Bornology BigOperators Uniformity Topology NNReal ENNReal @@ -192,10 +192,10 @@ theorem prod_edist_self (f : WithLp p (α × β)) : edist f f = 0 := by · simp [prod_edist_eq_add h, ENNReal.zero_rpow_of_pos h, ENNReal.zero_rpow_of_pos (inv_pos.2 <| h)] -open Classical in /-- 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 prod_edist_comm (f g : WithLp p (α × β)) : edist f g = edist g f := by + classical rcases p.trichotomy with (rfl | rfl | h) · simp [prod_edist_eq_card, eq_comm] · simp only [prod_edist_eq_sup, edist_comm] @@ -579,7 +579,7 @@ theorem infty_equiv_isometry [PseudoEMetricSpace α] [PseudoEMetricSpace β] : norm. -/ instance instProdSeminormedAddCommGroup [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] : SeminormedAddCommGroup (WithLp p (α × β)) where - dist_eq := fun x y => by + dist_eq x y := by rcases p.dichotomy with (rfl | h) · simp only [prod_dist_eq_sup, prod_norm_eq_sup, dist_eq_norm] rfl @@ -746,10 +746,9 @@ variable [NormedField 𝕜] [NormedSpace 𝕜 α] [NormedSpace 𝕜 β] /-- The product of two normed spaces is a normed space, with the `L^p` norm. -/ instance instProdNormedSpace : NormedSpace 𝕜 (WithLp p (α × β)) where - norm_smul_le := fun c f => by + norm_smul_le c f := by rcases p.dichotomy with (rfl | hp) - · letI : Module 𝕜 (WithLp ∞ (α × β)) := Prod.instModule - suffices ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊ by exact_mod_cast NNReal.coe_mono this.le + · suffices ‖c • f‖₊ = ‖c‖₊ * ‖f‖₊ by exact_mod_cast NNReal.coe_mono this.le simp only [prod_nnnorm_eq_sup, NNReal.mul_sup, ← nnnorm_smul] rfl · have : p.toReal * (1 / p.toReal) = 1 := mul_div_cancel' 1 (zero_lt_one.trans_le hp).ne' @@ -765,11 +764,11 @@ variable {𝕜 p α β} /-- The canonical map `WithLp.equiv` between `WithLp ∞ (α × β)` and `α × β` as a linear isometric equivalence. -/ -def prodEquivₗᵢ : WithLp ∞ (α × β) ≃ₗᵢ[𝕜] α × β := - { WithLp.equiv ∞ (α × β) with - map_add' := fun f g => rfl - map_smul' := fun c f => rfl - norm_map' := fun f => by simp [Norm.norm] } +def prodEquivₗᵢ : WithLp ∞ (α × β) ≃ₗᵢ[𝕜] α × β where + __ := WithLp.equiv ∞ (α × β) + map_add' f g := rfl + map_smul' c f := rfl + norm_map' f := by simp [Norm.norm] end NormedSpace From 14369773d677a2c65676b09e957c7809cc405beb Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Thu, 7 Sep 2023 16:11:01 +0800 Subject: [PATCH 34/39] Sebastien's suggestions --- .../Analysis/InnerProductSpace/ProdL2.lean | 43 ++++++++++- Mathlib/Analysis/NormedSpace/ProdLp.lean | 74 ++++++------------- 2 files changed, 64 insertions(+), 53 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean index 00c5a5c967b24..9c58f70c67dd1 100644 --- a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean @@ -14,11 +14,52 @@ The `L²` norm on product of two inner product spaces is compatible with an inne $$ \langle x, y\rangle = \langle x_1, y_1 \rangle + \langle x_2, y_2 \rangle. $$ -This is recorded in this file as an inner product space instance on `ProdLp 2`. +This is recorded in this file as an inner product space instance on `WithLp 2 (E × F)`. -/ +local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 + +open Real ENNReal NNReal + namespace WithLp +section norm_of + +variable {α β : Type*} +variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] + +theorem prod_norm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by + rw [prod_norm_eq_of_nat 2 (by norm_cast) _, Real.sqrt_eq_rpow] + norm_cast + +theorem prod_nnnorm_eq_of_L2 (x : WithLp 2 (α × β)) : + ‖x‖₊ = NNReal.sqrt (‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2) := + NNReal.eq <| by + push_cast + exact prod_norm_eq_of_L2 x + +theorem prod_norm_sq_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ ^ 2 = ‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2 := by + suffices ‖x‖₊ ^ 2 = ‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2 by + simpa only [NNReal.coe_sum] using congr_arg ((↑) : ℝ≥0 → ℝ) this + rw [prod_nnnorm_eq_of_L2, NNReal.sq_sqrt] + +theorem prod_dist_eq_of_L2 (x y : WithLp 2 (α × β)) : + dist x y = (dist x.fst y.fst ^ 2 + dist x.snd y.snd ^ 2).sqrt := by + simp_rw [dist_eq_norm, prod_norm_eq_of_L2] + rfl + +theorem prod_nndist_eq_of_L2 (x y : WithLp 2 (α × β)) : + nndist x y = NNReal.sqrt (nndist x.fst y.fst ^ 2 + nndist x.snd y.snd ^ 2) := + NNReal.eq <| by + push_cast + exact prod_dist_eq_of_L2 _ _ + +theorem prod_edist_eq_of_L2 (x y : WithLp 2 (α × β)) : + edist x y = (edist x.fst y.fst ^ 2 + edist x.snd y.snd ^ 2) ^ (1 / 2 : ℝ) := by + simp [prod_edist_eq_add] + +end norm_of + variable {𝕜 : Type*} (E F : Type*) variable [IsROrC 𝕜] [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 07608cba0d012..48e4f44ec07b0 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -128,7 +128,7 @@ end equiv section DistNorm /-! -### Definition of `edist`, `dist` and `norm` on `ProdLp` +### 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 @@ -142,7 +142,7 @@ variable [EDist α] [EDist β] open Classical in /-- Endowing the space `WithLp p (α × β)` with the `L^p` edistance. We register this instance -separate from `ProdLp.instPseudoEMetric` since the latter requires the type class hypothesis +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 @@ -182,7 +182,9 @@ section EDistProp variable {α β} variable [PseudoEMetricSpace α] [PseudoEMetricSpace β] -/-- This holds independent of `p` and does not require `[Fact (1 ≤ p)]`. We keep it separate +/-- 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 prod_edist_self (f : WithLp p (α × β)) : edist f f = 0 := by rcases p.trichotomy with (rfl | rfl | h) @@ -192,7 +194,9 @@ theorem prod_edist_self (f : WithLp p (α × β)) : edist f f = 0 := by · simp [prod_edist_eq_add h, ENNReal.zero_rpow_of_pos h, ENNReal.zero_rpow_of_pos (inv_pos.2 <| h)] -/-- This holds independent of `p` and does not require `[Fact (1 ≤ p)]`. We keep it separate +/-- 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`. -/ theorem prod_edist_comm (f g : WithLp p (α × β)) : edist f g = edist g f := by classical @@ -244,7 +248,7 @@ end Dist section Norm -variable [Norm α] [Zero α] [Norm β] [Zero β] +variable [Norm α] [Norm β] open Classical in /-- Endowing the space `WithLp p (α × β)` with the `L^p` norm. We register this instance @@ -256,7 +260,7 @@ Registering this separately allows for a future norm-like structure on `WithLp p instance instProdNorm : Norm (WithLp p (α × β)) where norm f := if _hp : p = 0 then - (if f.fst = 0 then 0 else 1) + (if f.snd = 0 then 0 else 1) + (if ‖f.fst‖ = 0 then 0 else 1) + (if ‖f.snd‖ = 0 then 0 else 1) else if p = ∞ then ‖f.fst‖ ⊔ ‖f.snd‖ else @@ -266,7 +270,7 @@ variable {p α β} @[simp] theorem prod_norm_eq_card (f : WithLp 0 (α × β)) [DecidableEq α] [DecidableEq β] : - ‖f‖ = (if f.fst = 0 then 0 else 1) + (if f.snd = 0 then 0 else 1) := by + ‖f‖ = (if ‖f.fst‖ = 0 then 0 else 1) + (if ‖f.snd‖ = 0 then 0 else 1) := by convert if_pos rfl theorem prod_norm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖ = ‖f.fst‖ ⊔ ‖f.snd‖ := by @@ -567,7 +571,7 @@ theorem prod_antilipschitzWith_equiv [PseudoEMetricSpace α] [PseudoEMetricSpace AntilipschitzWith ((2 : ℝ≥0) ^ (1 / p).toReal) (WithLp.equiv p (α × β)) := prod_antilipschitzWith_equiv_aux p α β -theorem infty_equiv_isometry [PseudoEMetricSpace α] [PseudoEMetricSpace β] : +theorem prod_infty_equiv_isometry [PseudoEMetricSpace α] [PseudoEMetricSpace β] : Isometry (WithLp.equiv ∞ (α × β)) := fun x y => le_antisymm (by simpa only [ENNReal.coe_one, one_mul] using prod_lipschitzWith_equiv ∞ α β x y) @@ -630,40 +634,6 @@ theorem prod_norm_eq_of_nat (n : ℕ) (h : p = n) (f : WithLp p (α × β)) : simp only [one_div, h, Real.rpow_nat_cast, ENNReal.toReal_nat, eq_self_iff_true, Finset.sum_congr, prod_norm_eq_add this] -theorem prod_norm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by - rw [prod_norm_eq_of_nat 2 (by norm_cast) _, Real.sqrt_eq_rpow] - norm_cast - -theorem prod_nnnorm_eq_of_L2 (x : WithLp 2 (α × β)) : - ‖x‖₊ = NNReal.sqrt (‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2) := - NNReal.eq <| by - push_cast - exact prod_norm_eq_of_L2 x - -variable (α β) - -theorem prod_norm_sq_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ ^ 2 = ‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2 := by - suffices ‖x‖₊ ^ 2 = ‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2 by - simpa only [NNReal.coe_sum] using congr_arg ((↑) : ℝ≥0 → ℝ) this - rw [prod_nnnorm_eq_of_L2, NNReal.sq_sqrt] - -variable {α β} - -theorem prod_dist_eq_of_L2 (x y : WithLp 2 (α × β)) : - dist x y = (dist x.fst y.fst ^ 2 + dist x.snd y.snd ^ 2).sqrt := by - simp_rw [dist_eq_norm, prod_norm_eq_of_L2] - rfl - -theorem prod_nndist_eq_of_L2 (x y : WithLp 2 (α × β)) : - nndist x y = NNReal.sqrt (nndist x.fst y.fst ^ 2 + nndist x.snd y.snd ^ 2) := - NNReal.eq <| by - push_cast - exact prod_dist_eq_of_L2 _ _ - -theorem prod_edist_eq_of_L2 (x y : WithLp 2 (α × β)) : - edist x y = (edist x.fst y.fst ^ 2 + edist x.snd y.snd ^ 2) ^ (1 / 2 : ℝ) := by - simp [prod_edist_eq_add] - end norm_of variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] @@ -701,42 +671,42 @@ theorem norm_equiv_symm_snd (y : β) : ‖(WithLp.equiv p (α × β)).symm (0, y congr_arg ((↑) : ℝ≥0 → ℝ) <| nnnorm_equiv_symm_snd p α β y @[simp] -theorem nndist_equiv_symm_single_fst (x₁ x₂ : α) : +theorem nndist_equiv_symm_fst (x₁ x₂ : α) : nndist ((WithLp.equiv p (α × β)).symm (x₁, 0)) ((WithLp.equiv p (α × β)).symm (x₂, 0)) = nndist x₁ x₂ := by rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← WithLp.equiv_symm_sub, Prod.mk_sub_mk, sub_zero, nnnorm_equiv_symm_fst] @[simp] -theorem nndist_equiv_symm_single_snd (y₁ y₂ : β) : +theorem nndist_equiv_symm_snd (y₁ y₂ : β) : nndist ((WithLp.equiv p (α × β)).symm (0, y₁)) ((WithLp.equiv p (α × β)).symm (0, y₂)) = nndist y₁ y₂ := by rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← WithLp.equiv_symm_sub, Prod.mk_sub_mk, sub_zero, nnnorm_equiv_symm_snd] @[simp] -theorem dist_equiv_symm_single_fst (x₁ x₂ : α) : +theorem dist_equiv_symm_fst (x₁ x₂ : α) : dist ((WithLp.equiv p (α × β)).symm (x₁, 0)) ((WithLp.equiv p (α × β)).symm (x₂, 0)) = dist x₁ x₂ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_fst p α β x₁ x₂ + congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_fst p α β x₁ x₂ @[simp] -theorem dist_equiv_symm_single_snd (y₁ y₂ : β) : +theorem dist_equiv_symm_snd (y₁ y₂ : β) : dist ((WithLp.equiv p (α × β)).symm (0, y₁)) ((WithLp.equiv p (α × β)).symm (0, y₂)) = dist y₁ y₂ := - congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_single_snd p α β y₁ y₂ + congr_arg ((↑) : ℝ≥0 → ℝ) <| nndist_equiv_symm_snd p α β y₁ y₂ @[simp] -theorem edist_equiv_symm_single_fst (x₁ x₂ : α) : +theorem edist_equiv_symm_fst (x₁ x₂ : α) : edist ((WithLp.equiv p (α × β)).symm (x₁, 0)) ((WithLp.equiv p (α × β)).symm (x₂, 0)) = edist x₁ x₂ := by - simp only [edist_nndist, nndist_equiv_symm_single_fst p α β x₁ x₂] + simp only [edist_nndist, nndist_equiv_symm_fst p α β x₁ x₂] @[simp] -theorem edist_equiv_symm_single_snd (y₁ y₂ : β) : +theorem edist_equiv_symm_snd (y₁ y₂ : β) : edist ((WithLp.equiv p (α × β)).symm (0, y₁)) ((WithLp.equiv p (α × β)).symm (0, y₂)) = edist y₁ y₂ := by - simp only [edist_nndist, nndist_equiv_symm_single_snd p α β y₁ y₂] + simp only [edist_nndist, nndist_equiv_symm_snd p α β y₁ y₂] end Single From 8abe9b0d12226777d23ab282d39ed925760b2451 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Thu, 7 Sep 2023 16:23:10 +0800 Subject: [PATCH 35/39] one easy typeclass golf --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 13 +++++++------ 1 file changed, 7 insertions(+), 6 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 48e4f44ec07b0..5d7a6612d13c5 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -617,6 +617,13 @@ example [NormedAddCommGroup α] [NormedAddCommGroup β] : section norm_of variable {p α β} + +theorem prod_norm_eq_of_nat [Norm α] [Norm β] (n : ℕ) (h : p = n) (f : WithLp p (α × β)) : + ‖f‖ = (‖f.fst‖ ^ n + ‖f.snd‖ ^ n) ^ (1 / (n : ℝ)) := by + have := p.toReal_pos_iff_ne_top.mpr (ne_of_eq_of_ne h <| ENNReal.nat_ne_top n) + simp only [one_div, h, Real.rpow_nat_cast, ENNReal.toReal_nat, eq_self_iff_true, Finset.sum_congr, + prod_norm_eq_add this] + variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] theorem prod_nnnorm_eq_add (hp : p ≠ ∞) (f : WithLp p (α × β)) : @@ -628,12 +635,6 @@ theorem prod_nnnorm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖₊ = ‖f.fst ext norm_cast -theorem prod_norm_eq_of_nat (n : ℕ) (h : p = n) (f : WithLp p (α × β)) : - ‖f‖ = (‖f.fst‖ ^ n + ‖f.snd‖ ^ n) ^ (1 / (n : ℝ)) := by - have := p.toReal_pos_iff_ne_top.mpr (ne_of_eq_of_ne h <| ENNReal.nat_ne_top n) - simp only [one_div, h, Real.rpow_nat_cast, ENNReal.toReal_nat, eq_self_iff_true, Finset.sum_congr, - prod_norm_eq_add this] - end norm_of variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] From c966caafc882233e3011e82f1a79380164c3e665 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 27 Sep 2023 22:30:25 +1000 Subject: [PATCH 36/39] move norm_of L2 section back to ProdLp --- .../Analysis/InnerProductSpace/ProdL2.lean | 41 ------------------- Mathlib/Analysis/NormedSpace/ProdLp.lean | 40 ++++++++++++++---- 2 files changed, 33 insertions(+), 48 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean index 9c58f70c67dd1..5d516f0c6577d 100644 --- a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean @@ -17,49 +17,8 @@ $$ This is recorded in this file as an inner product space instance on `WithLp 2 (E × F)`. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - -open Real ENNReal NNReal - namespace WithLp -section norm_of - -variable {α β : Type*} -variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] - -theorem prod_norm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by - rw [prod_norm_eq_of_nat 2 (by norm_cast) _, Real.sqrt_eq_rpow] - norm_cast - -theorem prod_nnnorm_eq_of_L2 (x : WithLp 2 (α × β)) : - ‖x‖₊ = NNReal.sqrt (‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2) := - NNReal.eq <| by - push_cast - exact prod_norm_eq_of_L2 x - -theorem prod_norm_sq_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ ^ 2 = ‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2 := by - suffices ‖x‖₊ ^ 2 = ‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2 by - simpa only [NNReal.coe_sum] using congr_arg ((↑) : ℝ≥0 → ℝ) this - rw [prod_nnnorm_eq_of_L2, NNReal.sq_sqrt] - -theorem prod_dist_eq_of_L2 (x y : WithLp 2 (α × β)) : - dist x y = (dist x.fst y.fst ^ 2 + dist x.snd y.snd ^ 2).sqrt := by - simp_rw [dist_eq_norm, prod_norm_eq_of_L2] - rfl - -theorem prod_nndist_eq_of_L2 (x y : WithLp 2 (α × β)) : - nndist x y = NNReal.sqrt (nndist x.fst y.fst ^ 2 + nndist x.snd y.snd ^ 2) := - NNReal.eq <| by - push_cast - exact prod_dist_eq_of_L2 _ _ - -theorem prod_edist_eq_of_L2 (x y : WithLp 2 (α × β)) : - edist x y = (edist x.fst y.fst ^ 2 + edist x.snd y.snd ^ 2) ^ (1 / 2 : ℝ) := by - simp [prod_edist_eq_add] - -end norm_of - variable {𝕜 : Type*} (E F : Type*) variable [IsROrC 𝕜] [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] [NormedAddCommGroup F] [InnerProductSpace 𝕜 F] diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 5d7a6612d13c5..d4a5a44ef4c4c 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -293,7 +293,7 @@ section Aux 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 Pi type (with the `L^∞` distance) is a uniform embedding, as it is both Lipschitz and +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 @@ -477,7 +477,7 @@ theorem prod_continuous_equiv_symm : Continuous (WithLp.equiv p (α × β)).symm variable [T0Space α] [T0Space β] instance instProdT0Space : T0Space (WithLp p (α × β)) := - instT0SpaceProdInstTopologicalSpaceProd + Prod.instT0Space end TopologicalSpace @@ -587,11 +587,7 @@ instance instProdSeminormedAddCommGroup [SeminormedAddCommGroup α] [SeminormedA rcases p.dichotomy with (rfl | h) · simp only [prod_dist_eq_sup, prod_norm_eq_sup, dist_eq_norm] rfl - · have : p ≠ ∞ := by - intro hp - rw [hp, ENNReal.top_toReal] at h - linarith - simp only [prod_dist_eq_add (zero_lt_one.trans_le h), + · simp only [prod_dist_eq_add (zero_lt_one.trans_le h), prod_norm_eq_add (zero_lt_one.trans_le h), dist_eq_norm] rfl @@ -635,6 +631,36 @@ theorem prod_nnnorm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖₊ = ‖f.fst ext norm_cast +theorem prod_norm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by + rw [prod_norm_eq_of_nat 2 (by norm_cast) _, Real.sqrt_eq_rpow] + norm_cast + +theorem prod_nnnorm_eq_of_L2 (x : WithLp 2 (α × β)) : + ‖x‖₊ = NNReal.sqrt (‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2) := + NNReal.eq <| by + push_cast + exact prod_norm_eq_of_L2 x + +theorem prod_norm_sq_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ ^ 2 = ‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2 := by + suffices ‖x‖₊ ^ 2 = ‖x.fst‖₊ ^ 2 + ‖x.snd‖₊ ^ 2 by + simpa only [NNReal.coe_sum] using congr_arg ((↑) : ℝ≥0 → ℝ) this + rw [prod_nnnorm_eq_of_L2, NNReal.sq_sqrt] + +theorem prod_dist_eq_of_L2 (x y : WithLp 2 (α × β)) : + dist x y = (dist x.fst y.fst ^ 2 + dist x.snd y.snd ^ 2).sqrt := by + simp_rw [dist_eq_norm, prod_norm_eq_of_L2] + rfl + +theorem prod_nndist_eq_of_L2 (x y : WithLp 2 (α × β)) : + nndist x y = NNReal.sqrt (nndist x.fst y.fst ^ 2 + nndist x.snd y.snd ^ 2) := + NNReal.eq <| by + push_cast + exact prod_dist_eq_of_L2 _ _ + +theorem prod_edist_eq_of_L2 (x y : WithLp 2 (α × β)) : + edist x y = (edist x.fst y.fst ^ 2 + edist x.snd y.snd ^ 2) ^ (1 / 2 : ℝ) := by + simp [prod_edist_eq_add] + end norm_of variable [SeminormedAddCommGroup α] [SeminormedAddCommGroup β] From dadffb950744c4ec3abe5f3ab4ea302b9fe88ad9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 27 Sep 2023 21:56:41 +0000 Subject: [PATCH 37/39] fix lint --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index d4a5a44ef4c4c..a33740eb37be8 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -269,7 +269,7 @@ instance instProdNorm : Norm (WithLp p (α × β)) where variable {p α β} @[simp] -theorem prod_norm_eq_card (f : WithLp 0 (α × β)) [DecidableEq α] [DecidableEq β] : +theorem prod_norm_eq_card (f : WithLp 0 (α × β)) : ‖f‖ = (if ‖f.fst‖ = 0 then 0 else 1) + (if ‖f.snd‖ = 0 then 0 else 1) := by convert if_pos rfl From 48c7e02e9a9c71e26b2185e4ab5af5fe45a01b1e Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Fri, 29 Sep 2023 18:20:21 +1000 Subject: [PATCH 38/39] coherence --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index d4a5a44ef4c4c..6a84d21a44df4 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -151,7 +151,7 @@ literature, but it is sometimes called a *quasi-metric* or *semi-metric*. -/ instance instProdEDist : EDist (WithLp p (α × β)) where edist f g := if _hp : p = 0 then - (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) + (if edist f.fst g.fst = 0 then 0 else 1) + (if edist f.snd g.snd = 0 then 0 else 1) else if p = ∞ then edist f.fst g.fst ⊔ edist f.snd g.snd else @@ -162,7 +162,8 @@ variable (x y : WithLp p (α × β)) (x' : α × β) @[simp] theorem prod_edist_eq_card [DecidableEq α] [DecidableEq β] (f g : WithLp 0 (α × β)) : - edist f g = (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := by + edist f g = + (if edist f.fst g.fst = 0 then 0 else 1) + (if edist f.snd g.snd = 0 then 0 else 1) := by convert if_pos rfl theorem prod_edist_eq_add (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : @@ -201,7 +202,7 @@ from `WithLp.instProdPseudoEMetricSpace` so it can be used also for `p < 1`. -/ theorem prod_edist_comm (f g : WithLp p (α × β)) : edist f g = edist g f := by classical rcases p.trichotomy with (rfl | rfl | h) - · simp [prod_edist_eq_card, eq_comm] + · simp only [prod_edist_eq_card, edist_comm] · simp only [prod_edist_eq_sup, edist_comm] · simp only [prod_edist_eq_add h, edist_comm] @@ -222,7 +223,7 @@ literature, but it is sometimes called a *quasi-metric* or *semi-metric*. -/ instance instProdDist : Dist (WithLp p (α × β)) where dist f g := if _hp : p = 0 then - (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) + (if dist f.fst g.fst = 0 then 0 else 1) + (if dist f.snd g.snd = 0 then 0 else 1) else if p = ∞ then dist f.fst g.fst ⊔ dist f.snd g.snd else @@ -231,7 +232,7 @@ instance instProdDist : Dist (WithLp p (α × β)) where variable {p α β} theorem prod_dist_eq_card [DecidableEq α] [DecidableEq β] (f g : WithLp 0 (α × β)) : dist f g = - (if f.fst = g.fst then 0 else 1) + (if f.snd = g.snd then 0 else 1) := by + (if dist f.fst g.fst = 0 then 0 else 1) + (if dist f.snd g.snd = 0 then 0 else 1) := by convert if_pos rfl theorem prod_dist_eq_add (hp : 0 < p.toReal) (f g : WithLp p (α × β)) : From 78ee3b085eda3ea7c956bdf3fd935260212210e7 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Fri, 29 Sep 2023 23:40:27 +1000 Subject: [PATCH 39/39] lint --- Mathlib/Analysis/NormedSpace/ProdLp.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index 43038c22524f2..ab393528aa5c8 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -161,7 +161,7 @@ variable {p α β} variable (x y : WithLp p (α × β)) (x' : α × β) @[simp] -theorem prod_edist_eq_card [DecidableEq α] [DecidableEq β] (f g : WithLp 0 (α × β)) : +theorem prod_edist_eq_card (f g : WithLp 0 (α × β)) : edist f g = (if edist f.fst g.fst = 0 then 0 else 1) + (if edist f.snd g.snd = 0 then 0 else 1) := by convert if_pos rfl @@ -231,7 +231,7 @@ instance instProdDist : Dist (WithLp p (α × β)) where variable {p α β} -theorem prod_dist_eq_card [DecidableEq α] [DecidableEq β] (f g : WithLp 0 (α × β)) : dist f g = +theorem prod_dist_eq_card (f g : WithLp 0 (α × β)) : dist f g = (if dist f.fst g.fst = 0 then 0 else 1) + (if dist f.snd g.snd = 0 then 0 else 1) := by convert if_pos rfl