refactor: Pauli matrices

This commit is contained in:
jstoobysmith 2024-10-29 11:10:26 +00:00
parent 7a50680794
commit d7d435a1f8
7 changed files with 650 additions and 700 deletions

View file

@ -386,66 +386,5 @@ lemma altRightMetric_expand_tree : {Fermion.altRightMetric | α β}ᵀ.tensor =
(fun | 0 => 1 | 1 => 0))))).tensor :=
altRightMetric_expand
/-- The expansion of the Pauli matrices `σ^μ^a^{dot a}` in terms of basis vectors. -/
lemma pauliMatrix_basis_expand : {PauliMatrix.asConsTensor | μ α β}ᵀ.tensor =
basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 0 | 1 => 0 | 2 => 0)
+ basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 0 | 1 => 1 | 2 => 1)
+ basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 1 | 1 => 0 | 2 => 1)
+ basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 1 | 1 => 1 | 2 => 0)
- I • basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 2 | 1 => 0 | 2 => 1)
+ I • basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 2 | 1 => 1 | 2 => 0)
+ basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 3 | 1 => 0 | 2 => 0)
- basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 3 | 1 => 1 | 2 => 1) := by
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, constThreeNode_tensor,
Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V, Fin.isValue]
erw [PauliMatrix.asConsTensor_apply_one, PauliMatrix.asTensor_expand]
simp only [Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, Action.instMonoidalCategory_tensorObj_V,
Fin.isValue, map_sub, map_add, _root_.map_smul]
congr 1
congr 1
congr 1
congr 1
congr 1
congr 1
congr 1
all_goals
erw [tripleIsoSep_tmul, basisVector]
apply congrArg
try apply congrArg
funext i
match i with
| (0 : Fin 3) =>
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, OverColor.mk_hom,
cons_val_zero, Fin.cases_zero]
change _ = Lorentz.complexContrBasisFin4 _
simp only [Fin.isValue, Lorentz.complexContrBasisFin4, Basis.coe_reindex, Function.comp_apply]
rfl
| (1 : Fin 3) => rfl
| (2 : Fin 3) => rfl
lemma pauliMatrix_basis_expand_tree : {PauliMatrix.asConsTensor | μ α β}ᵀ.tensor =
(TensorTree.add (tensorNode
(basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 0 | 1 => 0 | 2 => 0))) <|
TensorTree.add (tensorNode
(basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 0 | 1 => 1 | 2 => 1))) <|
TensorTree.add (tensorNode
(basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 1 | 1 => 0 | 2 => 1))) <|
TensorTree.add (tensorNode
(basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 1 | 1 => 1 | 2 => 0))) <|
TensorTree.add (smul (-I) (tensorNode
(basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 2 | 1 => 0 | 2 => 1)))) <|
TensorTree.add (smul I (tensorNode
(basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 2 | 1 => 1 | 2 => 0)))) <|
TensorTree.add (tensorNode
(basisVector ![Color.up, Color.upL, Color.upR] (fun | 0 => 3 | 1 => 0 | 2 => 0))) <|
(smul (-1) (tensorNode
(basisVector ![Color.up, Color.upL, Color.upR]
(fun | 0 => 3 | 1 => 1 | 2 => 1))))).tensor := by
rw [pauliMatrix_basis_expand]
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, add_tensor, tensorNode_tensor,
smul_tensor, neg_smul, one_smul]
rfl
end complexLorentzTensor
end