2024-10-11 15:47:35 +00:00
|
|
|
|
/-
|
|
|
|
|
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.Tensors.OverColor.Basic
|
2024-10-16 16:38:36 +00:00
|
|
|
|
import HepLean.Tensors.Tree.Dot
|
|
|
|
|
import HepLean.SpaceTime.WeylFermion.Contraction
|
|
|
|
|
import HepLean.SpaceTime.WeylFermion.Metric
|
|
|
|
|
import HepLean.SpaceTime.WeylFermion.Unit
|
|
|
|
|
import HepLean.SpaceTime.LorentzVector.Complex.Contraction
|
|
|
|
|
import HepLean.SpaceTime.LorentzVector.Complex.Metric
|
|
|
|
|
import HepLean.SpaceTime.LorentzVector.Complex.Unit
|
2024-10-11 15:47:35 +00:00
|
|
|
|
import HepLean.Mathematics.PiTensorProduct
|
2024-10-16 16:38:36 +00:00
|
|
|
|
import HepLean.SpaceTime.PauliMatrices.AsTensor
|
2024-10-11 15:47:35 +00:00
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
## Complex Lorentz tensors
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
open Matrix
|
|
|
|
|
open MatrixGroups
|
|
|
|
|
open Complex
|
|
|
|
|
open TensorProduct
|
|
|
|
|
open IndexNotation
|
|
|
|
|
open CategoryTheory
|
|
|
|
|
open MonoidalCategory
|
|
|
|
|
|
2024-10-25 13:50:19 +00:00
|
|
|
|
namespace complexLorentzTensor
|
|
|
|
|
|
2024-10-11 15:47:35 +00:00
|
|
|
|
/-- The colors associated with complex representations of SL(2, ℂ) of intrest to physics. -/
|
|
|
|
|
inductive Color
|
|
|
|
|
| upL : Color
|
|
|
|
|
| downL : Color
|
|
|
|
|
| upR : Color
|
|
|
|
|
| downR : Color
|
|
|
|
|
| up : Color
|
|
|
|
|
| down : Color
|
|
|
|
|
|
2024-10-17 11:43:33 +00:00
|
|
|
|
instance : DecidableEq Color := fun x y =>
|
|
|
|
|
match x, y with
|
|
|
|
|
| Color.upL, Color.upL => isTrue rfl
|
|
|
|
|
| Color.downL, Color.downL => isTrue rfl
|
|
|
|
|
| Color.upR, Color.upR => isTrue rfl
|
|
|
|
|
| Color.downR, Color.downR => isTrue rfl
|
|
|
|
|
| Color.up, Color.up => isTrue rfl
|
|
|
|
|
| Color.down, Color.down => isTrue rfl
|
|
|
|
|
/- The false -/
|
|
|
|
|
| Color.upL, Color.downL => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.upL, Color.upR => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.upL, Color.downR => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.upL, Color.up => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.upL, Color.down => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.downL, Color.upL => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.downL, Color.upR => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.downL, Color.downR => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.downL, Color.up => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.downL, Color.down => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.upR, Color.upL => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.upR, Color.downL => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.upR, Color.downR => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.upR, Color.up => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.upR, Color.down => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.downR, Color.upL => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.downR, Color.downL => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.downR, Color.upR => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.downR, Color.up => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.downR, Color.down => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.up, Color.upL => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.up, Color.downL => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.up, Color.upR => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.up, Color.downR => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.up, Color.down => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.down, Color.upL => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.down, Color.downL => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.down, Color.upR => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.down, Color.downR => isFalse fun h => Color.noConfusion h
|
|
|
|
|
| Color.down, Color.up => isFalse fun h => Color.noConfusion h
|
|
|
|
|
|
2024-10-25 13:50:19 +00:00
|
|
|
|
end complexLorentzTensor
|
2024-10-16 16:38:36 +00:00
|
|
|
|
|
2024-10-25 13:50:19 +00:00
|
|
|
|
noncomputable section
|
|
|
|
|
open complexLorentzTensor in
|
2024-10-16 16:38:36 +00:00
|
|
|
|
/-- The tensor structure for complex Lorentz tensors. -/
|
2024-10-21 12:24:17 +00:00
|
|
|
|
def complexLorentzTensor : TensorSpecies where
|
2024-10-25 13:50:19 +00:00
|
|
|
|
C := complexLorentzTensor.Color
|
2024-10-16 16:38:36 +00:00
|
|
|
|
G := SL(2, ℂ)
|
|
|
|
|
G_group := inferInstance
|
|
|
|
|
k := ℂ
|
|
|
|
|
k_commRing := inferInstance
|
|
|
|
|
FDiscrete := Discrete.functor fun c =>
|
|
|
|
|
match c with
|
|
|
|
|
| Color.upL => Fermion.leftHanded
|
|
|
|
|
| Color.downL => Fermion.altLeftHanded
|
|
|
|
|
| Color.upR => Fermion.rightHanded
|
|
|
|
|
| Color.downR => Fermion.altRightHanded
|
|
|
|
|
| Color.up => Lorentz.complexContr
|
|
|
|
|
| Color.down => Lorentz.complexCo
|
|
|
|
|
τ := fun c =>
|
|
|
|
|
match c with
|
|
|
|
|
| Color.upL => Color.downL
|
|
|
|
|
| Color.downL => Color.upL
|
|
|
|
|
| Color.upR => Color.downR
|
|
|
|
|
| Color.downR => Color.upR
|
|
|
|
|
| Color.up => Color.down
|
|
|
|
|
| Color.down => Color.up
|
|
|
|
|
τ_involution c := by
|
|
|
|
|
match c with
|
|
|
|
|
| Color.upL => rfl
|
|
|
|
|
| Color.downL => rfl
|
|
|
|
|
| Color.upR => rfl
|
|
|
|
|
| Color.downR => rfl
|
|
|
|
|
| Color.up => rfl
|
|
|
|
|
| Color.down => rfl
|
|
|
|
|
contr := Discrete.natTrans fun c =>
|
|
|
|
|
match c with
|
|
|
|
|
| Discrete.mk Color.upL => Fermion.leftAltContraction
|
|
|
|
|
| Discrete.mk Color.downL => Fermion.altLeftContraction
|
|
|
|
|
| Discrete.mk Color.upR => Fermion.rightAltContraction
|
|
|
|
|
| Discrete.mk Color.downR => Fermion.altRightContraction
|
|
|
|
|
| Discrete.mk Color.up => Lorentz.contrCoContraction
|
|
|
|
|
| Discrete.mk Color.down => Lorentz.coContrContraction
|
|
|
|
|
metric := Discrete.natTrans fun c =>
|
|
|
|
|
match c with
|
|
|
|
|
| Discrete.mk Color.upL => Fermion.leftMetric
|
|
|
|
|
| Discrete.mk Color.downL => Fermion.altLeftMetric
|
|
|
|
|
| Discrete.mk Color.upR => Fermion.rightMetric
|
|
|
|
|
| Discrete.mk Color.downR => Fermion.altRightMetric
|
|
|
|
|
| Discrete.mk Color.up => Lorentz.contrMetric
|
|
|
|
|
| Discrete.mk Color.down => Lorentz.coMetric
|
|
|
|
|
unit := Discrete.natTrans fun c =>
|
|
|
|
|
match c with
|
|
|
|
|
| Discrete.mk Color.upL => Fermion.altLeftLeftUnit
|
|
|
|
|
| Discrete.mk Color.downL => Fermion.leftAltLeftUnit
|
|
|
|
|
| Discrete.mk Color.upR => Fermion.altRightRightUnit
|
|
|
|
|
| Discrete.mk Color.downR => Fermion.rightAltRightUnit
|
|
|
|
|
| Discrete.mk Color.up => Lorentz.coContrUnit
|
|
|
|
|
| Discrete.mk Color.down => Lorentz.contrCoUnit
|
2024-10-21 16:21:29 +00:00
|
|
|
|
repDim := fun c =>
|
2024-10-16 16:38:36 +00:00
|
|
|
|
match c with
|
|
|
|
|
| Color.upL => 2
|
|
|
|
|
| Color.downL => 2
|
|
|
|
|
| Color.upR => 2
|
|
|
|
|
| Color.downR => 2
|
|
|
|
|
| Color.up => 4
|
|
|
|
|
| Color.down => 4
|
2024-10-22 06:42:06 +00:00
|
|
|
|
repDim_neZero := fun c =>
|
|
|
|
|
match c with
|
|
|
|
|
| Color.upL => inferInstance
|
|
|
|
|
| Color.downL => inferInstance
|
|
|
|
|
| Color.upR => inferInstance
|
|
|
|
|
| Color.downR => inferInstance
|
|
|
|
|
| Color.up => inferInstance
|
|
|
|
|
| Color.down => inferInstance
|
2024-10-21 16:21:29 +00:00
|
|
|
|
basis := fun c =>
|
|
|
|
|
match c with
|
|
|
|
|
| Color.upL => Fermion.leftBasis
|
|
|
|
|
| Color.downL => Fermion.altLeftBasis
|
|
|
|
|
| Color.upR => Fermion.rightBasis
|
|
|
|
|
| Color.downR => Fermion.altRightBasis
|
|
|
|
|
| Color.up => Lorentz.complexContrBasisFin4
|
|
|
|
|
| Color.down => Lorentz.complexCoBasisFin4
|
2024-10-21 12:20:43 +00:00
|
|
|
|
contr_tmul_symm := fun c =>
|
|
|
|
|
match c with
|
|
|
|
|
| Color.upL => Fermion.leftAltContraction_tmul_symm
|
|
|
|
|
| Color.downL => Fermion.altLeftContraction_tmul_symm
|
|
|
|
|
| Color.upR => Fermion.rightAltContraction_tmul_symm
|
|
|
|
|
| Color.downR => Fermion.altRightContraction_tmul_symm
|
|
|
|
|
| Color.up => Lorentz.contrCoContraction_tmul_symm
|
|
|
|
|
| Color.down => Lorentz.coContrContraction_tmul_symm
|
2024-10-24 15:52:56 +00:00
|
|
|
|
contr_unit := fun c =>
|
|
|
|
|
match c with
|
|
|
|
|
| Color.upL => Fermion.contr_altLeftLeftUnit
|
|
|
|
|
| Color.downL => Fermion.contr_leftAltLeftUnit
|
|
|
|
|
| Color.upR => Fermion.contr_altRightRightUnit
|
|
|
|
|
| Color.downR => Fermion.contr_rightAltRightUnit
|
|
|
|
|
| Color.up => Lorentz.contr_coContrUnit
|
|
|
|
|
| Color.down => Lorentz.contr_contrCoUnit
|
2024-10-24 16:04:05 +00:00
|
|
|
|
unit_symm := fun c =>
|
|
|
|
|
match c with
|
|
|
|
|
| Color.upL => Fermion.altLeftLeftUnit_symm
|
|
|
|
|
| Color.downL => Fermion.leftAltLeftUnit_symm
|
|
|
|
|
| Color.upR => Fermion.altRightRightUnit_symm
|
|
|
|
|
| Color.downR => Fermion.rightAltRightUnit_symm
|
|
|
|
|
| Color.up => Lorentz.coContrUnit_symm
|
|
|
|
|
| Color.down => Lorentz.contrCoUnit_symm
|
2024-10-24 16:42:25 +00:00
|
|
|
|
contr_metric := fun c =>
|
2024-10-24 16:35:15 +00:00
|
|
|
|
match c with
|
|
|
|
|
| Color.upL => by simpa using Fermion.leftAltContraction_apply_metric
|
|
|
|
|
| Color.downL => by simpa using Fermion.altLeftContraction_apply_metric
|
|
|
|
|
| Color.upR => by simpa using Fermion.rightAltContraction_apply_metric
|
|
|
|
|
| Color.downR => by simpa using Fermion.altRightContraction_apply_metric
|
|
|
|
|
| Color.up => by simpa using Lorentz.contrCoContraction_apply_metric
|
|
|
|
|
| Color.down => by simpa using Lorentz.coContrContraction_apply_metric
|
2024-10-25 13:50:19 +00:00
|
|
|
|
namespace complexLorentzTensor
|
|
|
|
|
instance : DecidableEq complexLorentzTensor.C := complexLorentzTensor.instDecidableEqColor
|
2024-10-17 11:43:33 +00:00
|
|
|
|
|
2024-10-23 08:01:23 +00:00
|
|
|
|
lemma basis_contr (c : complexLorentzTensor.C) (i : Fin (complexLorentzTensor.repDim c))
|
|
|
|
|
(j : Fin (complexLorentzTensor.repDim (complexLorentzTensor.τ c))) :
|
|
|
|
|
complexLorentzTensor.castToField
|
|
|
|
|
((complexLorentzTensor.contr.app {as := c}).hom
|
|
|
|
|
(complexLorentzTensor.basis c i ⊗ₜ complexLorentzTensor.basis (complexLorentzTensor.τ c) j)) =
|
|
|
|
|
if i.val = j.val then 1 else 0 :=
|
|
|
|
|
match c with
|
|
|
|
|
| Color.upL => Fermion.leftAltContraction_basis _ _
|
|
|
|
|
| Color.downL => Fermion.altLeftContraction_basis _ _
|
|
|
|
|
| Color.upR => Fermion.rightAltContraction_basis _ _
|
|
|
|
|
| Color.downR => Fermion.altRightContraction_basis _ _
|
|
|
|
|
| Color.up => Lorentz.contrCoContraction_basis _ _
|
|
|
|
|
| Color.down => Lorentz.coContrContraction_basis _ _
|
|
|
|
|
|
2024-10-29 11:23:08 +00:00
|
|
|
|
instance {n : ℕ} {c : Fin n → complexLorentzTensor.C} :
|
2024-10-29 10:37:18 +00:00
|
|
|
|
DecidableEq (OverColor.mk c).left := instDecidableEqFin n
|
|
|
|
|
|
2024-10-29 11:23:08 +00:00
|
|
|
|
instance {n : ℕ} {c : Fin n → complexLorentzTensor.C} :
|
2024-10-29 10:37:18 +00:00
|
|
|
|
Fintype (OverColor.mk c).left := Fin.fintype n
|
|
|
|
|
|
|
|
|
|
instance {n m : ℕ} {c : Fin n → complexLorentzTensor.C}
|
|
|
|
|
{c1 : Fin m → complexLorentzTensor.C} (σ σ' : OverColor.mk c ⟶ OverColor.mk c1) :
|
|
|
|
|
Decidable (σ = σ') :=
|
|
|
|
|
decidable_of_iff _ (OverColor.Hom.ext_iff σ σ')
|
|
|
|
|
|
2024-10-25 13:50:19 +00:00
|
|
|
|
end complexLorentzTensor
|
2024-10-11 15:47:35 +00:00
|
|
|
|
end
|