refactor: Text based Lint

This commit is contained in:
jstoobysmith 2024-10-29 11:23:08 +00:00
parent 319089ad54
commit 7010a1dae2
12 changed files with 54 additions and 52 deletions

View file

@ -63,7 +63,7 @@ lemma tensorNode_pauliContr : {pauliContr | μ α β}ᵀ.tensor =
rfl
/-- The definitional tensor node relation for `pauliCo`. -/
lemma tensorNode_pauliCo : {pauliCo | μ α β}ᵀ.tensor =
lemma tensorNode_pauliCo : {pauliCo | μ α β}ᵀ.tensor =
{Lorentz.coMetric | μ ν ⊗ PauliMatrix.asConsTensor | ν α β}ᵀ.tensor := by
rfl
@ -111,7 +111,7 @@ lemma pauliCoDown_eq_metric_mul_pauliCo :
rw [perm_perm]
rw [perm_tensor_eq <| contr_congr 1 2]
rw [perm_perm]
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| prod_comm _ _ _ _]
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| prod_comm _ _ _ _]
rw [perm_tensor_eq <| contr_tensor_eq <| perm_contr _ _]
rw [perm_tensor_eq <| contr_tensor_eq <| perm_tensor_eq <| contr_congr 5 0]
rw [perm_tensor_eq <| contr_tensor_eq <| perm_perm _ _ _]
@ -138,7 +138,7 @@ lemma pauliCoDown_eq_metric_mul_pauliCo :
set_option maxRecDepth 5000 in
/-- A rearanging of `pauliContrDown` to place the pauli matrices on the right. -/
lemma pauliContrDown_eq_metric_mul_pauliContr :
{pauliContrDown | μ α' β' = Fermion.altLeftMetric | α α' ⊗
{pauliContrDown | μ α' β' = Fermion.altLeftMetric | α α' ⊗
Fermion.altRightMetric | β β' ⊗ pauliContr | μ α β}ᵀ := by
conv =>
lhs
@ -161,7 +161,7 @@ lemma pauliContrDown_eq_metric_mul_pauliContr :
rw [perm_perm]
rw [perm_tensor_eq <| contr_congr 1 2]
rw [perm_perm]
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| prod_comm _ _ _ _]
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| prod_comm _ _ _ _]
rw [perm_tensor_eq <| contr_tensor_eq <| perm_contr _ _]
rw [perm_tensor_eq <| contr_tensor_eq <| perm_tensor_eq <| contr_congr 5 0]
rw [perm_tensor_eq <| contr_tensor_eq <| perm_perm _ _ _]

View file

@ -26,7 +26,6 @@ noncomputable section
namespace complexLorentzTensor
open Fermion
/-!
## Expanding pauliContr in a basis.
@ -94,7 +93,6 @@ lemma pauliContr_basis_expand_tree : {pauliContr | μ α β}ᵀ.tensor =
smul_tensor, neg_smul, one_smul]
rfl
/-- The map to colors one gets when contracting with Pauli matrices on the right. -/
abbrev pauliMatrixContrMap {n : } (c : Fin n → complexLorentzTensor.C) :=
(Sum.elim c ![Color.up, Color.upL, Color.upR] ∘ ⇑finSumFinEquiv.symm)
@ -325,14 +323,12 @@ lemma basis_contr_pauliMatrix_basis_tree_expand_tensor {n : } {c : Fin n →
simp_all only [Function.comp_apply, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue]
rfl
/-!
## Expanding pauliCo in a basis.
-/
/-- The map to color one gets when lowering the indices of pauli matrices. -/
def pauliCoMap := ((Sum.elim ![Color.down, Color.down] ![Color.up, Color.upL, Color.upR] ∘
⇑finSumFinEquiv.symm) ∘ Fin.succAbove 1 ∘ Fin.succAbove 1)