feat: Add action node

This commit is contained in:
jstoobysmith 2024-10-29 12:32:33 +00:00
parent f7499f8d86
commit 202df13a17
5 changed files with 175 additions and 0 deletions

View file

@ -185,4 +185,63 @@ lemma pauliContrDown_eq_metric_mul_pauliContr :
apply perm_congr _ rfl
decide
/-!
## Group actions
-/
/-- The tensor `pauliContr` is invariant under the action of `SL(2,)`. -/
lemma action_pauliContr (g : SL(2,)) : {g •ₐ pauliContr | μ α β}ᵀ.tensor =
{pauliContr | μ α β}ᵀ.tensor := by
rw [tensorNode_pauliContr, constThreeNodeE]
rw [← action_constThreeNode _ g]
rfl
/-- The tensor `pauliCo` is invariant under the action of `SL(2,)`. -/
lemma action_pauliCo (g : SL(2,)) : {g •ₐ pauliCo | μ α β}ᵀ.tensor =
{pauliCo | μ α β}ᵀ.tensor := by
conv =>
lhs
rw [action_tensor_eq <| tensorNode_pauliCo]
rw [(contr_action _ _).symm]
rw [contr_tensor_eq <| (prod_action _ _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| action_constTwoNode _ _]
rw [contr_tensor_eq <| prod_tensor_eq_snd <| action_constThreeNode _ _]
rfl
/-- The tensor `pauliCoDown` is invariant under the action of `SL(2,)`. -/
lemma action_pauliCoDown (g : SL(2,)) : {g •ₐ pauliCoDown | μ α β}ᵀ.tensor =
{pauliCoDown | μ α β}ᵀ.tensor := by
conv =>
lhs
rw [action_tensor_eq <| tensorNode_pauliCoDown]
rw [(contr_action _ _).symm]
rw [contr_tensor_eq <| (prod_action _ _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| (contr_action _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| (prod_action _ _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| prod_tensor_eq_fst <|
action_pauliCo _]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| prod_tensor_eq_snd <|
action_constTwoNode _ _]
rw [contr_tensor_eq <| prod_tensor_eq_snd <| action_constTwoNode _ _]
rfl
/-- The tensor `pauliContrDown` is invariant under the action of `SL(2,)`. -/
lemma action_pauliContrDown (g : SL(2,)) : {g •ₐ pauliContrDown | μ α β}ᵀ.tensor =
{pauliContrDown | μ α β}ᵀ.tensor := by
conv =>
lhs
rw [action_tensor_eq <| tensorNode_pauliContrDown]
rw [(contr_action _ _).symm]
rw [contr_tensor_eq <| (prod_action _ _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| (contr_action _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| (prod_action _ _ _).symm]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| prod_tensor_eq_fst <|
action_pauliContr _]
rw [contr_tensor_eq <| prod_tensor_eq_fst <| contr_tensor_eq <| prod_tensor_eq_snd <|
action_constTwoNode _ _]
rw [contr_tensor_eq <| prod_tensor_eq_snd <| action_constTwoNode _ _]
rfl
end complexLorentzTensor