refactor: add module docstrings, add copyright headers, and fix other typos

This commit is contained in:
kuotsanhsu 2025-01-12 19:40:15 +08:00
parent 4fa4a28d5d
commit 6d47ea18a0
2 changed files with 116 additions and 1 deletions

View file

@ -1,7 +1,29 @@
/-
Copyright (c) 2025 Gordon Hsu. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Gordon Hsu
-/
import Mathlib.LinearAlgebra.Eigenspace.Triangularizable
import Mathlib.LinearAlgebra.Matrix.Spectrum
open scoped InnerProductSpace
/-! # Schur triangulation
Schur triangulation is more commonly known as Schur decomposition or Schur triangularization, but
"triangulation" makes the API more readable. It states that a square matrix over an algebraically
closed field, e.g., ``, is unitarily similar to an upper triangular matrix.
## Main definitions
- `Matrix.schur_triangulation` : a matrix `A : Matrix n n 𝕜` with `𝕜` being algebraically closed can
be decomposed as `A = U * T * star U` where `U` is unitary and `T` is upper triangular.
- `Matrix.schurTriangulationUnitary` : the unitary matrix `U` as previously stated.
- `Matrix.schurTriangulation` : the upper triangular matrix `T` as previously stated.
- Some auxilary definitions are not meant to be used directly, but
`LinearMap.SchurTriangulationAux.of` contains the main algorithm for the triangulation procedure.
-/
/-- `subNat' i h` subtracts `m` from `i`. This is an alternative form of `Fin.subNat`. -/
@[inline] def Fin.subNat' (i : Fin (m + n)) (h : ¬ i < m) : Fin n :=
subNat m (Fin.cast (m.add_comm n) i) (Nat.ge_of_not_lt h)
@ -97,6 +119,35 @@ structure SchurTriangulationAux (f : Module.End 𝕜 E) where
end
/-! ## Schur's recursive triangulation procedure
Given a linear endomorphism `f` on a non-trivial finite-dimensional vector space `E` over an
algebraically closed field `𝕜`, one can always pick an eigenvalue `μ` of `f` whose corresponding
eigenspace `V` is non-trivial. Given that `E` is also an inner product space, let `bV` and `bW` be
othonormal bases for `V` and `Vᗮ` respectively. Then, the collection of vectors in `bV` and `bW`
forms an othornomal basis `bE` for `E`, as the direct sum of `V` and `Vᗮ` is an internal
decomposition of `E`. The matrix representation of `f` with respect to `bE` satisfies
$$
\sideset{_\mathrm{bE}}{_\mathrm{bE}}{[f]} =
\begin{bmatrix}
\sideset{_\mathrm{bV}}{_\mathrm{bV}}{[f]} &
\sideset{_\mathrm{bW}}{_\mathrm{bV}}{[f]} \\
\sideset{_\mathrm{bV}}{_\mathrm{bW}}{[f]} &
\sideset{_\mathrm{bW}}{_\mathrm{bW}}{[f]}
\end{bmatrix} =
\begin{bmatrix} \mu I & □ \\ 0 & \sideset{_\mathrm{bW}}{_\mathrm{bW}}{[f]} \end{bmatrix},
$$
which is upper triangular as long as $\sideset{_\mathrm{bW}}{_\mathrm{bW}}{[f]}$ is. Finally, one
observes that the recursion from $\sideset{_\mathrm{bE}}{_\mathrm{bE}}{[f]}$ to
$\sideset{_\mathrm{bW}}{_\mathrm{bW}}{[f]}$ is well-founded, as the dimension of `bW` is smaller
than that of `bE` because `bV` is non-trivial.
However, in order to leverage `DirectSum.IsInternal.collectedOrthonormalBasis`, the type
`Σ b, cond b (Fin m) (Fin n)` has to be used instead of the more natural `Fin m ⊕ Fin n` while their
equivalence is propositionally established by `Equiv.sumEquivSigmalCond`.
-/
variable [IsAlgClosed 𝕜]
/-- **Don't use this definition directly.** This is the key algorithm behind