mathlib3 documentation

analysis.inner_product_space.basic

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 #

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 #

The Coq code is available at the following address: http://www.lri.fr/~sboldo/elfic/index.html

@[class]
structure has_inner (𝕜 : Type u_4) (E : Type u_5) :
Type (max u_4 u_5)

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
@[class]
structure inner_product_space (𝕜 : Type u_4) (E : Type u_5) [is_R_or_C 𝕜] [normed_add_comm_group E] :
Type (max u_4 u_5)

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!

@[nolint, class]
structure inner_product_space.core (𝕜 : Type u_4) (F : Type u_5) [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] :
Type (max u_4 u_5)

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_spaces. 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
def inner_product_space.core.to_has_inner' {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] :
has_inner 𝕜 F

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
def inner_product_space.core.norm_sq {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x : F) :

The norm squared function for inner_product_space.core structure.

Equations
theorem inner_product_space.core.inner_conj_symm {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x y : F) :
theorem inner_product_space.core.inner_self_nonneg {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] {x : F} :
theorem inner_product_space.core.inner_self_im {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x : F) :
theorem inner_product_space.core.inner_add_left {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x y z : F) :
theorem inner_product_space.core.inner_add_right {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x y z : F) :
theorem inner_product_space.core.inner_smul_left {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x y : F) {r : 𝕜} :
theorem inner_product_space.core.inner_smul_right {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x y : F) {r : 𝕜} :
theorem inner_product_space.core.inner_zero_left {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x : F) :
theorem inner_product_space.core.inner_zero_right {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x : F) :
theorem inner_product_space.core.inner_self_eq_zero {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] {x : F} :
theorem inner_product_space.core.norm_sq_eq_zero {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] {x : F} :
theorem inner_product_space.core.inner_self_ne_zero {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] {x : F} :
theorem inner_product_space.core.inner_neg_left {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x y : F) :
theorem inner_product_space.core.inner_neg_right {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x y : F) :
theorem inner_product_space.core.inner_sub_left {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x y z : F) :
theorem inner_product_space.core.inner_sub_right {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x y z : F) :
theorem inner_product_space.core.inner_add_add_self {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x y : F) :

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.

noncomputable def inner_product_space.core.to_has_norm {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] :

Norm constructed from a inner_product_space.core structure, defined to be the square root of the scalar product.

Equations
theorem inner_product_space.core.norm_inner_le_norm {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] (x y : F) :

Cauchy–Schwarz inequality with norm

noncomputable def inner_product_space.core.to_normed_add_comm_group {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] :

Normed group structure constructed from an inner_product_space.core structure

Equations
def inner_product_space.core.to_normed_space {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] :

Normed space structure constructed from a inner_product_space.core structure

Equations
def inner_product_space.of_core {𝕜 : Type u_1} {F : Type u_3} [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] (c : inner_product_space.core 𝕜 F) :

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

Properties of inner product spaces #

@[simp]
theorem inner_conj_symm {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :
theorem inner_eq_zero_symm {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x y : E} :
@[simp]
theorem inner_self_im {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :
theorem inner_add_left {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y z : E) :
theorem inner_add_right {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y z : E) :
theorem inner_smul_left {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) (r : 𝕜) :
theorem inner_smul_real_left {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) (r : ) :
theorem inner_smul_right {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) (r : 𝕜) :
theorem inner_smul_real_right {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) (r : ) :
@[simp]
def sesq_form_of_inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] :
E →ₗ[𝕜] E →ₗ⋆[𝕜] 𝕜

The inner product as a sesquilinear form.

Note that in the case 𝕜 = ℝ this is a bilinear form.

Equations
theorem sum_inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_3} (s : finset ι) (f : ι E) (x : E) :
has_inner.inner (s.sum (λ (i : ι), f i)) x = s.sum (λ (i : ι), has_inner.inner (f i) x)

An inner product with a sum on the left.

theorem inner_sum {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_3} (s : finset ι) (f : ι E) (x : E) :
has_inner.inner x (s.sum (λ (i : ι), f i)) = s.sum (λ (i : ι), has_inner.inner x (f i))

An inner product with a sum on the right.

theorem finsupp.sum_inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_3} (l : ι →₀ 𝕜) (v : ι E) (x : E) :
has_inner.inner (l.sum (λ (i : ι) (a : 𝕜), a v i)) x = l.sum (λ (i : ι) (a : 𝕜), (star_ring_end 𝕜) a has_inner.inner (v i) x)

An inner product with a sum on the left, finsupp version.

theorem finsupp.inner_sum {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_3} (l : ι →₀ 𝕜) (v : ι E) (x : E) :
has_inner.inner x (l.sum (λ (i : ι) (a : 𝕜), a v i)) = l.sum (λ (i : ι) (a : 𝕜), a has_inner.inner x (v i))

An inner product with a sum on the right, finsupp version.

