fix: Slow builds with tensors
This commit is contained in:
parent
c9c9047a0c
commit
e24fd5b40b
3 changed files with 56 additions and 51 deletions
|
@ -63,17 +63,18 @@ lemma tensorNode_pauliContr : {pauliContr | μ α β}ᵀ.tensor =
|
|||
|
||||
/-- The definitional tensor node relation for `pauliCo`. -/
|
||||
lemma tensorNode_pauliCo : {pauliCo | μ α β}ᵀ.tensor =
|
||||
{η' | μ ν ⊗ PauliMatrix.asConsTensor | ν α β}ᵀ.tensor := by
|
||||
rfl
|
||||
{η' | μ ν ⊗ pauliContr | ν α β}ᵀ.tensor := by
|
||||
rw [pauliCo, tensorNode_tensor]
|
||||
|
||||
/-- The definitional tensor node relation for `pauliCoDown`. -/
|
||||
lemma tensorNode_pauliCoDown : {pauliCoDown | μ α β}ᵀ.tensor =
|
||||
{pauliCo | μ α β ⊗ εL' | α α' ⊗ εR' | β β'}ᵀ.tensor := by
|
||||
rfl
|
||||
rw [pauliCoDown, tensorNode_tensor]
|
||||
|
||||
/-- The definitional tensor node relation for `pauliContrDown`. -/
|
||||
lemma tensorNode_pauliContrDown : {pauliContrDown | μ α β}ᵀ.tensor =
|
||||
{pauliContr | μ α β ⊗ εL' | α α' ⊗ εR' | β β'}ᵀ.tensor := by
|
||||
rw [pauliContr, tensorNode_tensor]
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
@ -200,10 +201,15 @@ lemma action_pauliCo (g : SL(2,ℂ)) : {g •ₐ pauliCo | μ α β}ᵀ.tensor =
|
|||
conv =>
|
||||
lhs
|
||||
rw [action_tensor_eq <| tensorNode_pauliCo]
|
||||
rw [action_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_snd <| tensorNode_pauliContr]
|
||||
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 _ _]
|
||||
conv =>
|
||||
rhs
|
||||
rw [tensorNode_pauliCo]
|
||||
rw [contr_tensor_eq <| prod_tensor_eq_snd <| tensorNode_pauliContr]
|
||||
rfl
|
||||
|
||||
/-- The tensor `pauliCoDown` is invariant under the action of `SL(2,ℂ)`. -/
|
||||
|
@ -219,9 +225,11 @@ lemma action_pauliCoDown (g : SL(2,ℂ)) : {g •ₐ pauliCoDown | μ α β}ᵀ.
|
|||
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 _ _]
|
||||
erw [contr_tensor_eq <| prod_tensor_eq_snd <| action_constTwoNode _ _]
|
||||
rfl
|
||||
action_altLeftMetric _]
|
||||
rw [contr_tensor_eq <| prod_tensor_eq_snd <| action_altRightMetric _]
|
||||
conv =>
|
||||
rhs
|
||||
rw [tensorNode_pauliCoDown]
|
||||
|
||||
/-- The tensor `pauliContrDown` is invariant under the action of `SL(2,ℂ)`. -/
|
||||
lemma action_pauliContrDown (g : SL(2,ℂ)) : {g •ₐ pauliContrDown | μ α β}ᵀ.tensor =
|
||||
|
@ -236,8 +244,10 @@ lemma action_pauliContrDown (g : SL(2,ℂ)) : {g •ₐ pauliContrDown | μ α
|
|||
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 _ _]
|
||||
erw [contr_tensor_eq <| prod_tensor_eq_snd <| action_constTwoNode _ _]
|
||||
rfl
|
||||
action_altLeftMetric _]
|
||||
erw [contr_tensor_eq <| prod_tensor_eq_snd <| action_altRightMetric _]
|
||||
conv =>
|
||||
rhs
|
||||
rw [tensorNode_pauliContrDown]
|
||||
|
||||
end complexLorentzTensor
|
||||
|
|
|
@ -458,15 +458,10 @@ lemma pauliMatrix_contr_down_2 :
|
|||
conv =>
|
||||
lhs
|
||||
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
|
||||
congr 1
|
||||
· rw [basisVectorContrPauli]
|
||||
congr 2
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
· rw [basisVectorContrPauli]
|
||||
congr 2
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
rw [basisVectorContrPauli, basisVectorContrPauli]
|
||||
congr 3
|
||||
· decide
|
||||
· decide
|
||||
|
||||
lemma pauliMatrix_contr_down_3 :
|
||||
{(basisVector ![Color.down, Color.down] fun x => 3) | μ ν ⊗
|
||||
|
@ -503,16 +498,10 @@ lemma pauliMatrix_contr_down_3 :
|
|||
conv =>
|
||||
lhs
|
||||
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
|
||||
congr 1
|
||||
· rw [basisVectorContrPauli]
|
||||
congr 1
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
· rw [basisVectorContrPauli]
|
||||
congr 1
|
||||
congr 1
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
rw [basisVectorContrPauli, basisVectorContrPauli]
|
||||
congr 3
|
||||
· decide
|
||||
· decide
|
||||
|
||||
/-- The expansion of `pauliCo` in terms of a basis. -/
|
||||
lemma pauliCo_basis_expand : pauliCo
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue