Documentation

Mathlib.Analysis.InnerProductSpace.GramMatrix

Gram Matrices #

This file defines Gram matrices and proves their positive semidefiniteness. Results require RCLike ๐•œ.

Main definition #

Main results #

def Matrix.gram {E : Type u_1} {n : Type u_2} (๐•œ : Type u_5) [Inner ๐•œ E] (v : n โ†’ E) :
Matrix n n ๐•œ

The entries of a Gram matrix are inner products of vectors in an inner product space.

Equations
Instances For
    @[simp]
    theorem Matrix.gram_apply {E : Type u_1} {n : Type u_2} {๐•œ : Type u_4} [Inner ๐•œ E] (v : n โ†’ E) (i j : n) :
    gram ๐•œ v i j = inner ๐•œ (v i) (v j)
    @[simp]
    theorem Matrix.gram_zero {E : Type u_1} {n : Type u_2} {๐•œ : Type u_4} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] :
    gram ๐•œ 0 = 0
    @[simp]
    theorem Matrix.gram_single {E : Type u_1} {n : Type u_2} {๐•œ : Type u_4} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] [DecidableEq n] (i : n) (x : E) :
    gram ๐•œ (Pi.single i x) = single i i (inner ๐•œ x x)
    theorem Matrix.submatrix_gram {E : Type u_1} {n : Type u_2} {๐•œ : Type u_4} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] (v : n โ†’ E) {m : Set n} (f : โ†‘m โ†’ n) :
    (gram ๐•œ v).submatrix f f = gram ๐•œ (v โˆ˜ f)
    theorem Matrix.isHermitian_gram {E : Type u_1} {n : Type u_2} (๐•œ : Type u_4) [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] (v : n โ†’ E) :
    (gram ๐•œ v).IsHermitian

    A Gram matrix is Hermitian.

    theorem Matrix.star_dotProduct_gram_mulVec {E : Type u_1} {n : Type u_2} {๐•œ : Type u_4} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] [Fintype n] (v : n โ†’ E) (x y : n โ†’ ๐•œ) :
    star x โฌแตฅ (gram ๐•œ v).mulVec y = inner ๐•œ (โˆ‘ i : n, x i โ€ข v i) (โˆ‘ i : n, y i โ€ข v i)
    theorem Matrix.posSemidef_gram {E : Type u_1} {n : Type u_2} (๐•œ : Type u_4) [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] [Fintype n] (v : n โ†’ E) :
    (gram ๐•œ v).PosSemidef

    A Gram matrix is positive semidefinite.

    theorem Matrix.linearIndependent_of_posDef_gram {E : Type u_1} {n : Type u_2} {๐•œ : Type u_4} [RCLike ๐•œ] [SeminormedAddCommGroup E] [InnerProductSpace ๐•œ E] [Fintype n] {v : n โ†’ E} (h_gram : (gram ๐•œ v).PosDef) :
    LinearIndependent ๐•œ v

    In a normed space, positive definiteness of gram ๐•œ v implies linear independence of v.

    theorem Matrix.posDef_gram_of_linearIndependent {E : Type u_1} {n : Type u_2} {๐•œ : Type u_4} [RCLike ๐•œ] [NormedAddCommGroup E] [InnerProductSpace ๐•œ E] [Fintype n] {v : n โ†’ E} (h_li : LinearIndependent ๐•œ v) :
    (gram ๐•œ v).PosDef

    In a normed space, linear independence of v implies positive definiteness of gram ๐•œ v.

    theorem Matrix.posDef_gram_iff_linearIndependent {E : Type u_1} {n : Type u_2} {๐•œ : Type u_4} [RCLike ๐•œ] [NormedAddCommGroup E] [InnerProductSpace ๐•œ E] [Fintype n] {v : n โ†’ E} :
    (gram ๐•œ v).PosDef โ†” LinearIndependent ๐•œ v

    In a normed space, linear independence of v is equivalent to positive definiteness of gram ๐•œ v.