PhysLean/HepLean/Mathematics/SO3/Basic.lean

214 lines
7.9 KiB
Text
Raw Normal View History

2024-05-20 16:20:26 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
2024-07-12 16:39:44 -04:00
Released under Apache 2.0 license as described in the file LICENSE.
2024-05-20 16:20:26 -04:00
Authors: Joseph Tooby-Smith
-/
2024-07-26 16:32:54 -04:00
import Mathlib.LinearAlgebra.Matrix.GeneralLinearGroup.Defs
2024-05-22 16:50:53 -04:00
import Mathlib.LinearAlgebra.Eigenspace.Basic
2024-06-25 07:06:32 -04:00
import Mathlib.Analysis.InnerProductSpace.PiL2
2024-05-20 16:20:26 -04:00
/-!
2024-05-22 13:34:53 -04:00
# The group SO(3)
2024-05-20 16:20:26 -04:00
-/
namespace GroupTheory
open Matrix
2024-05-22 13:34:53 -04:00
/-- The group of `3×3` real matrices with determinant 1 and `A * Aᵀ = 1`. -/
2024-05-20 16:20:26 -04:00
def SO3 : Type := {A : Matrix (Fin 3) (Fin 3) // A.det = 1 ∧ A * Aᵀ = 1}
2024-11-11 16:55:15 +00:00
/-- The instance of a group on `SO3`. -/
2024-09-04 06:28:46 -04:00
@[simps! mul_coe one_coe inv div]
2024-05-20 16:20:26 -04:00
instance SO3Group : Group SO3 where
mul A B := ⟨A.1 * B.1,
by
simp only [det_mul, A.2.1, B.2.1, mul_one],
by
2024-12-10 13:44:39 +00:00
simp only [transpose_mul, ← Matrix.mul_assoc, Matrix.mul_assoc, B.2.2, mul_one, A.2.2]
trans A.1 * ((B.1 * (B.1)ᵀ) * (A.1)ᵀ)
· noncomm_ring
· simp [Matrix.mul_assoc, B.2.2, mul_one, A.2.2]⟩
2024-08-19 17:23:57 +02:00
mul_assoc A B C := Subtype.eq (Matrix.mul_assoc A.1 B.1 C.1)
one := ⟨1, det_one, by rw [transpose_one, mul_one]⟩
one_mul A := Subtype.eq (Matrix.one_mul A.1)
mul_one A := Subtype.eq (Matrix.mul_one A.1)
2024-09-04 07:31:34 -04:00
inv A := ⟨A.1ᵀ, by simp only [det_transpose, A.2],
by simp only [transpose_transpose, mul_eq_one_comm.mpr A.2.2]⟩
2024-09-04 06:28:46 -04:00
inv_mul_cancel A := Subtype.eq (mul_eq_one_comm.mpr A.2.2)
2024-05-20 16:20:26 -04:00
2024-05-22 13:34:53 -04:00
/-- Notation for the group `SO3`. -/
2024-05-20 16:20:26 -04:00
scoped[GroupTheory] notation (name := SO3_notation) "SO(3)" => SO3
/-- SO3 has the subtype topology. -/
instance : TopologicalSpace SO3 := instTopologicalSpaceSubtype
namespace SO3
2024-08-19 17:23:57 +02:00
lemma coe_inv (A : SO3) : (A⁻¹).1 = A.1⁻¹:=
(inv_eq_left_inv (mul_eq_one_comm.mpr A.2.2)).symm
2024-05-20 16:20:26 -04:00
2024-05-22 13:34:53 -04:00
/-- The inclusion of `SO(3)` into `GL (Fin 3) `. -/
def toGL : SO(3) →* GL (Fin 3) where
2024-05-20 16:20:26 -04:00
toFun A := ⟨A.1, (A⁻¹).1, A.2.2, mul_eq_one_comm.mpr A.2.2⟩
2024-08-30 10:11:55 -04:00
map_one' := (GeneralLinearGroup.ext_iff _ 1).mpr fun _=> congrFun rfl
map_mul' _ _ := (GeneralLinearGroup.ext_iff _ _).mpr fun _ => congrFun rfl
2024-05-20 16:20:26 -04:00
lemma subtype_val_eq_toGL : (Subtype.val : SO3 → Matrix (Fin 3) (Fin 3) ) =
2024-08-19 17:23:57 +02:00
Units.val ∘ toGL.toFun :=
2024-05-20 16:20:26 -04:00
rfl
2024-11-26 09:33:24 +00:00
/-- The inclusino of `SO(3)` into `GL(3,)` is an injection. -/
2024-05-20 16:20:26 -04:00
lemma toGL_injective : Function.Injective toGL := by
2024-08-19 17:23:57 +02:00
refine fun A B h ↦ Subtype.eq ?_
2024-05-20 16:20:26 -04:00
rw [@Units.ext_iff] at h
simpa using h
2024-05-22 13:34:53 -04:00
/-- The inclusion of `SO(3)` into the monoid of matrices times the opposite of
the monoid of matrices. -/
2024-05-20 16:20:26 -04:00
@[simps!]
def toProd : SO(3) →* (Matrix (Fin 3) (Fin 3) ) × (Matrix (Fin 3) (Fin 3) )ᵐᵒᵖ :=
MonoidHom.comp (Units.embedProduct _) toGL
2024-08-30 10:11:55 -04:00
lemma toProd_eq_transpose : toProd A = (A.1, ⟨A.1ᵀ⟩) := rfl
2024-05-20 16:20:26 -04:00
lemma toProd_injective : Function.Injective toProd := by
intro A B h
rw [toProd_eq_transpose, toProd_eq_transpose] at h
rw [@Prod.mk.inj_iff] at h
2024-08-19 17:23:57 +02:00
exact Subtype.eq h.1
lemma toProd_continuous : Continuous toProd :=
continuous_prod_mk.mpr ⟨continuous_iff_le_induced.mpr fun _ a ↦ a,
Continuous.comp' (MulOpposite.continuous_op)
(Continuous.matrix_transpose (continuous_iff_le_induced.mpr fun _ a ↦ a))⟩
2024-05-20 16:20:26 -04:00
2024-12-10 13:44:39 +00:00
open Topology
2024-05-22 13:34:53 -04:00
/-- The embedding of `SO(3)` into the monoid of matrices times the opposite of
the monoid of matrices. -/
2024-11-02 07:59:39 +00:00
lemma toProd_embedding : IsEmbedding toProd where
2024-12-10 13:44:39 +00:00
injective := toProd_injective
2024-11-02 07:59:39 +00:00
eq_induced := (isInducing_iff ⇑toProd).mp (IsInducing.of_comp toProd_continuous
continuous_fst ((isInducing_iff (Prod.fst ∘ ⇑toProd)).mpr rfl))
2024-05-20 16:20:26 -04:00
2024-05-22 13:34:53 -04:00
/-- The embedding of `SO(3)` into `GL (Fin 3) `. -/
2024-11-02 07:59:39 +00:00
lemma toGL_embedding : IsEmbedding toGL.toFun where
2024-12-10 13:44:39 +00:00
injective := toGL_injective
2024-11-02 07:59:39 +00:00
eq_induced := by
2024-05-20 16:20:26 -04:00
refine ((fun {X} {t t'} => TopologicalSpace.ext_iff.mpr) ?_).symm
intro s
2024-11-02 07:59:39 +00:00
rw [TopologicalSpace.ext_iff.mp toProd_embedding.eq_induced s]
2024-05-20 16:20:26 -04:00
rw [isOpen_induced_iff, isOpen_induced_iff]
apply Iff.intro ?_ ?_
· intro h
obtain ⟨U, hU1, hU2⟩ := h
rw [isOpen_induced_iff] at hU1
obtain ⟨V, hV1, hV2⟩ := hU1
use V
2024-09-04 07:31:34 -04:00
simp only [hV1, true_and]
2024-05-20 16:20:26 -04:00
rw [← hU2, ← hV2]
rfl
· intro h
obtain ⟨U, hU1, hU2⟩ := h
let t := (Units.embedProduct _) ⁻¹' U
use t
apply And.intro (isOpen_induced hU1)
exact hU2
2024-11-11 16:55:15 +00:00
/-- The instance of a topological group on `SO(3)`, defined through the embedding of `SO(3)`
into `GL(n)`. -/
instance : TopologicalGroup SO(3) :=
2024-11-02 07:59:39 +00:00
IsInducing.topologicalGroup toGL toGL_embedding.toIsInducing
2024-05-20 16:20:26 -04:00
2024-11-26 09:33:24 +00:00
/-- The determinant of an `SO(3)` matrix minus the identity is equal to zero. -/
2024-05-21 11:31:57 -04:00
lemma det_minus_id (A : SO(3)) : det (A.1 - 1) = 0 := by
have h1 : det (A.1 - 1) = - det (A.1 - 1) :=
2024-05-21 11:31:57 -04:00
calc
2024-07-12 11:23:02 -04:00
det (A.1 - 1) = det (A.1 - A.1 * A.1ᵀ) := by simp [A.2.2]
_ = det A.1 * det (1 - A.1ᵀ) := by rw [← det_mul, mul_sub, mul_one]
2024-07-19 17:00:32 -04:00
_ = det (1 - A.1ᵀ) := by simp [A.2.1]
2024-05-21 11:31:57 -04:00
_ = det (1 - A.1ᵀ)ᵀ := by rw [det_transpose]
2024-08-30 10:11:55 -04:00
_ = det (1 - A.1) := rfl
2024-05-21 11:31:57 -04:00
_ = det (- (A.1 - 1)) := by simp
_ = (- 1) ^ 3 * det (A.1 - 1) := by simp only [det_neg, Fintype.card_fin, neg_mul, one_mul]
_ = - det (A.1 - 1) := by simp [pow_three]
2024-09-04 06:28:46 -04:00
exact CharZero.eq_neg_self_iff.mp h1
2024-05-20 16:20:26 -04:00
2024-11-26 09:33:24 +00:00
/-- The determinant of the identity minus an `SO(3)` matrix is zero. -/
2024-05-22 09:18:12 -04:00
@[simp]
lemma det_id_minus (A : SO(3)) : det (1 - A.1) = 0 := by
have h1 : det (1 - A.1) = - det (A.1 - 1) := by
calc
det (1 - A.1) = det (- (A.1 - 1)) := by simp
2024-07-12 11:23:02 -04:00
_ = (- 1) ^ 3 * det (A.1 - 1) := by simp only [det_neg, Fintype.card_fin, neg_mul, one_mul]
_ = - det (A.1 - 1) := by simp [pow_three]
2024-05-22 09:18:12 -04:00
rw [h1, det_minus_id]
2024-08-30 10:11:55 -04:00
exact neg_zero
2024-05-20 16:20:26 -04:00
2024-11-26 09:33:24 +00:00
/-- For every matrix in `SO(3)`, the real number `1` is in its spectrum. -/
2024-05-22 09:18:12 -04:00
@[simp]
lemma one_in_spectrum (A : SO(3)) : 1 ∈ spectrum (A.1) := by
rw [spectrum.mem_iff]
2024-05-22 13:34:53 -04:00
simp only [_root_.map_one]
2024-05-22 09:18:12 -04:00
rw [Matrix.isUnit_iff_isUnit_det]
simp
2024-05-20 16:20:26 -04:00
2024-05-22 09:18:12 -04:00
noncomputable section action
open Module
2024-05-22 13:34:53 -04:00
/-- The endomorphism of `EuclideanSpace (Fin 3)` associated to a element of `SO(3)`. -/
2024-05-22 09:18:12 -04:00
@[simps!]
def toEnd (A : SO(3)) : End (EuclideanSpace (Fin 3)) :=
Matrix.toLin (EuclideanSpace.basisFun (Fin 3) ).toBasis
(EuclideanSpace.basisFun (Fin 3) ).toBasis A.1
2024-11-26 09:33:24 +00:00
/-- Every `SO(3)` matrix has an eigenvalue equal to `1`. -/
2024-05-22 09:18:12 -04:00
lemma one_is_eigenvalue (A : SO(3)) : A.toEnd.HasEigenvalue 1 := by
rw [End.hasEigenvalue_iff_mem_spectrum]
2024-07-12 16:22:06 -04:00
erw [AlgEquiv.spectrum_eq (Matrix.toLinAlgEquiv (EuclideanSpace.basisFun (Fin 3) ).toBasis) A.1]
2024-05-22 09:18:12 -04:00
exact one_in_spectrum A
2024-11-26 09:33:24 +00:00
/-- For every element of `SO(3)` there exists a vector which remains unchanged under the
action of that `SO(3)` element. -/
2024-05-22 09:18:12 -04:00
lemma exists_stationary_vec (A : SO(3)) :
∃ (v : EuclideanSpace (Fin 3)),
2024-07-12 16:22:06 -04:00
Orthonormal (({0} : Set (Fin 3)).restrict (fun _ => v))
2024-05-22 09:18:12 -04:00
∧ A.toEnd v = v := by
obtain ⟨v, hv⟩ := End.HasEigenvalue.exists_hasEigenvector $ one_is_eigenvalue A
have hvn : ‖v‖ ≠ 0 := norm_ne_zero_iff.mpr hv.2
use (1/‖v‖) • v
apply And.intro
· rw [@orthonormal_iff_ite]
intro v1 v2
have hv1 := v1.2
have hv2 := v2.2
simp_all only [one_div, Set.mem_singleton_iff]
have hveq : v1 = v2 := by
rw [@Subtype.ext_iff]
simp_all only
subst hveq
simp only [Set.restrict_apply, PiLp.smul_apply, smul_eq_mul,
_root_.map_mul, map_inv₀, conj_trivial, ↓reduceIte]
rw [inner_smul_right, inner_smul_left, real_inner_self_eq_norm_sq v]
field_simp
ring
· rw [_root_.map_smul, End.mem_eigenspace_iff.mp hv.1]
simp
2024-05-22 09:18:12 -04:00
2024-11-26 09:33:24 +00:00
/-- For every element of `SO(3)` there exists a basis indexed by `Fin 3` under which the first
element remains invariant. -/
2024-05-22 09:18:12 -04:00
lemma exists_basis_preserved (A : SO(3)) :
∃ (b : OrthonormalBasis (Fin 3) (EuclideanSpace (Fin 3))), A.toEnd (b 0) = b 0 := by
obtain ⟨v, hv⟩ := exists_stationary_vec A
2024-11-02 07:59:39 +00:00
have h3 : Module.finrank (EuclideanSpace (Fin 3)) = Fintype.card (Fin 3) := by
2024-05-22 09:18:12 -04:00
simp_all only [toEnd_apply, finrank_euclideanSpace, Fintype.card_fin]
obtain ⟨b, hb⟩ := Orthonormal.exists_orthonormalBasis_extension_of_card_eq h3 hv.1
simp only [Fin.isValue, Set.mem_singleton_iff, forall_eq] at hb
2024-05-22 09:18:12 -04:00
use b
rw [hb, hv.2]
end action
2024-05-20 16:20:26 -04:00
end SO3
end GroupTheory