refactor: Lint style
This commit is contained in:
parent
1053ccaa3a
commit
6a3bb431bf
4 changed files with 73 additions and 29 deletions
|
@ -6,11 +6,17 @@ namespace Lorentz
|
|||
open scoped Matrix
|
||||
open scoped ComplexConjugate
|
||||
|
||||
notation "ℂ²ˣ²" => Matrix (Fin 2) (Fin 2) ℂ
|
||||
/-- A notation for the type of complex 2-by-2 matrices. It would have been better to make it an
|
||||
abbreviation if it wasn't for Lean's inability to recognize `ℂ²ˣ²` as an identifier. -/
|
||||
scoped notation "ℂ²ˣ²" => Matrix (Fin 2) (Fin 2) ℂ
|
||||
|
||||
/-- A convenient abbreviation for the type of self-adjoint complex 2-by-2 matrices. -/
|
||||
noncomputable abbrev ℍ₂ := selfAdjoint ℂ²ˣ²
|
||||
|
||||
namespace SL2C
|
||||
|
||||
/-- Definitially equal to `Lorentz.SL2C.toSelfAdjointMap` but dropping the requirement that `M` be
|
||||
special linear. -/
|
||||
noncomputable def toSelfAdjointMap' (M : ℂ²ˣ²) : ℍ₂ →ₗ[ℝ] ℍ₂ where
|
||||
toFun | ⟨A, hA⟩ => ⟨M * A * Mᴴ, hA.conjugate M⟩
|
||||
map_add' | ⟨A, _⟩, ⟨B, _⟩ => Subtype.ext <|
|
||||
|
@ -47,12 +53,12 @@ theorem toSelfAdjointMap_det_one' {M : ℂ²ˣ²} (hM : M.IsUpperTriangular) (de
|
|||
Matrix.ext fun | 0, 0 | 1, 0 | 1, 1 => rfl | 0, 1 => by simp [hM10]
|
||||
|
||||
have detA_one : normSq x * normSq y = 1 := congrArg Complex.re <|
|
||||
calc ↑(normSq x * normSq y)
|
||||
calc ↑(normSq x * normSq y)
|
||||
_ = x * conj x * (y * conj y) := by simp [Complex.mul_conj]
|
||||
_ = x * y * (conj y * conj x) := by ring
|
||||
_ = x * y * conj (x * y) := congrArg _ (star_mul ..).symm
|
||||
_ = 1 := suffices x * y = 1 by simp [this]
|
||||
calc x * y
|
||||
calc x * y
|
||||
_ = !![x, _; 0, y].det := by simp
|
||||
_ = M.det := congrArg _ he.symm
|
||||
_ = 1 := detM
|
||||
|
@ -67,7 +73,7 @@ theorem toSelfAdjointMap_det_one' {M : ℂ²ˣ²} (hM : M.IsUpperTriangular) (de
|
|||
have hD : D = !![z.re, -z.im; z.im, z.re] := Matrix.ext fun
|
||||
| 0, 0 => congrArg Complex.re k₀ | 1, 0 => congrArg Complex.im k₀
|
||||
| 0, 1 => congrArg Complex.re k₁ | 1, 1 => congrArg Complex.im k₁
|
||||
calc D.det
|
||||
calc D.det
|
||||
_ = normSq z := by simp [hD, z.normSq_apply]
|
||||
_ = normSq x * normSq y := by simp [x.normSq_mul]
|
||||
_ = 1 := detA_one
|
||||
|
@ -95,23 +101,26 @@ theorem toSelfAdjointMap_det_one' {M : ℂ²ˣ²} (hM : M.IsUpperTriangular) (de
|
|||
show (M * E₂ * Mᴴ) 1 1 = 0 by rw [he', he] ; simp [E₂]
|
||||
have hB11 : B 1 1 = 0 := congrArg Complex.re <|
|
||||
show (M * E₃ * Mᴴ) 1 1 = 0 by rw [he', he] ; simp [E₃]
|
||||
calc A 1 1 - (B * ⅟D * C) 1 1
|
||||
calc A 1 1 - (B * ⅟D * C) 1 1
|
||||
_ = A 1 1 - B 1 ⬝ᵥ ((⅟D * C) · 1) := by noncomm_ring
|
||||
_ = normSq y := by simp [hB10, hB11, hA11]
|
||||
calc LinearMap.det (toSelfAdjointMap' M)
|
||||
calc LinearMap.det (toSelfAdjointMap' M)
|
||||
_ = F.det := (LinearMap.det_toMatrix ..).symm
|
||||
_ = D.det * (A - B * ⅟D * C).det := F.fromBlocks_toBlocks ▸ Matrix.det_fromBlocks₂₂ ..
|
||||
_ = 1 := by rw [hE] ; simp [detD_one, detA_one]
|
||||
|
||||
/-- This promotes `Lorentz.SL2C.toSelfAdjointMap M` and its definitional equivalence,
|
||||
`Lorentz.SL2C.toSelfAdjointMap' M`, to a linear equivalence by recognising the linear inverse to be
|
||||
`Lorentz.SL2C.toSelfAdjointMap M⁻¹`, i.e., `Lorentz.SL2C.toSelfAdjointMap' M⁻¹`. -/
|
||||
noncomputable def toSelfAdjointEquiv (M : ℂ²ˣ²) [Invertible M] : ℍ₂ ≃ₗ[ℝ] ℍ₂ where
|
||||
toLinearMap := toSelfAdjointMap' M
|
||||
invFun := toSelfAdjointMap' M⁻¹
|
||||
left_inv | ⟨A, _⟩ => Subtype.ext <|
|
||||
calc M⁻¹ * (M * A * Mᴴ) * M⁻¹ᴴ
|
||||
calc M⁻¹ * (M * A * Mᴴ) * M⁻¹ᴴ
|
||||
_ = M⁻¹ * ↑M * A * (M⁻¹ * M)ᴴ := by noncomm_ring [Matrix.conjTranspose_mul]
|
||||
_ = A := by simp
|
||||
right_inv | ⟨A, _⟩ => Subtype.ext <|
|
||||
calc M * (M⁻¹ * A * M⁻¹ᴴ) * Mᴴ
|
||||
calc M * (M⁻¹ * A * M⁻¹ᴴ) * Mᴴ
|
||||
_ = M * M⁻¹ * A * (M * M⁻¹)ᴴ := by noncomm_ring [Matrix.conjTranspose_mul]
|
||||
_ = A := by simp
|
||||
|
||||
|
@ -120,7 +129,7 @@ theorem toSelfAdjointMap_mul (M N : ℂ²ˣ²)
|
|||
LinearMap.ext fun A => Subtype.ext <|
|
||||
show M * N * A * (M * N)ᴴ = M * (N * A * Nᴴ) * Mᴴ by noncomm_ring [Matrix.conjTranspose_mul]
|
||||
|
||||
theorem toSelfAdjointMap_similar_det (M N : ℂ²ˣ²) [Invertible M]
|
||||
theorem toSelfAdjointMap_similar_det (M N : ℂ²ˣ²) [Invertible M]
|
||||
: LinearMap.det (toSelfAdjointMap' (M * N * M⁻¹)) = LinearMap.det (toSelfAdjointMap' N) :=
|
||||
let e := toSelfAdjointEquiv M
|
||||
let f := toSelfAdjointMap' N
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue