Documentation

Mathlib.Analysis.InnerProductSpace.Orthonormal

Orthonormal sets #

This file defines orthonormal sets in inner product spaces.

Main results #

For the existence of orthonormal bases, Hilbert bases, etc., see the file Analysis.InnerProductSpace.projection.

def Orthonormal (๐•œ : Type u_1) {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} (v : ฮน โ†’ E) :

An orthonormal set of vectors in an InnerProductSpace

Equations
Instances For
    theorem orthonormal_iff_ite {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} [DecidableEq ฮน] {v : ฮน โ†’ E} :
    Orthonormal ๐•œ v โ†” โˆ€ (i j : ฮน), inner (v i) (v j) = if i = j then 1 else 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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] [DecidableEq E] {s : Set E} :
    Orthonormal ๐•œ Subtype.val โ†” โˆ€ v โˆˆ s, โˆ€ w โˆˆ s, inner v w = if v = w then 1 else 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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (l : ฮน โ†’โ‚€ ๐•œ) (i : ฮน) :
    inner (v i) ((Finsupp.linearCombination ๐•œ 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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (l : ฮน โ†’ ๐•œ) {s : Finset ฮน} {i : ฮน} (hi : i โˆˆ s) :
    inner (v i) (โˆ‘ i โˆˆ s, 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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} [Fintype ฮน] {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (l : ฮน โ†’ ๐•œ) (i : ฮน) :
    inner (v i) (โˆ‘ 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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (l : ฮน โ†’โ‚€ ๐•œ) (i : ฮน) :
    inner ((Finsupp.linearCombination ๐•œ v) l) (v i) = (starRingEnd ๐•œ) (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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (l : ฮน โ†’ ๐•œ) {s : Finset ฮน} {i : ฮน} (hi : i โˆˆ s) :
    inner (โˆ‘ i โˆˆ s, l i โ€ข v i) (v i) = (starRingEnd ๐•œ) (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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} [Fintype ฮน] {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (l : ฮน โ†’ ๐•œ) (i : ฮน) :
    inner (โˆ‘ i : ฮน, l i โ€ข v i) (v i) = (starRingEnd ๐•œ) (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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (lโ‚ lโ‚‚ : ฮน โ†’โ‚€ ๐•œ) :
    inner ((Finsupp.linearCombination ๐•œ v) lโ‚) ((Finsupp.linearCombination ๐•œ v) lโ‚‚) = lโ‚.sum fun (i : ฮน) (y : ๐•œ) => (starRingEnd ๐•œ) 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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (lโ‚ lโ‚‚ : ฮน โ†’โ‚€ ๐•œ) :
    inner ((Finsupp.linearCombination ๐•œ v) lโ‚) ((Finsupp.linearCombination ๐•œ v) lโ‚‚) = lโ‚‚.sum fun (i : ฮน) (y : ๐•œ) => (starRingEnd ๐•œ) (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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (lโ‚ lโ‚‚ : ฮน โ†’ ๐•œ) (s : Finset ฮน) :
    inner (โˆ‘ i โˆˆ s, lโ‚ i โ€ข v i) (โˆ‘ i โˆˆ s, lโ‚‚ i โ€ข v i) = โˆ‘ i โˆˆ s, (starRingEnd ๐•œ) (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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {s : Finset ฮน} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) {a : ฮน โ†’ ฮน โ†’ ๐•œ} :
    โˆ‘ i โˆˆ s, โˆ‘ j โˆˆ s, a i j โ€ข inner (v j) (v i) = โˆ‘ k โˆˆ s, 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.linearIndependent {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) :
    LinearIndependent ๐•œ v

    An orthonormal set is linearly independent.

    theorem Orthonormal.comp {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {ฮน' : Type u_5} {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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Function.Injective v) :
    Orthonormal ๐•œ Subtype.val โ†” Orthonormal ๐•œ v

    An injective family v : ฮน โ†’ E is orthonormal if and only if Subtype.val : (range v) โ†’ E is orthonormal.

    theorem Orthonormal.toSubtypeRange {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) :
    Orthonormal ๐•œ Subtype.val

    If v : ฮน โ†’ E is an orthonormal family, then Subtype.val : (range v) โ†’ E is an orthonormal family.

    theorem Orthonormal.inner_finsupp_eq_zero {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) {s : Set ฮน} {i : ฮน} (hi : i โˆ‰ s) {l : ฮน โ†’โ‚€ ๐•œ} (hl : l โˆˆ Finsupp.supported ๐•œ ๐•œ s) :
    inner ((Finsupp.linearCombination ๐•œ 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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v w : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (hw : โˆ€ (i : ฮน), w i = v i โˆจ w i = -v i) :
    Orthonormal ๐•œ w

    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) [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] :
    Orthonormal ๐•œ fun (x : โ†‘โˆ…) => โ†‘x
    theorem orthonormal_iUnion_of_directed {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮท : Type u_5} {s : ฮท โ†’ Set E} (hs : Directed (fun (x1 x2 : Set E) => x1 โŠ† x2) s) (h : โˆ€ (i : ฮท), Orthonormal ๐•œ fun (x : โ†‘(s i)) => โ†‘x) :
    Orthonormal ๐•œ fun (x : โ†‘(โ‹ƒ (i : ฮท), s i)) => โ†‘x
    theorem orthonormal_sUnion_of_directed {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {s : Set (Set E)} (hs : DirectedOn (fun (x1 x2 : Set E) => x1 โŠ† x2) s) (h : โˆ€ a โˆˆ s, Orthonormal ๐•œ fun (x : โ†‘a) => โ†‘x) :
    Orthonormal ๐•œ fun (x : โ†‘(โ‹ƒโ‚€ s)) => โ†‘x
    theorem exists_maximal_orthonormal {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {s : Set E} (hs : Orthonormal ๐•œ Subtype.val) :
    โˆƒ w โŠ‡ s, Orthonormal ๐•œ Subtype.val โˆง โˆ€ u โŠ‡ w, Orthonormal ๐•œ Subtype.val โ†’ u = w

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

    def basisOfOrthonormalOfCardEqFinrank {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} [Fintype ฮน] [Nonempty ฮน] {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (card_eq : Fintype.card ฮน = Module.finrank ๐•œ E) :
    Basis ฮน ๐•œ E

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

    Equations
    Instances For
      @[simp]
      theorem coe_basisOfOrthonormalOfCardEqFinrank {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} [Fintype ฮน] [Nonempty ฮน] {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (card_eq : Fintype.card ฮน = Module.finrank ๐•œ E) :
      โ‡‘(basisOfOrthonormalOfCardEqFinrank hv card_eq) = v
      theorem Orthonormal.ne_zero {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (i : ฮน) :
      v i โ‰  0
      theorem LinearIsometry.orthonormal_comp_iff {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] {v : ฮน โ†’ E} (f : E โ†’โ‚—แตข[๐•œ] E') :
      Orthonormal ๐•œ (โ‡‘f โˆ˜ v) โ†” Orthonormal ๐•œ v

      A linear isometry preserves the property of being orthonormal.

      theorem Orthonormal.comp_linearIsometry {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (f : E โ†’โ‚—แตข[๐•œ] E') :
      Orthonormal ๐•œ (โ‡‘f โˆ˜ v)

      A linear isometry preserves the property of being orthonormal.

      theorem Orthonormal.comp_linearIsometryEquiv {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) (f : E โ‰ƒโ‚—แตข[๐•œ] E') :
      Orthonormal ๐•œ (โ‡‘f โˆ˜ v)

      A linear isometric equivalence preserves the property of being orthonormal.

      theorem Orthonormal.mapLinearIsometryEquiv {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) (f : E โ‰ƒโ‚—แตข[๐•œ] E') :
      Orthonormal ๐•œ โ‡‘(v.map f.toLinearEquiv)

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

      def LinearMap.isometryOfOrthonormal {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] (f : E โ†’โ‚—[๐•œ] E') {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) (hf : Orthonormal ๐•œ (โ‡‘f โˆ˜ โ‡‘v)) :

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

      Equations
      • f.isometryOfOrthonormal hv hf = f.isometryOfInner โ‹ฏ
      Instances For
        @[simp]
        theorem LinearMap.coe_isometryOfOrthonormal {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] (f : E โ†’โ‚—[๐•œ] E') {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) (hf : Orthonormal ๐•œ (โ‡‘f โˆ˜ โ‡‘v)) :
        โ‡‘(f.isometryOfOrthonormal hv hf) = โ‡‘f
        @[simp]
        theorem LinearMap.isometryOfOrthonormal_toLinearMap {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] (f : E โ†’โ‚—[๐•œ] E') {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) (hf : Orthonormal ๐•œ (โ‡‘f โˆ˜ โ‡‘v)) :
        (f.isometryOfOrthonormal hv hf).toLinearMap = f
        def LinearEquiv.isometryOfOrthonormal {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] (f : E โ‰ƒโ‚—[๐•œ] E') {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) (hf : Orthonormal ๐•œ (โ‡‘f โˆ˜ โ‡‘v)) :

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

        Equations
        • f.isometryOfOrthonormal hv hf = f.isometryOfInner โ‹ฏ
        Instances For
          @[simp]
          theorem LinearEquiv.coe_isometryOfOrthonormal {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] (f : E โ‰ƒโ‚—[๐•œ] E') {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) (hf : Orthonormal ๐•œ (โ‡‘f โˆ˜ โ‡‘v)) :
          โ‡‘(f.isometryOfOrthonormal hv hf) = โ‡‘f
          @[simp]
          theorem LinearEquiv.isometryOfOrthonormal_toLinearEquiv {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] (f : E โ‰ƒโ‚—[๐•œ] E') {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) (hf : Orthonormal ๐•œ (โ‡‘f โˆ˜ โ‡‘v)) :
          (f.isometryOfOrthonormal hv hf).toLinearEquiv = f
          def Orthonormal.equiv {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {ฮน' : Type u_5} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) {v' : Basis ฮน' ๐•œ E'} (hv' : Orthonormal ๐•œ โ‡‘v') (e : ฮน โ‰ƒ ฮน') :

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

          Equations
          • hv.equiv hv' e = (v.equiv v' e).isometryOfOrthonormal hv โ‹ฏ
          Instances For
            @[simp]
            theorem Orthonormal.equiv_toLinearEquiv {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {ฮน' : Type u_5} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) {v' : Basis ฮน' ๐•œ E'} (hv' : Orthonormal ๐•œ โ‡‘v') (e : ฮน โ‰ƒ ฮน') :
            (hv.equiv hv' e).toLinearEquiv = v.equiv v' e
            @[simp]
            theorem Orthonormal.equiv_apply {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] {ฮน' : Type u_9} {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_trans {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {ฮน' : Type u_5} {ฮน'' : Type u_6} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] {E'' : Type u_8} [SeminormedAddCommGroup E''] [InnerProductSpace ๐•œ 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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {ฮน' : Type u_5} {E' : Type u_7} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) {v' : Basis ฮน' ๐•œ E'} (hv' : Orthonormal ๐•œ โ‡‘v') (e : ฮน โ‰ƒ ฮน') :
            v.map (hv.equiv hv' e).toLinearEquiv = v'.reindex e.symm
            @[simp]
            theorem Orthonormal.equiv_refl {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) :
            hv.equiv hv (Equiv.refl ฮน) = LinearIsometryEquiv.refl ๐•œ E
            @[simp]
            theorem Orthonormal.equiv_symm {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} {ฮน' : Type u_5} {E' : Type u_6} [SeminormedAddCommGroup E'] [InnerProductSpace ๐•œ E'] {v : Basis ฮน ๐•œ E} (hv : Orthonormal ๐•œ โ‡‘v) {v' : Basis ฮน' ๐•œ E'} (hv' : Orthonormal ๐•œ โ‡‘v') (e : ฮน โ‰ƒ ฮน') :
            (hv.equiv hv' e).symm = hv'.equiv hv e.symm
            theorem Orthonormal.sum_inner_products_le {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} (x : E) {v : ฮน โ†’ E} {s : Finset ฮน} (hv : Orthonormal ๐•œ v) :
            โˆ‘ i โˆˆ s, โ€–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} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} (x : E) {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) :
            โˆ‘' (i : ฮน), โ€–inner (v i) xโ€– ^ 2 โ‰ค โ€–xโ€– ^ 2

            Bessel's inequality.

            theorem Orthonormal.inner_products_summable {๐•œ : Type u_1} {E : Type u_2} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] {ฮน : Type u_4} (x : E) {v : ฮน โ†’ E} (hv : Orthonormal ๐•œ v) :
            Summable fun (i : ฮน) => โ€–inner (v i) xโ€– ^ 2

            The sum defined in Bessel's inequality is summable.