feat: Metric, unit, contract of complex Lorentz vec

This commit is contained in:
jstoobysmith 2024-10-15 13:19:46 +00:00
parent 12dd1fbbac
commit 255ea5ffd7
10 changed files with 677 additions and 44 deletions

View file

@ -0,0 +1,88 @@
/-
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.SpaceTime.LorentzVector.Complex.Two
import HepLean.SpaceTime.MinkowskiMetric
/-!
# Metric for complex Lorentz vectors
-/
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open SpaceTime
open CategoryTheory.MonoidalCategory
namespace Lorentz
/-- The metric `ηᵃᵃ` as an element of `(complexContr ⊗ complexContr).V`. -/
def contrMetricVal : (complexContr ⊗ complexContr).V :=
contrContrToMatrix.symm ((@minkowskiMatrix 3).map ofReal)
/-- The metric `ηᵃᵃ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ complexContr`,
making its invariance under the action of `SL(2,)`. -/
def contrMetric : 𝟙_ (Rep SL(2,)) ⟶ complexContr ⊗ complexContr where
hom := {
toFun := fun a =>
let a' : := a
a' • contrMetricVal,
map_add' := fun x y => by
simp only [add_smul],
map_smul' := fun m x => by
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • contrMetricVal =
(TensorProduct.map (complexContr.ρ M) (complexContr.ρ M)) (x' • contrMetricVal)
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp only [Action.instMonoidalCategory_tensorObj_V, contrMetricVal]
erw [contrContrToMatrix_ρ_symm]
apply congrArg
simp only [LorentzGroup.toComplex_mul_minkowskiMatrix_mul_transpose]
/-- The metric `ηᵢᵢ` as an element of `(complexCo ⊗ complexCo).V`. -/
def coMetricVal : (complexCo ⊗ complexCo).V :=
coCoToMatrix.symm ((@minkowskiMatrix 3).map ofReal)
/-- The metric `ηᵢᵢ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ complexCo ⊗ complexCo`,
making its invariance under the action of `SL(2,)`. -/
def coMetric : 𝟙_ (Rep SL(2,)) ⟶ complexCo ⊗ complexCo where
hom := {
toFun := fun a =>
let a' : := a
a' • coMetricVal,
map_add' := fun x y => by
simp only [add_smul],
map_smul' := fun m x => by
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • coMetricVal =
(TensorProduct.map (complexCo.ρ M) (complexCo.ρ M)) (x' • coMetricVal)
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp only [Action.instMonoidalCategory_tensorObj_V, coMetricVal]
erw [coCoToMatrix_ρ_symm]
apply congrArg
rw [LorentzGroup.toComplex_inv]
simp only [lorentzGroupIsGroup_inv, SL2C.toLorentzGroup_apply_coe,
LorentzGroup.toComplex_transpose_mul_minkowskiMatrix_mul_self]
end Lorentz
end