feat: Join of Wick contractions

This commit is contained in:
jstoobysmith 2025-01-31 16:02:02 +00:00
parent ab7f479fdc
commit 12d36dc1d9
11 changed files with 1536 additions and 1 deletions

View file

@ -106,4 +106,24 @@ lemma ofFinset_union_disjoint (q : 𝓕 → FieldStatistic) (φs : List 𝓕) (a
rw [ofFinset_union, Finset.disjoint_iff_inter_eq_empty.mp h]
simp
lemma ofFinset_filter_mul_neg (q : 𝓕 → FieldStatistic) (φs : List 𝓕) (a : Finset (Fin φs.length))
(p : Fin φs.length → Prop) [DecidablePred p] :
ofFinset q φs.get (Finset.filter p a) *
ofFinset q φs.get (Finset.filter (fun i => ¬ p i) a) = ofFinset q φs.get a := by
rw [ofFinset_union_disjoint]
congr
exact Finset.filter_union_filter_neg_eq p a
exact Finset.disjoint_filter_filter_neg a a p
lemma ofFinset_filter (q : 𝓕 → FieldStatistic) (φs : List 𝓕) (a : Finset (Fin φs.length))
(p : Fin φs.length → Prop) [DecidablePred p] :
ofFinset q φs.get (Finset.filter p a) = ofFinset q φs.get (Finset.filter (fun i => ¬ p i) a) *
ofFinset q φs.get a := by
rw [← ofFinset_filter_mul_neg q φs a p]
conv_rhs =>
rhs
rw [mul_comm]
rw [← mul_assoc]
simp
end FieldStatistic