PhysLean/HepLean/SpaceTime/LorentzVector/Modules.lean
2024-10-17 11:52:49 +00:00

174 lines
5.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 as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Meta.Informal
import HepLean.SpaceTime.SL2C.Basic
import Mathlib.RepresentationTheory.Rep
import Mathlib.Logic.Equiv.TransferInstance
/-!
## Modules associated with Lorentz vectors
These have not yet been fully-implmented.
We define these modules to prevent casting between different types of Lorentz vectors.
-/
namespace Lorentz
noncomputable section
open Matrix
open MatrixGroups
open Complex
/-- The module for contravariant (up-index) complex Lorentz vectors. -/
structure ContrModule where
/-- The underlying value as a vector `Fin 1 ⊕ Fin 3 → `. -/
val : Fin 1 ⊕ Fin 3 →
namespace ContrModule
/-- The equivalence between `ContrModule` and `Fin 1 ⊕ Fin 3 → `. -/
def toFin13Fun : ContrModule ≃ (Fin 1 ⊕ Fin 3 → ) where
toFun v := v.val
invFun f := ⟨f⟩
left_inv _ := rfl
right_inv _ := rfl
/-- The instance of `AddCommMonoid` on `ContrModule` defined via its equivalence
with `Fin 1 ⊕ Fin 3 → `. -/
instance : AddCommMonoid ContrModule := Equiv.addCommMonoid toFin13Fun
/-- The instance of `AddCommGroup` on `ContrModule` defined via its equivalence
with `Fin 1 ⊕ Fin 3 → `. -/
instance : AddCommGroup ContrModule := Equiv.addCommGroup toFin13Fun
/-- The instance of `Module` on `ContrModule` defined via its equivalence
with `Fin 1 ⊕ Fin 3 → `. -/
instance : Module ContrModule := Equiv.module toFin13Fun
@[ext]
lemma ext (ψ ψ' : ContrModule) (h : ψ.val = ψ'.val) : ψ = ψ' := by
cases ψ
cases ψ'
subst h
simp_all only
@[simp]
lemma val_add (ψ ψ' : ContrModule) : (ψ + ψ').val = ψ.val + ψ'.val := rfl
@[simp]
lemma val_smul (r : ) (ψ : ContrModule) : (r • ψ).val = r • ψ.val := rfl
/-- The linear equivalence between `ContrModule` and `(Fin 1 ⊕ Fin 3 → )`. -/
@[simps!]
def toFin13Equiv : ContrModule ≃ₗ[] (Fin 1 ⊕ Fin 3 → ) where
toFun := toFin13Fun
map_add' := fun _ _ => rfl
map_smul' := fun _ _ => rfl
invFun := toFin13Fun.symm
left_inv := fun _ => rfl
right_inv := fun _ => rfl
/-- The underlying element of `Fin 1 ⊕ Fin 3 → ` of a element in `ContrModule` defined
through the linear equivalence `toFin13Equiv`. -/
abbrev toFin13 (ψ : ContrModule) := toFin13Equiv ψ
/-- The representation of the Lorentz group on `ContrModule`. -/
def lorentzGroupRep : Representation (LorentzGroup 3) ContrModule where
toFun M := {
toFun := fun v => toFin13Equiv.symm (LorentzGroup.toComplex M *ᵥ v.toFin13),
map_add' := by
intro ψ ψ'
simp [mulVec_add]
map_smul' := by
intro r ψ
simp [mulVec_smul]}
map_one' := by
ext i
simp
map_mul' M N := by
ext i
simp
/-- The representation of the SL(2, ) on `ContrModule` induced by the representation of the
Lorentz group. -/
def SL2CRep : Representation SL(2, ) ContrModule :=
MonoidHom.comp lorentzGroupRep SpaceTime.SL2C.toLorentzGroup
end ContrModule
/-- The module for covariant (up-index) complex Lorentz vectors. -/
structure CoModule where
/-- The underlying value as a vector `Fin 1 ⊕ Fin 3 → `. -/
val : Fin 1 ⊕ Fin 3 →
namespace CoModule
/-- The equivalence between `CoModule` and `Fin 1 ⊕ Fin 3 → `. -/
def toFin13Fun : CoModule ≃ (Fin 1 ⊕ Fin 3 → ) where
toFun v := v.val
invFun f := ⟨f⟩
left_inv _ := rfl
right_inv _ := rfl
/-- The instance of `AddCommMonoid` on `CoModule` defined via its equivalence
with `Fin 1 ⊕ Fin 3 → `. -/
instance : AddCommMonoid CoModule := Equiv.addCommMonoid toFin13Fun
/-- The instance of `AddCommGroup` on `CoModule` defined via its equivalence
with `Fin 1 ⊕ Fin 3 → `. -/
instance : AddCommGroup CoModule := Equiv.addCommGroup toFin13Fun
/-- The instance of `Module` on `CoModule` defined via its equivalence
with `Fin 1 ⊕ Fin 3 → `. -/
instance : Module CoModule := Equiv.module toFin13Fun
/-- The linear equivalence between `CoModule` and `(Fin 1 ⊕ Fin 3 → )`. -/
@[simps!]
def toFin13Equiv : CoModule ≃ₗ[] (Fin 1 ⊕ Fin 3 → ) where
toFun := toFin13Fun
map_add' := fun _ _ => rfl
map_smul' := fun _ _ => rfl
invFun := toFin13Fun.symm
left_inv := fun _ => rfl
right_inv := fun _ => rfl
/-- The underlying element of `Fin 1 ⊕ Fin 3 → ` of a element in `CoModule` defined
through the linear equivalence `toFin13Equiv`. -/
abbrev toFin13 (ψ : CoModule) := toFin13Equiv ψ
/-- The representation of the Lorentz group on `CoModule`. -/
def lorentzGroupRep : Representation (LorentzGroup 3) CoModule where
toFun M := {
toFun := fun v => toFin13Equiv.symm ((LorentzGroup.toComplex M)⁻¹ᵀ *ᵥ v.toFin13),
map_add' := by
intro ψ ψ'
simp [mulVec_add]
map_smul' := by
intro r ψ
simp [mulVec_smul]}
map_one' := by
ext i
simp
map_mul' M N := by
ext1 x
simp only [SpecialLinearGroup.coe_mul, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.mul_apply,
LinearEquiv.apply_symm_apply, mulVec_mulVec, EmbeddingLike.apply_eq_iff_eq]
refine (congrFun (congrArg _ ?_) _)
simp only [_root_.map_mul]
rw [Matrix.mul_inv_rev]
exact transpose_mul _ _
/-- The representation of the SL(2, ) on `ContrModule` induced by the representation of the
Lorentz group. -/
def SL2CRep : Representation SL(2, ) CoModule :=
MonoidHom.comp lorentzGroupRep SpaceTime.SL2C.toLorentzGroup
end CoModule
end
end Lorentz