refactor: Lint

This commit is contained in:
jstoobysmith 2024-10-24 12:08:35 +00:00
parent 1e8efdb16a
commit c9c7b25ea8
9 changed files with 946 additions and 696 deletions

View file

@ -79,6 +79,15 @@ lemma perm_basisVector {n m : } {c : Fin n → complexLorentzTensor.C}
eqToIso.hom, Functor.mapIso_inv, eqToIso.inv, LinearEquiv.ofLinear_apply]
rw [basis_eq_FDiscrete]
lemma perm_basisVector_tree {n m : } {c : Fin n → complexLorentzTensor.C}
{c1 : Fin m → complexLorentzTensor.C} (σ : OverColor.mk c ⟶ OverColor.mk c1)
(b : Π j, Fin (complexLorentzTensor.repDim (c j))) :
(perm σ (tensorNode (basisVector c b))).tensor =
(tensorNode (basisVector c1 (fun i => Fin.cast (perm_basisVector_cast σ i)
(b ((OverColor.Hom.toEquiv σ).symm i))))).tensor := by
exact perm_basisVector _ _
/-- The scalar determining if contracting two basis vectors together gives zero or not. -/
def contrBasisVectorMul {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
(i : Fin n.succ.succ) (j : Fin n.succ)
(b : Π k, Fin (complexLorentzTensor.repDim (c k))) : :=
@ -89,7 +98,7 @@ lemma contrBasisVectorMul_neg {n : } {c : Fin n.succ.succ → complexLorentzT
(h : ¬ (b i).val = (b (i.succAbove j)).val := by decide) :
contrBasisVectorMul i j b = 0 := by
rw [contrBasisVectorMul]
simp
simp only [ite_eq_else, one_ne_zero, imp_false]
exact h
lemma contrBasisVectorMul_pos {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
@ -97,7 +106,7 @@ lemma contrBasisVectorMul_pos {n : } {c : Fin n.succ.succ → complexLorentzT
(h : (b i).val = (b (i.succAbove j)).val := by decide) :
contrBasisVectorMul i j b = 1 := by
rw [contrBasisVectorMul]
simp
simp only [ite_eq_then, zero_ne_one, imp_false, Decidable.not_not]
exact h
lemma contr_basisVector {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
@ -119,37 +128,41 @@ lemma contr_basisVector {n : } {c : Fin n.succ.succ → complexLorentzTensor.
erw [basis_contr]
rfl
lemma contr_basisVector_tree {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
lemma contr_basisVector_tree {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
{i : Fin n.succ.succ} {j : Fin n.succ} {h : c (i.succAbove j) = complexLorentzTensor.τ (c i)}
(b : Π k, Fin (complexLorentzTensor.repDim (c k))) :
(contr i j h (tensorNode (basisVector c b))).tensor =
(smul (contrBasisVectorMul i j b) (tensorNode ( basisVector (c ∘ Fin.succAbove i ∘ Fin.succAbove j)
(fun k => b (i.succAbove (j.succAbove k)))) )).tensor := by
(smul (contrBasisVectorMul i j b) (tensorNode
(basisVector (c ∘ Fin.succAbove i ∘ Fin.succAbove j)
(fun k => b (i.succAbove (j.succAbove k)))))).tensor := by
exact contr_basisVector _
lemma contr_basisVector_tree_pos {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
lemma contr_basisVector_tree_pos {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
{i : Fin n.succ.succ} {j : Fin n.succ} {h : c (i.succAbove j) = complexLorentzTensor.τ (c i)}
(b : Π k, Fin (complexLorentzTensor.repDim (c k))) (hn : (b i).val = (b (i.succAbove j)).val := by decide) :
(b : Π k, Fin (complexLorentzTensor.repDim (c k)))
(hn : (b i).val = (b (i.succAbove j)).val := by decide) :
(contr i j h (tensorNode (basisVector c b))).tensor =
((tensorNode ( basisVector (c ∘ Fin.succAbove i ∘ Fin.succAbove j)
((tensorNode (basisVector (c ∘ Fin.succAbove i ∘ Fin.succAbove j)
(fun k => b (i.succAbove (j.succAbove k)))))).tensor := by
rw [contr_basisVector_tree, contrBasisVectorMul]
rw [if_pos hn]
simp [smul_tensor]
lemma contr_basisVector_tree_neg {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
lemma contr_basisVector_tree_neg {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
{i : Fin n.succ.succ} {j : Fin n.succ} {h : c (i.succAbove j) = complexLorentzTensor.τ (c i)}
(b : Π k, Fin (complexLorentzTensor.repDim (c k))) (hn : ¬ (b i).val = (b (i.succAbove j)).val := by decide) :
(b : Π k, Fin (complexLorentzTensor.repDim (c k)))
(hn : ¬ (b i).val = (b (i.succAbove j)).val := by decide) :
(contr i j h (tensorNode (basisVector c b))).tensor =
(tensorNode 0).tensor := by
rw [contr_basisVector_tree, contrBasisVectorMul]
rw [if_neg hn]
simp [smul_tensor]
simp only [Nat.succ_eq_add_one, smul_tensor, tensorNode_tensor, zero_smul]
/-- Equivalence of Fin types appearing in the product of two basis vectors. -/
def prodBasisVecEquiv {n m : } {c : Fin n → complexLorentzTensor.C}
{c1 : Fin m → complexLorentzTensor.C} (i : Fin n ⊕ Fin m) :
Sum.elim (fun i => Fin (complexLorentzTensor.repDim (c i))) (fun i => Fin (complexLorentzTensor.repDim (c1 i)))
Sum.elim (fun i => Fin (complexLorentzTensor.repDim (c i))) (fun i =>
Fin (complexLorentzTensor.repDim (c1 i)))
i ≃ Fin (complexLorentzTensor.repDim ((Sum.elim c c1 i))) :=
match i with
| Sum.inl _ => Equiv.refl _
@ -184,17 +197,17 @@ lemma prod_basisVector {n m : } {c : Fin n → complexLorentzTensor.C}
| Sum.inl k => rfl
| Sum.inr k => rfl
lemma prod_basisVector_tree {n m : } {c : Fin n → complexLorentzTensor.C}
lemma prod_basisVector_tree {n m : } {c : Fin n → complexLorentzTensor.C}
{c1 : Fin m → complexLorentzTensor.C}
(b : Π k, Fin (complexLorentzTensor.repDim (c k)))
(b1 : Π k, Fin (complexLorentzTensor.repDim (c1 k))) :
(prod (tensorNode (basisVector c b)) (tensorNode (basisVector c1 b1))).tensor =
(prod (tensorNode (basisVector c b)) (tensorNode (basisVector c1 b1))).tensor =
(tensorNode (basisVector (Sum.elim c c1 ∘ finSumFinEquiv.symm) (fun i =>
prodBasisVecEquiv (finSumFinEquiv.symm i)
((HepLean.PiTensorProduct.elimPureTensor b b1) (finSumFinEquiv.symm i))))).tensor := by
exact prod_basisVector _ _
lemma eval_basisVector {n : } {c : Fin n.succ → complexLorentzTensor.C}
lemma eval_basisVector {n : } {c : Fin n.succ → complexLorentzTensor.C}
{i : Fin n.succ} (j : Fin (complexLorentzTensor.repDim (c i)))
(b : Π k, Fin (complexLorentzTensor.repDim (c k))) :
(eval i j (tensorNode (basisVector c b))).tensor = (if j = b i then (1 : ) else 0) •
@ -282,7 +295,7 @@ lemma contrMatrix_basis_expand_tree : {Lorentz.contrMetric | μ ν}ᵀ.tensor =
contrMatrix_basis_expand
lemma leftMetric_expand : {Fermion.leftMetric | α β}ᵀ.tensor =
- basisVector ![Color.upL, Color.upL] (fun | 0 => 0 | 1 => 1)
- basisVector ![Color.upL, Color.upL] (fun | 0 => 0 | 1 => 1)
+ basisVector ![Color.upL, Color.upL] (fun | 0 => 1 | 1 => 0) := by
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, constTwoNode_tensor,
Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V, Fin.isValue]
@ -299,7 +312,8 @@ lemma leftMetric_expand : {Fermion.leftMetric | α β}ᵀ.tensor =
· rfl
lemma leftMetric_expand_tree : {Fermion.leftMetric | α β}ᵀ.tensor =
(TensorTree.add (smul (-1) (tensorNode (basisVector ![Color.upL, Color.upL] (fun | 0 => 0 | 1 => 1)))) <|
(TensorTree.add (smul (-1) (tensorNode (basisVector ![Color.upL, Color.upL]
(fun | 0 => 0 | 1 => 1)))) <|
(tensorNode (basisVector ![Color.upL, Color.upL] (fun | 0 => 1 | 1 => 0)))).tensor :=
leftMetric_expand
@ -320,8 +334,10 @@ lemma altLeftMetric_expand : {Fermion.altLeftMetric | α β}ᵀ.tensor =
· rfl
lemma altLeftMetric_expand_tree : {Fermion.altLeftMetric | α β}ᵀ.tensor =
(TensorTree.add (tensorNode (basisVector ![Color.downL, Color.downL] (fun | 0 => 0 | 1 => 1))) <|
(smul (-1) (tensorNode (basisVector ![Color.downL, Color.downL] (fun | 0 => 1 | 1 => 0))))).tensor :=
(TensorTree.add (tensorNode (basisVector ![Color.downL, Color.downL]
(fun | 0 => 0 | 1 => 1))) <|
(smul (-1) (tensorNode (basisVector ![Color.downL, Color.downL]
(fun | 0 => 1 | 1 => 0))))).tensor :=
altLeftMetric_expand
lemma rightMetric_expand : {Fermion.rightMetric | α β}ᵀ.tensor =
@ -342,7 +358,8 @@ lemma rightMetric_expand : {Fermion.rightMetric | α β}ᵀ.tensor =
· rfl
lemma rightMetric_expand_tree : {Fermion.rightMetric | α β}ᵀ.tensor =
(TensorTree.add (smul (-1) (tensorNode (basisVector ![Color.upR, Color.upR] (fun | 0 => 0 | 1 => 1)))) <|
(TensorTree.add (smul (-1) (tensorNode (basisVector ![Color.upR, Color.upR]
(fun | 0 => 0 | 1 => 1)))) <|
(tensorNode (basisVector ![Color.upR, Color.upR] (fun | 0 => 1 | 1 => 0)))).tensor :=
rightMetric_expand
@ -363,8 +380,10 @@ lemma altRightMetric_expand : {Fermion.altRightMetric | α β}ᵀ.tensor =
· rfl
lemma altRightMetric_expand_tree : {Fermion.altRightMetric | α β}ᵀ.tensor =
(TensorTree.add (tensorNode (basisVector ![Color.downR, Color.downR] (fun | 0 => 0 | 1 => 1))) <|
(smul (-1) (tensorNode (basisVector ![Color.downR, Color.downR] (fun | 0 => 1 | 1 => 0))))).tensor :=
(TensorTree.add (tensorNode (basisVector
![Color.downR, Color.downR] (fun | 0 => 0 | 1 => 1))) <|
(smul (-1) (tensorNode (basisVector ![Color.downR, Color.downR]
(fun | 0 => 1 | 1 => 0))))).tensor :=
altRightMetric_expand
/-- The expansion of the Pauli matrices `σ^μ^a^{dot a}` in terms of basis vectors. -/
@ -406,13 +425,13 @@ lemma pauliMatrix_basis_expand : {PauliMatrix.asConsTensor | μ α β}ᵀ.tensor
| (2 : Fin 3) => rfl
lemma pauliMatrix_basis_expand_tree : {PauliMatrix.asConsTensor | μ α β}ᵀ.tensor =
(TensorTree.add (tensorNode
(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
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)))) <|
@ -421,7 +440,8 @@ lemma pauliMatrix_basis_expand_tree : {PauliMatrix.asConsTensor | μ α β}ᵀ.t
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
(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]