Docs: Documentation related to Wick contraction (#287)
* refactor: Fix field struct defn. * rename: FieldStruct to FieldSpecification * feat: Add examples of field specifications * docs: Slight improvement of module docs * refactor: Rename CreateAnnihilate * docs: Algebras * Update CrAnStates.lean
This commit is contained in:
parent
b5c987180a
commit
2fcafb1796
5 changed files with 64 additions and 31 deletions
|
@ -8,6 +8,16 @@ import HepLean.PerturbationTheory.FieldStatistics.Basic
|
|||
|
||||
# Exchange sign for field statistics
|
||||
|
||||
Suppose we have two fields `φ` and `ψ`, and the term `φψ`, if we swap them
|
||||
`ψφ`, we may pick up a sign. This sign is called the exchange sign.
|
||||
This sign is `-1` if both fields `ψ` and `φ` are fermionic and `1` otherwise.
|
||||
|
||||
In this module we define the exchange sign for general field statistics,
|
||||
and prove some properties of it. Importantly:
|
||||
- It is symmetric `exchangeSign_symm`.
|
||||
- When multiplied with itself it is `1` `exchangeSign_mul_self`.
|
||||
- It is a cocycle `exchangeSign_cocycle`.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStatistic
|
||||
|
@ -47,16 +57,17 @@ def exchangeSign : FieldStatistic →* FieldStatistic →* ℂ where
|
|||
Defined to be `-1` if both field statistics are `fermionic` and `1` otherwise. -/
|
||||
scoped[FieldStatistic] notation "𝓢(" a "," b ")" => exchangeSign a b
|
||||
|
||||
/-- The exchange sign is symmetric. -/
|
||||
lemma exchangeSign_symm (a b : FieldStatistic) : 𝓢(a, b) = 𝓢(b, a) := by
|
||||
fin_cases a <;> fin_cases b <;> rfl
|
||||
|
||||
@[simp]
|
||||
lemma exchangeSign_bosonic (a : FieldStatistic) : 𝓢(a, bosonic) = 1 := by
|
||||
fin_cases a <;> rfl
|
||||
|
||||
@[simp]
|
||||
lemma bosonic_exchangeSign (a : FieldStatistic) : 𝓢(bosonic, a) = 1 := by
|
||||
fin_cases a <;> rfl
|
||||
|
||||
lemma exchangeSign_symm (a b : FieldStatistic) : 𝓢(a, b) = 𝓢(b, a) := by
|
||||
fin_cases a <;> fin_cases b <;> rfl
|
||||
rw [exchangeSign_symm, exchangeSign_bosonic]
|
||||
|
||||
lemma exchangeSign_eq_if (a b : FieldStatistic) :
|
||||
𝓢(a, b) = if a = fermionic ∧ b = fermionic then - 1 else 1 := by
|
||||
|
@ -75,6 +86,7 @@ lemma exchangeSign_ofList_cons (a : FieldStatistic)
|
|||
𝓢(a, ofList s (φ :: φs)) = 𝓢(a, s φ) * 𝓢(a, ofList s φs) := by
|
||||
rw [ofList_cons_eq_mul, map_mul]
|
||||
|
||||
/-- The exchange sign is a cocycle. -/
|
||||
lemma exchangeSign_cocycle (a b c : FieldStatistic) :
|
||||
𝓢(a, b * c) * 𝓢(b, c) = 𝓢(a, b) * 𝓢(a * b, c) := by
|
||||
fin_cases a <;> fin_cases b <;> fin_cases c <;> simp
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue