refactor: Some basic golfings
This commit is contained in:
parent
c739a5eeb8
commit
6db7f86471
5 changed files with 44 additions and 70 deletions
|
@ -342,28 +342,28 @@ lemma pauliMatrix_contr_down_0 :
|
|||
lhs
|
||||
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
|
||||
enter [1, 1, 1, 1, 1, 1, 1, 1]
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
enter [1, 1, 1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
enter [1, 1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
enter [1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; rhs; rhs; lhs
|
||||
enter [1, 1, 1, 1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; rhs; rhs; lhs
|
||||
enter [1, 1, 1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; rhs; lhs;
|
||||
enter [1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs; rhs; rhs; lhs;
|
||||
enter [1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs
|
||||
|
@ -416,12 +416,10 @@ lemma pauliMatrix_contr_down_1 :
|
|||
congr 1
|
||||
· rw [basisVectorContrPauli]
|
||||
congr 1
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
· rw [basisVectorContrPauli]
|
||||
congr 1
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
|
||||
lemma pauliMatrix_contr_down_2 :
|
||||
{(basisVector ![Color.down, Color.down] fun _ => 2) | μ ν ⊗
|
||||
|
|
|
@ -92,28 +92,28 @@ lemma pauliMatrix_contr_lower_0_1_1 :
|
|||
lhs
|
||||
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
|
||||
enter [1, 1, 1, 1, 1, 1, 1, 1]
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
enter [1, 1, 1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
enter [1, 1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
enter [1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; rhs; rhs; lhs
|
||||
enter [1, 1, 1, 1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; rhs; rhs; lhs
|
||||
enter [1, 1, 1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs; lhs; rhs; lhs;
|
||||
enter [1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs; rhs; rhs; lhs;
|
||||
enter [1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
conv =>
|
||||
lhs
|
||||
|
@ -122,11 +122,9 @@ lemma pauliMatrix_contr_lower_0_1_1 :
|
|||
/- Simplifying. -/
|
||||
congr 1
|
||||
· congr 1
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
· congr 1
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
|
||||
lemma pauliMatrix_contr_lower_1_0_1 :
|
||||
{(basisVector pauliCoMap (fun | 0 => 1 | 1 => 0 | 2 => 1)) | μ α β ⊗
|
||||
|
@ -212,12 +210,9 @@ lemma pauliMatrix_contr_lower_1_1_0 :
|
|||
/- Simplifying. -/
|
||||
congr 1
|
||||
· congr 1
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
· congr 1
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
|
||||
decide
|
||||
lemma pauliMatrix_contr_lower_2_0_1 :
|
||||
{(basisVector pauliCoMap (fun | 0 => 2 | 1 => 0 | 2 => 1)) | μ α β ⊗
|
||||
pauliContr | μ α' β'}ᵀ.tensor =
|
||||
|
@ -258,11 +253,9 @@ lemma pauliMatrix_contr_lower_2_0_1 :
|
|||
/- Simplifying. -/
|
||||
congr 1
|
||||
· congr 2
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
· congr 2
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
|
||||
lemma pauliMatrix_contr_lower_2_1_0 :
|
||||
{(basisVector pauliCoMap (fun | 0 => 2 | 1 => 1 | 2 => 0)) | μ α β ⊗
|
||||
|
@ -304,11 +297,9 @@ lemma pauliMatrix_contr_lower_2_1_0 :
|
|||
/- Simplifying. -/
|
||||
congr 1
|
||||
· congr 2
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
· congr 2
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
|
||||
lemma pauliMatrix_contr_lower_3_0_0 :
|
||||
{(basisVector pauliCoMap (fun | 0 => 3 | 1 => 0 | 2 => 0)) | μ α β ⊗
|
||||
|
@ -350,11 +341,9 @@ lemma pauliMatrix_contr_lower_3_0_0 :
|
|||
/- Simplifying. -/
|
||||
congr 1
|
||||
· congr 2
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
· congr 2
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
|
||||
lemma pauliMatrix_contr_lower_3_1_1 :
|
||||
{(basisVector pauliCoMap (fun | 0 => 3 | 1 => 1 | 2 => 1)) | μ α β ⊗
|
||||
|
@ -396,11 +385,9 @@ lemma pauliMatrix_contr_lower_3_1_1 :
|
|||
/- Simplifying. -/
|
||||
congr 1
|
||||
· congr 2
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
· congr 2
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
|
||||
/-! TODO: Work out why `pauliCo_prod_basis_expand'` is needed. -/
|
||||
/-- This lemma is exactly the same as `pauliCo_prod_basis_expand`.
|
||||
|
@ -549,13 +536,9 @@ theorem pauliCo_contr_pauliContr :
|
|||
abel
|
||||
rw [h1]
|
||||
congr
|
||||
· funext i
|
||||
fin_cases i <;> rfl
|
||||
· funext i
|
||||
fin_cases i <;> rfl
|
||||
· funext i
|
||||
fin_cases i <;> rfl
|
||||
· funext i
|
||||
fin_cases i <;> rfl
|
||||
· decide
|
||||
· decide
|
||||
· decide
|
||||
· decide
|
||||
|
||||
end complexLorentzTensor
|
||||
|
|
|
@ -336,9 +336,7 @@ lemma le_all_mul_koszulOrder_ofListLift_expand {I : Type} {f : I → Type} [∀
|
|||
enter [2, 2]
|
||||
intro a
|
||||
simp [optionEraseZ]
|
||||
rhs
|
||||
rhs
|
||||
lhs
|
||||
enter [2, 2, 1]
|
||||
rw [← CreateAnnilateSect.eraseIdx_toList]
|
||||
erw [CreateAnnilateSect.extractEquiv_symm_eraseIdx]
|
||||
rw [← Finset.sum_mul]
|
||||
|
|
|
@ -8,7 +8,9 @@ import Mathlib.Algebra.Lie.OfAssociative
|
|||
import Mathlib.Analysis.Complex.Basic
|
||||
/-!
|
||||
|
||||
# Koszul sign insert
|
||||
# Grade
|
||||
|
||||
The grade `0` (for boson) or `1` (for fermion) of a list of fields.
|
||||
|
||||
-/
|
||||
|
||||
|
|
|
@ -55,29 +55,22 @@ lemma static_wick_cons {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
|
|||
rw [Subalgebra.mem_center_iff] at hi
|
||||
rw [hi, mul_assoc, ← map_mul, hb, add_mul, map_add]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
lhs
|
||||
rw [ofList_eq_smul_one]
|
||||
rw [Algebra.smul_mul_assoc]
|
||||
rw [ofList_singleton]
|
||||
enter [2, 1]
|
||||
rw [ofList_eq_smul_one, Algebra.smul_mul_assoc, ofList_singleton]
|
||||
rw [mul_koszulOrder_le]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
lhs
|
||||
enter [2, 1]
|
||||
rw [← map_smul, ← Algebra.smul_mul_assoc]
|
||||
rw [← ofList_singleton, ← ofList_eq_smul_one]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
enter [2, 2]
|
||||
rw [ofList_eq_smul_one, Algebra.smul_mul_assoc, map_smul]
|
||||
rw [le_all_mul_koszulOrder_ofListLift_expand]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
enter [2, 2]
|
||||
rw [smul_add, Finset.smul_sum]
|
||||
rw [← map_smul, ← map_smul, ← Algebra.smul_mul_assoc, ← ofList_eq_smul_one]
|
||||
rhs
|
||||
rhs
|
||||
enter [2, 2]
|
||||
intro n
|
||||
rw [← Algebra.smul_mul_assoc, smul_comm, ← map_smul, ← LinearMap.map_smul₂,
|
||||
← ofList_eq_smul_one]
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue