/- 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 /-! # Unit for complex Lorentz vectors -/ noncomputable section open Matrix open MatrixGroups open Complex open TensorProduct open SpaceTime open CategoryTheory.MonoidalCategory namespace Lorentz /-- The contra-co unit for complex lorentz vectors. Usually denoted `ฮดโฑแตข`. -/ def contrCoUnitVal : (complexContr โŠ— complexCo).V := contrCoToMatrix.symm 1 /-- The contra-co unit for complex lorentz vectors as a morphism `๐Ÿ™_ (Rep โ„‚ SL(2,โ„‚)) โŸถ complexContr โŠ— complexCo`, manifesting the invaraince under the `SL(2, โ„‚)` action. -/ def contrCoUnit : ๐Ÿ™_ (Rep โ„‚ SL(2,โ„‚)) โŸถ complexContr โŠ— complexCo where hom := { toFun := fun a => let a' : โ„‚ := a a' โ€ข contrCoUnitVal, 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' โ€ข contrCoUnitVal = (TensorProduct.map (complexContr.ฯ M) (complexCo.ฯ M)) (x' โ€ข contrCoUnitVal) simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul] apply congrArg simp only [Action.instMonoidalCategory_tensorObj_V, contrCoUnitVal] erw [contrCoToMatrix_ฯ_symm] apply congrArg simp /-- The co-contra unit for complex lorentz vectors. Usually denoted `ฮดแตขโฑ`. -/ def coContrUnitVal : (complexCo โŠ— complexContr).V := coContrToMatrix.symm 1 /-- The co-contra unit for complex lorentz vectors as a morphism `๐Ÿ™_ (Rep โ„‚ SL(2,โ„‚)) โŸถ complexCo โŠ— complexContr`, manifesting the invaraince under the `SL(2, โ„‚)` action. -/ def coContrUnit : ๐Ÿ™_ (Rep โ„‚ SL(2,โ„‚)) โŸถ complexCo โŠ— complexContr where hom := { toFun := fun a => let a' : โ„‚ := a a' โ€ข coContrUnitVal, 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' โ€ข coContrUnitVal = (TensorProduct.map (complexCo.ฯ M) (complexContr.ฯ M)) (x' โ€ข coContrUnitVal) simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul] apply congrArg simp only [Action.instMonoidalCategory_tensorObj_V, coContrUnitVal] erw [coContrToMatrix_ฯ_symm] apply congrArg symm refine transpose_eq_one.mp ?h.h.h.a simp end Lorentz end