Hilbert sum of a family of inner product spaces #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
Given a family (G : ι → Type*) [Π i, inner_product_space 𝕜 (G i)] of inner product spaces, this
file equips lp G 2 with an inner product space structure, where lp G 2 consists of those
dependent functions f : Π i, G i for which ∑' i, ‖f i‖ ^ 2, the sum of the norms-squared, is
summable. This construction is sometimes called the Hilbert sum of the family G. By choosing
G to be ι → 𝕜, the Hilbert space ℓ²(ι, 𝕜) may be seen as a special case of this construction.
We also define a predicate is_hilbert_sum 𝕜 G V, where V : Π i, G i →ₗᵢ[𝕜] E, expressing that
V is an orthogonal_family and that the associated map lp G 2 →ₗᵢ[𝕜] E is surjective.
Main definitions #
-
orthogonal_family.linear_isometry: Given a Hilbert spaceE, a familyGof inner product spaces and a familyV : Π i, G i →ₗᵢ[𝕜] Eof isometric embeddings of theG iintoEwith mutually-orthogonal images, there is an induced isometric embedding of the Hilbert sum ofGintoE. -
is_hilbert_sum: Given a Hilbert spaceE, a familyGof inner product spaces and a familyV : Π i, G i →ₗᵢ[𝕜] Eof isometric embeddings of theG iintoE,is_hilbert_sum 𝕜 G Vmeans thatVis anorthogonal_familyand that the above linear isometry is surjective. -
is_hilbert_sum.linear_isometry_equiv: If a Hilbert spaceEis a Hilbert sum of the inner product spacesG iwith respect to the familyV : Π i, G i →ₗᵢ[𝕜] E, then the correspondingorthogonal_family.linear_isometrycan be upgraded to alinear_isometry_equiv. -
hilbert_basis: We define a Hilbert basis of a Hilbert spaceEto be a structure whose single fieldhilbert_basis.repris an isometric isomorphism ofEwithℓ²(ι, 𝕜)(i.e., the Hilbert sum ofιcopies of𝕜). This parallels the definition ofbasis, inlinear_algebra.basis, as an isomorphism of anR-module withι →₀ R. -
hilbert_basis.has_coe_to_fun: More conventionally a Hilbert basis is thought of as a familyι → Eof vectors inEsatisfying certain properties (orthonormality, completeness). We obtain this interpretation of a Hilbert basisbby defining⇑b, of typeι → E, to be the image underb.reproflp.single 2 i (1:𝕜). This parallels the definitionbasis.has_coe_to_funinlinear_algebra.basis. -
hilbert_basis.mk: Make a Hilbert basis ofEfrom an orthonormal familyv : ι → Eof vectors inEwhose span is dense. This parallels the definitionbasis.mkinlinear_algebra.basis. -
hilbert_basis.mk_of_orthogonal_eq_bot: Make a Hilbert basis ofEfrom an orthonormal familyv : ι → Eof vectors inEwhose span has trivial orthogonal complement.
Main results #
-
lp.inner_product_space: Construction of the inner product space instance on the Hilbert sumlp G 2. Note that from the fileanalysis.normed_space.lp_space, the spacelp G 2already held a normed space instance (lp.normed_space), and if eachG iis a Hilbert space (i.e., complete), thenlp G 2was already known to be complete (lp.complete_space). So the work here is to define the inner product and show it is compatible. -
orthogonal_family.range_linear_isometry: Given a familyGof inner product spaces and a familyV : Π i, G i →ₗᵢ[𝕜] Eof isometric embeddings of theG iintoEwith mutually-orthogonal images, the image of the embeddingorthogonal_family.linear_isometryof the Hilbert sum ofGintoEis the closure of the span of the images of theG i. -
hilbert_basis.repr_apply_apply: Given a Hilbert basisbofE, the entryb.repr x iofx's representation inℓ²(ι, 𝕜)is the inner product⟪b i, x⟫. -
hilbert_basis.has_sum_repr: Given a Hilbert basisbofE, a vectorxinEcan be expressed as the "infinite linear combination"∑' i, b.repr x i • b iof the basis vectorsb i, with coefficients given by the entriesb.repr x iofx's representation inℓ²(ι, 𝕜). -
exists_hilbert_basis: A Hilbert space admits a Hilbert basis.
Keywords #
Hilbert space, Hilbert sum, l2, Hilbert basis, unitary equivalence, isometric isomorphism
Equations
- lp.inner_product_space = {to_normed_space := {to_module := normed_space.to_module lp.normed_space, norm_smul_le := _}, to_has_inner := {inner := λ (f g : ↥(lp G 2)), ∑' (i : ι), has_inner.inner (⇑f i) (⇑g i)}, norm_sq_eq_inner := _, conj_symm := _, add_left := _, smul_left := _}
Identification of a general Hilbert space E with a Hilbert sum #
A mutually orthogonal family of subspaces of E induce a linear isometry from lp 2 of the
subspaces into E.
The canonical linear isometry from the lp 2 of a mutually orthogonal family of subspaces of
E into E, has range the closure of the span of the subspaces.
- orthogonal_family : orthogonal_family 𝕜 G V
- surjective_isometry : function.surjective ⇑(_.linear_isometry)
Given a family of Hilbert spaces G : ι → Type*, a Hilbert sum of G consists of a Hilbert
space E and an orthogonal family V : Π i, G i →ₗᵢ[𝕜] E such that the induced isometry
Φ : lp G 2 → E is surjective.
Keeping in mind that lp G 2 is "the" external Hilbert sum of G : ι → Type*, this is analogous
to direct_sum.is_internal, except that we don't express it in terms of actual submodules.
If V : Π i, G i →ₗᵢ[𝕜] E is an orthogonal family such that the supremum of the ranges of
V i is dense, then (E, V) is a Hilbert sum of G.
This is orthogonal_family.is_hilbert_sum in the case of actual inclusions from subspaces.
A Hilbert sum (E, V) of G is canonically isomorphic to the Hilbert sum of G,
i.e lp G 2.
Note that this goes in the opposite direction from orthogonal_family.linear_isometry.
Equations
In the canonical isometric isomorphism between a Hilbert sum E of G and lp G 2,
a vector w : lp G 2 is the image of the infinite sum of the associated elements in E.
In the canonical isometric isomorphism between a Hilbert sum E of G and lp G 2,
a vector w : lp G 2 is the image of the infinite sum of the associated elements in E, and this
sum indeed converges.
In the canonical isometric isomorphism between a Hilbert sum E of G : ι → Type* and
lp G 2, an "elementary basis vector" in lp G 2 supported at i : ι is the image of the
associated element in E.
In the canonical isometric isomorphism between a Hilbert sum E of G : ι → Type* and
lp G 2, a finitely-supported vector in lp G 2 is the image of the associated finite sum of
elements of E.
In the canonical isometric isomorphism between a Hilbert sum E of G : ι → Type* and
lp G 2, a finitely-supported vector in lp G 2 is the image of the associated finite sum of
elements of E.
Given a total orthonormal family v : ι → E, E is a Hilbert sum of λ i : ι, 𝕜 relative to
the family of linear isometries λ i, λ k, k • v i.
Hilbert bases #
A Hilbert basis on ι for an inner product space E is an identification of E with the lp
space ℓ²(ι, 𝕜).
Instances for hilbert_basis
- hilbert_basis.has_sizeof_inst
- hilbert_basis.inhabited
- hilbert_basis.has_coe_to_fun
Equations
- hilbert_basis.inhabited = {default := {repr := linear_isometry_equiv.refl 𝕜 ↥(lp (λ (i : ι), 𝕜) 2) normed_space.to_module}}
b i is the ith basis vector.
Equations
- hilbert_basis.has_coe_to_fun = {coe := λ (b : hilbert_basis ι 𝕜 E) (i : ι), ⇑(b.repr.symm) (lp.single 2 i 1)}
A finite Hilbert basis is an orthonormal basis.
Equations
An orthonormal family of vectors whose span is dense in the whole module is a Hilbert basis.
Equations
- hilbert_basis.mk hv hsp = {repr := _.linear_isometry_equiv}
An orthonormal family of vectors whose span has trivial orthogonal complement is a Hilbert basis.
Equations
- hilbert_basis.mk_of_orthogonal_eq_bot hv hsp = hilbert_basis.mk hv _
An orthonormal basis is an Hilbert basis.
Equations
A Hilbert space admits a Hilbert basis extending a given orthonormal subset.
A Hilbert space admits a Hilbert basis.