theorem dfinsupp.sum_inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_3} [dec : decidable_eq ι] {α : ι Type u_4} [Π (i : ι), add_zero_class (α i)] [Π (i : ι) (x : α i), decidable (x 0)] (f : Π (i : ι), α i E) (l : Π₀ (i : ι), α i) (x : E) :
has_inner.inner (l.sum f) x = l.sum (λ (i : ι) (a : α i), has_inner.inner (f i a) x)
theorem dfinsupp.inner_sum {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_3} [dec : decidable_eq ι] {α : ι Type u_4} [Π (i : ι), add_zero_class (α i)] [Π (i : ι) (x : α i), decidable (x 0)] (f : Π (i : ι), α i E) (l : Π₀ (i : ι), α i) (x : E) :
has_inner.inner x (l.sum f) = l.sum (λ (i : ι) (a : α i), has_inner.inner x (f i a))
@[simp]
theorem inner_zero_left {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :
theorem inner_re_zero_left {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :
@[simp]
theorem inner_zero_right {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :
theorem inner_re_zero_right {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :
theorem inner_self_nonneg {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x : E} :
@[simp]
theorem inner_self_re_to_K {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :
theorem inner_self_eq_norm_sq_to_K {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :
@[simp]
theorem inner_self_eq_zero {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x : E} :
theorem inner_self_ne_zero {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x : E} :
@[simp]
theorem inner_self_nonpos {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x : E} :
theorem norm_inner_symm {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :
@[simp]
theorem inner_neg_left {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :
@[simp]
theorem inner_neg_right {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :
theorem inner_neg_neg {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :
@[simp]
theorem inner_self_conj {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :
theorem inner_sub_left {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y z : E) :
theorem inner_sub_right {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y z : E) :
theorem inner_add_add_self {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :

Expand ⟪x + y, x + y⟫

Expand ⟪x + y, x + y⟫_ℝ

theorem inner_sub_sub_self {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :

Expand ⟪x - y, x - y⟫_ℝ

theorem ext_inner_left (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x y : E} (h : (v : E), has_inner.inner v x = has_inner.inner v y) :
x = y
theorem ext_inner_right (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x y : E} (h : (v : E), has_inner.inner x v = has_inner.inner y v) :
x = y
theorem parallelogram_law {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x y : E} :
has_inner.inner (x + y) (x + y) + has_inner.inner (x - y) (x - y) = 2 * (has_inner.inner x x + has_inner.inner y y)

Parallelogram law

Cauchy–Schwarz inequality.

Cauchy–Schwarz inequality for real inner products.

theorem linear_independent_of_ne_zero_of_inner_eq_zero {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_3} {v : ι E} (hz : (i : ι), v i 0) (ho : (i j : ι), i j has_inner.inner (v i) (v j) = 0) :

A family of vectors is linearly independent if they are nonzero and orthogonal.

def orthonormal (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} (v : ι E) :
Prop

An orthonormal set of vectors in an inner_product_space

Equations
theorem orthonormal_iff_ite {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} [dec_ι : decidable_eq ι] {v : ι E} :
orthonormal 𝕜 v (i j : ι), has_inner.inner (v i) (v j) = ite (i = j) 1 0

if ... then ... else characterization of an indexed set of vectors being orthonormal. (Inner product equals Kronecker delta.)

theorem orthonormal_subtype_iff_ite {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] [dec_E : decidable_eq E] {s : set E} :
orthonormal 𝕜 coe (v : E), v s (w : E), w s has_inner.inner v w = ite (v = w) 1 0

if ... then ... else characterization of a set of vectors being orthonormal. (Inner product equals Kronecker delta.)

theorem orthonormal.inner_right_finsupp {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) (l : ι →₀ 𝕜) (i : ι) :
has_inner.inner (v i) ((finsupp.total ι E 𝕜 v) l) = l i

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.

theorem orthonormal.inner_right_sum {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) (l : ι 𝕜) {s : finset ι} {i : ι} (hi : i s) :
has_inner.inner (v i) (s.sum (λ (i : ι), l i v i)) = l i

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.

theorem orthonormal.inner_right_fintype {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} [fintype ι] {v : ι E} (hv : orthonormal 𝕜 v) (l : ι 𝕜) (i : ι) :
has_inner.inner (v i) (finset.univ.sum (λ (i : ι), l i v i)) = l i

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.

theorem orthonormal.inner_left_finsupp {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) (l : ι →₀ 𝕜) (i : ι) :
has_inner.inner ((finsupp.total ι E 𝕜 v) l) (v i) = (star_ring_end 𝕜) (l i)

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.

theorem orthonormal.inner_left_sum {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) (l : ι 𝕜) {s : finset ι} {i : ι} (hi : i s) :
has_inner.inner (s.sum (λ (i : ι), l i v i)) (v i) = (star_ring_end 𝕜) (l i)

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.

theorem orthonormal.inner_left_fintype {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} [fintype ι] {v : ι E} (hv : orthonormal 𝕜 v) (l : ι 𝕜) (i : ι) :
has_inner.inner (finset.univ.sum (λ (i : ι), l i v i)) (v i) = (star_ring_end 𝕜) (l i)

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.

theorem orthonormal.inner_finsupp_eq_sum_left {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) (l₁ l₂ : ι →₀ 𝕜) :
has_inner.inner ((finsupp.total ι E 𝕜 v) l₁) ((finsupp.total ι E 𝕜 v) l₂) = l₁.sum (λ (i : ι) (y : 𝕜), (star_ring_end 𝕜) y * l₂ i)

The inner product of two linear combinations of a set of orthonormal vectors, expressed as a sum over the first finsupp.

theorem orthonormal.inner_finsupp_eq_sum_right {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) (l₁ l₂ : ι →₀ 𝕜) :
has_inner.inner ((finsupp.total ι E 𝕜 v) l₁) ((finsupp.total ι E 𝕜 v) l₂) = l₂.sum (λ (i : ι) (y : 𝕜), (star_ring_end 𝕜) (l₁ i) * y)

The inner product of two linear combinations of a set of orthonormal vectors, expressed as a sum over the second finsupp.

theorem orthonormal.inner_sum {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) (l₁ l₂ : ι 𝕜) (s : finset ι) :
has_inner.inner (s.sum (λ (i : ι), l₁ i v i)) (s.sum (λ (i : ι), l₂ i v i)) = s.sum (λ (i : ι), (star_ring_end 𝕜) (l₁ i) * l₂ i)

The inner product of two linear combinations of a set of orthonormal vectors, expressed as a sum.

theorem orthonormal.inner_left_right_finset {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {s : finset ι} {v : ι E} (hv : orthonormal 𝕜 v) {a : ι ι 𝕜} :
s.sum (λ (i : ι), s.sum (λ (j : ι), a i j has_inner.inner (v j) (v i))) = s.sum (λ (k : ι), a k k)

The double sum of weighted inner products of pairs of vectors from an orthonormal sequence is the sum of the weights.

theorem orthonormal.linear_independent {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) :

An orthonormal set is linearly independent.

theorem orthonormal.comp {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {ι' : Type u_3} {v : ι E} (hv : orthonormal 𝕜 v) (f : ι' ι) (hf : function.injective f) :
orthonormal 𝕜 (v f)

A subfamily of an orthonormal family (i.e., a composition with an injective map) is an orthonormal family.

theorem orthonormal_subtype_range {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : function.injective v) :

An injective family v : ι → E is orthonormal if and only if coe : (range v) → E is orthonormal.

theorem orthonormal.to_subtype_range {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) :

If v : ι → E is an orthonormal family, then coe : (range v) → E is an orthonormal family.

theorem orthonormal.inner_finsupp_eq_zero {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) {s : set ι} {i : ι} (hi : i s) {l : ι →₀ 𝕜} (hl : l finsupp.supported 𝕜 𝕜 s) :
has_inner.inner ((finsupp.total ι E 𝕜 v) l) (v i) = 0

A linear combination of some subset of an orthonormal set is orthogonal to other members of the set.

theorem orthonormal.orthonormal_of_forall_eq_or_eq_neg {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v w : ι E} (hv : orthonormal 𝕜 v) (hw : (i : ι), w i = v i w i = -v i) :

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.

theorem orthonormal_empty (𝕜 : Type u_1) (E : Type u_2) [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] :
orthonormal 𝕜 (λ (x : ), x)
theorem orthonormal_Union_of_directed {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {η : Type u_3} {s : η set E} (hs : directed has_subset.subset s) (h : (i : η), orthonormal 𝕜 (λ (x : (s i)), x)) :
orthonormal 𝕜 (λ (x : (i : η), s i), x)
theorem orthonormal_sUnion_of_directed {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {s : set (set E)} (hs : directed_on has_subset.subset s) (h : (a : set E), a s orthonormal 𝕜 (λ (x : a), x)) :
orthonormal 𝕜 (λ (x : ⋃₀ s), x)
theorem exists_maximal_orthonormal {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {s : set E} (hs : orthonormal 𝕜 coe) :
(w : set E) (H : w s), orthonormal 𝕜 coe (u : set E), u w orthonormal 𝕜 coe u = w

Given an orthonormal set v of vectors in E, there exists a maximal orthonormal set containing it.

theorem orthonormal.ne_zero {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) (i : ι) :
v i 0
noncomputable def basis_of_orthonormal_of_card_eq_finrank {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} [fintype ι] [nonempty ι] {v : ι E} (hv : orthonormal 𝕜 v) (card_eq : fintype.card ι = finite_dimensional.finrank 𝕜 E) :
basis ι 𝕜 E

A family of orthonormal vectors with the correct cardinality forms a basis.

Equations
@[simp]
theorem coe_basis_of_orthonormal_of_card_eq_finrank {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} [fintype ι] [nonempty ι] {v : ι E} (hv : orthonormal 𝕜 v) (card_eq : fintype.card ι = finite_dimensional.finrank 𝕜 E) :
theorem norm_eq_sqrt_inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :
theorem inner_self_eq_norm_sq {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :
theorem norm_add_sq {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :

Expand the square

theorem norm_add_pow_two {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :

Alias of norm_add_sq.

theorem norm_add_sq_real {F : Type u_3} [normed_add_comm_group F] [inner_product_space F] (x y : F) :
x + y ^ 2 = x ^ 2 + 2 * has_inner.inner x y + y ^ 2

Expand the square

theorem norm_add_mul_self {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :

Expand the square

Expand the square

theorem norm_sub_sq {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :

Expand the square

theorem norm_sub_pow_two {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :

Alias of norm_sub_sq.

theorem norm_sub_sq_real {F : Type u_3} [normed_add_comm_group F] [inner_product_space F] (x y : F) :
x - y ^ 2 = x ^ 2 - 2 * has_inner.inner x y + y ^ 2

Expand the square

theorem norm_sub_mul_self {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :

Expand the square

Expand the square

theorem norm_inner_le_norm {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :

Cauchy–Schwarz inequality with norm

theorem re_inner_le_norm {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :

Cauchy–Schwarz inequality with norm

Cauchy–Schwarz inequality with norm

theorem parallelogram_law_with_norm (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :

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.

theorem dist_div_norm_sq_smul {F : Type u_3} [normed_add_comm_group F] [inner_product_space F] {x y : F} (hx : x 0) (hy : y 0) (R : ) :
has_dist.dist ((R / x) ^ 2 x) ((R / y) ^ 2 y) = R ^ 2 / (x * y) * has_dist.dist x y

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

theorem inner_map_self_eq_zero {V : Type u_4} [normed_add_comm_group V] [inner_product_space V] (T : V →ₗ[] V) :
( (x : V), has_inner.inner (T x) x = 0) T = 0

A linear map T is zero, if and only if the identity ⟪T x, x⟫_ℂ = 0 holds for all x.

theorem ext_inner_map {V : Type u_4} [normed_add_comm_group V] [inner_product_space V] (S T : V →ₗ[] V) :
( (x : V), has_inner.inner (S x) x = has_inner.inner (T x) x) S = T

Two linear maps S and T are equal, if and only if the identity ⟪S x, x⟫_ℂ = ⟪T x, x⟫_ℂ holds for all x.

@[simp]
theorem linear_isometry.inner_map_map {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E →ₗᵢ[𝕜] E') (x y : E) :

A linear isometry preserves the inner product.

@[simp]
theorem linear_isometry_equiv.inner_map_map {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E ≃ₗᵢ[𝕜] E') (x y : E) :

A linear isometric equivalence preserves the inner product.

def linear_map.isometry_of_inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E →ₗ[𝕜] E') (h : (x y : E), has_inner.inner (f x) (f y) = has_inner.inner x y) :
E →ₗᵢ[𝕜] E'

A linear map that preserves the inner product is a linear isometry.

Equations
@[simp]
theorem linear_map.coe_isometry_of_inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E →ₗ[𝕜] E') (h : (x y : E), has_inner.inner (f x) (f y) = has_inner.inner x y) :
@[simp]
theorem linear_map.isometry_of_inner_to_linear_map {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E →ₗ[𝕜] E') (h : (x y : E), has_inner.inner (f x) (f y) = has_inner.inner x y) :
def linear_equiv.isometry_of_inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E ≃ₗ[𝕜] E') (h : (x y : E), has_inner.inner (f x) (f y) = has_inner.inner x y) :
E ≃ₗᵢ[𝕜] E'

A linear equivalence that preserves the inner product is a linear isometric equivalence.

Equations
@[simp]
theorem linear_equiv.coe_isometry_of_inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E ≃ₗ[𝕜] E') (h : (x y : E), has_inner.inner (f x) (f y) = has_inner.inner x y) :
@[simp]
theorem linear_equiv.isometry_of_inner_to_linear_equiv {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E ≃ₗ[𝕜] E') (h : (x y : E), has_inner.inner (f x) (f y) = has_inner.inner x y) :
theorem linear_isometry.orthonormal_comp_iff {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] {v : ι E} (f : E →ₗᵢ[𝕜] E') :
orthonormal 𝕜 (f v) orthonormal 𝕜 v

A linear isometry preserves the property of being orthonormal.

theorem orthonormal.comp_linear_isometry {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] {v : ι E} (hv : orthonormal 𝕜 v) (f : E →ₗᵢ[𝕜] E') :
orthonormal 𝕜 (f v)

A linear isometry preserves the property of being orthonormal.

theorem orthonormal.comp_linear_isometry_equiv {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] {v : ι E} (hv : orthonormal 𝕜 v) (f : E ≃ₗᵢ[𝕜] E') :
orthonormal 𝕜 (f v)

A linear isometric equivalence preserves the property of being orthonormal.

theorem orthonormal.map_linear_isometry_equiv {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) (f : E ≃ₗᵢ[𝕜] E') :

A linear isometric equivalence, applied with basis.map, preserves the property of being orthonormal.

def linear_map.isometry_of_orthonormal {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E →ₗ[𝕜] E') {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) (hf : orthonormal 𝕜 (f v)) :
E →ₗᵢ[𝕜] E'

A linear map that sends an orthonormal basis to orthonormal vectors is a linear isometry.

Equations
@[simp]
theorem linear_map.coe_isometry_of_orthonormal {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E →ₗ[𝕜] E') {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) (hf : orthonormal 𝕜 (f v)) :
@[simp]
theorem linear_map.isometry_of_orthonormal_to_linear_map {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E →ₗ[𝕜] E') {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) (hf : orthonormal 𝕜 (f v)) :
def linear_equiv.isometry_of_orthonormal {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E ≃ₗ[𝕜] E') {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) (hf : orthonormal 𝕜 (f v)) :
E ≃ₗᵢ[𝕜] E'

A linear equivalence that sends an orthonormal basis to orthonormal vectors is a linear isometric equivalence.

Equations
@[simp]
theorem linear_equiv.coe_isometry_of_orthonormal {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E ≃ₗ[𝕜] E') {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) (hf : orthonormal 𝕜 (f v)) :
@[simp]
theorem linear_equiv.isometry_of_orthonormal_to_linear_equiv {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E ≃ₗ[𝕜] E') {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) (hf : orthonormal 𝕜 (f v)) :
noncomputable def orthonormal.equiv {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {ι' : Type u_5} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) {v' : basis ι' 𝕜 E'} (hv' : orthonormal 𝕜 v') (e : ι ι') :
E ≃ₗᵢ[𝕜] E'

A linear isometric equivalence that sends an orthonormal basis to a given orthonormal basis.

Equations
@[simp]
theorem orthonormal.equiv_to_linear_equiv {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {ι' : Type u_5} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) {v' : basis ι' 𝕜 E'} (hv' : orthonormal 𝕜 v') (e : ι ι') :
(hv.equiv hv' e).to_linear_equiv = v.equiv v' e
@[simp]
theorem orthonormal.equiv_apply {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] {ι' : Type u_3} {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) {v' : basis ι' 𝕜 E'} (hv' : orthonormal 𝕜 v') (e : ι ι') (i : ι) :
(hv.equiv hv' e) (v i) = v' (e i)
@[simp]
theorem orthonormal.equiv_refl {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) :
@[simp]
theorem orthonormal.equiv_symm {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {ι' : Type u_5} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) {v' : basis ι' 𝕜 E'} (hv' : orthonormal 𝕜 v') (e : ι ι') :
(hv.equiv hv' e).symm = hv'.equiv hv e.symm
@[simp]
theorem orthonormal.equiv_trans {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {ι' : Type u_5} {ι'' : Type u_6} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] {E'' : Type u_8} [normed_add_comm_group E''] [inner_product_space 𝕜 E''] {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) {v' : basis ι' 𝕜 E'} (hv' : orthonormal 𝕜 v') (e : ι ι') {v'' : basis ι'' 𝕜 E''} (hv'' : orthonormal 𝕜 v'') (e' : ι' ι'') :
(hv.equiv hv' e).trans (hv'.equiv hv'' e') = hv.equiv hv'' (e.trans e')
theorem orthonormal.map_equiv {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {ι' : Type u_5} {E' : Type u_7} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) {v' : basis ι' 𝕜 E'} (hv' : orthonormal 𝕜 v') (e : ι ι') :

Polarization identity: The real inner product, in terms of the norm.

Polarization identity: The real inner product, in terms of the norm.

Pythagorean theorem, if-and-only-if vector inner product form.

Pythagorean theorem, if-and-if vector inner product form using square roots.

Pythagorean theorem, vector inner product form.

Pythagorean theorem, vector inner product form.

Pythagorean theorem, subtracting vectors, if-and-only-if vector inner product form.

Pythagorean theorem, subtracting vectors, if-and-if vector inner product form using square roots.

Pythagorean theorem, subtracting vectors, vector inner product form.

The sum and difference of two vectors are orthogonal if and only if they have the same norm.

theorem norm_sub_eq_norm_add {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {v w : E} (h : has_inner.inner v w = 0) :
w - v = w + v

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.

theorem norm_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x : E} {r : 𝕜} (hx : x 0) (hr : r 0) :

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.

theorem real_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_pos_mul {F : Type u_3} [normed_add_comm_group F] [inner_product_space F] {x : F} {r : } (hx : x 0) (hr : 0 < r) :

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.

theorem norm_inner_eq_norm_tfae {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :
[has_inner.inner x y = x * y, x = 0 y = (has_inner.inner x y / has_inner.inner x x) x, x = 0 (r : 𝕜), y = r x, x = 0 y submodule.span 𝕜 {x}].tfae
theorem norm_inner_eq_norm_iff {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x y : E} (hx₀ : x 0) (hy₀ : y 0) :
has_inner.inner x y = x * y (r : 𝕜), r 0 y = r x

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‖.

theorem norm_inner_div_norm_mul_norm_eq_one_iff {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :
has_inner.inner x y / (x * y) = 1 x 0 (r : 𝕜), r 0 y = r x

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.

theorem inner_eq_norm_mul_iff_div {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x y : E} (h₀ : x 0) :
theorem inner_eq_norm_mul_iff {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x y : E} :

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.

theorem inner_eq_one_iff_of_norm_one {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {x y : E} (hx : x = 1) (hy : y = 1) :

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.

theorem inner_lt_one_iff_real_of_norm_one {F : Type u_3} [normed_add_comm_group F] [inner_product_space F] {x y : F} (hx : x = 1) (hy : y = 1) :

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.

theorem inner_sum_smul_sum_smul_of_sum_eq_zero {F : Type u_3} [normed_add_comm_group F] [inner_product_space F] {ι₁ : Type u_1} {s₁ : finset ι₁} {w₁ : ι₁ } (v₁ : ι₁ F) (h₁ : s₁.sum (λ (i : ι₁), w₁ i) = 0) {ι₂ : Type u_2} {s₂ : finset ι₂} {w₂ : ι₂ } (v₂ : ι₂ F) (h₂ : s₂.sum (λ (i : ι₂), w₂ i) = 0) :
has_inner.inner (s₁.sum (λ (i₁ : ι₁), w₁ i₁ v₁ i₁)) (s₂.sum (λ (i₂ : ι₂), w₂ i₂ v₂ i₂)) = -s₁.sum (λ (i₁ : ι₁), s₂.sum (λ (i₂ : ι₂), w₁ i₁ * w₂ i₂ * (v₁ i₁ - v₂ i₂ * v₁ i₁ - v₂ i₂))) / 2

The inner product of two weighted sums, where the weights in each sum add to 0, in terms of the norms of pairwise differences.

def innerₛₗ (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] :
E →ₗ⋆[𝕜] E →ₗ[𝕜] 𝕜

The inner product as a sesquilinear map.

Equations
@[simp]
theorem innerₛₗ_apply_coe (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (v : E) :
((innerₛₗ 𝕜) v) = λ (w : E), has_inner.inner v w
@[simp]
theorem innerₛₗ_apply (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (v w : E) :
noncomputable def innerSL (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] :
E →L⋆[𝕜] E →L[𝕜] 𝕜

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
@[simp]
theorem innerSL_apply_coe (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (v : E) :
((innerSL 𝕜) v) = λ (w : E), has_inner.inner v w
@[simp]
theorem innerSL_apply (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (v w : E) :
@[simp]
theorem innerSL_apply_norm (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :

innerSL is an isometry. Note that the associated linear_isometry is defined in inner_product_space.dual as to_dual_map.

noncomputable def innerSL_flip (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] :
E →L[𝕜] E →L⋆[𝕜] 𝕜

The inner product as a continuous sesquilinear map, with the two arguments flipped.

Equations
@[simp]
theorem innerSL_flip_apply (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x y : E) :
noncomputable def continuous_linear_map.to_sesq_form {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {E' : Type u_4} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] :
(E →L[𝕜] E') →L[𝕜] E' →L⋆[𝕜] E →L[𝕜] 𝕜

Given f : E →L[𝕜] E', construct the continuous sesquilinear form λ x y, ⟪x, A y⟫, given as a continuous linear map.

Equations
@[simp]
theorem continuous_linear_map.to_sesq_form_apply_coe {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {E' : Type u_4} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] (f : E →L[𝕜] E') (x : E') :

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].

theorem orthonormal.sum_inner_products_le {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} (x : E) {v : ι E} {s : finset ι} (hv : orthonormal 𝕜 v) :
s.sum (λ (i : ι), has_inner.inner (v i) x ^ 2) x ^ 2

Bessel's inequality for finite sums.

theorem orthonormal.tsum_inner_products_le {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} (x : E) {v : ι E} (hv : orthonormal 𝕜 v) :
∑' (i : ι), has_inner.inner (v i) x ^ 2 x ^ 2

Bessel's inequality.

theorem orthonormal.inner_products_summable {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} (x : E) {v : ι E} (hv : orthonormal 𝕜 v) :
summable (λ (i : ι), has_inner.inner (v i) x ^ 2)

The sum defined in Bessel's inequality is summable.

@[protected, instance]
def is_R_or_C.inner_product_space {𝕜 : Type u_1} [is_R_or_C 𝕜] :

A field 𝕜 satisfying is_R_or_C is itself a 𝕜-inner product space.

Equations
@[simp]
theorem is_R_or_C.inner_apply {𝕜 : Type u_1} [is_R_or_C 𝕜] (x y : 𝕜) :

Inner product space structure on subspaces #

@[protected, instance]
def submodule.inner_product_space {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (W : submodule 𝕜 E) :

Induced inner product on a submodule.

Equations
@[simp]
theorem submodule.coe_inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (W : submodule 𝕜 E) (x y : W) :

The inner product on submodules is the same as on the ambient space.

theorem orthonormal.cod_restrict {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_3} {v : ι E} (hv : orthonormal 𝕜 v) (s : submodule 𝕜 E) (hvs : (i : ι), v i s) :
theorem orthonormal_span {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_3} {v : ι E} (hv : orthonormal 𝕜 v) :
orthonormal 𝕜 (λ (i : ι), v i, _⟩)

Families of mutually-orthogonal subspaces of an inner product space #

def orthogonal_family (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} (G : ι Type u_3) [Π (i : ι), normed_add_comm_group (G i)] [Π (i : ι), inner_product_space 𝕜 (G i)] (V : Π (i : ι), G i →ₗᵢ[𝕜] E) :
Prop

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
theorem orthonormal.orthogonal_family {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {v : ι E} (hv : orthonormal 𝕜 v) :
orthogonal_family 𝕜 (λ (i : ι), 𝕜) (λ (i : ι), linear_isometry.to_span_singleton 𝕜 E _)
theorem orthogonal_family.eq_ite {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} [dec_ι : decidable_eq ι] {G : ι Type u_5} [Π (i : ι), normed_add_comm_group (G i)] [Π (i : ι), inner_product_space 𝕜 (G i)] {V : Π (i : ι), G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 G V) {i j : ι} (v : G i) (w : G j) :
has_inner.inner ((V i) v) ((V j) w) = ite (i = j) (has_inner.inner ((V i) v) ((V j) w)) 0
theorem orthogonal_family.inner_right_dfinsupp {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} [dec_ι : decidable_eq ι] {G : ι Type u_5} [Π (i : ι), normed_add_comm_group (G i)] [Π (i : ι), inner_product_space 𝕜 (G i)] {V : Π (i : ι), G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 G V) [dec_V : Π (i : ι) (x : G i), decidable (x 0)] (l : direct_sum ι (λ (i : ι), G i)) (i : ι) (v : G i) :
has_inner.inner ((V i) v) (dfinsupp.sum l (λ (j : ι), (V j))) = has_inner.inner v (l i)
theorem orthogonal_family.inner_right_fintype {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {G : ι Type u_5} [Π (i : ι), normed_add_comm_group (G i)] [Π (i : ι), inner_product_space 𝕜 (G i)] {V : Π (i : ι), G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 G V) [fintype ι] (l : Π (i : ι), G i) (i : ι) (v : G i) :
has_inner.inner ((V i) v) (finset.univ.sum (λ (j : ι), (V j) (l j))) = has_inner.inner v (l i)
theorem orthogonal_family.inner_sum {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {G : ι Type u_5} [Π (i : ι), normed_add_comm_group (G i)] [Π (i : ι), inner_product_space 𝕜 (G i)] {V : Π (i : ι), G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 G V) (l₁ l₂ : Π (i : ι), G i) (s : finset ι) :
has_inner.inner (s.sum (λ (i : ι), (V i) (l₁ i))) (s.sum (λ (j : ι), (V j) (l₂ j))) = s.sum (λ (i : ι), has_inner.inner (l₁ i) (l₂ i))
theorem orthogonal_family.norm_sum {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {G : ι Type u_5} [Π (i : ι), normed_add_comm_group (G i)] [Π (i : ι), inner_product_space 𝕜 (G i)] {V : Π (i : ι), G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 G V) (l : Π (i : ι), G i) (s : finset ι) :
s.sum (λ (i : ι), (V i) (l i)) ^ 2 = s.sum (λ (i : ι), l i ^ 2)
theorem orthogonal_family.comp {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {G : ι Type u_5} [Π (i : ι), normed_add_comm_group (G i)] [Π (i : ι), inner_product_space 𝕜 (G i)] {V : Π (i : ι), G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 G V) {γ : Type u_3} {f : γ ι} (hf : function.injective f) :
orthogonal_family 𝕜 (λ (g : γ), G (f g)) (λ (g : γ), V (f g))

The composition of an orthogonal family of subspaces with an injective function is also an orthogonal family.

theorem orthogonal_family.orthonormal_sigma_orthonormal {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {G : ι Type u_5} [Π (i : ι), normed_add_comm_group (G i)] [Π (i : ι), inner_product_space 𝕜 (G i)] {V : Π (i : ι), G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 G V) {α : ι Type u_3} {v_family : Π (i : ι), α i G i} (hv_family : (i : ι), orthonormal 𝕜 (v_family i)) :
orthonormal 𝕜 (λ (a : Σ (i : ι), α i), (V a.fst) (v_family a.fst a.snd))
theorem orthogonal_family.norm_sq_diff_sum {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} [dec_ι : decidable_eq ι] {G : ι Type u_5} [Π (i : ι), normed_add_comm_group (G i)] [Π (i : ι), inner_product_space 𝕜 (G i)] {V : Π (i : ι), G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 G V) (f : Π (i : ι), G i) (s₁ s₂ : finset ι) :
s₁.sum (λ (i : ι), (V i) (f i)) - s₂.sum (λ (i : ι), (V i) (f i)) ^ 2 = (s₁ \ s₂).sum (λ (i : ι), f i ^ 2) + (s₂ \ s₁).sum (λ (i : ι), f i ^ 2)
theorem orthogonal_family.summable_iff_norm_sq_summable {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {G : ι Type u_5} [Π (i : ι), normed_add_comm_group (G i)] [Π (i : ι), inner_product_space 𝕜 (G i)] {V : Π (i : ι), G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 G V) [complete_space E] (f : Π (i : ι), G i) :
summable (λ (i : ι), (V i) (f i)) summable (λ (i : ι), f i ^ 2)

A family f of mutually-orthogonal elements of E is summable, if and only if (λ i, ‖f i‖ ^ 2) is summable.

theorem orthogonal_family.independent {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} {V : ι submodule 𝕜 E} (hV : orthogonal_family 𝕜 (λ (i : ι), (V i)) (λ (i : ι), (V i).subtypeₗᵢ)) :

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.

theorem direct_sum.is_internal.collected_basis_orthonormal {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {ι : Type u_4} [dec_ι : decidable_eq ι] {V : ι submodule 𝕜 E} (hV : orthogonal_family 𝕜 (λ (i : ι), (V i)) (λ (i : ι), (V i).subtypeₗᵢ)) (hV_sum : direct_sum.is_internal (λ (i : ι), V i)) {α : ι Type u_3} {v_family : Π (i : ι), basis (α i) 𝕜 (V i)} (hv_family : (i : ι), orthonormal 𝕜 (v_family i)) :
orthonormal 𝕜 (hV_sum.collected_basis v_family)

A general inner product implies a real inner product. This is not registered as an instance since it creates problems with the case 𝕜 = ℝ.

Equations

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
theorem real_inner_I_smul_self (𝕜 : Type u_1) {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (x : E) :
@[protected, instance]

A complex inner product implies a real inner product

Equations
@[protected, simp]
theorem complex.inner (w z : ) :

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 #

theorem continuous_inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] :
continuous (λ (p : E × E), has_inner.inner p.fst p.snd)
theorem filter.tendsto.inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {α : Type u_4} {f g : α E} {l : filter α} {x y : E} (hf : filter.tendsto f l (nhds x)) (hg : filter.tendsto g l (nhds y)) :
filter.tendsto (λ (t : α), has_inner.inner (f t) (g t)) l (nhds (has_inner.inner x y))
theorem continuous_within_at.inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {α : Type u_4} [topological_space α] {f g : α E} {x : α} {s : set α} (hf : continuous_within_at f s x) (hg : continuous_within_at g s x) :
continuous_within_at (λ (t : α), has_inner.inner (f t) (g t)) s x
theorem continuous_at.inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {α : Type u_4} [topological_space α] {f g : α E} {x : α} (hf : continuous_at f x) (hg : continuous_at g x) :
continuous_at (λ (t : α), has_inner.inner (f t) (g t)) x
theorem continuous_on.inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {α : Type u_4} [topological_space α] {f g : α E} {s : set α} (hf : continuous_on f s) (hg : continuous_on g s) :
continuous_on (λ (t : α), has_inner.inner (f t) (g t)) s
@[continuity]
theorem continuous.inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {α : Type u_4} [topological_space α] {f g : α E} (hf : continuous f) (hg : continuous g) :
continuous (λ (t : α), has_inner.inner (f t) (g t))
def continuous_linear_map.re_apply_inner_self {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (T : E →L[𝕜] E) (x : E) :

Extract a real bilinear form from an operator T, by taking the pairing λ x, re ⟪T x, x⟫.

Equations
theorem continuous_linear_map.re_apply_inner_self_smul {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] (T : E →L[𝕜] E) (x : E) {c : 𝕜} :
@[protected, instance]
noncomputable def uniform_space.completion.has_inner {𝕜' : Type u_1} {E' : Type u_2} [topological_space 𝕜'] [uniform_space E'] [has_inner 𝕜' E'] :
Equations
@[simp]
@[protected]
theorem uniform_space.completion.continuous.inner {𝕜 : Type u_1} {E : Type u_2} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {α : Type u_3} [topological_space α] {f g : α uniform_space.completion E} (hf : continuous f) (hg : continuous g) :
continuous (λ (x : α), has_inner.inner (f x) (g x))