|
| 1 | +/- |
| 2 | +Copyright (c) 2020 Alexander Bentkamp. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Alexander Bentkamp |
| 5 | +
|
| 6 | +! This file was ported from Lean 3 source module linear_algebra.eigenspace.minpoly |
| 7 | +! leanprover-community/mathlib commit 6b0169218d01f2837d79ea2784882009a0da1aa1 |
| 8 | +! Please do not edit these lines, except to modify the commit id |
| 9 | +! if you have ported upstream changes. |
| 10 | +-/ |
| 11 | +import Mathlib.LinearAlgebra.Eigenspace.Basic |
| 12 | +import Mathlib.FieldTheory.Minpoly.Field |
| 13 | + |
| 14 | +/-! |
| 15 | +# Eigenvalues are the roots of the minimal polynomial. |
| 16 | +
|
| 17 | +## Tags |
| 18 | +
|
| 19 | +eigenvalue, minimal polynomial |
| 20 | +-/ |
| 21 | + |
| 22 | + |
| 23 | +universe u v w |
| 24 | + |
| 25 | +namespace Module |
| 26 | + |
| 27 | +namespace End |
| 28 | + |
| 29 | +open Polynomial FiniteDimensional |
| 30 | + |
| 31 | +open scoped Polynomial |
| 32 | + |
| 33 | +variable {K : Type v} {V : Type w} [Field K] [AddCommGroup V] [Module K V] |
| 34 | + |
| 35 | +theorem eigenspace_aeval_polynomial_degree_1 (f : End K V) (q : K[X]) (hq : degree q = 1) : |
| 36 | + eigenspace f (-q.coeff 0 / q.leadingCoeff) = LinearMap.ker (aeval f q) := |
| 37 | + calc |
| 38 | + eigenspace f (-q.coeff 0 / q.leadingCoeff) |
| 39 | + _ = LinearMap.ker (q.leadingCoeff • f - algebraMap K (End K V) (-q.coeff 0)) := by |
| 40 | + rw [eigenspace_div] |
| 41 | + intro h |
| 42 | + rw [leadingCoeff_eq_zero_iff_deg_eq_bot.1 h] at hq |
| 43 | + cases hq |
| 44 | + _ = LinearMap.ker (aeval f (C q.leadingCoeff * X + C (q.coeff 0))) := by |
| 45 | + rw [C_mul', aeval_def]; simp [algebraMap, Algebra.toRingHom] |
| 46 | + _ = LinearMap.ker (aeval f q) := by rwa [← eq_X_add_C_of_degree_eq_one] |
| 47 | +#align module.End.eigenspace_aeval_polynomial_degree_1 Module.End.eigenspace_aeval_polynomial_degree_1 |
| 48 | + |
| 49 | +theorem ker_aeval_ring_hom'_unit_polynomial (f : End K V) (c : K[X]ˣ) : |
| 50 | + LinearMap.ker (aeval f (c : K[X])) = ⊥ := by |
| 51 | + rw [Polynomial.eq_C_of_degree_eq_zero (degree_coe_units c)] |
| 52 | + simp only [aeval_def, eval₂_C] |
| 53 | + apply ker_algebraMap_end |
| 54 | + apply coeff_coe_units_zero_ne_zero c |
| 55 | +#align module.End.ker_aeval_ring_hom'_unit_polynomial Module.End.ker_aeval_ring_hom'_unit_polynomial |
| 56 | + |
| 57 | +theorem aeval_apply_of_hasEigenvector {f : End K V} {p : K[X]} {μ : K} {x : V} |
| 58 | + (h : f.HasEigenvector μ x) : aeval f p x = p.eval μ • x := by |
| 59 | + refine' p.induction_on _ _ _ |
| 60 | + · intro a; simp [Module.algebraMap_end_apply] |
| 61 | + · intro p q hp hq; simp [hp, hq, add_smul] |
| 62 | + · intro n a hna |
| 63 | + rw [mul_comm, pow_succ, mul_assoc, AlgHom.map_mul, LinearMap.mul_apply, mul_comm, hna] |
| 64 | + simp only [mem_eigenspace_iff.1 h.1, smul_smul, aeval_X, eval_mul, eval_C, eval_pow, eval_X, |
| 65 | + LinearMap.map_smulₛₗ, RingHom.id_apply, mul_comm] |
| 66 | +#align module.End.aeval_apply_of_has_eigenvector Module.End.aeval_apply_of_hasEigenvector |
| 67 | + |
| 68 | +theorem isRoot_of_hasEigenvalue {f : End K V} {μ : K} (h : f.HasEigenvalue μ) : |
| 69 | + (minpoly K f).IsRoot μ := by |
| 70 | + rcases(Submodule.ne_bot_iff _).1 h with ⟨w, ⟨H, ne0⟩⟩ |
| 71 | + refine' Or.resolve_right (smul_eq_zero.1 _) ne0 |
| 72 | + simp [← aeval_apply_of_hasEigenvector ⟨H, ne0⟩, minpoly.aeval K f] |
| 73 | +#align module.End.is_root_of_has_eigenvalue Module.End.isRoot_of_hasEigenvalue |
| 74 | + |
| 75 | +variable [FiniteDimensional K V] (f : End K V) |
| 76 | + |
| 77 | +variable {f} {μ : K} |
| 78 | + |
| 79 | +theorem hasEigenvalue_of_isRoot (h : (minpoly K f).IsRoot μ) : f.HasEigenvalue μ := by |
| 80 | + cases' dvd_iff_isRoot.2 h with p hp |
| 81 | + rw [HasEigenvalue, eigenspace] |
| 82 | + intro con |
| 83 | + cases' (LinearMap.isUnit_iff_ker_eq_bot _).2 con with u hu |
| 84 | + have p_ne_0 : p ≠ 0 := by |
| 85 | + intro con |
| 86 | + apply minpoly.ne_zero f.isIntegral |
| 87 | + rw [hp, con, MulZeroClass.mul_zero] |
| 88 | + have : (aeval f) p = 0 := by |
| 89 | + have h_aeval := minpoly.aeval K f |
| 90 | + revert h_aeval |
| 91 | + simp [hp, ← hu] |
| 92 | + have h_deg := minpoly.degree_le_of_ne_zero K f p_ne_0 this |
| 93 | + rw [hp, degree_mul, degree_X_sub_C, Polynomial.degree_eq_natDegree p_ne_0] at h_deg |
| 94 | + norm_cast at h_deg |
| 95 | + linarith |
| 96 | +#align module.End.has_eigenvalue_of_is_root Module.End.hasEigenvalue_of_isRoot |
| 97 | + |
| 98 | +theorem hasEigenvalue_iff_isRoot : f.HasEigenvalue μ ↔ (minpoly K f).IsRoot μ := |
| 99 | + ⟨isRoot_of_hasEigenvalue, hasEigenvalue_of_isRoot⟩ |
| 100 | +#align module.End.has_eigenvalue_iff_is_root Module.End.hasEigenvalue_iff_isRoot |
| 101 | + |
| 102 | +/-- An endomorphism of a finite-dimensional vector space has finitely many eigenvalues. -/ |
| 103 | +noncomputable instance (f : End K V) : Fintype f.Eigenvalues := |
| 104 | + -- Porting note: added `show` to avoid unfolding `Set K` to `K → Prop` |
| 105 | + show Fintype { μ : K | f.HasEigenvalue μ } from |
| 106 | + Set.Finite.fintype |
| 107 | + (by |
| 108 | + have h : minpoly K f ≠ 0 := minpoly.ne_zero f.isIntegral |
| 109 | + convert (minpoly K f).rootSet_finite K |
| 110 | + ext μ |
| 111 | + -- Porting note: was |
| 112 | + -- have : μ ∈ {μ : K | f.eigenspace μ = ⊥ → False} ↔ ¬f.eigenspace μ = ⊥ := by tauto |
| 113 | + -- convert rfl.mpr this |
| 114 | + -- simp only [Polynomial.rootSet_def, Polynomial.mem_roots h, ← hasEigenvalue_iff_isRoot, |
| 115 | + -- HasEigenvalue] |
| 116 | + -- which didn't work, but worked with |
| 117 | + -- simp only [Polynomial.rootSet_def, Polynomial.mem_roots h, ← hasEigenvalue_iff_isRoot, |
| 118 | + -- HasEigenvalue, (Multiset.mem_toFinset), Algebra.id.map_eq_id, iff_self, Ne.def, |
| 119 | + -- Polynomial.map_id, Finset.mem_coe] |
| 120 | + -- but the code below is simpler. |
| 121 | + rw [Set.mem_setOf_eq, hasEigenvalue_iff_isRoot, mem_rootSet_of_ne h, IsRoot, |
| 122 | + coe_aeval_eq_eval]) |
| 123 | + |
| 124 | +end End |
| 125 | + |
| 126 | +end Module |
0 commit comments