diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index a88341d32c461..ca1dd6790cb0d 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -5,6 +5,7 @@ Authors: Heather Macbeth -/ import analysis.inner_product_space.projection import analysis.normed_space.lp_space +import analysis.inner_product_space.pi_L2 /-! # Hilbert sum of a family of inner product spaces @@ -403,6 +404,22 @@ protected lemma tsum_inner_mul_inner (b : hilbert_basis Ī¹ š•œ E) (x y : E) : āˆ‘' i, āŸŖx, b iāŸ« * āŸŖb i, yāŸ« = āŸŖx, yāŸ« := (b.has_sum_inner_mul_inner x y).tsum_eq +-- Note : this should be `b.repr` composed with an identification of `lp (Ī» i : Ī¹, š•œ) p` with +-- `pi_Lp p (Ī» i : Ī¹, š•œ)` (in this case with `p = 2`), but we don't have this yet (July 2022). +/-- A finite Hilbert basis is an orthonormal basis. -/ +protected def to_orthonormal_basis [fintype Ī¹] (b : hilbert_basis Ī¹ š•œ E) : + orthonormal_basis Ī¹ š•œ E := +orthonormal_basis.mk b.orthonormal +begin + have := (span š•œ (finset.univ.image b : set E)).closed_of_finite_dimensional, + simpa only [finset.coe_image, finset.coe_univ, set.image_univ, hilbert_basis.dense_span] using + this.submodule_topological_closure_eq.symm +end + +@[simp] lemma coe_to_orthonormal_basis [fintype Ī¹] (b : hilbert_basis Ī¹ š•œ E) : + (b.to_orthonormal_basis : Ī¹ ā†’ E) = b := +orthonormal_basis.coe_mk _ _ + variables {v : Ī¹ ā†’ E} (hv : orthonormal š•œ v) include hv cplt @@ -437,6 +454,19 @@ hilbert_basis.coe_mk hv _ omit hv +-- Note : this should be `b.repr` composed with an identification of `lp (Ī» i : Ī¹, š•œ) p` with +-- `pi_Lp p (Ī» i : Ī¹, š•œ)` (in this case with `p = 2`), but we don't have this yet (July 2022). +/-- An orthonormal basis is an Hilbert basis. -/ +protected def _root_.orthonormal_basis.to_hilbert_basis [fintype Ī¹] (b : orthonormal_basis Ī¹ š•œ E) : + hilbert_basis Ī¹ š•œ E := +hilbert_basis.mk b.orthonormal $ +by simpa only [ā† orthonormal_basis.coe_to_basis, b.to_basis.span_eq, eq_top_iff] + using @subset_closure E _ _ + +@[simp] lemma _root_.orthonormal_basis.coe_to_hilbert_basis [fintype Ī¹] + (b : orthonormal_basis Ī¹ š•œ E) : (b.to_hilbert_basis : Ī¹ ā†’ E) = b := +hilbert_basis.coe_mk _ _ + /-- A Hilbert space admits a Hilbert basis extending a given orthonormal subset. -/ lemma _root_.orthonormal.exists_hilbert_basis_extension {s : set E} (hs : orthonormal š•œ (coe : s ā†’ E)) :