2025-01-30 12:45:00 +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
|
|
|
|
|
-/
|
2025-02-03 10:47:18 +00:00
|
|
|
|
import HepLean.PerturbationTheory.WickContraction.Sign.Basic
|
2025-02-03 12:12:36 +00:00
|
|
|
|
import HepLean.PerturbationTheory.FieldOpAlgebra.TimeContraction
|
2025-01-30 12:45:00 +00:00
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
# Time contractions
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
open FieldSpecification
|
|
|
|
|
variable {𝓕 : FieldSpecification}
|
|
|
|
|
|
|
|
|
|
namespace WickContraction
|
|
|
|
|
variable {n : ℕ} (c : WickContraction n)
|
|
|
|
|
open HepLean.List
|
|
|
|
|
open FieldOpAlgebra
|
2025-02-07 09:56:37 +00:00
|
|
|
|
|
2025-02-10 10:21:57 +00:00
|
|
|
|
/-- For a list `φs` of `𝓕.FieldOp` and a Wick contraction `φsΛ`, the
|
2025-02-07 09:56:37 +00:00
|
|
|
|
element of the center of `𝓕.FieldOpAlgebra`, `φsΛ.staticContract` is defined as the product
|
2025-02-10 10:21:57 +00:00
|
|
|
|
of `[anPart φs[j], φs[k]]ₛ` over contracted pairs `{j, k}` in `φsΛ`
|
2025-02-07 09:56:37 +00:00
|
|
|
|
with `j < k`. -/
|
2025-02-03 11:28:14 +00:00
|
|
|
|
noncomputable def staticContract {φs : List 𝓕.FieldOp}
|
2025-01-30 12:45:00 +00:00
|
|
|
|
(φsΛ : WickContraction φs.length) :
|
|
|
|
|
Subalgebra.center ℂ 𝓕.FieldOpAlgebra :=
|
|
|
|
|
∏ (a : φsΛ.1), ⟨[anPart (φs.get (φsΛ.fstFieldOfContract a)),
|
|
|
|
|
ofFieldOp (φs.get (φsΛ.sndFieldOfContract a))]ₛ,
|
|
|
|
|
superCommute_anPart_ofFieldOp_mem_center _ _⟩
|
|
|
|
|
|
2025-02-07 09:56:37 +00:00
|
|
|
|
/-- For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
|
2025-02-10 10:21:57 +00:00
|
|
|
|
`𝓕.FieldOp`, and a `i ≤ φs.length`, then the following relation holds:
|
2025-02-07 09:56:37 +00:00
|
|
|
|
|
|
|
|
|
`(φsΛ ↩Λ φ i none).staticContract = φsΛ.staticContract`
|
|
|
|
|
|
|
|
|
|
The prove of this result ultimately a consequence of definitions.
|
|
|
|
|
-/
|
2025-01-30 12:45:00 +00:00
|
|
|
|
@[simp]
|
2025-02-05 11:52:55 +00:00
|
|
|
|
lemma staticContract_insert_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
2025-01-30 12:45:00 +00:00
|
|
|
|
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) :
|
|
|
|
|
(φsΛ ↩Λ φ i none).staticContract = φsΛ.staticContract := by
|
|
|
|
|
rw [staticContract, insertAndContract_none_prod_contractions]
|
|
|
|
|
congr
|
|
|
|
|
ext a
|
|
|
|
|
simp
|
|
|
|
|
|
2025-02-07 09:56:37 +00:00
|
|
|
|
/--
|
|
|
|
|
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
|
2025-02-07 10:34:48 +00:00
|
|
|
|
`𝓕.FieldOp`, a `i ≤ φs.length` and a `k` in `φsΛ.uncontracted`, then
|
2025-02-07 09:56:37 +00:00
|
|
|
|
`(φsΛ ↩Λ φ i (some k)).staticContract` is equal to the product of
|
|
|
|
|
- `[anPart φ, φs[k]]ₛ` if `i ≤ k` or `[anPart φs[k], φ]ₛ` if `k < i`
|
|
|
|
|
- `φsΛ.staticContract`.
|
|
|
|
|
|
|
|
|
|
The proof of this result ultimately a consequence of definitions.
|
|
|
|
|
-/
|
|
|
|
|
lemma staticContract_insert_some
|
2025-02-03 11:28:14 +00:00
|
|
|
|
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
2025-01-30 12:45:00 +00:00
|
|
|
|
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
|
|
|
|
(φsΛ ↩Λ φ i (some j)).staticContract =
|
|
|
|
|
(if i < i.succAbove j then
|
|
|
|
|
⟨[anPart φ, ofFieldOp φs[j.1]]ₛ, superCommute_anPart_ofFieldOp_mem_center _ _⟩
|
2025-02-03 05:39:48 +00:00
|
|
|
|
else ⟨[anPart φs[j.1], ofFieldOp φ]ₛ, superCommute_anPart_ofFieldOp_mem_center _ _⟩) *
|
2025-01-30 12:45:00 +00:00
|
|
|
|
φsΛ.staticContract := by
|
|
|
|
|
rw [staticContract, insertAndContract_some_prod_contractions]
|
|
|
|
|
congr 1
|
|
|
|
|
· simp only [Nat.succ_eq_add_one, insertAndContract_fstFieldOfContract_some_incl, finCongr_apply,
|
|
|
|
|
List.get_eq_getElem, insertAndContract_sndFieldOfContract_some_incl, Fin.getElem_fin]
|
|
|
|
|
split
|
|
|
|
|
· simp
|
|
|
|
|
· simp
|
|
|
|
|
· congr
|
|
|
|
|
ext a
|
|
|
|
|
simp
|
|
|
|
|
|
|
|
|
|
open FieldStatistic
|
|
|
|
|
|
2025-02-05 11:52:55 +00:00
|
|
|
|
lemma staticContract_insert_some_of_lt
|
2025-02-03 11:28:14 +00:00
|
|
|
|
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
2025-01-30 12:45:00 +00:00
|
|
|
|
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
|
|
|
|
(hik : i < i.succAbove k) :
|
|
|
|
|
(φsΛ ↩Λ φ i (some k)).staticContract =
|
|
|
|
|
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (φsΛ.uncontracted.filter (fun x => x < k))⟩)
|
2025-02-03 11:28:14 +00:00
|
|
|
|
• (contractStateAtIndex φ [φsΛ]ᵘᶜ ((uncontractedFieldOpEquiv φs φsΛ) (some k)) *
|
2025-01-30 12:45:00 +00:00
|
|
|
|
φsΛ.staticContract) := by
|
2025-02-07 09:56:37 +00:00
|
|
|
|
rw [staticContract_insert_some]
|
2025-01-30 12:45:00 +00:00
|
|
|
|
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
|
2025-02-03 11:28:14 +00:00
|
|
|
|
contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
|
2025-01-30 12:45:00 +00:00
|
|
|
|
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
|
|
|
|
|
List.getElem_map, uncontractedList_getElem_uncontractedIndexEquiv_symm, List.get_eq_getElem,
|
|
|
|
|
Algebra.smul_mul_assoc, uncontractedListGet]
|
|
|
|
|
· simp only [hik, ↓reduceIte, MulMemClass.coe_mul]
|
|
|
|
|
trans (1 : ℂ) • ((superCommute (anPart φ)) (ofFieldOp φs[k.1]) * ↑φsΛ.staticContract)
|
|
|
|
|
· simp
|
|
|
|
|
simp only [smul_smul]
|
|
|
|
|
congr 1
|
2025-02-06 10:06:05 +00:00
|
|
|
|
have h1 : ofList 𝓕.fieldOpStatistic (List.take (↑(φsΛ.uncontractedIndexEquiv.symm k))
|
2025-01-30 12:45:00 +00:00
|
|
|
|
(List.map φs.get φsΛ.uncontractedList))
|
|
|
|
|
= (𝓕 |>ₛ ⟨φs.get, (Finset.filter (fun x => x < k) φsΛ.uncontracted)⟩) := by
|
|
|
|
|
simp only [ofFinset]
|
|
|
|
|
congr
|
|
|
|
|
rw [← List.map_take]
|
|
|
|
|
congr
|
|
|
|
|
rw [take_uncontractedIndexEquiv_symm]
|
|
|
|
|
rw [filter_uncontractedList]
|
|
|
|
|
rw [h1]
|
|
|
|
|
simp only [exchangeSign_mul_self]
|
|
|
|
|
|
2025-02-03 11:28:14 +00:00
|
|
|
|
lemma staticContract_of_not_gradingCompliant (φs : List 𝓕.FieldOp)
|
2025-01-30 12:45:00 +00:00
|
|
|
|
(φsΛ : WickContraction φs.length) (h : ¬ GradingCompliant φs φsΛ) :
|
|
|
|
|
φsΛ.staticContract = 0 := by
|
|
|
|
|
rw [staticContract]
|
|
|
|
|
simp only [GradingCompliant, Fin.getElem_fin, Subtype.forall, not_forall] at h
|
|
|
|
|
obtain ⟨a, ha⟩ := h
|
|
|
|
|
obtain ⟨ha, ha2⟩ := ha
|
|
|
|
|
apply Finset.prod_eq_zero (i := ⟨a, ha⟩)
|
|
|
|
|
simp only [Finset.univ_eq_attach, Finset.mem_attach]
|
|
|
|
|
apply Subtype.eq
|
|
|
|
|
simp only [List.get_eq_getElem, ZeroMemClass.coe_zero]
|
2025-02-03 11:13:23 +00:00
|
|
|
|
rw [superCommute_anPart_ofFieldOpF_diff_grade_zero]
|
2025-01-30 12:45:00 +00:00
|
|
|
|
simp [ha2]
|
|
|
|
|
|
|
|
|
|
end WickContraction
|