Inner product space #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file defines inner product spaces and proves the basic properties. We do not formally
define Hilbert spaces, but they can be obtained using the set of assumptions
[normed_add_comm_group E] [inner_product_space 𝕜 E] [complete_space E]
.
An inner product space is a vector space endowed with an inner product. It generalizes the notion of
dot product in ℝ^n
and provides the means of defining the length of a vector and the angle between
two vectors. In particular vectors x
and y
are orthogonal if their inner product equals zero.
We define both the real and complex cases at the same time using the is_R_or_C
typeclass.
This file proves general results on inner product spaces. For the specific construction of an inner
product structure on n → 𝕜
for 𝕜 = ℝ
or ℂ
, see euclidean_space
in
analysis.inner_product_space.pi_L2
.
Main results #
- We define the class
inner_product_space 𝕜 E
extendingnormed_space 𝕜 E
with a number of basic properties, most notably the Cauchy-Schwarz inequality. Here𝕜
is understood to be eitherℝ
orℂ
, through theis_R_or_C
typeclass. - We show that the inner product is continuous,
continuous_inner
, and bundle it as the the continuous sesquilinear mapinnerSL
(see alsoinnerₛₗ
for the non-continuous version). - We define
orthonormal
, a predicate on a functionv : ι → E
, and prove the existence of a maximal orthonormal set,exists_maximal_orthonormal
. Bessel's inequality,orthonormal.tsum_inner_products_le
, states that given an orthonormal setv
and a vectorx
, the sum of the norm-squares of the inner products⟪v i, x⟫
is no more than the norm-square ofx
. For the existence of orthonormal bases, Hilbert bases, etc., see the fileanalysis.inner_product_space.projection
.
Notation #
We globally denote the real and complex inner products by ⟪·, ·⟫_ℝ
and ⟪·, ·⟫_ℂ
respectively.
We also provide two notation namespaces: real_inner_product_space
, complex_inner_product_space
,
which respectively introduce the plain notation ⟪·, ·⟫
for the real and complex inner product.
Implementation notes #
We choose the convention that inner products are conjugate linear in the first argument and linear in the second.
Tags #
inner product space, Hilbert space, norm
References #
- [Clément & Martin, The Lax-Milgram Theorem. A detailed proof to be formalized in Coq]
- [Clément & Martin, A Coq formal proof of the Lax–Milgram theorem]
The Coq code is available at the following address: http://www.lri.fr/~sboldo/elfic/index.html
Syntactic typeclass for types endowed with an inner product
Instances of this typeclass
Instances of other typeclasses for has_inner
- has_inner.has_sizeof_inst
- to_normed_space : normed_space 𝕜 E
- to_has_inner : has_inner 𝕜 E
- norm_sq_eq_inner : ∀ (x : E), ‖x‖ ^ 2 = ⇑is_R_or_C.re (has_inner.inner x x)
- conj_symm : ∀ (x y : E), ⇑(star_ring_end 𝕜) (has_inner.inner y x) = has_inner.inner x y
- add_left : ∀ (x y z : E), has_inner.inner (x + y) z = has_inner.inner x z + has_inner.inner y z
- smul_left : ∀ (x y : E) (r : 𝕜), has_inner.inner (r • x) y = ⇑(star_ring_end 𝕜) r * has_inner.inner x y
An inner product space is a vector space with an additional operation called inner product.
The norm could be derived from the inner product, instead we require the existence of a norm and
the fact that ‖x‖^2 = re ⟪x, x⟫
to be able to put instances on 𝕂
or product
spaces.
To construct a norm from an inner product, see inner_product_space.of_core
.
Instances of this typeclass
Instances of other typeclasses for inner_product_space
- inner_product_space.has_sizeof_inst
Constructing a normed space structure from an inner product #
In the definition of an inner product space, we require the existence of a norm, which is equal
(but maybe not defeq) to the square root of the scalar product. This makes it possible to put
an inner product space structure on spaces with a preexisting norm (for instance ℝ
), with good
properties. However, sometimes, one would like to define the norm starting only from a well-behaved
scalar product. This is what we implement in this paragraph, starting from a structure
inner_product_space.core
stating that we have a nice scalar product.
Our goal here is not to develop a whole theory with all the supporting API, as this will be done
below for inner_product_space
. Instead, we implement the bare minimum to go as directly as
possible to the construction of the norm and the proof of the triangular inequality.
Warning: Do not use this core
structure if the space you are interested in already has a norm
instance defined on it, otherwise this will create a second non-defeq norm instance!
- to_has_inner : has_inner 𝕜 F
- conj_symm : ∀ (x y : F), ⇑(star_ring_end 𝕜) (has_inner.inner y x) = has_inner.inner x y
- nonneg_re : ∀ (x : F), 0 ≤ ⇑is_R_or_C.re (has_inner.inner x x)
- definite : ∀ (x : F), has_inner.inner x x = 0 → x = 0
- add_left : ∀ (x y z : F), has_inner.inner (x + y) z = has_inner.inner x z + has_inner.inner y z
- smul_left : ∀ (x y : F) (r : 𝕜), has_inner.inner (r • x) y = ⇑(star_ring_end 𝕜) r * has_inner.inner x y
A structure requiring that a scalar product is positive definite and symmetric, from which one
can construct an inner_product_space
instance in inner_product_space.of_core
.
Instances for inner_product_space.core
- inner_product_space.core.has_sizeof_inst
Define inner_product_space.core
from inner_product_space
. Defined to reuse lemmas about
inner_product_space.core
for inner_product_space
s. Note that the has_norm
instance provided by
inner_product_space.core.has_norm
is propositionally but not definitionally equal to the original
norm.
Equations
- inner_product_space.to_core = {to_has_inner := inner_product_space.to_has_inner c, conj_symm := _, nonneg_re := _, definite := _, add_left := _, smul_left := _}
Inner product defined by the inner_product_space.core
structure. We can't reuse
inner_product_space.core.to_has_inner
because it takes inner_product_space.core
as an explicit
argument.
Equations
The norm squared function for inner_product_space.core
structure.
Equations
Expand inner (x + y) (x + y)
An auxiliary equality useful to prove the Cauchy–Schwarz inequality: the square of the norm
of ⟪x, y⟫ • x - ⟪x, x⟫ • y
is equal to ‖x‖ ^ 2 * (‖x‖ ^ 2 * ‖y‖ ^ 2 - ‖⟪x, y⟫‖ ^ 2)
. We use
inner_product_space.of_core.norm_sq x
etc (defeq to is_R_or_C.re ⟪x, x⟫
) instead of ‖x‖ ^ 2
etc to avoid extra rewrites when applying it to an inner_product_space
.
Cauchy–Schwarz inequality.
We need this for the core
structure to prove the triangle inequality below when
showing the core is a normed group.
Norm constructed from a inner_product_space.core
structure, defined to be the square root
of the scalar product.
Equations
- inner_product_space.core.to_has_norm = {norm := λ (x : F), √ (⇑is_R_or_C.re (has_inner.inner x x))}
Cauchy–Schwarz inequality with norm
Normed group structure constructed from an inner_product_space.core
structure
Equations
- inner_product_space.core.to_normed_add_comm_group = {to_fun := λ (x : F), √ (⇑is_R_or_C.re (has_inner.inner x x)), map_zero' := _, add_le' := _, neg' := _, eq_zero_of_map_eq_zero' := _}.to_normed_add_comm_group
Normed space structure constructed from a inner_product_space.core
structure
Equations
- inner_product_space.core.to_normed_space = {to_module := _inst_3, norm_smul_le := _}
Given a inner_product_space.core
structure on a space, one can use it to turn
the space into an inner product space. The normed_add_comm_group
structure is expected
to already be defined with inner_product_space.of_core.to_normed_add_comm_group
.
Equations
- inner_product_space.of_core c = let _inst : normed_space 𝕜 F := inner_product_space.core.to_normed_space in {to_normed_space := _inst, to_has_inner := c.to_has_inner, norm_sq_eq_inner := _, conj_symm := _, add_left := _, smul_left := _}
Properties of inner product spaces #
The inner product as a sesquilinear form.
Note that in the case 𝕜 = ℝ
this is a bilinear form.
Equations
- sesq_form_of_inner = linear_map.mk₂'ₛₗ (ring_hom.id 𝕜) (star_ring_end 𝕜) (λ (x y : E), has_inner.inner y x) sesq_form_of_inner._proof_3 sesq_form_of_inner._proof_4 sesq_form_of_inner._proof_5 sesq_form_of_inner._proof_6
The real inner product as a bilinear form.
Equations
- bilin_form_of_real_inner = {bilin := has_inner.inner inner_product_space.to_has_inner, bilin_add_left := _, bilin_smul_left := _, bilin_add_right := _, bilin_smul_right := _}
An inner product with a sum on the left.
An inner product with a sum on the right.
An inner product with a sum on the left, finsupp
version.
An inner product with a sum on the right, finsupp
version.
Expand ⟪x + y, x + y⟫
Expand ⟪x + y, x + y⟫_ℝ
Expand ⟪x - y, x - y⟫_ℝ
Parallelogram law
Cauchy–Schwarz inequality.
Cauchy–Schwarz inequality for real inner products.
A family of vectors is linearly independent if they are nonzero and orthogonal.
An orthonormal set of vectors in an inner_product_space
if ... then ... else
characterization of an indexed set of vectors being orthonormal. (Inner
product equals Kronecker delta.)
if ... then ... else
characterization of a set of vectors being orthonormal. (Inner product
equals Kronecker delta.)
The inner product of a linear combination of a set of orthonormal vectors with one of those vectors picks out the coefficient of that vector.
The inner product of a linear combination of a set of orthonormal vectors with one of those vectors picks out the coefficient of that vector.
The inner product of a linear combination of a set of orthonormal vectors with one of those vectors picks out the coefficient of that vector.
The inner product of a linear combination of a set of orthonormal vectors with one of those vectors picks out the coefficient of that vector.
The inner product of a linear combination of a set of orthonormal vectors with one of those vectors picks out the coefficient of that vector.
The inner product of a linear combination of a set of orthonormal vectors with one of those vectors picks out the coefficient of that vector.
The inner product of two linear combinations of a set of orthonormal vectors, expressed as
a sum over the first finsupp
.
The inner product of two linear combinations of a set of orthonormal vectors, expressed as
a sum over the second finsupp
.
The inner product of two linear combinations of a set of orthonormal vectors, expressed as a sum.
The double sum of weighted inner products of pairs of vectors from an orthonormal sequence is the sum of the weights.
An orthonormal set is linearly independent.
A subfamily of an orthonormal family (i.e., a composition with an injective map) is an orthonormal family.
An injective family v : ι → E
is orthonormal if and only if coe : (range v) → E
is
orthonormal.
If v : ι → E
is an orthonormal family, then coe : (range v) → E
is an orthonormal
family.
A linear combination of some subset of an orthonormal set is orthogonal to other members of the set.
Given an orthonormal family, a second family of vectors is orthonormal if every vector equals the corresponding vector in the original family or its negation.
Given an orthonormal set v
of vectors in E
, there exists a maximal orthonormal set
containing it.
A family of orthonormal vectors with the correct cardinality forms a basis.
Equations
- basis_of_orthonormal_of_card_eq_finrank hv card_eq = basis_of_linear_independent_of_card_eq_finrank _ card_eq
Expand the square
Alias of norm_add_sq
.
Alias of norm_add_sq_real
.
Expand the square
Alias of norm_sub_sq
.
Alias of norm_sub_sq_real
.
Cauchy–Schwarz inequality with norm
Cauchy–Schwarz inequality with norm
Cauchy–Schwarz inequality with norm
Polarization identity: The real part of the inner product, in terms of the norm.
Polarization identity: The real part of the inner product, in terms of the norm.
Polarization identity: The real part of the inner product, in terms of the norm.
Polarization identity: The imaginary part of the inner product, in terms of the norm.
Polarization identity: The inner product, in terms of the norm.
Formula for the distance between the images of two nonzero points under an inversion with center
zero. See also euclidean_geometry.dist_inversion_inversion
for inversions around a general
point.
A complex polarization identity, with a linear map
A linear map T
is zero, if and only if the identity ⟪T x, x⟫_ℂ = 0
holds for all x
.
Two linear maps S
and T
are equal, if and only if the identity ⟪S x, x⟫_ℂ = ⟪T x, x⟫_ℂ
holds
for all x
.
A linear isometry preserves the inner product.
A linear isometric equivalence preserves the inner product.
A linear map that preserves the inner product is a linear isometry.
Equations
- f.isometry_of_inner h = {to_linear_map := f, norm_map' := _}
A linear equivalence that preserves the inner product is a linear isometric equivalence.
Equations
- f.isometry_of_inner h = {to_linear_equiv := f, norm_map' := _}
A linear isometry preserves the property of being orthonormal.
A linear isometry preserves the property of being orthonormal.
A linear isometric equivalence preserves the property of being orthonormal.
A linear isometric equivalence, applied with basis.map
, preserves the property of being
orthonormal.
A linear map that sends an orthonormal basis to orthonormal vectors is a linear isometry.
Equations
- f.isometry_of_orthonormal hv hf = f.isometry_of_inner _
A linear equivalence that sends an orthonormal basis to orthonormal vectors is a linear isometric equivalence.
Equations
- f.isometry_of_orthonormal hv hf = f.isometry_of_inner _
A linear isometric equivalence that sends an orthonormal basis to a given orthonormal basis.
Equations
- hv.equiv hv' e = (v.equiv v' e).isometry_of_orthonormal hv _
Polarization identity: The real inner product, in terms of the norm.
Polarization identity: The real inner product, in terms of the norm.
Pythagorean theorem, subtracting vectors, if-and-only-if vector inner product form.
The sum and difference of two vectors are orthogonal if and only if they have the same norm.
Given two orthogonal vectors, their sum and difference have equal norms.
The real inner product of two vectors, divided by the product of their norms, has absolute value at most 1.
The inner product of a vector with a multiple of itself.
The inner product of a vector with a multiple of itself.
The inner product of a nonzero vector with a nonzero multiple of itself, divided by the product of their norms, has absolute value 1.
The inner product of a nonzero vector with a nonzero multiple of itself, divided by the product of their norms, has absolute value 1.
The inner product of a nonzero vector with a positive multiple of itself, divided by the product of their norms, has value 1.
The inner product of a nonzero vector with a negative multiple of itself, divided by the product of their norms, has value -1.
If the inner product of two vectors is equal to the product of their norms, then the two vectors
are multiples of each other. One form of the equality case for Cauchy-Schwarz.
Compare inner_eq_norm_mul_iff
, which takes the stronger hypothesis ⟪x, y⟫ = ‖x‖ * ‖y‖
.
The inner product of two vectors, divided by the product of their norms, has absolute value 1 if and only if they are nonzero and one is a multiple of the other. One form of equality case for Cauchy-Schwarz.
The inner product of two vectors, divided by the product of their norms, has absolute value 1 if and only if they are nonzero and one is a multiple of the other. One form of equality case for Cauchy-Schwarz.
If the inner product of two vectors is equal to the product of their norms (i.e.,
⟪x, y⟫ = ‖x‖ * ‖y‖
), then the two vectors are nonnegative real multiples of each other. One form
of the equality case for Cauchy-Schwarz.
Compare norm_inner_eq_norm_iff
, which takes the weaker hypothesis abs ⟪x, y⟫ = ‖x‖ * ‖y‖
.
If the inner product of two vectors is equal to the product of their norms (i.e.,
⟪x, y⟫ = ‖x‖ * ‖y‖
), then the two vectors are nonnegative real multiples of each other. One form
of the equality case for Cauchy-Schwarz.
Compare norm_inner_eq_norm_iff
, which takes the weaker hypothesis abs ⟪x, y⟫ = ‖x‖ * ‖y‖
.
The inner product of two vectors, divided by the product of their norms, has value 1 if and only if they are nonzero and one is a positive multiple of the other.
The inner product of two vectors, divided by the product of their norms, has value -1 if and only if they are nonzero and one is a negative multiple of the other.
If the inner product of two unit vectors is 1
, then the two vectors are equal. One form of
the equality case for Cauchy-Schwarz.
If the inner product of two unit vectors is strictly less than 1
, then the two vectors are
distinct. One form of the equality case for Cauchy-Schwarz.
The inner product of two weighted sums, where the weights in each sum add to 0, in terms of the norms of pairwise differences.
The inner product as a sesquilinear map.
Equations
- innerₛₗ 𝕜 = linear_map.mk₂'ₛₗ (star_ring_end 𝕜) (ring_hom.id 𝕜) (λ (v w : E), has_inner.inner v w) inner_add_left _ inner_add_right _
The inner product as a continuous sesquilinear map. Note that to_dual_map
(resp. to_dual
)
in inner_product_space.dual
is a version of this given as a linear isometry (resp. linear
isometric equivalence).
Equations
- innerSL 𝕜 = (innerₛₗ 𝕜).mk_continuous₂ 1 _
innerSL
is an isometry. Note that the associated linear_isometry
is defined in
inner_product_space.dual
as to_dual_map
.
The inner product as a continuous sesquilinear map, with the two arguments flipped.
Equations
- innerSL_flip 𝕜 = ⇑(continuous_linear_map.flipₗᵢ' E E 𝕜 (ring_hom.id 𝕜) (star_ring_end 𝕜)) (innerSL 𝕜)
Given f : E →L[𝕜] E'
, construct the continuous sesquilinear form λ x y, ⟪x, A y⟫
, given
as a continuous linear map.
Equations
- continuous_linear_map.to_sesq_form = ↑((continuous_linear_map.flipₗᵢ' E E' 𝕜 (star_ring_end 𝕜) (ring_hom.id 𝕜)).to_continuous_linear_equiv).comp (⇑(continuous_linear_map.compSL E E' (E' →L⋆[𝕜] 𝕜) (ring_hom.id 𝕜) (ring_hom.id 𝕜)) (innerSL_flip 𝕜))
When an inner product space E
over 𝕜
is considered as a real normed space, its inner
product satisfies is_bounded_bilinear_map
.
In order to state these results, we need a normed_space ℝ E
instance. We will later establish
such an instance by restriction-of-scalars, inner_product_space.is_R_or_C_to_real 𝕜 E
, but this
instance may be not definitionally equal to some other “natural” instance. So, we assume
[normed_space ℝ E]
.
Bessel's inequality for finite sums.
Bessel's inequality.
The sum defined in Bessel's inequality is summable.
A field 𝕜
satisfying is_R_or_C
is itself a 𝕜
-inner product space.
Equations
- is_R_or_C.inner_product_space = {to_normed_space := normed_field.to_normed_space densely_normed_field.to_normed_field, to_has_inner := {inner := λ (x y : 𝕜), ⇑(star_ring_end 𝕜) x * y}, norm_sq_eq_inner := _, conj_symm := _, add_left := _, smul_left := _}
Inner product space structure on subspaces #
Induced inner product on a submodule.
Equations
- W.inner_product_space = {to_normed_space := {to_module := normed_space.to_module W.normed_space, norm_smul_le := _}, to_has_inner := {inner := λ (x y : ↥W), has_inner.inner ↑x ↑y}, norm_sq_eq_inner := _, conj_symm := _, add_left := _, smul_left := _}
The inner product on submodules is the same as on the ambient space.
Families of mutually-orthogonal subspaces of an inner product space #
An indexed family of mutually-orthogonal subspaces of an inner product space E
.
The simple way to express this concept would be as a condition on V : ι → submodule 𝕜 E
. We
We instead implement it as a condition on a family of inner product spaces each equipped with an
isometric embedding into E
, thus making it a property of morphisms rather than subobjects.
The connection to the subobject spelling is shown in orthogonal_family_iff_pairwise
.
This definition is less lightweight, but allows for better definitional properties when the inner
product space structure on each of the submodules is important -- for example, when considering
their Hilbert sum (pi_lp V 2
). For example, given an orthonormal set of vectors v : ι → E
,
we have an associated orthogonal family of one-dimensional subspaces of E
, which it is convenient
to be able to discuss using ι → 𝕜
rather than Π i : ι, span 𝕜 (v i)
.
Equations
- orthogonal_family 𝕜 G V = ∀ ⦃i j : ι⦄, i ≠ j → ∀ (v : G i) (w : G j), has_inner.inner (⇑(V i) v) (⇑(V j) w) = 0
The composition of an orthogonal family of subspaces with an injective function is also an orthogonal family.
A family f
of mutually-orthogonal elements of E
is summable, if and only if
(λ i, ‖f i‖ ^ 2)
is summable.
An orthogonal family forms an independent family of subspaces; that is, any collection of elements each from a different subspace in the family is linearly independent. In particular, the pairwise intersections of elements of the family are 0.
A general inner product implies a real inner product. This is not registered as an instance
since it creates problems with the case 𝕜 = ℝ
.
Equations
- has_inner.is_R_or_C_to_real 𝕜 E = {inner := λ (x y : E), ⇑is_R_or_C.re (has_inner.inner x y)}
A general inner product space structure implies a real inner product structure. This is not
registered as an instance since it creates problems with the case 𝕜 = ℝ
, but in can be used in a
proof to obtain a real inner product space structure from a given 𝕜
-inner product space
structure.
Equations
- inner_product_space.is_R_or_C_to_real 𝕜 E = {to_normed_space := {to_module := normed_space.to_module (normed_space.restrict_scalars ℝ 𝕜 E), norm_smul_le := _}, to_has_inner := {inner := has_inner.inner (has_inner.is_R_or_C_to_real 𝕜 E)}, norm_sq_eq_inner := _, conj_symm := _, add_left := _, smul_left := _}
A complex inner product implies a real inner product
The inner product on an inner product space of dimension 2 can be evaluated in terms of a complex-number representation of the space.
Continuity of the inner product #
Extract a real bilinear form from an operator T
, by taking the pairing λ x, re ⟪T x, x⟫
.
Equations
- T.re_apply_inner_self x = ⇑is_R_or_C.re (has_inner.inner (⇑T x) x)
Equations
- uniform_space.completion.has_inner = {inner := function.curry (uniform_space.completion.has_inner._proof_1.extend (function.uncurry has_inner.inner))}
Equations
- uniform_space.completion.inner_product_space = {to_normed_space := uniform_space.completion.normed_space 𝕜 E inner_product_space.to_normed_space, to_has_inner := uniform_space.completion.has_inner inner_product_space.to_has_inner, norm_sq_eq_inner := _, conj_symm := _, add_left := _, smul_left := _}