feat: Wick's theorem for normal order

This commit is contained in:
jstoobysmith 2025-02-01 11:51:06 +00:00
parent 12d36dc1d9
commit 006e29fd08
11 changed files with 1055 additions and 2 deletions

View file

@ -98,6 +98,43 @@ lemma subContraction_uncontractedList_get {S : Finset (Finset (Fin φs.length))}
erw [← getElem_uncontractedListEmd]
rfl
@[simp]
lemma subContraction_fstFieldOfContract {S : Finset (Finset (Fin φs.length))} {hs : S ⊆ φsΛ.1}
(a : (subContraction S hs).1) :
(subContraction S hs).fstFieldOfContract a =
φsΛ.fstFieldOfContract ⟨a.1, mem_of_mem_subContraction a.2⟩:= by
apply eq_fstFieldOfContract_of_mem _ _ _ (φsΛ.sndFieldOfContract ⟨a.1, mem_of_mem_subContraction a.2⟩)
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _ ⟨a.1, mem_of_mem_subContraction a.2⟩
simp at ha
conv_lhs =>
rw [ha]
simp
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _ ⟨a.1, mem_of_mem_subContraction a.2⟩
simp at ha
conv_lhs =>
rw [ha]
simp
· exact fstFieldOfContract_lt_sndFieldOfContract φsΛ ⟨↑a, mem_of_mem_subContraction a.property⟩
@[simp]
lemma subContraction_sndFieldOfContract {S : Finset (Finset (Fin φs.length))} {hs : S ⊆ φsΛ.1}
(a : (subContraction S hs).1) :
(subContraction S hs).sndFieldOfContract a =
φsΛ.sndFieldOfContract ⟨a.1, mem_of_mem_subContraction a.2⟩:= by
apply eq_sndFieldOfContract_of_mem _ _ (φsΛ.fstFieldOfContract ⟨a.1, mem_of_mem_subContraction a.2⟩)
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _ ⟨a.1, mem_of_mem_subContraction a.2⟩
simp at ha
conv_lhs =>
rw [ha]
simp
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _ ⟨a.1, mem_of_mem_subContraction a.2⟩
simp at ha
conv_lhs =>
rw [ha]
simp
· exact fstFieldOfContract_lt_sndFieldOfContract φsΛ ⟨↑a, mem_of_mem_subContraction a.property⟩
@[simp]
lemma quotContraction_fstFieldOfContract_uncontractedListEmd {S : Finset (Finset (Fin φs.length))}
{hs : S ⊆ φsΛ.1} (a : (quotContraction S hs).1) :
@ -133,5 +170,18 @@ lemma quotContraction_gradingCompliant {S : Finset (Finset (Fin φs.length))} {h
simp
apply hsΛ
lemma mem_quotContraction_iff {S : Finset (Finset (Fin φs.length))} {hs : S ⊆ φsΛ.1}
{a : Finset (Fin [φsΛ.subContraction S hs]ᵘᶜ.length)} :
a ∈ (quotContraction S hs).1 ↔ a.map uncontractedListEmd ∈ φsΛ.1
∧ a.map uncontractedListEmd ∉ (subContraction S hs).1 := by
apply Iff.intro
· intro h
apply And.intro
· exact mem_of_mem_quotContraction h
· exact uncontractedListEmd_finset_not_mem _
· intro h
have h2 := mem_subContraction_or_quotContraction (S := S) (hs := hs) h.1
simp_all
end WickContraction