PhysLean/HepLean/SpaceTime/SL2C/Basic.lean
2024-06-26 11:54:02 -04:00

123 lines
3.7 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzGroup.Basic
import Mathlib.RepresentationTheory.Basic
/-!
# The group SL(2, ) and it's relation to the Lorentz group
The aim of this file is to give the relationship between `SL(2, )` and the Lorentz group.
-/
namespace SpaceTime
open Matrix
open MatrixGroups
open Complex
namespace SL2C
open SpaceTime
noncomputable section
/-!
## Representation of SL(2, ) on spacetime
Through the correspondence between spacetime and self-adjoint matrices,
we can define a representation a representation of `SL(2, )` on spacetime.
-/
/-- Given an element `M ∈ SL(2, )` the linear map from `selfAdjoint (Matrix (Fin 2) (Fin 2) )` to
itself defined by `A ↦ M * A * Mᴴ`. -/
@[simps!]
def toLinearMapSelfAdjointMatrix (M : SL(2, )) :
selfAdjoint (Matrix (Fin 2) (Fin 2) ) →ₗ[] selfAdjoint (Matrix (Fin 2) (Fin 2) ) where
toFun A := ⟨M.1 * A.1 * Matrix.conjTranspose M,
by
noncomm_ring [selfAdjoint.mem_iff, star_eq_conjTranspose,
conjTranspose_mul, conjTranspose_conjTranspose,
(star_eq_conjTranspose A.1).symm.trans $ selfAdjoint.mem_iff.mp A.2]⟩
map_add' A B := by
noncomm_ring [AddSubmonoid.coe_add, AddSubgroup.coe_toAddSubmonoid, AddSubmonoid.mk_add_mk,
Subtype.mk.injEq]
map_smul' r A := by
noncomm_ring [selfAdjoint.val_smul, Algebra.mul_smul_comm, Algebra.smul_mul_assoc,
RingHom.id_apply]
/-- The representation of `SL(2, )` on `selfAdjoint (Matrix (Fin 2) (Fin 2) )` given by
`M A ↦ M * A * Mᴴ`. -/
@[simps!]
def repSelfAdjointMatrix : Representation SL(2, ) $ selfAdjoint (Matrix (Fin 2) (Fin 2) ) where
toFun := toLinearMapSelfAdjointMatrix
map_one' := by
noncomm_ring [toLinearMapSelfAdjointMatrix, SpecialLinearGroup.coe_one, one_mul,
conjTranspose_one, mul_one, Subtype.coe_eta]
map_mul' M N := by
ext x i j : 3
noncomm_ring [toLinearMapSelfAdjointMatrix, SpecialLinearGroup.coe_mul, mul_assoc,
conjTranspose_mul, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.mul_apply]
/-- The representation of `SL(2, )` on `spaceTime` obtained from `toSelfAdjointMatrix` and
`repSelfAdjointMatrix`. -/
def repSpaceTime : Representation SL(2, ) SpaceTime where
toFun M := toSelfAdjointMatrix.symm.comp ((repSelfAdjointMatrix M).comp
toSelfAdjointMatrix.toLinearMap)
map_one' := by
ext
simp
map_mul' M N := by
ext x : 3
simp
/-!
## Homomorphism to the Lorentz group
There is a group homomorphism from `SL(2, )` to the Lorentz group `𝓛`.
The purpose of this section is to define this homomorphism.
In the next section we will restrict this homomorphism to the restricted Lorentz group.
-/
/-- Given an element `M ∈ SL(2, )` the corresponding element of the Lorentz group. -/
@[simps!]
def toLorentzGroupElem (M : SL(2, )) : 𝓛 :=
⟨LinearMap.toMatrix stdBasis stdBasis (repSpaceTime M) ,
by simp [repSpaceTime, PreservesηLin.iff_det_selfAdjoint]⟩
/-- The group homomorphism from ` SL(2, )` to the Lorentz group `𝓛`. -/
@[simps!]
def toLorentzGroup : SL(2, ) →* 𝓛 where
toFun := toLorentzGroupElem
map_one' := by
simp only [toLorentzGroupElem, _root_.map_one, LinearMap.toMatrix_one]
rfl
map_mul' M N := by
apply Subtype.eq
simp only [toLorentzGroupElem, _root_.map_mul, LinearMap.toMatrix_mul,
lorentzGroupIsGroup_mul_coe]
/-!
## Homomorphism to the restricted Lorentz group
The homomorphism `toLorentzGroup` restricts to a homomorphism to the restricted Lorentz group.
In this section we will define this homomorphism.
### TODO
Complete this section.
-/
end
end SL2C
end SpaceTime