2024-12-19 14:25:09 +00:00
|
|
|
|
/-
|
|
|
|
|
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
|
|
|
|
|
-/
|
2024-12-20 11:09:09 +00:00
|
|
|
|
import HepLean.PerturbationTheory.FieldStatistics
|
2024-12-19 14:25:09 +00:00
|
|
|
|
/-!
|
|
|
|
|
|
2024-12-20 11:24:37 +00:00
|
|
|
|
# Super commutation coefficent.
|
|
|
|
|
|
|
|
|
|
This is a complex number which is `-1` when commuting two fermionic operators and `1` otherwise.
|
2024-12-19 14:25:09 +00:00
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
namespace Wick
|
2024-12-20 11:09:09 +00:00
|
|
|
|
open FieldStatistic
|
|
|
|
|
|
|
|
|
|
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic)
|
2024-12-19 14:25:09 +00:00
|
|
|
|
|
2024-12-19 15:40:04 +00:00
|
|
|
|
/-- Given two lists `la` and `lb` returns `-1` if they are both of grade `1` and
|
|
|
|
|
`1` otherwise. This corresponds to the sign associated with the super commutator
|
|
|
|
|
when commuting `la` and `lb` in the free algebra.
|
|
|
|
|
In terms of physics it is `-1` if commuting two fermionic operators and `1` otherwise. -/
|
2024-12-20 11:09:09 +00:00
|
|
|
|
def superCommuteCoef (la lb : List 𝓕) : ℂ :=
|
|
|
|
|
if FieldStatistic.ofList q la = fermionic ∧
|
|
|
|
|
FieldStatistic.ofList q lb = fermionic then - 1 else 1
|
2024-12-19 14:25:09 +00:00
|
|
|
|
|
2024-12-20 11:09:09 +00:00
|
|
|
|
lemma superCommuteCoef_comm (la lb : List 𝓕) :
|
2024-12-19 14:25:09 +00:00
|
|
|
|
superCommuteCoef q la lb = superCommuteCoef q lb la := by
|
|
|
|
|
simp only [superCommuteCoef, Fin.isValue]
|
|
|
|
|
congr 1
|
|
|
|
|
exact Eq.propIntro (fun a => id (And.symm a)) fun a => id (And.symm a)
|
|
|
|
|
|
2024-12-19 15:40:04 +00:00
|
|
|
|
/-- Given a list `l : List (Σ i, f i)` and a list `r : List I` returns `-1` if the
|
|
|
|
|
grade of `l` is `1` and the grade of `r` is `1` and `1` otherwise. This corresponds
|
|
|
|
|
to the sign associated with the super commutator when commuting
|
|
|
|
|
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. -/
|
2024-12-20 11:09:09 +00:00
|
|
|
|
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)
|
2024-12-19 14:25:09 +00:00
|
|
|
|
|
2024-12-20 11:09:09 +00:00
|
|
|
|
lemma superCommuteLiftCoef_empty {f : 𝓕 → Type} (l : List (Σ i, f i)) :
|
2024-12-19 14:25:09 +00:00
|
|
|
|
superCommuteLiftCoef q l [] = 1 := by
|
|
|
|
|
simp [superCommuteLiftCoef]
|
|
|
|
|
|
2024-12-20 11:09:09 +00:00
|
|
|
|
lemma superCommuteCoef_perm_snd (la lb lb' : List 𝓕)
|
2024-12-19 14:25:09 +00:00
|
|
|
|
(h : lb.Perm lb') :
|
|
|
|
|
superCommuteCoef q la lb = superCommuteCoef q la lb' := by
|
2024-12-20 11:09:09 +00:00
|
|
|
|
rw [superCommuteCoef, superCommuteCoef, FieldStatistic.ofList_perm q h]
|
2024-12-19 14:25:09 +00:00
|
|
|
|
|
2024-12-20 11:09:09 +00:00
|
|
|
|
lemma superCommuteCoef_mul_self (l lb : List 𝓕) :
|
2024-12-19 14:25:09 +00:00
|
|
|
|
superCommuteCoef q l lb * superCommuteCoef q l lb = 1 := by
|
|
|
|
|
simp only [superCommuteCoef, Fin.isValue, mul_ite, mul_neg, mul_one]
|
2024-12-20 11:09:09 +00:00
|
|
|
|
have ha (a b : FieldStatistic) : (if a = fermionic ∧ b = fermionic then
|
|
|
|
|
-if a = fermionic ∧ b = fermionic then -1 else 1
|
|
|
|
|
else if a = fermionic ∧ b = fermionic then -1 else 1) = (1 : ℂ) := by
|
2024-12-19 14:25:09 +00:00
|
|
|
|
fin_cases a <;> fin_cases b
|
|
|
|
|
any_goals rfl
|
|
|
|
|
simp
|
2024-12-20 11:09:09 +00:00
|
|
|
|
exact ha (FieldStatistic.ofList q l) (FieldStatistic.ofList q lb)
|
2024-12-19 14:25:09 +00:00
|
|
|
|
|
2024-12-20 11:09:09 +00:00
|
|
|
|
lemma superCommuteCoef_empty (la : List 𝓕) :
|
2024-12-19 14:25:09 +00:00
|
|
|
|
superCommuteCoef q la [] = 1 := by
|
2024-12-20 11:09:09 +00:00
|
|
|
|
simp only [superCommuteCoef, ofList_empty, reduceCtorEq, and_false, ↓reduceIte]
|
2024-12-19 14:25:09 +00:00
|
|
|
|
|
2024-12-20 11:09:09 +00:00
|
|
|
|
lemma superCommuteCoef_append (la lb lc : List 𝓕) :
|
2024-12-19 14:25:09 +00:00
|
|
|
|
superCommuteCoef q la (lb ++ lc) = superCommuteCoef q la lb * superCommuteCoef q la lc := by
|
2024-12-20 11:09:09 +00:00
|
|
|
|
simp only [superCommuteCoef, Fin.isValue, ofList_append, ite_eq_right_iff, zero_ne_one, imp_false,
|
2024-12-19 14:25:09 +00:00
|
|
|
|
mul_ite, mul_neg, mul_one]
|
2024-12-20 11:09:09 +00:00
|
|
|
|
by_cases hla : ofList q la = fermionic
|
|
|
|
|
· by_cases hlb : ofList q lb = fermionic
|
|
|
|
|
· by_cases hlc : ofList q lc = fermionic
|
2024-12-19 14:25:09 +00:00
|
|
|
|
· simp [hlc, hlb, hla]
|
2024-12-20 11:09:09 +00:00
|
|
|
|
· have hc : ofList q lc = bosonic := by
|
|
|
|
|
exact (neq_fermionic_iff_eq_bosonic (ofList q lc)).mp hlc
|
2024-12-19 14:25:09 +00:00
|
|
|
|
simp [hc, hlb, hla]
|
2024-12-20 11:09:09 +00:00
|
|
|
|
· have hb : ofList q lb = bosonic := by
|
|
|
|
|
exact (neq_fermionic_iff_eq_bosonic (ofList q lb)).mp hlb
|
|
|
|
|
by_cases hlc : ofList q lc = fermionic
|
2024-12-19 14:25:09 +00:00
|
|
|
|
· simp [hlc, hb]
|
2024-12-20 11:09:09 +00:00
|
|
|
|
· have hc : ofList q lc = bosonic := by
|
|
|
|
|
exact (neq_fermionic_iff_eq_bosonic (ofList q lc)).mp hlc
|
2024-12-19 14:25:09 +00:00
|
|
|
|
simp [hc, hb]
|
2024-12-20 11:09:09 +00:00
|
|
|
|
· have ha : ofList q la = bosonic := by
|
|
|
|
|
exact (neq_fermionic_iff_eq_bosonic (ofList q la)).mp hla
|
2024-12-19 14:25:09 +00:00
|
|
|
|
simp [ha]
|
|
|
|
|
|
2024-12-20 11:09:09 +00:00
|
|
|
|
lemma superCommuteCoef_cons (i : 𝓕) (la lb : List 𝓕) :
|
2024-12-19 14:25:09 +00:00
|
|
|
|
superCommuteCoef q la (i :: lb) = superCommuteCoef q la [i] * superCommuteCoef q la lb := by
|
|
|
|
|
trans superCommuteCoef q la ([i] ++ lb)
|
|
|
|
|
simp only [List.singleton_append]
|
|
|
|
|
rw [superCommuteCoef_append]
|
|
|
|
|
|
|
|
|
|
end Wick
|