PhysLean/HepLean/PerturbationTheory/FieldOpAlgebra/StaticWickTheorem.lean

113 lines
4.6 KiB
Text
Raw Normal View History

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
-/
import HepLean.PerturbationTheory.WickContraction.StaticContract
2025-02-03 12:12:36 +00:00
import HepLean.PerturbationTheory.FieldOpAlgebra.WicksTheorem
2025-01-30 12:45:00 +00:00
import HepLean.Meta.Remark.Basic
/-!
# Static Wick's theorem
-/
namespace FieldSpecification
variable {𝓕 : FieldSpecification}
2025-02-03 11:05:43 +00:00
open FieldOpFreeAlgebra
2025-01-30 12:45:00 +00:00
open HepLean.List
open WickContraction
open FieldStatistic
namespace FieldOpAlgebra
lemma static_wick_theorem_nil : ofFieldOpList [] = ∑ (φsΛ : WickContraction [].length),
φsΛ.sign (𝓕 := 𝓕) • φsΛ.staticContract * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ) := by
simp only [ofFieldOpList, ofFieldOpListF_nil, map_one, List.length_nil]
2025-01-30 12:45:00 +00:00
rw [sum_WickContraction_nil, uncontractedListGet, nil_zero_uncontractedList]
simp [sign, empty, staticContract]
2025-02-03 10:47:18 +00:00
/--
The static Wicks theorem states that
`φ₀…φₙ` is equal to the sum of
`φsΛ.1.sign • φsΛ.1.staticContract * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)`
over all Wick contraction `φsΛ`.
This is compared to the ordinary Wicks theorem in which `staticContract` is replaced with
`timeContract`.
-/
2025-02-03 11:28:14 +00:00
theorem static_wick_theorem : (φs : List 𝓕.FieldOp) →
2025-01-30 12:45:00 +00:00
ofFieldOpList φs = ∑ (φsΛ : WickContraction φs.length),
φsΛ.sign • φsΛ.staticContract * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ)
| [] => static_wick_theorem_nil
| φ :: φs => by
2025-02-03 10:47:18 +00:00
rw [ofFieldOpList_cons, static_wick_theorem φs]
2025-02-03 05:39:48 +00:00
rw [show (φ :: φs) = φs.insertIdx (⟨0, Nat.zero_lt_succ φs.length⟩ : Fin φs.length.succ) φ
2025-01-30 12:45:00 +00:00
from rfl]
2025-02-03 05:39:48 +00:00
conv_rhs => rw [insertLift_sum]
2025-01-30 12:45:00 +00:00
rw [Finset.mul_sum]
apply Finset.sum_congr rfl
intro c _
2025-02-03 05:39:48 +00:00
trans (sign φs c • ↑c.staticContract * (ofFieldOp φ * normalOrder (ofFieldOpList [c]ᵘᶜ)))
2025-01-30 12:45:00 +00:00
· have ht := Subalgebra.mem_center_iff.mp (Subalgebra.smul_mem (Subalgebra.center _)
2025-02-03 05:39:48 +00:00
(c.staticContract).2 c.sign)
2025-01-30 12:45:00 +00:00
conv_rhs => rw [← mul_assoc, ← ht]
simp [mul_assoc]
rw [ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum]
rw [Finset.mul_sum]
2025-02-03 11:28:14 +00:00
rw [uncontractedFieldOpEquiv_list_sum]
2025-01-30 12:45:00 +00:00
refine Finset.sum_congr rfl (fun n _ => ?_)
match n with
| none =>
2025-02-03 11:28:14 +00:00
simp only [contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
2025-01-30 12:46:16 +00:00
Equiv.coe_trans, Option.map_none', one_mul, Algebra.smul_mul_assoc, Nat.succ_eq_add_one,
Fin.zero_eta, Fin.val_zero, List.insertIdx_zero, staticContract_insertAndContract_none,
insertAndContract_uncontractedList_none_zero]
2025-01-30 12:45:00 +00:00
erw [sign_insert_none_zero]
rfl
| some n =>
2025-01-30 12:46:16 +00:00
simp only [Algebra.smul_mul_assoc, Nat.succ_eq_add_one, Fin.zero_eta, Fin.val_zero,
List.insertIdx_zero]
2025-01-30 12:45:00 +00:00
rw [normalOrder_uncontracted_some]
2025-02-03 05:39:48 +00:00
simp only [← mul_assoc]
2025-01-30 12:45:00 +00:00
rw [← smul_mul_assoc]
conv_rhs => rw [← smul_mul_assoc]
congr 1
rw [staticConract_insertAndContract_some_eq_mul_contractStateAtIndex_lt]
swap
· simp
rw [smul_smul]
2025-02-03 05:39:48 +00:00
by_cases hn : GradingCompliant φs c ∧ (𝓕|>ₛφ) = (𝓕|>ₛ φs[n.1])
2025-01-30 12:45:00 +00:00
· congr 1
swap
· have h1 := c.staticContract.2
rw [@Subalgebra.mem_center_iff] at h1
rw [h1]
erw [sign_insert_some]
rw [mul_assoc, mul_comm c.sign, ← mul_assoc]
rw [signInsertSome_mul_filter_contracted_of_not_lt]
2025-01-30 12:46:16 +00:00
simp only [instCommGroup.eq_1, Fin.zero_succAbove, Fin.not_lt_zero, Finset.filter_False,
ofFinset_empty, map_one, one_mul]
simp only [Fin.zero_succAbove, Fin.not_lt_zero, not_false_eq_true]
2025-01-30 12:45:00 +00:00
exact hn
2025-02-03 05:39:48 +00:00
· simp only [Fin.getElem_fin, not_and] at hn
2025-01-30 12:45:00 +00:00
by_cases h0 : ¬ GradingCompliant φs c
· rw [staticContract_of_not_gradingCompliant]
2025-01-30 12:46:16 +00:00
simp only [ZeroMemClass.coe_zero, zero_mul, smul_zero, instCommGroup.eq_1, mul_zero]
2025-01-30 12:45:00 +00:00
exact h0
2025-02-03 05:39:48 +00:00
· simp_all only [Finset.mem_univ, not_not, instCommGroup.eq_1, forall_const]
have h1 : contractStateAtIndex φ [c]ᵘᶜ
2025-02-03 11:28:14 +00:00
((uncontractedFieldOpEquiv φs c) (some n)) = 0 := by
simp only [contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
2025-01-30 12:45:00 +00:00
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply,
instCommGroup.eq_1, Fin.coe_cast, Fin.getElem_fin, smul_eq_zero]
right
2025-02-03 05:39:48 +00:00
simp only [uncontractedListGet, List.getElem_map,
uncontractedList_getElem_uncontractedIndexEquiv_symm, List.get_eq_getElem]
2025-02-03 11:13:23 +00:00
rw [superCommute_anPart_ofFieldOpF_diff_grade_zero]
2025-01-30 12:45:00 +00:00
exact hn
rw [h1]
simp
end FieldOpAlgebra
end FieldSpecification