Merge pull request #271 from HEPLean/WickContract
refactor: renaming variables and golfing
This commit is contained in:
commit
508850fd4e
12 changed files with 180 additions and 226 deletions
|
@ -104,7 +104,6 @@ import HepLean.Mathematics.LinearMaps
|
|||
import HepLean.Mathematics.List
|
||||
import HepLean.Mathematics.PiTensorProduct
|
||||
import HepLean.Mathematics.SO3.Basic
|
||||
import HepLean.Mathematics.SuperAlgebra.Basic
|
||||
import HepLean.Meta.AllFilePaths
|
||||
import HepLean.Meta.Basic
|
||||
import HepLean.Meta.Informal.Basic
|
||||
|
|
|
@ -9,15 +9,11 @@ import HepLean.Lorentz.ComplexTensor.PauliMatrices.Basic
|
|||
## Pauli matrices and the basis of complex Lorentz tensors
|
||||
|
||||
-/
|
||||
open IndexNotation
|
||||
open CategoryTheory
|
||||
open MonoidalCategory
|
||||
open Matrix
|
||||
open MatrixGroups
|
||||
open Complex
|
||||
open TensorProduct
|
||||
open IndexNotation
|
||||
open CategoryTheory
|
||||
open TensorTree
|
||||
open OverColor.Discrete
|
||||
noncomputable section
|
||||
|
@ -342,40 +338,38 @@ lemma pauliMatrix_contr_down_0 :
|
|||
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
|
||||
conv =>
|
||||
enter [1, 1, 1, 1, 1, 1, 1, 1]
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
rw [contrBasisVectorMul_pos _]
|
||||
conv =>
|
||||
enter [1, 1, 1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
rw [contrBasisVectorMul_pos _]
|
||||
conv =>
|
||||
enter [1, 1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
enter [1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
enter [1, 1, 1, 1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
enter [1, 1, 1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
enter [1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
enter [1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
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
|
||||
decide
|
||||
· rw [basisVectorContrPauli]
|
||||
congr 1
|
||||
funext k
|
||||
fin_cases k <;> rfl
|
||||
decide
|
||||
|
||||
lemma pauliMatrix_contr_down_1 :
|
||||
{(basisVector ![Color.down, Color.down] fun _ => 1) | ν μ ⊗
|
||||
|
@ -386,29 +380,29 @@ lemma pauliMatrix_contr_down_1 :
|
|||
lhs
|
||||
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
enter [1, 1, 1, 1, 1, 1, 1]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
enter [1, 1, 1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
enter [1, 1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_pos _]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
enter [1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_pos _]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; rhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
enter [1, 1, 1, 1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; rhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
enter [1, 1, 1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; rhs; lhs;
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
enter [1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; rhs; rhs; lhs;
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
enter [1, 2, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs
|
||||
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
|
||||
|
@ -429,29 +423,29 @@ lemma pauliMatrix_contr_down_2 :
|
|||
lhs
|
||||
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
enter [1, 1, 1, 1, 1, 1, 1, 1]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
enter [1, 1, 1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
enter [1, 1, 1, 1, 1, 1, 2, 1]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; rhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
rw [contrBasisVectorMul_pos _]
|
||||
conv =>
|
||||
lhs; lhs; lhs; rhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
rw [contrBasisVectorMul_pos _]
|
||||
conv =>
|
||||
lhs; lhs; rhs; lhs;
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; rhs; rhs; lhs;
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs
|
||||
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
|
||||
|
@ -470,28 +464,28 @@ lemma pauliMatrix_contr_down_3 :
|
|||
rw [basis_contr_pauliMatrix_basis_tree_expand_tensor]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; lhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; lhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; lhs; rhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; lhs; rhs; rhs; lhs
|
||||
rw [contrBasisVectorMul_neg (by decide)]
|
||||
rw [contrBasisVectorMul_neg (Nat.ne_of_beq_eq_false _)]
|
||||
conv =>
|
||||
lhs; lhs; rhs; lhs;
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
rw [contrBasisVectorMul_pos _]
|
||||
conv =>
|
||||
lhs; rhs; rhs; lhs;
|
||||
rw [contrBasisVectorMul_pos (by decide)]
|
||||
rw [contrBasisVectorMul_pos _]
|
||||
conv =>
|
||||
lhs
|
||||
simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add]
|
||||
|
@ -535,16 +529,16 @@ lemma pauliCo_basis_expand : pauliCo
|
|||
simp only [tensorNode_tensor, add_tensor, smul_tensor]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, neg_smul, one_smul]
|
||||
conv =>
|
||||
lhs; lhs;
|
||||
enter [1, 1]
|
||||
rw [pauliMatrix_contr_down_0]
|
||||
conv =>
|
||||
lhs; rhs; lhs; rhs;
|
||||
enter [1, 2, 1, 1]
|
||||
rw [pauliMatrix_contr_down_1]
|
||||
conv =>
|
||||
lhs; rhs; rhs; lhs; rhs;
|
||||
enter [1, 2, 2, 1, 1]
|
||||
rw [pauliMatrix_contr_down_2]
|
||||
conv =>
|
||||
lhs; rhs; rhs; rhs; rhs;
|
||||
enter [1, 2, 2, 2, 1]
|
||||
rw [pauliMatrix_contr_down_3]
|
||||
simp only [neg_smul, one_smul]
|
||||
abel
|
||||
|
|
|
@ -126,10 +126,7 @@ lemma inclCongrRealLorentz_ρ (M : SL(2, ℂ)) (v : ContrMod 3) :
|
|||
apply Lorentz.ContrℂModule.ext
|
||||
rw [complexContrBasis_ρ_val, inclCongrRealLorentz_val, inclCongrRealLorentz_val]
|
||||
rw [LorentzGroup.toComplex_mulVec_ofReal]
|
||||
apply congrArg
|
||||
simp only [SL2C.toLorentzGroup_apply_coe]
|
||||
change _ = ContrMod.toFin1dℝ ((SL2C.toLorentzGroup M) *ᵥ v)
|
||||
simp only [SL2C.toLorentzGroup_apply_coe, ContrMod.mulVec_toFin1dℝ]
|
||||
rfl
|
||||
|
||||
/-! TODO: Rename. -/
|
||||
lemma SL2CRep_ρ_basis (M : SL(2, ℂ)) (i : Fin 1 ⊕ Fin 3) :
|
||||
|
|
|
@ -1,31 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import Mathlib.RingTheory.GradedAlgebra.Basic
|
||||
import HepLean.Meta.Informal.Basic
|
||||
/-!
|
||||
|
||||
# Super Algebras
|
||||
|
||||
A super algebra is a special type of graded algebra.
|
||||
|
||||
It is used in physics to model the commutator of fermionic operators among themselves,
|
||||
aswell as among bosonic operators.
|
||||
|
||||
-/
|
||||
|
||||
informal_definition SuperAlgebra where
|
||||
math :≈ "A super algebra is a graded algebra A with a ℤ₂ grading."
|
||||
physics :≈ "A super algebra is used to model the commutator of fermionic operators among
|
||||
themselves, aswell as among bosonic operators."
|
||||
ref :≈ "https://en.wikipedia.org/wiki/Superalgebra"
|
||||
|
||||
namespace SuperAlgebra
|
||||
|
||||
informal_definition superCommuator where
|
||||
math :≈ "The commutator which for `a ∈ Aᵢ` and `b ∈ Aⱼ` is defined as `ab - (-1)^(i * j) ba`."
|
||||
deps :≈ [``SuperAlgebra]
|
||||
|
||||
end SuperAlgebra
|
|
@ -89,9 +89,9 @@ lemma ofList_freeMonoid (s : 𝓕 → FieldStatistic) (φ : 𝓕) : ofList s (Fr
|
|||
lemma ofList_empty (s : 𝓕 → FieldStatistic) : ofList s [] = bosonic := rfl
|
||||
|
||||
@[simp]
|
||||
lemma ofList_append (s : 𝓕 → FieldStatistic) (l r : List 𝓕) :
|
||||
ofList s (l ++ r) = if ofList s l = ofList s r then bosonic else fermionic := by
|
||||
induction l with
|
||||
lemma ofList_append (s : 𝓕 → FieldStatistic) (φs φs' : List 𝓕) :
|
||||
ofList s (φs ++ φs') = if ofList s φs = ofList s φs' then bosonic else fermionic := by
|
||||
induction φs with
|
||||
| nil =>
|
||||
simp only [List.nil_append, ofList_empty, Fin.isValue, eq_self_if_bosonic_eq]
|
||||
| cons a l ih =>
|
||||
|
|
|
@ -78,51 +78,57 @@ section InsertSign
|
|||
|
||||
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic)
|
||||
|
||||
/-- The sign associated with inserting `r0` into `r` at the position `n`.
|
||||
That is the sign associated with commuting `r0` with `List.take n r`. -/
|
||||
def insertSign (n : ℕ) (r0 : 𝓕) (r : List 𝓕) : ℂ :=
|
||||
superCommuteCoef q [r0] (List.take n r)
|
||||
/-- The sign associated with inserting a field `φ` into a list of fields `φs` at
|
||||
the `n`th position. -/
|
||||
def insertSign (n : ℕ) (φ : 𝓕) (φs : List 𝓕) : ℂ :=
|
||||
superCommuteCoef q [φ] (List.take n φs)
|
||||
|
||||
lemma insertSign_insert (n : ℕ) (r0 : 𝓕) (r : List 𝓕) :
|
||||
insertSign q n r0 r = insertSign q n r0 (List.insertIdx n r0 r) := by
|
||||
/-- If `φ` is bosonic, there is no sign associated with inserting it into a list of fields. -/
|
||||
lemma insertSign_bosonic (n : ℕ) (φ : 𝓕) (φs : List 𝓕) (hφ : q φ = bosonic) :
|
||||
insertSign q n φ φs = 1 := by
|
||||
simp only [insertSign, superCommuteCoef, ofList_singleton, hφ, reduceCtorEq, false_and,
|
||||
↓reduceIte]
|
||||
|
||||
lemma insertSign_insert (n : ℕ) (φ : 𝓕) (φs : List 𝓕) :
|
||||
insertSign q n φ φs = insertSign q n φ (List.insertIdx n φ φs) := by
|
||||
simp only [insertSign]
|
||||
congr 1
|
||||
rw [take_insert_same]
|
||||
|
||||
lemma insertSign_eraseIdx (n : ℕ) (r0 : 𝓕) (r : List 𝓕) :
|
||||
insertSign q n r0 (r.eraseIdx n) = insertSign q n r0 r := by
|
||||
lemma insertSign_eraseIdx (n : ℕ) (φ : 𝓕) (φs : List 𝓕) :
|
||||
insertSign q n φ (φs.eraseIdx n) = insertSign q n φ φs := by
|
||||
simp only [insertSign]
|
||||
congr 1
|
||||
rw [take_eraseIdx_same]
|
||||
|
||||
lemma insertSign_zero (r0 : 𝓕) (r : List 𝓕) : insertSign q 0 r0 r = 1 := by
|
||||
lemma insertSign_zero (φ : 𝓕) (φs : List 𝓕) : insertSign q 0 φ φs = 1 := by
|
||||
simp [insertSign, superCommuteCoef]
|
||||
|
||||
lemma insertSign_succ_cons (n : ℕ) (r0 r1 : 𝓕) (r : List 𝓕) : insertSign q (n + 1) r0 (r1 :: r) =
|
||||
superCommuteCoef q [r0] [r1] * insertSign q n r0 r := by
|
||||
lemma insertSign_succ_cons (n : ℕ) (φ0 φ1 : 𝓕) (φs : List 𝓕) : insertSign q (n + 1) φ0 (φ1 :: φs) =
|
||||
superCommuteCoef q [φ0] [φ1] * insertSign q n φ0 φs := by
|
||||
simp only [insertSign, List.take_succ_cons]
|
||||
rw [superCommuteCoef_cons]
|
||||
|
||||
lemma insertSign_insert_gt (n m : ℕ) (r0 r1 : 𝓕) (r : List 𝓕) (hn : n < m) :
|
||||
insertSign q n r0 (List.insertIdx m r1 r) = insertSign q n r0 r := by
|
||||
lemma insertSign_insert_gt (n m : ℕ) (φ0 φ1 : 𝓕) (φs : List 𝓕) (hn : n < m) :
|
||||
insertSign q n φ0 (List.insertIdx m φ1 φs) = insertSign q n φ0 φs := by
|
||||
rw [insertSign, insertSign]
|
||||
congr 1
|
||||
exact take_insert_gt r1 n m hn r
|
||||
exact take_insert_gt φ1 n m hn φs
|
||||
|
||||
lemma insertSign_insert_lt_eq_insertSort (n m : ℕ) (r0 r1 : 𝓕) (r : List 𝓕) (hn : m ≤ n)
|
||||
(hm : m ≤ r.length) :
|
||||
insertSign q (n + 1) r0 (List.insertIdx m r1 r) = insertSign q (n + 1) r0 (r1 :: r) := by
|
||||
lemma insertSign_insert_lt_eq_insertSort (n m : ℕ) (φ0 φ1 : 𝓕) (φs : List 𝓕) (hn : m ≤ n)
|
||||
(hm : m ≤ φs.length) :
|
||||
insertSign q (n + 1) φ0 (List.insertIdx m φ1 φs) = insertSign q (n + 1) φ0 (φ1 :: φs) := by
|
||||
rw [insertSign, insertSign]
|
||||
apply superCommuteCoef_perm_snd
|
||||
simp only [List.take_succ_cons]
|
||||
refine take_insert_let r1 n m hn r hm
|
||||
refine take_insert_let φ1 n m hn φs hm
|
||||
|
||||
lemma insertSign_insert_lt (n m : ℕ) (r0 r1 : 𝓕) (r : List 𝓕) (hn : m ≤ n) (hm : m ≤ r.length) :
|
||||
insertSign q (n + 1) r0 (List.insertIdx m r1 r) = superCommuteCoef q [r0] [r1] *
|
||||
insertSign q n r0 r := by
|
||||
lemma insertSign_insert_lt (n m : ℕ) (φ0 φ1 : 𝓕) (φs : List 𝓕) (hn : m ≤ n) (hm : m ≤ φs.length) :
|
||||
insertSign q (n + 1) φ0 (List.insertIdx m φ1 φs) = superCommuteCoef q [φ0] [φ1] *
|
||||
insertSign q n φ0 φs := by
|
||||
rw [insertSign_insert_lt_eq_insertSort, insertSign_succ_cons]
|
||||
exact hn
|
||||
exact hm
|
||||
· exact hn
|
||||
· exact hm
|
||||
|
||||
end InsertSign
|
||||
|
||||
|
|
|
@ -31,20 +31,18 @@ lemma koszulSign_mul_self (l : List 𝓕) : koszulSign q le l * koszulSign q le
|
|||
simp only [koszulSign]
|
||||
trans (koszulSignInsert q le a l * koszulSignInsert q le a l) *
|
||||
(koszulSign q le l * koszulSign q le l)
|
||||
ring
|
||||
rw [ih]
|
||||
rw [koszulSignInsert_mul_self, mul_one]
|
||||
· ring
|
||||
· rw [ih, koszulSignInsert_mul_self, mul_one]
|
||||
|
||||
@[simp]
|
||||
lemma koszulSign_freeMonoid_of (i : 𝓕) : koszulSign q le (FreeMonoid.of i) = 1 := by
|
||||
change koszulSign q le [i] = 1
|
||||
lemma koszulSign_freeMonoid_of (φ : 𝓕) : koszulSign q le (FreeMonoid.of φ) = 1 := by
|
||||
simp only [koszulSign, mul_one]
|
||||
rfl
|
||||
|
||||
lemma koszulSignInsert_erase_boson {𝓕 : Type} (q : 𝓕 → FieldStatistic)
|
||||
(le : 𝓕 → 𝓕 → Prop) [DecidableRel le] (r0 : 𝓕) :
|
||||
(r : List 𝓕) → (n : Fin r.length) → (heq : q (r.get n) = bosonic) →
|
||||
koszulSignInsert q le r0 (r.eraseIdx n) = koszulSignInsert q le r0 r
|
||||
(le : 𝓕 → 𝓕 → Prop) [DecidableRel le] (φ : 𝓕) :
|
||||
(φs : List 𝓕) → (n : Fin φs.length) → (heq : q (φs.get n) = bosonic) →
|
||||
koszulSignInsert q le φ (φs.eraseIdx n) = koszulSignInsert q le φ φs
|
||||
| [], _, _ => by
|
||||
simp
|
||||
| r1 :: r, ⟨0, h⟩, hr => by
|
||||
|
@ -56,44 +54,44 @@ lemma koszulSignInsert_erase_boson {𝓕 : Type} (q : 𝓕 → FieldStatistic)
|
|||
| r1 :: r, ⟨n + 1, h⟩, hr => by
|
||||
simp only [List.eraseIdx_cons_succ]
|
||||
rw [koszulSignInsert, koszulSignInsert]
|
||||
rw [koszulSignInsert_erase_boson q le r0 r ⟨n, Nat.succ_lt_succ_iff.mp h⟩ hr]
|
||||
rw [koszulSignInsert_erase_boson q le φ r ⟨n, Nat.succ_lt_succ_iff.mp h⟩ hr]
|
||||
|
||||
lemma koszulSign_erase_boson {𝓕 : Type} (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Prop)
|
||||
[DecidableRel le] :
|
||||
(r : List 𝓕) → (n : Fin r.length) → (heq : q (r.get n) = bosonic) →
|
||||
koszulSign q le (r.eraseIdx n) = koszulSign q le r
|
||||
(φs : List 𝓕) → (n : Fin φs.length) → (heq : q (φs.get n) = bosonic) →
|
||||
koszulSign q le (φs.eraseIdx n) = koszulSign q le φs
|
||||
| [], _ => by
|
||||
simp
|
||||
| r0 :: r, ⟨0, h⟩ => by
|
||||
| φ :: φs, ⟨0, h⟩ => by
|
||||
simp only [List.length_cons, Fin.zero_eta, List.get_eq_getElem, Fin.val_zero,
|
||||
List.getElem_cons_zero, Fin.isValue, List.eraseIdx_zero, List.tail_cons, koszulSign]
|
||||
intro h
|
||||
rw [koszulSignInsert_boson]
|
||||
simp only [one_mul]
|
||||
exact h
|
||||
| r0 :: r, ⟨n + 1, h⟩ => by
|
||||
| φ :: φs, ⟨n + 1, h⟩ => by
|
||||
simp only [List.length_cons, List.get_eq_getElem, List.getElem_cons_succ, Fin.isValue,
|
||||
List.eraseIdx_cons_succ]
|
||||
intro h'
|
||||
rw [koszulSign, koszulSign, koszulSign_erase_boson q le r ⟨n, Nat.succ_lt_succ_iff.mp h⟩]
|
||||
rw [koszulSign, koszulSign, koszulSign_erase_boson q le φs ⟨n, Nat.succ_lt_succ_iff.mp h⟩]
|
||||
congr 1
|
||||
rw [koszulSignInsert_erase_boson q le r0 r ⟨n, Nat.succ_lt_succ_iff.mp h⟩ h']
|
||||
rw [koszulSignInsert_erase_boson q le φ φs ⟨n, Nat.succ_lt_succ_iff.mp h⟩ h']
|
||||
exact h'
|
||||
|
||||
lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (i : 𝓕) :
|
||||
(r : List 𝓕) → (n : ℕ) → (hn : n ≤ r.length) →
|
||||
koszulSign q le (List.insertIdx n i r) = insertSign q n i r * koszulSign q le r *
|
||||
insertSign q (insertionSortEquiv le (List.insertIdx n i r) ⟨n, by
|
||||
lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) :
|
||||
(φs : List 𝓕) → (n : ℕ) → (hn : n ≤ φs.length) →
|
||||
koszulSign q le (List.insertIdx n φ φs) = insertSign q n φ φs * koszulSign q le φs *
|
||||
insertSign q (insertionSortEquiv le (List.insertIdx n φ φs) ⟨n, by
|
||||
rw [List.length_insertIdx _ _ hn]
|
||||
omega⟩) i (List.insertionSort le (List.insertIdx n i r))
|
||||
omega⟩) φ (List.insertionSort le (List.insertIdx n φ φs))
|
||||
| [], 0, h => by
|
||||
simp [koszulSign, insertSign, superCommuteCoef, koszulSignInsert]
|
||||
| [], n + 1, h => by
|
||||
simp at h
|
||||
| r0 :: r, 0, h => by
|
||||
| φ1 :: φs, 0, h => by
|
||||
simp only [List.insertIdx_zero, List.insertionSort, List.length_cons, Fin.zero_eta]
|
||||
rw [koszulSign]
|
||||
trans koszulSign q le (r0 :: r) * koszulSignInsert q le i (r0 :: r)
|
||||
trans koszulSign q le (φ1 :: φs) * koszulSignInsert q le φ (φ1 :: φs)
|
||||
ring
|
||||
simp only [insertionSortEquiv, List.length_cons, Nat.succ_eq_add_one, List.insertionSort,
|
||||
orderedInsertEquiv, OrderIso.toEquiv_symm, Fin.symm_castOrderIso, HepLean.Fin.equivCons_trans,
|
||||
|
@ -101,18 +99,17 @@ lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (i : 𝓕) :
|
|||
Fin.isValue, HepLean.Fin.finExtractOne_symm_inl_apply, RelIso.coe_fn_toEquiv,
|
||||
Fin.castOrderIso_apply, Fin.cast_mk, Fin.eta]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rhs
|
||||
enter [2, 4]
|
||||
rw [orderedInsert_eq_insertIdx_orderedInsertPos]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [← insertSign_insert]
|
||||
change insertSign q (↑(orderedInsertPos le ((List.insertionSort le (r0 :: r))) i)) i
|
||||
(List.insertionSort le (r0 :: r))
|
||||
change insertSign q (↑(orderedInsertPos le ((List.insertionSort le (φ1 :: φs))) φ)) φ
|
||||
(List.insertionSort le (φ1 :: φs))
|
||||
rw [← koszulSignInsert_eq_insertSign q le]
|
||||
rw [insertSign_zero]
|
||||
simp
|
||||
| r0 :: r, n + 1, h => by
|
||||
| φ1 :: φs, n + 1, h => by
|
||||
conv_lhs =>
|
||||
rw [List.insertIdx_succ_cons]
|
||||
rw [koszulSign]
|
||||
|
@ -137,41 +134,39 @@ lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (i : 𝓕) :
|
|||
conv_rhs =>
|
||||
rw [mul_assoc, mul_assoc]
|
||||
congr 1
|
||||
let rs := (List.insertionSort le (List.insertIdx n i r))
|
||||
have hnsL : n < (List.insertIdx n i r).length := by
|
||||
let rs := (List.insertionSort le (List.insertIdx n φ φs))
|
||||
have hnsL : n < (List.insertIdx n φ φs).length := by
|
||||
rw [List.length_insertIdx _ _]
|
||||
simp only [List.length_cons, add_le_add_iff_right] at h
|
||||
omega
|
||||
exact Nat.le_of_lt_succ h
|
||||
let ni : Fin rs.length := (insertionSortEquiv le (List.insertIdx n i r))
|
||||
let ni : Fin rs.length := (insertionSortEquiv le (List.insertIdx n φ φs))
|
||||
⟨n, hnsL⟩
|
||||
let nro : Fin (rs.length + 1) :=
|
||||
⟨↑(orderedInsertPos le rs r0), orderedInsertPos_lt_length le rs r0⟩
|
||||
⟨↑(orderedInsertPos le rs φ1), orderedInsertPos_lt_length le rs φ1⟩
|
||||
rw [koszulSignInsert_insertIdx, koszulSignInsert_cons]
|
||||
trans koszulSignInsert q le r0 r * (koszulSignCons q le r0 i *insertSign q ni i rs)
|
||||
trans koszulSignInsert q le φ1 φs * (koszulSignCons q le φ1 φ *insertSign q ni φ rs)
|
||||
· simp only [rs, ni]
|
||||
ring
|
||||
trans koszulSignInsert q le r0 r * (superCommuteCoef q [i] [r0] *
|
||||
insertSign q (nro.succAbove ni) i (List.insertIdx nro r0 rs))
|
||||
trans koszulSignInsert q le φ1 φs * (superCommuteCoef q [φ] [φ1] *
|
||||
insertSign q (nro.succAbove ni) φ (List.insertIdx nro φ1 rs))
|
||||
swap
|
||||
· simp only [rs, nro, ni]
|
||||
ring
|
||||
congr 1
|
||||
simp only [Fin.succAbove]
|
||||
have hns : rs.get ni = i := by
|
||||
have hns : rs.get ni = φ := by
|
||||
simp only [Fin.eta, rs]
|
||||
rw [← insertionSortEquiv_get]
|
||||
simp only [Function.comp_apply, Equiv.symm_apply_apply, List.get_eq_getElem, ni]
|
||||
simp_all only [List.length_cons, add_le_add_iff_right, List.getElem_insertIdx_self]
|
||||
have hc1 : ni.castSucc < nro → ¬ le r0 i := by
|
||||
intro hninro
|
||||
have hc1 (hninro : ni.castSucc < nro) : ¬ le φ1 φ := by
|
||||
rw [← hns]
|
||||
exact lt_orderedInsertPos_rel le r0 rs ni hninro
|
||||
have hc2 : ¬ ni.castSucc < nro → le r0 i := by
|
||||
intro hninro
|
||||
exact lt_orderedInsertPos_rel le φ1 rs ni hninro
|
||||
have hc2 (hninro : ¬ ni.castSucc < nro) : le φ1 φ := by
|
||||
rw [← hns]
|
||||
refine gt_orderedInsertPos_rel le r0 rs ?_ ni hninro
|
||||
exact List.sorted_insertionSort le (List.insertIdx n i r)
|
||||
refine gt_orderedInsertPos_rel le φ1 rs ?_ ni hninro
|
||||
exact List.sorted_insertionSort le (List.insertIdx n φ φs)
|
||||
by_cases hn : ni.castSucc < nro
|
||||
· simp only [hn, ↓reduceIte, Fin.coe_castSucc]
|
||||
rw [insertSign_insert_gt]
|
||||
|
@ -187,7 +182,7 @@ lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (i : 𝓕) :
|
|||
rw [superCommuteCoef_mul_self, koszulSignCons]
|
||||
simp only [hc2 hn, ↓reduceIte]
|
||||
exact Nat.le_of_not_lt hn
|
||||
exact Nat.le_of_lt_succ (orderedInsertPos_lt_length le rs r0)
|
||||
exact Nat.le_of_lt_succ (orderedInsertPos_lt_length le rs φ1)
|
||||
· exact Nat.le_of_lt_succ h
|
||||
· exact Nat.le_of_lt_succ h
|
||||
|
||||
|
|
|
@ -28,7 +28,7 @@ def koszulSignInsert {𝓕 : Type} (q : 𝓕 → FieldStatistic) (le : 𝓕 →
|
|||
|
||||
/-- When inserting a boson the `koszulSignInsert` is always `1`. -/
|
||||
lemma koszulSignInsert_boson (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Prop) [DecidableRel le]
|
||||
(a : 𝓕) (ha : q a = bosonic) : (l : List 𝓕) → koszulSignInsert q le a l = 1
|
||||
(a : 𝓕) (ha : q a = bosonic) : (φs : List 𝓕) → koszulSignInsert q le a φs = 1
|
||||
| [] => by
|
||||
simp [koszulSignInsert]
|
||||
| b :: l => by
|
||||
|
@ -38,7 +38,7 @@ lemma koszulSignInsert_boson (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕
|
|||
|
||||
@[simp]
|
||||
lemma koszulSignInsert_mul_self (a : 𝓕) :
|
||||
(l : List 𝓕) → koszulSignInsert q le a l * koszulSignInsert q le a l = 1
|
||||
(φs : List 𝓕) → koszulSignInsert q le a φs * koszulSignInsert q le a φs = 1
|
||||
| [] => by
|
||||
simp [koszulSignInsert]
|
||||
| b :: l => by
|
||||
|
@ -64,9 +64,9 @@ lemma koszulSignInsert_le_forall (a : 𝓕) (l : List 𝓕) (hi : ∀ b, le a b)
|
|||
intro h
|
||||
exact False.elim (h (hi j))
|
||||
|
||||
lemma koszulSignInsert_ge_forall_append (l : List 𝓕) (j i : 𝓕) (hi : ∀ j, le j i) :
|
||||
koszulSignInsert q le j l = koszulSignInsert q le j (l ++ [i]) := by
|
||||
induction l with
|
||||
lemma koszulSignInsert_ge_forall_append (φs : List 𝓕) (j i : 𝓕) (hi : ∀ j, le j i) :
|
||||
koszulSignInsert q le j φs = koszulSignInsert q le j (φs ++ [i]) := by
|
||||
induction φs with
|
||||
| nil => simp [koszulSignInsert, hi]
|
||||
| cons b l ih =>
|
||||
simp only [koszulSignInsert, Fin.isValue, List.append_eq]
|
||||
|
@ -74,15 +74,15 @@ lemma koszulSignInsert_ge_forall_append (l : List 𝓕) (j i : 𝓕) (hi : ∀ j
|
|||
· rw [if_pos hr, if_pos hr, ih]
|
||||
· rw [if_neg hr, if_neg hr, ih]
|
||||
|
||||
lemma koszulSignInsert_eq_filter (r0 : 𝓕) : (r : List 𝓕) →
|
||||
koszulSignInsert q le r0 r =
|
||||
koszulSignInsert q le r0 (List.filter (fun i => decide (¬ le r0 i)) r)
|
||||
lemma koszulSignInsert_eq_filter (φ : 𝓕) : (φs : List 𝓕) →
|
||||
koszulSignInsert q le φ φs =
|
||||
koszulSignInsert q le φ (List.filter (fun i => decide (¬ le φ i)) φs)
|
||||
| [] => by
|
||||
simp [koszulSignInsert]
|
||||
| r1 :: r => by
|
||||
| φ1 :: φs => by
|
||||
dsimp only [koszulSignInsert, Fin.isValue]
|
||||
simp only [Fin.isValue, List.filter, decide_not]
|
||||
by_cases h : le r0 r1
|
||||
by_cases h : le φ φ1
|
||||
· simp only [h, ↓reduceIte, decide_True, Bool.not_true]
|
||||
rw [koszulSignInsert_eq_filter]
|
||||
congr
|
||||
|
@ -95,22 +95,22 @@ lemma koszulSignInsert_eq_filter (r0 : 𝓕) : (r : List 𝓕) →
|
|||
simp only [decide_not]
|
||||
simp
|
||||
|
||||
lemma koszulSignInsert_eq_cons [IsTotal 𝓕 le] (r0 : 𝓕) (r : List 𝓕) :
|
||||
koszulSignInsert q le r0 r = koszulSignInsert q le r0 (r0 :: r) := by
|
||||
lemma koszulSignInsert_eq_cons [IsTotal 𝓕 le] (φ : 𝓕) (φs : List 𝓕) :
|
||||
koszulSignInsert q le φ φs = koszulSignInsert q le φ (φ :: φs) := by
|
||||
simp only [koszulSignInsert, Fin.isValue, and_self]
|
||||
have h1 : le r0 r0 := by
|
||||
simpa using IsTotal.total (r := le) r0 r0
|
||||
have h1 : le φ φ := by
|
||||
simpa only [or_self] using IsTotal.total (r := le) φ φ
|
||||
simp [h1]
|
||||
|
||||
lemma koszulSignInsert_eq_grade (r0 : 𝓕) (r : List 𝓕) :
|
||||
koszulSignInsert q le r0 r = if ofList q [r0] = fermionic ∧
|
||||
ofList q (List.filter (fun i => decide (¬ le r0 i)) r) = fermionic then -1 else 1 := by
|
||||
induction r with
|
||||
lemma koszulSignInsert_eq_grade (φ : 𝓕) (φs : List 𝓕) :
|
||||
koszulSignInsert q le φ φs = if ofList q [φ] = fermionic ∧
|
||||
ofList q (List.filter (fun i => decide (¬ le φ i)) φs) = fermionic then -1 else 1 := by
|
||||
induction φs with
|
||||
| nil =>
|
||||
simp [koszulSignInsert]
|
||||
| cons r1 r ih =>
|
||||
| cons φ1 φs ih =>
|
||||
rw [koszulSignInsert_eq_filter]
|
||||
by_cases hr1 : ¬ le r0 r1
|
||||
by_cases hr1 : ¬ le φ φ1
|
||||
· rw [List.filter_cons_of_pos]
|
||||
· dsimp only [koszulSignInsert, Fin.isValue, decide_not]
|
||||
rw [if_neg hr1]
|
||||
|
@ -123,7 +123,7 @@ lemma koszulSignInsert_eq_grade (r0 : 𝓕) (r : List 𝓕) :
|
|||
fin_cases a <;> fin_cases b <;> fin_cases c
|
||||
any_goals rfl
|
||||
simp
|
||||
rw [← ha (q r0) (q r1) (ofList q (List.filter (fun a => !decide (le r0 a)) r))]
|
||||
rw [← ha (q φ) (q φ1) (ofList q (List.filter (fun a => !decide (le φ a)) φs))]
|
||||
congr
|
||||
· rw [koszulSignInsert_eq_filter] at ih
|
||||
simpa [ofList] using ih
|
||||
|
@ -136,27 +136,26 @@ lemma koszulSignInsert_eq_grade (r0 : 𝓕) (r : List 𝓕) :
|
|||
simpa [ofList] using ih
|
||||
simpa using hr1
|
||||
|
||||
lemma koszulSignInsert_eq_perm (r r' : List 𝓕) (a : 𝓕) (h : r.Perm r') :
|
||||
koszulSignInsert q le a r = koszulSignInsert q le a r' := by
|
||||
rw [koszulSignInsert_eq_grade]
|
||||
rw [koszulSignInsert_eq_grade]
|
||||
lemma koszulSignInsert_eq_perm (φs φs' : List 𝓕) (a : 𝓕) (h : φs.Perm φs') :
|
||||
koszulSignInsert q le a φs = koszulSignInsert q le a φs' := by
|
||||
rw [koszulSignInsert_eq_grade, koszulSignInsert_eq_grade]
|
||||
congr 1
|
||||
simp only [Fin.isValue, decide_not, eq_iff_iff, and_congr_right_iff]
|
||||
intro h'
|
||||
have hg : ofList q (List.filter (fun i => !decide (le a i)) r) =
|
||||
ofList q (List.filter (fun i => !decide (le a i)) r') := by
|
||||
have hg : ofList q (List.filter (fun i => !decide (le a i)) φs) =
|
||||
ofList q (List.filter (fun i => !decide (le a i)) φs') := by
|
||||
apply ofList_perm
|
||||
exact List.Perm.filter (fun i => !decide (le a i)) h
|
||||
rw [hg]
|
||||
|
||||
lemma koszulSignInsert_eq_sort (r : List 𝓕) (a : 𝓕) :
|
||||
koszulSignInsert q le a r = koszulSignInsert q le a (List.insertionSort le r) := by
|
||||
lemma koszulSignInsert_eq_sort (φs : List 𝓕) (a : 𝓕) :
|
||||
koszulSignInsert q le a φs = koszulSignInsert q le a (List.insertionSort le φs) := by
|
||||
apply koszulSignInsert_eq_perm
|
||||
exact List.Perm.symm (List.perm_insertionSort le r)
|
||||
exact List.Perm.symm (List.perm_insertionSort le φs)
|
||||
|
||||
lemma koszulSignInsert_eq_insertSign [IsTotal 𝓕 le] [IsTrans 𝓕 le] (r0 : 𝓕) (r : List 𝓕) :
|
||||
koszulSignInsert q le r0 r = insertSign q (orderedInsertPos le (List.insertionSort le r) r0)
|
||||
r0 (List.insertionSort le r) := by
|
||||
lemma koszulSignInsert_eq_insertSign [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) (φs : List 𝓕) :
|
||||
koszulSignInsert q le φ φs = insertSign q (orderedInsertPos le (List.insertionSort le φs) φ)
|
||||
φ (List.insertionSort le φs) := by
|
||||
rw [koszulSignInsert_eq_cons, koszulSignInsert_eq_sort, koszulSignInsert_eq_filter,
|
||||
koszulSignInsert_eq_grade, insertSign, superCommuteCoef]
|
||||
congr
|
||||
|
@ -164,10 +163,10 @@ lemma koszulSignInsert_eq_insertSign [IsTotal 𝓕 le] [IsTrans 𝓕 le] (r0 :
|
|||
rw [List.insertionSort]
|
||||
nth_rewrite 1 [List.orderedInsert_eq_take_drop]
|
||||
rw [List.filter_append]
|
||||
have h1 : List.filter (fun a => decide ¬le r0 a)
|
||||
(List.takeWhile (fun b => decide ¬le r0 b) (List.insertionSort le r))
|
||||
= (List.takeWhile (fun b => decide ¬le r0 b) (List.insertionSort le r)) := by
|
||||
induction r with
|
||||
have h1 : List.filter (fun a => decide ¬le φ a)
|
||||
(List.takeWhile (fun b => decide ¬le φ b) (List.insertionSort le φs))
|
||||
= (List.takeWhile (fun b => decide ¬le φ b) (List.insertionSort le φs)) := by
|
||||
induction φs with
|
||||
| nil => simp
|
||||
| cons r1 r ih =>
|
||||
simp only [decide_not, List.insertionSort, List.filter_eq_self, Bool.not_eq_eq_eq_not,
|
||||
|
@ -177,23 +176,22 @@ lemma koszulSignInsert_eq_insertSign [IsTotal 𝓕 le] [IsTrans 𝓕 le] (r0 :
|
|||
simp_all
|
||||
rw [h1]
|
||||
rw [List.filter_cons]
|
||||
simp only [decide_not, (IsTotal.to_isRefl le).refl r0, not_true_eq_false, decide_False,
|
||||
simp only [decide_not, (IsTotal.to_isRefl le).refl φ, not_true_eq_false, decide_False,
|
||||
Bool.false_eq_true, ↓reduceIte]
|
||||
rw [orderedInsertPos_take]
|
||||
simp only [decide_not, List.append_right_eq_self, List.filter_eq_nil_iff, Bool.not_eq_eq_eq_not,
|
||||
Bool.not_true, decide_eq_false_iff_not, Decidable.not_not]
|
||||
intro a ha
|
||||
refine List.Sorted.rel_of_mem_take_of_mem_drop
|
||||
(k := (orderedInsertPos le (List.insertionSort le r) r0).1 + 1)
|
||||
(List.sorted_insertionSort le (r0 :: r)) ?_ ?_
|
||||
(k := (orderedInsertPos le (List.insertionSort le φs) φ).1 + 1)
|
||||
(List.sorted_insertionSort le (φ :: φs)) ?_ ?_
|
||||
· simp only [List.insertionSort, List.orderedInsert_eq_take_drop, decide_not]
|
||||
rw [List.take_append_eq_append_take]
|
||||
rw [List.take_of_length_le]
|
||||
· simp [orderedInsertPos]
|
||||
· simp [orderedInsertPos]
|
||||
· simp only [List.insertionSort, List.orderedInsert_eq_take_drop, decide_not]
|
||||
rw [List.drop_append_eq_append_drop]
|
||||
rw [List.drop_of_length_le]
|
||||
rw [List.drop_append_eq_append_drop, List.drop_of_length_le]
|
||||
· simpa [orderedInsertPos] using ha
|
||||
· simp [orderedInsertPos]
|
||||
|
||||
|
@ -204,24 +202,24 @@ lemma koszulSignInsert_insertIdx (i j : 𝓕) (r : List 𝓕) (n : ℕ) (hn : n
|
|||
|
||||
/-- The difference in `koszulSignInsert` on inserting `r0` into `r` compared to
|
||||
into `r1 :: r` for any `r`. -/
|
||||
def koszulSignCons (r0 r1 : 𝓕) : ℂ :=
|
||||
if le r0 r1 then 1 else
|
||||
if q r0 = fermionic ∧ q r1 = fermionic then -1 else 1
|
||||
def koszulSignCons (φ0 φ1 : 𝓕) : ℂ :=
|
||||
if le φ0 φ1 then 1 else
|
||||
if q φ0 = fermionic ∧ q φ1 = fermionic then -1 else 1
|
||||
|
||||
lemma koszulSignCons_eq_superComuteCoef (r0 r1 : 𝓕) : koszulSignCons q le r0 r1 =
|
||||
if le r0 r1 then 1 else superCommuteCoef q [r0] [r1] := by
|
||||
lemma koszulSignCons_eq_superComuteCoef (φ0 φ1 : 𝓕) : koszulSignCons q le φ0 φ1 =
|
||||
if le φ0 φ1 then 1 else superCommuteCoef q [φ0] [φ1] := by
|
||||
simp only [koszulSignCons, Fin.isValue, superCommuteCoef, ofList, ite_eq_right_iff, zero_ne_one,
|
||||
imp_false]
|
||||
congr 1
|
||||
by_cases h0 : q r0 = fermionic
|
||||
· by_cases h1 : q r1 = fermionic
|
||||
by_cases h0 : q φ0 = fermionic
|
||||
· by_cases h1 : q φ1 = fermionic
|
||||
· simp [h0, h1]
|
||||
· have h1 : q r1 = bosonic := (neq_fermionic_iff_eq_bosonic (q r1)).mp h1
|
||||
· have h1 : q φ1 = bosonic := (neq_fermionic_iff_eq_bosonic (q φ1)).mp h1
|
||||
simp [h0, h1]
|
||||
· have h0 : q r0 = bosonic := (neq_fermionic_iff_eq_bosonic (q r0)).mp h0
|
||||
by_cases h1 : q r1 = fermionic
|
||||
· have h0 : q φ0 = bosonic := (neq_fermionic_iff_eq_bosonic (q φ0)).mp h0
|
||||
by_cases h1 : q φ1 = fermionic
|
||||
· simp [h0, h1]
|
||||
· have h1 : q r1 = bosonic := (neq_fermionic_iff_eq_bosonic (q r1)).mp h1
|
||||
· have h1 : q φ1 = bosonic := (neq_fermionic_iff_eq_bosonic (q φ1)).mp h1
|
||||
simp [h0, h1]
|
||||
|
||||
lemma koszulSignInsert_cons (r0 r1 : 𝓕) (r : List 𝓕) :
|
||||
|
|
|
@ -37,12 +37,12 @@ lemma superCommuteCoef_comm (la lb : List 𝓕) :
|
|||
the lift of `l` and `r` (by summing over fibers) in the
|
||||
free algebra over `Σ i, f i`.
|
||||
In terms of physics it is `-1` if commuting two fermionic operators and `1` otherwise. -/
|
||||
def superCommuteLiftCoef {f : 𝓕 → Type} (l : List (Σ i, f i)) (r : List 𝓕) : ℂ :=
|
||||
(if FieldStatistic.ofList (fun i => q i.fst) l = fermionic ∧
|
||||
FieldStatistic.ofList q r = fermionic then -1 else 1)
|
||||
def superCommuteLiftCoef {f : 𝓕 → Type} (φs : List (Σ i, f i)) (φs' : List 𝓕) : ℂ :=
|
||||
(if FieldStatistic.ofList (fun i => q i.fst) φs = fermionic ∧
|
||||
FieldStatistic.ofList q φs' = fermionic then -1 else 1)
|
||||
|
||||
lemma superCommuteLiftCoef_empty {f : 𝓕 → Type} (l : List (Σ i, f i)) :
|
||||
superCommuteLiftCoef q l [] = 1 := by
|
||||
lemma superCommuteLiftCoef_empty {f : 𝓕 → Type} (φs : List (Σ i, f i)) :
|
||||
superCommuteLiftCoef q φs [] = 1 := by
|
||||
simp [superCommuteLiftCoef]
|
||||
|
||||
lemma superCommuteCoef_perm_snd (la lb lb' : List 𝓕)
|
||||
|
|
|
@ -273,8 +273,7 @@ lemma contr_contr (t : TensorTree S c) :
|
|||
rw [contrMapFst_contrMapSnd_swap]
|
||||
simp only [Nat.succ_eq_add_one, contrMapFst, contrMapSnd, Action.comp_hom, ModuleCat.coe_comp,
|
||||
Function.comp_apply, swap]
|
||||
apply congrArg
|
||||
apply congrArg
|
||||
refine congrArg _ (congrArg _ ?_)
|
||||
apply congrArg
|
||||
rfl
|
||||
|
||||
|
|
|
@ -31,8 +31,7 @@ def permProdLeft := (equivToIso finSumFinEquiv).inv ≫ σ ▷ OverColor.mk c2
|
|||
|
||||
@[simp]
|
||||
lemma permProdLeft_toEquiv : Hom.toEquiv (permProdLeft c2 σ) = finSumFinEquiv.symm.trans
|
||||
(((Hom.toEquiv σ).sumCongr (Equiv.refl (Fin n2))).trans finSumFinEquiv) := by
|
||||
simp [permProdLeft]
|
||||
(((Hom.toEquiv σ).sumCongr (Equiv.refl (Fin n2))).trans finSumFinEquiv) := rfl
|
||||
|
||||
/-- The permutation that arises when moving a `perm` node in the right entry through a `prod` node.
|
||||
This permutation is defined using left-whiskering and composition with `finSumFinEquiv`
|
||||
|
@ -42,8 +41,7 @@ def permProdRight := (equivToIso finSumFinEquiv).inv ≫ OverColor.mk c2 ◁ σ
|
|||
|
||||
@[simp]
|
||||
lemma permProdRight_toEquiv : Hom.toEquiv (permProdRight c2 σ) = finSumFinEquiv.symm.trans
|
||||
(((Equiv.refl (Fin n2)).sumCongr (Hom.toEquiv σ)).trans finSumFinEquiv) := by
|
||||
simp [permProdRight]
|
||||
(((Equiv.refl (Fin n2)).sumCongr (Hom.toEquiv σ)).trans finSumFinEquiv) := rfl
|
||||
|
||||
/-- When a `prod` acts on a `perm` node in the left entry, the `perm` node can be moved through
|
||||
the `prod` node via right-whiskering. -/
|
||||
|
|
|
@ -33,8 +33,7 @@ def braidPerm : OverColor.mk (Sum.elim c2 c ∘ ⇑finSumFinEquiv.symm) ⟶
|
|||
|
||||
@[simp]
|
||||
lemma braidPerm_toEquiv : Hom.toEquiv (braidPerm c c2) = finSumFinEquiv.symm.trans
|
||||
((Equiv.sumComm (Fin n2) (Fin n)).trans finSumFinEquiv) := by
|
||||
simp [braidPerm]
|
||||
((Equiv.sumComm (Fin n2) (Fin n)).trans finSumFinEquiv) := rfl
|
||||
|
||||
lemma finSumFinEquiv_comp_braidPerm :
|
||||
(equivToIso finSumFinEquiv).hom ≫ braidPerm c c2 =
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue