2025-02-03 10:47:18 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
|
|
|
|
import HepLean.PerturbationTheory.WickContraction.InsertAndContract
|
|
|
|
|
|
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
# Sign associated with a contraction
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
open FieldSpecification
|
|
|
|
|
variable {𝓕 : FieldSpecification}
|
|
|
|
|
|
|
|
|
|
namespace WickContraction
|
|
|
|
|
variable {n : ℕ} (c : WickContraction n)
|
|
|
|
|
open HepLean.List
|
|
|
|
|
open FieldStatistic
|
|
|
|
|
|
|
|
|
|
/-- Given a Wick contraction `c : WickContraction n` and `i1 i2 : Fin n` the finite set
|
|
|
|
|
of elements of `Fin n` between `i1` and `i2` which are either uncontracted
|
2025-02-10 10:51:44 +00:00
|
|
|
|
or are contracted but are contracted with an element occurring after `i1`.
|
2025-02-12 06:14:11 +00:00
|
|
|
|
In other words, the elements of `Fin n` between `i1` and `i2` which are not
|
|
|
|
|
contracted with before `i1`.
|
2025-02-03 10:47:18 +00:00
|
|
|
|
One should assume `i1 < i2` otherwise this finite set is empty. -/
|
|
|
|
|
def signFinset (c : WickContraction n) (i1 i2 : Fin n) : Finset (Fin n) :=
|
|
|
|
|
Finset.univ.filter (fun i => i1 < i ∧ i < i2 ∧
|
|
|
|
|
(c.getDual? i = none ∨ ∀ (h : (c.getDual? i).isSome), i1 < (c.getDual? i).get h))
|
|
|
|
|
|
2025-02-07 06:58:41 +00:00
|
|
|
|
/-- For a list `φs` of `𝓕.FieldOp`, and a Wick contraction `φsΛ` of `φs`,
|
|
|
|
|
the complex number `φsΛ.sign` is defined to be the sign (`1` or `-1`) corresponding
|
2025-02-10 10:21:57 +00:00
|
|
|
|
to the number of `fermionic`-`fermionic` exchanges that must be done to put
|
|
|
|
|
contracted pairs within `φsΛ` next to one another, starting recursively
|
|
|
|
|
from the contracted pair
|
2025-02-13 10:44:15 +00:00
|
|
|
|
whose first element occurs at the left-most position.
|
|
|
|
|
|
|
|
|
|
As an example, if `[φ1, φ2, φ3, φ4]` correspond to fermionic fields then the sign
|
|
|
|
|
associated with
|
|
|
|
|
- `{{0, 1}}` is `1`
|
|
|
|
|
- `{{0, 1}, {2, 3}}` is `1`
|
|
|
|
|
- `{{0, 2}, {1, 3}}` is `-1`
|
|
|
|
|
-/
|
2025-02-03 11:28:14 +00:00
|
|
|
|
def sign (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) : ℂ :=
|
2025-02-03 10:47:18 +00:00
|
|
|
|
∏ (a : φsΛ.1), 𝓢(𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a],
|
|
|
|
|
𝓕 |>ₛ ⟨φs.get, φsΛ.signFinset (φsΛ.fstFieldOfContract a) (φsΛ.sndFieldOfContract a)⟩)
|
|
|
|
|
|
2025-02-03 11:28:14 +00:00
|
|
|
|
lemma sign_empty (φs : List 𝓕.FieldOp) :
|
2025-02-03 10:47:18 +00:00
|
|
|
|
sign φs empty = 1 := by
|
|
|
|
|
rw [sign]
|
|
|
|
|
simp [empty]
|
|
|
|
|
|
2025-02-03 11:28:14 +00:00
|
|
|
|
lemma sign_congr {φs φs' : List 𝓕.FieldOp} (h : φs = φs') (φsΛ : WickContraction φs.length) :
|
2025-02-03 10:47:18 +00:00
|
|
|
|
sign φs' (congr (by simp [h]) φsΛ) = sign φs φsΛ := by
|
|
|
|
|
subst h
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
end WickContraction
|