feat: Add Pauli-matrices as tensor.

This commit is contained in:
jstoobysmith 2024-10-16 10:39:11 +00:00
parent a60ade65f0
commit 691b7e112e
13 changed files with 989 additions and 154 deletions

View file

@ -27,7 +27,7 @@ open Complex
open TensorProduct
/-- The vector space ^2 carrying the fundamental representation of SL(2,C).
In index notation corresponds to a Weyl fermion with indices ψ_a. -/
In index notation corresponds to a Weyl fermion with indices ψ^a. -/
def leftHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : LeftHandedModule) =>
@ -60,7 +60,7 @@ lemma leftBasis_ρ_apply (M : SL(2,)) (i j : Fin 2) :
simp only [mulVec_single, mul_one]
/-- The vector space ^2 carrying the representation of SL(2,C) given by
M → (M⁻¹)ᵀ. In index notation corresponds to a Weyl fermion with indices ψ^a. -/
M → (M⁻¹)ᵀ. In index notation corresponds to a Weyl fermion with indices ψ_a. -/
def altLeftHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : AltLeftHandedModule) =>
@ -95,7 +95,7 @@ lemma altLeftBasis_ρ_apply (M : SL(2,)) (i j : Fin 2) :
simp only [mulVec_single, transpose_apply, mul_one]
/-- The vector space ^2 carrying the conjugate representation of SL(2,C).
In index notation corresponds to a Weyl fermion with indices ψ_{dot a}. -/
In index notation corresponds to a Weyl fermion with indices ψ^{dot a}. -/
def rightHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : RightHandedModule) =>
@ -128,7 +128,7 @@ lemma rightBasis_ρ_apply (M : SL(2,)) (i j : Fin 2) :
/-- The vector space ^2 carrying the representation of SL(2,C) given by
M → (M⁻¹)^†.
In index notation this corresponds to a Weyl fermion with index `ψ^{dot a}`. -/
In index notation this corresponds to a Weyl fermion with index `ψ_{dot a}`. -/
def altRightHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : AltRightHandedModule) =>