refactor: Text based Lint
This commit is contained in:
parent
319089ad54
commit
7010a1dae2
12 changed files with 54 additions and 52 deletions
|
@ -209,10 +209,10 @@ lemma basis_contr (c : complexLorentzTensor.C) (i : Fin (complexLorentzTensor.re
|
|||
| Color.up => Lorentz.contrCoContraction_basis _ _
|
||||
| Color.down => Lorentz.coContrContraction_basis _ _
|
||||
|
||||
instance {n : ℕ} {c : Fin n → complexLorentzTensor.C} :
|
||||
instance {n : ℕ} {c : Fin n → complexLorentzTensor.C} :
|
||||
DecidableEq (OverColor.mk c).left := instDecidableEqFin n
|
||||
|
||||
instance {n : ℕ} {c : Fin n → complexLorentzTensor.C} :
|
||||
instance {n : ℕ} {c : Fin n → complexLorentzTensor.C} :
|
||||
Fintype (OverColor.mk c).left := Fin.fintype n
|
||||
|
||||
instance {n m : ℕ} {c : Fin n → complexLorentzTensor.C}
|
||||
|
@ -220,6 +220,5 @@ instance {n m : ℕ} {c : Fin n → complexLorentzTensor.C}
|
|||
Decidable (σ = σ') :=
|
||||
decidable_of_iff _ (OverColor.Hom.ext_iff σ σ')
|
||||
|
||||
|
||||
end complexLorentzTensor
|
||||
end
|
||||
|
|
|
@ -76,7 +76,7 @@ lemma tensorNode_contrBispinorDown (p : complexContr) :
|
|||
rw [contrBispinorDown, tensorNode_tensor]
|
||||
|
||||
/-- The definitional tensor node relation for `coBispinorUp`. -/
|
||||
lemma tensorNode_coBispinorUp (p : complexCo) :
|
||||
lemma tensorNode_coBispinorUp (p : complexCo) :
|
||||
{coBispinorUp p | α β}ᵀ.tensor = {pauliContr | μ α β ⊗ p | μ}ᵀ.tensor := by
|
||||
rw [coBispinorUp, tensorNode_tensor]
|
||||
|
||||
|
@ -94,23 +94,26 @@ lemma tensorNode_coBispinorDown (p : complexCo) :
|
|||
-/
|
||||
|
||||
lemma contrBispinorDown_expand (p : complexContr) :
|
||||
{contrBispinorDown p | α β}ᵀ.tensor = {Fermion.altLeftMetric | α α' ⊗ Fermion.altRightMetric | β β' ⊗
|
||||
{contrBispinorDown p | α β}ᵀ.tensor =
|
||||
{Fermion.altLeftMetric | α α' ⊗ Fermion.altRightMetric | β β' ⊗
|
||||
(pauliCo | μ α β ⊗ p | μ)}ᵀ.tensor := by
|
||||
rw [tensorNode_contrBispinorDown p]
|
||||
rw [contr_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_snd <| tensorNode_contrBispinorUp p]
|
||||
|
||||
lemma coBispinorDown_expand (p : complexCo) :
|
||||
{coBispinorDown p | α β}ᵀ.tensor = {Fermion.altLeftMetric | α α' ⊗ Fermion.altRightMetric | β β' ⊗
|
||||
{coBispinorDown p | α β}ᵀ.tensor =
|
||||
{Fermion.altLeftMetric | α α' ⊗ Fermion.altRightMetric | β β' ⊗
|
||||
(pauliContr | μ α β ⊗ p | μ)}ᵀ.tensor := by
|
||||
rw [tensorNode_coBispinorDown p]
|
||||
rw [contr_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_snd <| tensorNode_coBispinorUp p]
|
||||
|
||||
set_option maxRecDepth 5000 in
|
||||
lemma contrBispinorDown_eq_pauliCoDown_contr (p : complexContr) :
|
||||
{contrBispinorDown p | α β = pauliCoDown | μ α β ⊗ p | μ}ᵀ := by
|
||||
{contrBispinorDown p | α β = pauliCoDown | μ α β ⊗ p | μ}ᵀ := by
|
||||
conv =>
|
||||
rhs
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_fst <| pauliCoDown_eq_metric_mul_pauliCo]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_fst <|
|
||||
pauliCoDown_eq_metric_mul_pauliCo]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| prod_perm_left _ _ _ _]
|
||||
rw [perm_tensor_eq <| perm_contr_congr 2 2]
|
||||
rw [perm_perm]
|
||||
|
@ -118,12 +121,14 @@ lemma contrBispinorDown_eq_pauliCoDown_contr (p : complexContr) :
|
|||
rw [perm_tensor_eq <| perm_contr_congr 2 2]
|
||||
rw [perm_perm]
|
||||
apply (perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| perm_eq_id _ rfl _).trans
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_prod _ _ _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_prod _ _ _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| perm_contr_congr 1 3]
|
||||
rw [perm_tensor_eq <| perm_contr_congr 2 2]
|
||||
rw [perm_perm]
|
||||
erw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| perm_eq_id _ rfl _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| prod_assoc' _ _ _ _ _ _]
|
||||
erw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <|
|
||||
perm_eq_id _ rfl _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <|
|
||||
prod_assoc' _ _ _ _ _ _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| perm_contr_congr 0 4]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| perm_contr_congr 1 3]
|
||||
rw [perm_tensor_eq <| perm_contr_congr 2 2]
|
||||
|
@ -144,10 +149,11 @@ lemma contrBispinorDown_eq_pauliCoDown_contr (p : complexContr) :
|
|||
|
||||
set_option maxRecDepth 5000 in
|
||||
lemma coBispinorDown_eq_pauliContrDown_contr (p : complexCo) :
|
||||
{coBispinorDown p | α β = pauliContrDown | μ α β ⊗ p | μ}ᵀ := by
|
||||
{coBispinorDown p | α β = pauliContrDown | μ α β ⊗ p | μ}ᵀs := by
|
||||
conv =>
|
||||
rhs
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_fst <| pauliContrDown_eq_metric_mul_pauliContr]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| prod_tensor_eq_fst <|
|
||||
pauliContrDown_eq_metric_mul_pauliContr]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| prod_perm_left _ _ _ _]
|
||||
rw [perm_tensor_eq <| perm_contr_congr 2 2]
|
||||
rw [perm_perm]
|
||||
|
@ -155,12 +161,14 @@ lemma coBispinorDown_eq_pauliContrDown_contr (p : complexCo) :
|
|||
rw [perm_tensor_eq <| perm_contr_congr 2 2]
|
||||
rw [perm_perm]
|
||||
apply (perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| perm_eq_id _ rfl _).trans
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_prod _ _ _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_prod _ _ _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| perm_contr_congr 1 3]
|
||||
rw [perm_tensor_eq <| perm_contr_congr 2 2]
|
||||
rw [perm_perm]
|
||||
erw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| perm_eq_id _ rfl _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| prod_assoc' _ _ _ _ _ _]
|
||||
erw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <|
|
||||
perm_eq_id _ rfl _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <|
|
||||
prod_assoc' _ _ _ _ _ _]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| contr_tensor_eq <| perm_contr_congr 0 4]
|
||||
rw [perm_tensor_eq <| contr_tensor_eq <| perm_contr_congr 1 3]
|
||||
rw [perm_tensor_eq <| perm_contr_congr 2 2]
|
||||
|
|
|
@ -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 _ _ _]
|
||||
|
|
|
@ -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)
|
||||
|
|
|
@ -116,15 +116,16 @@ lemma perm_eq_of_eq_perm {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C}
|
|||
simp only [Iso.map_hom_inv_id, Action.id_hom, ModuleCat.id_apply]
|
||||
|
||||
lemma perm_eq_iff_eq_perm {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C}
|
||||
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1))
|
||||
(σ : (OverColor.mk c) ⟶ (OverColor.mk c1))
|
||||
{t : TensorTree S c} {t2 : TensorTree S c1} :
|
||||
(perm σ t).tensor = t2.tensor ↔ t.tensor =
|
||||
(perm (equivToHomEq (Hom.toEquiv σ).symm (fun x => Hom.toEquiv_comp_apply σ x)) t2).tensor := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· simp [perm_tensor, ← h]
|
||||
change _ = (S.F.map _ ≫ S.F.map _).hom _
|
||||
rw [← S.F.map_comp]
|
||||
have h1 : (σ ≫ equivToHomEq (Hom.toEquiv σ).symm (fun x => Hom.toEquiv_comp_apply σ x)) = 𝟙 _ := by
|
||||
have h1 : (σ ≫ equivToHomEq (Hom.toEquiv σ).symm
|
||||
(fun x => Hom.toEquiv_comp_apply σ x)) = 𝟙 _ := by
|
||||
apply Hom.ext
|
||||
ext x
|
||||
change (Hom.toEquiv σ).symm ((Hom.toEquiv σ) x) = x
|
||||
|
@ -134,7 +135,8 @@ lemma perm_eq_iff_eq_perm {n m : ℕ} {c : Fin n → S.C} {c1 : Fin m → S.C}
|
|||
· rw [perm_tensor, h]
|
||||
change (S.F.map _ ≫ S.F.map _).hom _ = _
|
||||
rw [← S.F.map_comp]
|
||||
have h1 : (equivToHomEq (Hom.toEquiv σ).symm (fun x => Hom.toEquiv_comp_apply σ x) ≫ σ) = 𝟙 _ := by
|
||||
have h1 : (equivToHomEq (Hom.toEquiv σ).symm
|
||||
(fun x => Hom.toEquiv_comp_apply σ x) ≫ σ) = 𝟙 _ := by
|
||||
apply Hom.ext
|
||||
ext x
|
||||
change (Hom.toEquiv σ) ((Hom.toEquiv σ).symm x) = x
|
||||
|
|
|
@ -30,21 +30,20 @@ variable {n m : ℕ}
|
|||
|
||||
lemma perm_congr {c1 : Fin n → S.C} {c2 : Fin m → S.C} {T T' : TensorTree S c1}
|
||||
{σ σ' : OverColor.mk c1 ⟶ OverColor.mk c2}
|
||||
(h : σ = σ') (hT : T.tensor = T'.tensor):
|
||||
(h : σ = σ') (hT : T.tensor = T'.tensor) :
|
||||
(perm σ T).tensor = (perm σ' T').tensor := by
|
||||
rw [h]
|
||||
simp only [perm_tensor, hT]
|
||||
|
||||
lemma perm_update {c1 : Fin n → S.C} {c2 : Fin m → S.C} {T : TensorTree S c1}
|
||||
lemma perm_update {c1 : Fin n → S.C} {c2 : Fin m → S.C} {T : TensorTree S c1}
|
||||
{σ : OverColor.mk c1 ⟶ OverColor.mk c2} (σ' : OverColor.mk c1 ⟶ OverColor.mk c2)
|
||||
(h : σ = σ') :
|
||||
(perm σ T).tensor = (perm σ' T).tensor := by rw [h]
|
||||
|
||||
lemma contr_congr {n : ℕ} {c : Fin n.succ.succ → S.C} {i : Fin n.succ.succ}
|
||||
(i' : Fin n.succ.succ) {j : Fin n.succ} (j' : Fin n.succ)
|
||||
{h : c (i.succAbove j) = S.τ (c i)} {t : TensorTree S c}
|
||||
(hi : i = i' := by decide) (hj : j = j' := by decide)
|
||||
:(contr i j h t).tensor =
|
||||
(i' : Fin n.succ.succ) {j : Fin n.succ} (j' : Fin n.succ){h : c (i.succAbove j) = S.τ (c i)}
|
||||
{t : TensorTree S c} (hi : i = i' := by decide) (hj : j = j' := by decide) :
|
||||
(contr i j h t).tensor =
|
||||
(perm (mkIso (by rw [hi, hj])).hom (contr i' j' (by rw [← hi, ← hj, h]) t)).tensor := by
|
||||
subst hi
|
||||
subst hj
|
||||
|
|
|
@ -285,15 +285,13 @@ def contrContrPerm {n : ℕ} {c : Fin n.succ.succ.succ.succ → S.C} {i : Fin n.
|
|||
{j : Fin n.succ.succ.succ} {k : Fin n.succ.succ} {l : Fin n.succ}
|
||||
(hij : c (i.succAbove j) = S.τ (c i)) (hkl : (c ∘ i.succAbove ∘ j.succAbove) (k.succAbove l) =
|
||||
S.τ ((c ∘ i.succAbove ∘ j.succAbove) k)) :
|
||||
OverColor.mk
|
||||
((c ∘
|
||||
(ContrQuartet.mk i j k l hij hkl).swapI.succAbove ∘
|
||||
(ContrQuartet.mk i j k l hij hkl).swapJ.succAbove) ∘
|
||||
OverColor.mk ((c ∘ (ContrQuartet.mk i j k l hij hkl).swapI.succAbove ∘
|
||||
(ContrQuartet.mk i j k l hij hkl).swapJ.succAbove) ∘
|
||||
(ContrQuartet.mk i j k l hij hkl).swapK.succAbove ∘
|
||||
(ContrQuartet.mk i j k l hij hkl).swapL.succAbove) ⟶
|
||||
OverColor.mk
|
||||
((c ∘ i.succAbove ∘ j.succAbove) ∘ k.succAbove ∘ l.succAbove)
|
||||
:= (ContrQuartet.mk i j k l hij hkl).contrSwapHom
|
||||
(ContrQuartet.mk i j k l hij hkl).swapL.succAbove) ⟶
|
||||
OverColor.mk
|
||||
((c ∘ i.succAbove ∘ j.succAbove) ∘ k.succAbove ∘ l.succAbove) :=
|
||||
(ContrQuartet.mk i j k l hij hkl).contrSwapHom
|
||||
|
||||
/-- Contraction nodes commute on adjusting indices. -/
|
||||
theorem contr_contr {n : ℕ} {c : Fin n.succ.succ.succ.succ → S.C} {i : Fin n.succ.succ.succ.succ}
|
||||
|
|
|
@ -264,8 +264,8 @@ lemma perm_contr_congr_mkIso_cond {n : ℕ} {c : Fin n.succ.succ → S.C} {c1 :
|
|||
{i' : Fin n.succ.succ} {j' : Fin n.succ}
|
||||
(hi : i' = ((Hom.toEquiv σ).symm i))
|
||||
(hj : j' = (((Hom.toEquiv (extractOne i σ))).symm j)) :
|
||||
c ∘ i'.succAbove ∘ j'.succAbove =
|
||||
c ∘ Fin.succAbove ((Hom.toEquiv σ).symm i) ∘ Fin.succAbove ((Hom.toEquiv (extractOne i σ)).symm j) := by
|
||||
c ∘ i'.succAbove ∘ j'.succAbove = c ∘ Fin.succAbove ((Hom.toEquiv σ).symm i) ∘
|
||||
Fin.succAbove ((Hom.toEquiv (extractOne i σ)).symm j) := by
|
||||
rw [hi, hj]
|
||||
|
||||
lemma perm_contr_congr_contr_cond {n : ℕ} {c : Fin n.succ.succ → S.C} {c1 : Fin n.succ.succ → S.C}
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue