Documentation

Mathlib.LinearAlgebra.Matrix.Charpoly.Coeff

Characteristic polynomials #

We give methods for computing coefficients of the characteristic polynomial.

Main definitions #

theorem charmatrix_apply_natDegree {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] {M : Matrix n n R} [Nontrivial R] (i : n) (j : n) :
Polynomial.natDegree (charmatrix M i j) = if i = j then 1 else 0
theorem charmatrix_apply_natDegree_le {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] {M : Matrix n n R} (i : n) (j : n) :
Polynomial.natDegree (charmatrix M i j) if i = j then 1 else 0
theorem Matrix.charpoly_sub_diagonal_degree_lt {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] (M : Matrix n n R) :
Polynomial.degree (Matrix.charpoly M - Finset.prod Finset.univ fun i => Polynomial.X - Polynomial.C (M i i)) < ↑(Fintype.card n - 1)
theorem Matrix.charpoly_coeff_eq_prod_coeff_of_le {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] (M : Matrix n n R) {k : } (h : Fintype.card n - 1 k) :
Polynomial.coeff (Matrix.charpoly M) k = Polynomial.coeff (Finset.prod Finset.univ fun i => Polynomial.X - Polynomial.C (M i i)) k
theorem Matrix.det_of_card_zero {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] (h : Fintype.card n = 0) (M : Matrix n n R) :
theorem Matrix.matPolyEquiv_eval {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] (M : Matrix n n (Polynomial R)) (r : R) (i : n) (j : n) :
Polynomial.eval ((fun x => Matrix n n R) r) Matrix.semiring (↑(Matrix.scalar n) r) (matPolyEquiv M) i j = Polynomial.eval r (M i j)
theorem Matrix.eval_det {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] (M : Matrix n n (Polynomial R)) (r : R) :
Polynomial.eval r (Matrix.det M) = Matrix.det (Polynomial.eval (↑(Matrix.scalar n) r) (matPolyEquiv M))
theorem matPolyEquiv_eq_x_pow_sub_c {n : Type v} [DecidableEq n] [Fintype n] {K : Type u_1} (k : ) [Field K] (M : Matrix n n K) :
matPolyEquiv (↑(RingHom.mapMatrix ↑(Polynomial.expand K k)) (charmatrix (M ^ k))) = Polynomial.X ^ k - Polynomial.C (M ^ k)
theorem Matrix.aeval_eq_aeval_mod_charpoly {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] (M : Matrix n n R) (p : Polynomial R) :

Any matrix polynomial p is equivalent under evaluation to p %ₘ M.charpoly; that is, p is equivalent to a polynomial with degree less than the dimension of the matrix.

theorem Matrix.pow_eq_aeval_mod_charpoly {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] (M : Matrix n n R) (k : ) :
M ^ k = ↑(Polynomial.aeval M) (Polynomial.X ^ k %ₘ Matrix.charpoly M)

Any matrix power can be computed as the sum of matrix powers less than Fintype.card n.

TODO: add the statement for negative powers phrased with zpow.

theorem coeff_charpoly_mem_ideal_pow {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] {M : Matrix n n R} {I : Ideal R} (h : ∀ (i j : n), M i j I) (k : ) :
def Matrix.charpolyRev {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] (M : Matrix n n R) :

The reverse of the characteristic polynomial of a matrix.

It has some advantages over the characteristic polynomial, including the fact that it can be extended to infinite dimensions (for appropriate operators). In such settings it is known as the "characteristic power series".

Equations
Instances For
    @[simp]
    theorem Matrix.eval_charpolyRev {R : Type u} [CommRing R] {n : Type v} [DecidableEq n] [Fintype n] {M : Matrix n n R} :