refactor: Pauli matrices

This commit is contained in:
jstoobysmith 2024-10-28 15:29:58 +00:00
parent e761463e8f
commit 83ff8f5358
2 changed files with 60 additions and 141 deletions

View file

@ -33,19 +33,19 @@ open Fermion
-/
/- The Pauli matrices as `σ^μ^α^{dot β}`. -/
/- The Pauli matrices as the complex Lorentz tensor `σ^μ^α^{dot β}`. -/
def pauliContr := {PauliMatrix.asConsTensor | ν α β}ᵀ.tensor
/-- The Pauli matrices as `σ_μ^α^{dot β}`. -/
/-- The Pauli matrices as the complex Lorentz tensor `σ_μ^α^{dot β}`. -/
def pauliCo := {Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | ν α β}ᵀ.tensor
/-- The Pauli matrices as `σ_μ_α_{dot β}`. -/
def pauliCoDown := {Fermion.altLeftMetric | α α' ⊗
Fermion.altRightMetric | β β' ⊗ pauliCo | μ α β}ᵀ.tensor
/-- The Pauli matrices as the complex Lorentz tensor `σ_μ_α_{dot β}`. -/
def pauliCoDown := {pauliCo | μ α β ⊗ Fermion.altLeftMetric | α α' ⊗
Fermion.altRightMetric | β β'}ᵀ.tensor
/-- The Pauli matrices as `σ^μ_α_{dot β}`. -/
def pauliContrDown := {Fermion.altLeftMetric | α α' ⊗
Fermion.altRightMetric | β β' ⊗ pauliContr | μ α β}ᵀ.tensor
/-- The Pauli matrices as the complex Lorentz tensor `σ^μ_α_{dot β}`. -/
def pauliContrDown := {pauliContr | μ α β ⊗ Fermion.altLeftMetric | α α' ⊗
Fermion.altRightMetric | β β'}ᵀ.tensor
/-!
@ -65,14 +65,14 @@ lemma tensorNode_pauliCo : {pauliCo | μ α β}ᵀ.tensor =
/-- The definitional tensor node relation for `pauliCoDown`. -/
lemma tensorNode_pauliCoDown : {pauliCoDown | μ α β}ᵀ.tensor =
{Fermion.altLeftMetric | α α' ⊗
Fermion.altRightMetric | β β' ⊗ pauliCo | μ α β}ᵀ.tensor := by
{pauliCo | μ α β ⊗ Fermion.altLeftMetric | α α' ⊗
Fermion.altRightMetric | β β'}ᵀ.tensor := by
rfl
/-- The definitional tensor node relation for `pauliContrDown`. -/
lemma tensorNode_pauliContrDown : {pauliContrDown | μ α β}ᵀ.tensor =
{Fermion.altLeftMetric | α α' ⊗
Fermion.altRightMetric | β β' ⊗ pauliContr | μ α β}ᵀ.tensor := by
{pauliContr | μ α β ⊗ Fermion.altLeftMetric | α α' ⊗
Fermion.altRightMetric | β β'}ᵀ.tensor := by
rfl
end complexLorentzTensor