PhysLean/HepLean/PerturbationTheory/FieldOpAlgebra/StaticWickTheorem.lean

71 lines
2.2 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
-/
2025-02-05 10:01:48 +00:00
import HepLean.PerturbationTheory.FieldOpAlgebra.StaticWickTerm
2025-01-30 12:45:00 +00:00
/-!
# 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
2025-02-03 10:47:18 +00:00
/--
2025-02-07 09:56:37 +00:00
For a list `φs` of `𝓕.FieldOp`, the static version of Wick's theorem states that
`φs = ∑ φsΛ, φsΛ.staticWickTerm`
2025-02-07 09:56:37 +00:00
where the sum is over all Wick contraction `φsΛ`.
The proof is via induction on `φs`.
- The base case `φs = []` is handled by `staticWickTerm_empty_nil`.
2025-02-05 08:52:14 +00:00
The inductive step works as follows:
2025-02-07 09:56:37 +00:00
For the LHS:
1. The proof considers `φ₀…φₙ` as `φ₀(φ₁…φₙ)` and use the induction hypothesis on `φ₁…φₙ`.
2025-02-07 10:34:48 +00:00
2. This gives terms of the form `φ * φsΛ.staticWickTerm` on which
2025-02-07 09:56:37 +00:00
`mul_staticWickTerm_eq_sum` is used where `φsΛ` is a Wick contraction of `φ₁…φₙ`,
to rewrite terms as a sum over optional uncontracted elements of `φsΛ`
On the LHS we now have a sum over Wick contractions `φsΛ` of `φ₁…φₙ` (from 1) and optional
uncontracted elements of `φsΛ` (from 2)
For the RHS:
1. The sum over Wick contractions of `φ₀…φₙ` on the RHS
is split via `insertLift_sum` into a sum over Wick contractions `φsΛ` of `φ₁…φₙ` and
sum over optional uncontracted elements of `φsΛ`.
Both side now are sums over the same thing and their terms equate by the nature of the
lemmas used.
2025-02-03 10:47:18 +00:00
-/
2025-02-03 11:28:14 +00:00
theorem static_wick_theorem : (φs : List 𝓕.FieldOp) →
2025-02-05 10:01:48 +00:00
ofFieldOpList φs = ∑ (φsΛ : WickContraction φs.length), φsΛ.staticWickTerm
| [] => by
simp only [ofFieldOpList, ofFieldOpListF_nil, map_one, List.length_nil]
rw [sum_WickContraction_nil]
2025-02-07 09:56:37 +00:00
rw [staticWickTerm_empty_nil]
2025-01-30 12:45:00 +00:00
| φ :: φ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-05 10:01:48 +00:00
rw [mul_staticWickTerm_eq_sum]
rfl
2025-01-30 12:45:00 +00:00
end FieldOpAlgebra
end FieldSpecification