Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 20981be

Browse files
feat(linear_algebra/charpoly): add linear_map.charpoly (#9279)
We add `linear_map.charpoly`, the characteristic polynomial of an endomorphism of a finite free module, and a basic API.
1 parent 5b3b71a commit 20981be

File tree

5 files changed

+160
-15
lines changed

5 files changed

+160
-15
lines changed

src/linear_algebra/charpoly.lean

Lines changed: 138 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,138 @@
1+
/-
2+
Copyright (c) 2021 Riccardo Brasca. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Riccardo Brasca
5+
-/
6+
7+
import linear_algebra.free_module
8+
import linear_algebra.charpoly.coeff
9+
import linear_algebra.matrix.basis
10+
11+
/-!
12+
13+
# Characteristic polynomial
14+
15+
We define the characteristic polynomial of `f : M →ₗ[R] M`, where `M` is a finite and
16+
free `R`-module.
17+
18+
## Main definition
19+
20+
* `linear_map.charpoly f` : the characteristic polynomial of `f : M →ₗ[R] M`.
21+
22+
-/
23+
24+
universes u v w
25+
26+
variables {R : Type u} {M : Type v} [comm_ring R] [nontrivial R]
27+
variables [add_comm_group M] [module R M] [module.free R M] [module.finite R M] (f : M →ₗ[R] M)
28+
29+
open_locale classical matrix
30+
31+
noncomputable theory
32+
33+
open module.free polynomial matrix
34+
35+
namespace linear_map
36+
37+
section basic
38+
39+
/-- The characteristic polynomial of `f : M →ₗ[R] M`. -/
40+
def charpoly : polynomial R :=
41+
(to_matrix (choose_basis R M) (choose_basis R M) f).charpoly
42+
43+
lemma charpoly_def :
44+
f.charpoly = (to_matrix (choose_basis R M) (choose_basis R M) f).charpoly := rfl
45+
46+
/-- `charpoly f` is the characteristic polynomial of the matrix of `f` in any basis. -/
47+
@[simp] lemma charpoly_to_matrix {ι : Type w} [fintype ι] (b : basis ι R M) :
48+
(to_matrix b b f).charpoly = f.charpoly :=
49+
begin
50+
set A := to_matrix b b f,
51+
set b' := choose_basis R M,
52+
set ι' := choose_basis_index R M,
53+
set A' := to_matrix b' b' f,
54+
set e := basis.index_equiv b b',
55+
set φ := reindex_linear_equiv R R e e,
56+
set φ₁ := reindex_linear_equiv R R e (equiv.refl ι'),
57+
set φ₂ := reindex_linear_equiv R R (equiv.refl ι') (equiv.refl ι'),
58+
set φ₃ := reindex_linear_equiv R R (equiv.refl ι') e,
59+
set P := b.to_matrix b',
60+
set Q := b'.to_matrix b,
61+
62+
have hPQ : C.map_matrix (φ₁ P) ⬝ (C.map_matrix (φ₃ Q)) = 1,
63+
{ rw [ring_hom.map_matrix_apply, ring_hom.map_matrix_apply, ← matrix.map_mul, ←
64+
@reindex_linear_equiv_mul _ ι' _ _ _ _ R R, basis.to_matrix_mul_to_matrix_flip,
65+
reindex_linear_equiv_one, ← ring_hom.map_matrix_apply, ring_hom.map_one] },
66+
67+
calc A.charpoly = (reindex e e A).charpoly : (charpoly_reindex _ _).symm
68+
... = (scalar ι' X - C.map_matrix (φ A)).det : rfl
69+
... = (scalar ι' X - C.map_matrix (φ (P ⬝ A' ⬝ Q))).det :
70+
by rw [basis_to_matrix_mul_linear_map_to_matrix_mul_basis_to_matrix]
71+
... = (scalar ι' X - C.map_matrix (φ₁ P ⬝ φ₂ A' ⬝ φ₃ Q)).det :
72+
by rw [reindex_linear_equiv_mul R R _ _ e, reindex_linear_equiv_mul R R e _ _]
73+
... = (scalar ι' X - (C.map_matrix (φ₁ P) ⬝ C.map_matrix A' ⬝ C.map_matrix (φ₃ Q))).det : by simp
74+
... = (scalar ι' X ⬝ C.map_matrix (φ₁ P) ⬝ (C.map_matrix (φ₃ Q)) -
75+
(C.map_matrix (φ₁ P) ⬝ C.map_matrix A' ⬝ C.map_matrix (φ₃ Q))).det :
76+
by { rw [matrix.mul_assoc ((scalar ι') X), hPQ, matrix.mul_one] }
77+
... = (C.map_matrix (φ₁ P) ⬝ scalar ι' X ⬝ (C.map_matrix (φ₃ Q)) -
78+
(C.map_matrix (φ₁ P) ⬝ C.map_matrix A' ⬝ C.map_matrix (φ₃ Q))).det : by simp
79+
... = (C.map_matrix (φ₁ P) ⬝ (scalar ι' X - C.map_matrix A') ⬝ C.map_matrix (φ₃ Q)).det :
80+
by rw [← matrix.sub_mul, ← matrix.mul_sub]
81+
... = (C.map_matrix (φ₁ P)).det * (scalar ι' X - C.map_matrix A').det *
82+
(C.map_matrix (φ₃ Q)).det : by rw [det_mul, det_mul]
83+
... = (C.map_matrix (φ₁ P)).det * (C.map_matrix (φ₃ Q)).det *
84+
(scalar ι' X - C.map_matrix A').det : by ring
85+
... = (scalar ι' X - C.map_matrix A').det : by rw [← det_mul, hPQ, det_one, one_mul]
86+
... = f.charpoly : rfl
87+
end
88+
89+
end basic
90+
91+
section coeff
92+
93+
lemma charpoly_monic : f.charpoly.monic := charpoly_monic _
94+
95+
end coeff
96+
97+
section cayley_hamilton
98+
99+
/-- The Cayley-Hamilton Theorem, that the characteristic polynomial of a linear map, applied to
100+
the linear map itself, is zero. -/
101+
lemma aeval_self_charpoly : aeval f f.charpoly = 0 :=
102+
begin
103+
apply (linear_equiv.map_eq_zero_iff (alg_equiv_matrix _).to_linear_equiv).1,
104+
rw [alg_equiv.to_linear_equiv_apply, ← alg_equiv.coe_alg_hom,
105+
← polynomial.aeval_alg_hom_apply _ _ _, charpoly_def],
106+
exact aeval_self_charpoly _,
107+
end
108+
109+
lemma is_integral : is_integral R f := ⟨f.charpoly, ⟨charpoly_monic f, aeval_self_charpoly f⟩⟩
110+
111+
lemma minpoly_dvd_charpoly {K : Type u} {M : Type v} [field K] [add_comm_group M] [module K M]
112+
[finite_dimensional K M] (f : M →ₗ[K] M) : minpoly K f ∣ f.charpoly :=
113+
minpoly.dvd _ _ (aeval_self_charpoly f)
114+
115+
variable {f}
116+
117+
lemma charpoly_coeff_zero_of_injective (hf : function.injective f) : (minpoly R f).coeff 00 :=
118+
begin
119+
intro h,
120+
obtain ⟨P, hP⟩ := X_dvd_iff.2 h,
121+
have hdegP : P.degree < (minpoly R f).degree,
122+
{ rw [hP, mul_comm],
123+
refine degree_lt_degree_mul_X (λ h, _),
124+
rw [h, mul_zero] at hP,
125+
exact minpoly.ne_zero (is_integral f) hP },
126+
have hPmonic : P.monic,
127+
{ suffices : (minpoly R f).monic,
128+
{ rwa [monic.def, hP, mul_comm, leading_coeff_mul_X, ← monic.def] at this },
129+
exact minpoly.monic (is_integral f) },
130+
have hzero : aeval f (minpoly R f) = 0 := minpoly.aeval _ _,
131+
simp only [hP, mul_eq_comp, ext_iff, hf, aeval_X, map_eq_zero_iff, coe_comp, alg_hom.map_mul,
132+
zero_apply] at hzero,
133+
exact not_le.2 hdegP (minpoly.min _ _ hPmonic (ext hzero)),
134+
end
135+
136+
end cayley_hamilton
137+
138+
end linear_map

src/linear_algebra/charpoly/basic.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -7,6 +7,7 @@ import tactic.apply_fun
77
import ring_theory.matrix_algebra
88
import ring_theory.polynomial_algebra
99
import linear_algebra.matrix.nonsingular_inverse
10+
import linear_algebra.matrix.reindex
1011
import tactic.squeeze
1112

1213
/-!
@@ -67,12 +68,27 @@ begin
6768
split_ifs; simp [h], }
6869
end
6970

71+
lemma charmatrix_reindex {m : Type v} [decidable_eq m] [fintype m] (e : n ≃ m)
72+
(M : matrix n n R) : charmatrix (reindex e e M) = reindex e e (charmatrix M) :=
73+
begin
74+
ext i j x,
75+
by_cases h : i = j,
76+
all_goals { simp [h] }
77+
end
78+
7079
/--
7180
The characteristic polynomial of a matrix `M` is given by $\det (t I - M)$.
7281
-/
7382
def matrix.charpoly (M : matrix n n R) : polynomial R :=
7483
(charmatrix M).det
7584

85+
lemma matrix.charpoly_reindex {m : Type v} [decidable_eq m] [fintype m] (e : n ≃ m)
86+
(M : matrix n n R) : (reindex e e M).charpoly = M.charpoly :=
87+
begin
88+
unfold matrix.charpoly,
89+
rw [charmatrix_reindex, matrix.det_reindex_self]
90+
end
91+
7692
/--
7793
The **Cayley-Hamilton Theorem**, that the characteristic polynomial of a matrix,
7894
applied to the matrix itself, is zero.

src/linear_algebra/dimension.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -462,6 +462,11 @@ begin
462462
exact le_antisymm w₁ w₂, }
463463
end
464464

465+
/-- Given two basis indexed by `ι` and `ι'` of an `R`-module, where `R` satisfies the invariant
466+
basis number property, an equiv `ι ≃ ι' `. -/
467+
def basis.index_equiv (v : basis ι R M) (v' : basis ι' R M) : ι ≃ ι' :=
468+
nonempty.some (cardinal.lift_mk_eq.1 (cardinal.lift_max.2 (mk_eq_mk_of_basis v v')))
469+
465470
theorem mk_eq_mk_of_basis' {ι' : Type w} (v : basis ι R M) (v' : basis ι' R M) :
466471
#ι = #ι' :=
467472
cardinal.lift_inj.1 $ mk_eq_mk_of_basis v v'

src/linear_algebra/eigenspace.lean

Lines changed: 1 addition & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@ import field_theory.is_alg_closed.basic
88
import linear_algebra.finsupp
99
import linear_algebra.matrix.to_lin
1010
import order.preorder_hom
11+
import linear_algebra.charpoly
1112

1213
/-!
1314
# Eigenvectors and eigenvalues
@@ -130,9 +131,6 @@ end
130131

131132
variables [finite_dimensional K V] (f : End K V)
132133

133-
protected theorem is_integral : is_integral K f :=
134-
is_integral_of_noetherian (by apply_instance) f
135-
136134
variables {f} {μ : K}
137135

138136
theorem has_eigenvalue_of_is_root (h : (minpoly K f).is_root μ) :
@@ -518,7 +516,3 @@ end
518516

519517
end End
520518
end module
521-
variables {K V : Type*} [field K] [add_comm_group V] [module K V] [finite_dimensional K V]
522-
523-
protected lemma linear_map.is_integral (f : V →ₗ[K] V) : is_integral K f :=
524-
module.End.is_integral f

src/linear_algebra/free_module_pid.lean

Lines changed: 0 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -308,14 +308,6 @@ lemma basis.card_le_card_of_le
308308
b.card_le_card_of_linear_independent
309309
(b'.linear_independent.map' (submodule.of_le hNO) (N.ker_of_le O _))
310310

311-
/-- If we have two bases on the same space, their indices are in bijection. -/
312-
noncomputable def basis.index_equiv {R ι ι' : Type*} [integral_domain R] [module R M]
313-
[fintype ι] [fintype ι'] (b : basis ι R M) (b' : basis ι' R M) :
314-
ι ≃ ι' :=
315-
(fintype.card_eq.mp (le_antisymm
316-
(b'.card_le_card_of_linear_independent b.linear_independent)
317-
(b.card_le_card_of_linear_independent b'.linear_independent))).some
318-
319311
/-- If `N` is a submodule in a free, finitely generated module,
320312
do induction on adjoining a linear independent element to a submodule. -/
321313
def submodule.induction_on_rank [fintype ι] (b : basis ι R M) (P : submodule R M → Sort*)

0 commit comments

Comments
 (0)