feat: Join of Wick contractions
This commit is contained in:
parent
ab7f479fdc
commit
12d36dc1d9
11 changed files with 1536 additions and 1 deletions
|
@ -52,6 +52,66 @@ lemma normalOrderF_one : normalOrderF (𝓕 := 𝓕) 1 = 1 := by
|
|||
rw [← ofCrAnList_nil, normalOrderF_ofCrAnList, normalOrderSign_nil, normalOrderList_nil,
|
||||
ofCrAnList_nil, one_smul]
|
||||
|
||||
lemma normalOrderF_normalOrderF_mid (a b c : 𝓕.CrAnAlgebra) : 𝓝ᶠ(a * b * c) = 𝓝ᶠ(a * 𝓝ᶠ(b) * c) := by
|
||||
let pc (c : 𝓕.CrAnAlgebra) (hc : c ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓝ᶠ(a * b * c) = 𝓝ᶠ(a * 𝓝ᶠ(b) * c)
|
||||
change pc c (Basis.mem_span _ c)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pc]
|
||||
let pb (b : 𝓕.CrAnAlgebra) (hb : b ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓝ᶠ(a * b * ofCrAnList φs) = 𝓝ᶠ(a * 𝓝ᶠ(b) * ofCrAnList φs)
|
||||
change pb b (Basis.mem_span _ b)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pb]
|
||||
let pa (a : 𝓕.CrAnAlgebra) (ha : a ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓝ᶠ(a * ofCrAnList φs' * ofCrAnList φs) = 𝓝ᶠ(a * 𝓝ᶠ(ofCrAnList φs') * ofCrAnList φs)
|
||||
change pa a (Basis.mem_span _ a)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs'', rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pa]
|
||||
rw [normalOrderF_ofCrAnList]
|
||||
simp only [← ofCrAnList_append, Algebra.mul_smul_comm,
|
||||
Algebra.smul_mul_assoc, map_smul]
|
||||
rw [normalOrderF_ofCrAnList, normalOrderF_ofCrAnList, smul_smul]
|
||||
congr 1
|
||||
· simp only [normalOrderSign, normalOrderList]
|
||||
rw [Wick.koszulSign_of_append_eq_insertionSort, mul_comm]
|
||||
· congr 1
|
||||
simp only [normalOrderList]
|
||||
rw [HepLean.List.insertionSort_append_insertionSort_append]
|
||||
· simp [pa]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pa, add_mul]
|
||||
· intro x hx h
|
||||
simp_all [pa]
|
||||
· simp [pb]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pb, mul_add, add_mul]
|
||||
· intro x hx h
|
||||
simp_all [pb]
|
||||
· simp [pc]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pc, mul_add]
|
||||
· intro x hx h hp
|
||||
simp_all [pc]
|
||||
|
||||
lemma normalOrderF_normalOrderF_right (a b : 𝓕.CrAnAlgebra) : 𝓝ᶠ(a * b) = 𝓝ᶠ(a * 𝓝ᶠ(b)) := by
|
||||
trans 𝓝ᶠ(a * b * 1)
|
||||
· simp
|
||||
· rw [normalOrderF_normalOrderF_mid]
|
||||
simp
|
||||
|
||||
lemma normalOrderF_normalOrderF_left (a b : 𝓕.CrAnAlgebra) : 𝓝ᶠ(a * b) = 𝓝ᶠ(𝓝ᶠ(a) * b) := by
|
||||
trans 𝓝ᶠ(1 * a * b)
|
||||
· simp
|
||||
· rw [normalOrderF_normalOrderF_mid]
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## Normal ordering with a creation operator on the left or annihilation on the right
|
||||
|
|
|
@ -261,6 +261,43 @@ lemma ofCrAnFieldOpList_eq_normalOrder (φs : List 𝓕.CrAnStates) :
|
|||
rw [normalOrder_ofCrAnFieldOpList, smul_smul, normalOrderSign, Wick.koszulSign_mul_self,
|
||||
one_smul]
|
||||
|
||||
lemma normalOrder_normalOrder_mid (a b c : 𝓕.FieldOpAlgebra) :
|
||||
𝓝(a * b * c) = 𝓝(a * 𝓝(b) * c) := by
|
||||
obtain ⟨a, rfl⟩ := ι_surjective a
|
||||
obtain ⟨b, rfl⟩ := ι_surjective b
|
||||
obtain ⟨c, rfl⟩ := ι_surjective c
|
||||
rw [normalOrder_eq_ι_normalOrderF]
|
||||
simp [← map_mul]
|
||||
rw [normalOrder_eq_ι_normalOrderF]
|
||||
rw [normalOrderF_normalOrderF_mid]
|
||||
rfl
|
||||
|
||||
lemma normalOrder_normalOrder_left (a b : 𝓕.FieldOpAlgebra) :
|
||||
𝓝(a * b) = 𝓝(𝓝(a) * b) := by
|
||||
obtain ⟨a, rfl⟩ := ι_surjective a
|
||||
obtain ⟨b, rfl⟩ := ι_surjective b
|
||||
rw [normalOrder_eq_ι_normalOrderF]
|
||||
simp [← map_mul]
|
||||
rw [normalOrder_eq_ι_normalOrderF]
|
||||
rw [normalOrderF_normalOrderF_left]
|
||||
rfl
|
||||
|
||||
lemma normalOrder_normalOrder_right (a b : 𝓕.FieldOpAlgebra) :
|
||||
𝓝(a * b) = 𝓝(a * 𝓝(b)) := by
|
||||
obtain ⟨a, rfl⟩ := ι_surjective a
|
||||
obtain ⟨b, rfl⟩ := ι_surjective b
|
||||
rw [normalOrder_eq_ι_normalOrderF]
|
||||
simp [← map_mul]
|
||||
rw [normalOrder_eq_ι_normalOrderF]
|
||||
rw [normalOrderF_normalOrderF_right]
|
||||
rfl
|
||||
|
||||
lemma normalOrder_normalOrder (a : 𝓕.FieldOpAlgebra) : 𝓝(𝓝(a)) = 𝓝(a) := by
|
||||
trans 𝓝(𝓝(a) * 1)
|
||||
· simp
|
||||
· rw [← normalOrder_normalOrder_left]
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## mul anpart and crpart
|
||||
|
|
|
@ -81,6 +81,19 @@ lemma timeContract_zero_of_diff_grade (φ ψ : 𝓕.States) (h : (𝓕 |>ₛ φ)
|
|||
have ht := IsTotal.total (r := 𝓕.timeOrderRel) φ ψ
|
||||
simp_all
|
||||
|
||||
lemma normalOrder_timeContract (φ ψ : 𝓕.States) :
|
||||
𝓝(timeContract φ ψ) = 0 := by
|
||||
by_cases h : timeOrderRel φ ψ
|
||||
· rw [timeContract_of_timeOrderRel _ _ h]
|
||||
simp
|
||||
· rw [timeContract_of_not_timeOrderRel _ _ h]
|
||||
simp
|
||||
have h1 : timeOrderRel ψ φ := by
|
||||
have ht : timeOrderRel φ ψ ∨ timeOrderRel ψ φ := IsTotal.total (r := 𝓕.timeOrderRel) φ ψ
|
||||
simp_all
|
||||
rw [timeContract_of_timeOrderRel _ _ h1]
|
||||
simp
|
||||
|
||||
end FieldOpAlgebra
|
||||
|
||||
end
|
||||
|
|
|
@ -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
|
||||
|
|
|
@ -38,9 +38,17 @@ namespace WickContraction
|
|||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
|
||||
/-- Wick contractions are decidable. -/
|
||||
instance : DecidableEq (WickContraction n) := Subtype.instDecidableEq
|
||||
|
||||
/-- The contraction consisting of no contracted pairs. -/
|
||||
def empty : WickContraction n := ⟨∅, by simp, by simp⟩
|
||||
|
||||
@[simp]
|
||||
lemma card_zero_iff_empty (c : WickContraction n) : c.1.card = 0 ↔ c = empty := by
|
||||
rw [Subtype.eq_iff]
|
||||
simp [empty]
|
||||
|
||||
/-- The equivalence between `WickContraction n` and `WickContraction m`
|
||||
derived from a propositional equality of `n` and `m`. -/
|
||||
def congr : {n m : ℕ} → (h : n = m) → WickContraction n ≃ WickContraction m
|
||||
|
@ -48,9 +56,14 @@ def congr : {n m : ℕ} → (h : n = m) → WickContraction n ≃ WickContractio
|
|||
|
||||
@[simp]
|
||||
lemma congr_refl : c.congr rfl = c := by
|
||||
cases c
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma card_congr {n m : ℕ} (h : n = m) (c : WickContraction n) :
|
||||
(congr h c).1.card = c.1.card := by
|
||||
subst h
|
||||
simp
|
||||
|
||||
lemma congr_contractions {n m : ℕ} (h : n = m) (c : WickContraction n) :
|
||||
((congr h) c).1 = Finset.map (Finset.mapEmbedding (finCongr h)).toEmbedding c.1 := by
|
||||
subst h
|
||||
|
@ -83,6 +96,11 @@ lemma congr_trans_apply {n m o : ℕ} (h1 : n = m) (h2 : m = o) (c : WickContrac
|
|||
subst h1 h2
|
||||
simp
|
||||
|
||||
lemma mem_congr_iff {n m : ℕ} (h : n = m) {c : WickContraction n } {a : Finset (Fin m)} :
|
||||
a ∈ (congr h c).1 ↔ Finset.map (finCongr h.symm).toEmbedding a ∈ c.1 := by
|
||||
subst h
|
||||
simp
|
||||
|
||||
/-- Given a contracted pair in `c : WickContraction n` the contracted pair
|
||||
in `congr h c`. -/
|
||||
def congrLift {n m : ℕ} (h : n = m) {c : WickContraction n} (a : c.1) : (congr h c).1 :=
|
||||
|
@ -112,6 +130,18 @@ lemma congrLift_bijective {n m : ℕ} {c : WickContraction n} (h : n = m) :
|
|||
Function.Bijective (c.congrLift h) := by
|
||||
exact ⟨c.congrLift_injective h, c.congrLift_surjective h⟩
|
||||
|
||||
/-- Given a contracted pair in `c : WickContraction n` the contracted pair
|
||||
in `congr h c`. -/
|
||||
def congrLiftInv {n m : ℕ} (h : n = m) {c : WickContraction n} (a : (congr h c).1 ) : c.1 :=
|
||||
⟨a.1.map (finCongr h.symm).toEmbedding, by
|
||||
subst h
|
||||
simp⟩
|
||||
|
||||
lemma congrLiftInv_rfl {n : ℕ} {c : WickContraction n} :
|
||||
c.congrLiftInv rfl = id := by
|
||||
funext a
|
||||
simp [congrLiftInv]
|
||||
|
||||
lemma eq_filter_mem_self : c.1 = Finset.filter (fun x => x ∈ c.1) Finset.univ := by
|
||||
exact Eq.symm (Finset.filter_univ_mem c.1)
|
||||
|
||||
|
@ -481,6 +511,12 @@ lemma prod_finset_eq_mul_fst_snd (c : WickContraction n) (a : c.1)
|
|||
def GradingCompliant (φs : List 𝓕.States) (φsΛ : WickContraction φs.length) :=
|
||||
∀ (a : φsΛ.1), (𝓕 |>ₛ φs[φsΛ.fstFieldOfContract a]) = (𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a])
|
||||
|
||||
lemma gradingCompliant_congr {φs φs' : List 𝓕.States} (h : φs = φs')
|
||||
(φsΛ : WickContraction φs.length) :
|
||||
GradingCompliant φs φsΛ ↔ GradingCompliant φs' (congr (by simp [h]) φsΛ) := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
/-- An equivalence from the sigma type `(a : c.1) × a` to the subtype of `Fin n` consisting of
|
||||
those positions which are contracted. -/
|
||||
def sigmaContractedEquiv : (a : c.1) × a ≃ {x : Fin n // (c.getDual? x).isSome} where
|
||||
|
|
978
HepLean/PerturbationTheory/WickContraction/Join.lean
Normal file
978
HepLean/PerturbationTheory/WickContraction/Join.lean
Normal file
|
@ -0,0 +1,978 @@
|
|||
/-
|
||||
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.TimeContract
|
||||
import HepLean.PerturbationTheory.WickContraction.StaticContract
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.TimeContraction
|
||||
import HepLean.PerturbationTheory.WickContraction.SubContraction
|
||||
import HepLean.PerturbationTheory.WickContraction.Singleton
|
||||
|
||||
/-!
|
||||
|
||||
# Join of contractions
|
||||
|
||||
-/
|
||||
|
||||
open FieldSpecification
|
||||
variable {𝓕 : FieldSpecification}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
open FieldOpAlgebra
|
||||
|
||||
def join {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) : WickContraction φs.length :=
|
||||
⟨φsΛ.1 ∪ φsucΛ.1.map (Finset.mapEmbedding uncontractedListEmd).toEmbedding, by
|
||||
intro a ha
|
||||
simp at ha
|
||||
rcases ha with ha | ha
|
||||
· exact φsΛ.2.1 a ha
|
||||
· obtain ⟨a, ha, rfl⟩ := ha
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
simp only [Finset.card_map]
|
||||
exact φsucΛ.2.1 a ha, by
|
||||
intro a ha b hb
|
||||
simp at ha hb
|
||||
rcases ha with ha | ha <;> rcases hb with hb | hb
|
||||
· exact φsΛ.2.2 a ha b hb
|
||||
· obtain ⟨b, hb, rfl⟩ := hb
|
||||
right
|
||||
symm
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
apply uncontractedListEmd_finset_disjoint_left
|
||||
exact ha
|
||||
· obtain ⟨a, ha, rfl⟩ := ha
|
||||
right
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
apply uncontractedListEmd_finset_disjoint_left
|
||||
exact hb
|
||||
· obtain ⟨a, ha, rfl⟩ := ha
|
||||
obtain ⟨b, hb, rfl⟩ := hb
|
||||
simp only [EmbeddingLike.apply_eq_iff_eq]
|
||||
rw [Finset.mapEmbedding_apply, Finset.mapEmbedding_apply]
|
||||
rw [Finset.disjoint_map]
|
||||
exact φsucΛ.2.2 a ha b hb⟩
|
||||
|
||||
lemma join_congr {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} {φsΛ' : WickContraction φs.length}
|
||||
(h1 : φsΛ = φsΛ') :
|
||||
join φsΛ φsucΛ = join φsΛ' (congr (by simp [h1]) φsucΛ):= by
|
||||
subst h1
|
||||
rfl
|
||||
|
||||
|
||||
|
||||
def joinLiftLeft {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : φsΛ.1 → (join φsΛ φsucΛ).1 :=
|
||||
fun a => ⟨a, by simp [join]⟩
|
||||
|
||||
lemma jointLiftLeft_injective {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} :
|
||||
Function.Injective (@joinLiftLeft _ _ φsΛ φsucΛ) := by
|
||||
intro a b h
|
||||
simp only [joinLiftLeft] at h
|
||||
rw [Subtype.mk_eq_mk] at h
|
||||
refine Subtype.eq h
|
||||
|
||||
def joinLiftRight {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : φsucΛ.1 → (join φsΛ φsucΛ).1 :=
|
||||
fun a => ⟨a.1.map uncontractedListEmd, by
|
||||
simp only [join, Finset.le_eq_subset, Finset.mem_union, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding]
|
||||
right
|
||||
use a.1
|
||||
simp only [Finset.coe_mem, true_and]
|
||||
rfl⟩
|
||||
|
||||
lemma joinLiftRight_injective {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} :
|
||||
Function.Injective (@joinLiftRight _ _ φsΛ φsucΛ) := by
|
||||
intro a b h
|
||||
simp only [joinLiftRight] at h
|
||||
rw [Subtype.mk_eq_mk] at h
|
||||
simp only [Finset.map_inj] at h
|
||||
refine Subtype.eq h
|
||||
|
||||
lemma jointLiftLeft_disjoint_joinLiftRight {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} (a : φsΛ.1) (b : φsucΛ.1) :
|
||||
Disjoint (@joinLiftLeft _ _ _ φsucΛ a).1 (joinLiftRight b).1 := by
|
||||
simp only [joinLiftLeft, joinLiftRight]
|
||||
symm
|
||||
apply uncontractedListEmd_finset_disjoint_left
|
||||
exact a.2
|
||||
|
||||
lemma jointLiftLeft_neq_joinLiftRight {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} (a : φsΛ.1) (b : φsucΛ.1) :
|
||||
joinLiftLeft a ≠ joinLiftRight b := by
|
||||
by_contra hn
|
||||
have h1 := jointLiftLeft_disjoint_joinLiftRight a b
|
||||
rw [hn] at h1
|
||||
simp at h1
|
||||
have hj := (join φsΛ φsucΛ).2.1 (joinLiftRight b).1 (joinLiftRight b).2
|
||||
rw [h1] at hj
|
||||
simp at hj
|
||||
|
||||
def joinLift {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : φsΛ.1 ⊕ φsucΛ.1 → (join φsΛ φsucΛ).1 := fun a =>
|
||||
match a with
|
||||
| Sum.inl a => joinLiftLeft a
|
||||
| Sum.inr a => joinLiftRight a
|
||||
|
||||
lemma joinLift_injective {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : Function.Injective (@joinLift _ _ φsΛ φsucΛ) := by
|
||||
intro a b h
|
||||
match a, b with
|
||||
| Sum.inl a, Sum.inl b =>
|
||||
simp
|
||||
exact jointLiftLeft_injective h
|
||||
| Sum.inr a, Sum.inr b =>
|
||||
simp
|
||||
exact joinLiftRight_injective h
|
||||
| Sum.inl a, Sum.inr b =>
|
||||
simp [joinLift] at h
|
||||
have h1 := jointLiftLeft_neq_joinLiftRight a b
|
||||
simp_all
|
||||
| Sum.inr a, Sum.inl b =>
|
||||
simp [joinLift] at h
|
||||
have h1 := jointLiftLeft_neq_joinLiftRight b a
|
||||
simp_all
|
||||
|
||||
lemma joinLift_surjective {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : Function.Surjective (@joinLift _ _ φsΛ φsucΛ) := by
|
||||
intro a
|
||||
have ha2 := a.2
|
||||
simp [- Finset.coe_mem, join] at ha2
|
||||
rcases ha2 with ha2 | ⟨a2, ha3⟩
|
||||
· use Sum.inl ⟨a, ha2⟩
|
||||
simp [joinLift, joinLiftLeft]
|
||||
· rw [Finset.mapEmbedding_apply] at ha3
|
||||
use Sum.inr ⟨a2, ha3.1⟩
|
||||
simp [joinLift, joinLiftRight]
|
||||
refine Subtype.eq ?_
|
||||
exact ha3.2
|
||||
|
||||
lemma joinLift_bijective {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : Function.Bijective (@joinLift _ _ φsΛ φsucΛ) := by
|
||||
apply And.intro
|
||||
· exact joinLift_injective
|
||||
· exact joinLift_surjective
|
||||
|
||||
lemma prod_join {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (f : (join φsΛ φsucΛ).1 → M) [ CommMonoid M]:
|
||||
∏ (a : (join φsΛ φsucΛ).1), f a = (∏ (a : φsΛ.1), f (joinLiftLeft a)) *
|
||||
∏ (a : φsucΛ.1), f (joinLiftRight a) := by
|
||||
let e1 := Equiv.ofBijective (@joinLift _ _ φsΛ φsucΛ) joinLift_bijective
|
||||
rw [← e1.prod_comp]
|
||||
simp only [Fintype.prod_sum_type, Finset.univ_eq_attach]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma join_fstFieldOfContract_joinLiftRight {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (a : φsucΛ.1) :
|
||||
(join φsΛ φsucΛ).fstFieldOfContract (joinLiftRight a) =
|
||||
uncontractedListEmd (φsucΛ.fstFieldOfContract a) := by
|
||||
apply eq_fstFieldOfContract_of_mem _ _ _ (uncontractedListEmd (φsucΛ.sndFieldOfContract a))
|
||||
· simp [joinLiftRight]
|
||||
· simp [joinLiftRight]
|
||||
· apply uncontractedListEmd_strictMono
|
||||
exact fstFieldOfContract_lt_sndFieldOfContract φsucΛ a
|
||||
|
||||
@[simp]
|
||||
lemma join_sndFieldOfContract_joinLiftRight {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (a : φsucΛ.1) :
|
||||
(join φsΛ φsucΛ).sndFieldOfContract (joinLiftRight a) =
|
||||
uncontractedListEmd (φsucΛ.sndFieldOfContract a) := by
|
||||
apply eq_sndFieldOfContract_of_mem _ _ (uncontractedListEmd (φsucΛ.fstFieldOfContract a))
|
||||
· simp [joinLiftRight]
|
||||
· simp [joinLiftRight]
|
||||
· apply uncontractedListEmd_strictMono
|
||||
exact fstFieldOfContract_lt_sndFieldOfContract φsucΛ a
|
||||
|
||||
@[simp]
|
||||
lemma join_fstFieldOfContract_joinLiftLeft {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (a : φsΛ.1) :
|
||||
(join φsΛ φsucΛ).fstFieldOfContract (joinLiftLeft a) =
|
||||
(φsΛ.fstFieldOfContract a) := by
|
||||
apply eq_fstFieldOfContract_of_mem _ _ _ (φsΛ.sndFieldOfContract a)
|
||||
· simp [joinLiftLeft]
|
||||
· simp [joinLiftLeft]
|
||||
· exact fstFieldOfContract_lt_sndFieldOfContract φsΛ a
|
||||
|
||||
@[simp]
|
||||
lemma join_sndFieldOfContract_joinLift {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (a : φsΛ.1) :
|
||||
(join φsΛ φsucΛ).sndFieldOfContract (joinLiftLeft a) =
|
||||
(φsΛ.sndFieldOfContract a) := by
|
||||
apply eq_sndFieldOfContract_of_mem _ _ (φsΛ.fstFieldOfContract a) (φsΛ.sndFieldOfContract a)
|
||||
· simp [joinLiftLeft]
|
||||
· simp [joinLiftLeft]
|
||||
· exact fstFieldOfContract_lt_sndFieldOfContract φsΛ a
|
||||
|
||||
|
||||
lemma join_card {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} :
|
||||
(join φsΛ φsucΛ).1.card = φsΛ.1.card + φsucΛ.1.card := by
|
||||
simp [join]
|
||||
rw [Finset.card_union_of_disjoint]
|
||||
simp
|
||||
rw [@Finset.disjoint_left]
|
||||
intro a ha
|
||||
simp
|
||||
intro x hx
|
||||
by_contra hn
|
||||
have hdis : Disjoint (Finset.map uncontractedListEmd x) a := by
|
||||
exact uncontractedListEmd_finset_disjoint_left x a ha
|
||||
rw [Finset.mapEmbedding_apply] at hn
|
||||
rw [hn] at hdis
|
||||
simp at hdis
|
||||
have hcard := φsΛ.2.1 a ha
|
||||
simp_all
|
||||
|
||||
@[simp]
|
||||
lemma empty_join {φs : List 𝓕.States} (φsΛ : WickContraction [empty (n := φs.length)]ᵘᶜ.length) :
|
||||
join empty φsΛ = congr (by simp) φsΛ := by
|
||||
apply Subtype.ext
|
||||
simp [join]
|
||||
ext a
|
||||
conv_lhs =>
|
||||
left
|
||||
left
|
||||
rw [empty]
|
||||
simp
|
||||
rw [mem_congr_iff]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
obtain ⟨a, ha, rfl⟩ := h
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
rw [Finset.map_map]
|
||||
apply Set.mem_of_eq_of_mem _ ha
|
||||
trans Finset.map (Equiv.refl _).toEmbedding a
|
||||
rfl
|
||||
simp
|
||||
· intro h
|
||||
use Finset.map (finCongr (by simp)).toEmbedding a
|
||||
simp [h]
|
||||
trans Finset.map (Equiv.refl _).toEmbedding a
|
||||
rw [Finset.mapEmbedding_apply, Finset.map_map]
|
||||
rfl
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma join_empty {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
join φsΛ empty = φsΛ := by
|
||||
apply Subtype.ext
|
||||
ext a
|
||||
simp [join, empty]
|
||||
|
||||
lemma join_timeContract {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).timeContract = φsΛ.timeContract * φsucΛ.timeContract := by
|
||||
simp only [timeContract, List.get_eq_getElem]
|
||||
rw [prod_join]
|
||||
congr 1
|
||||
congr
|
||||
funext a
|
||||
simp
|
||||
|
||||
lemma mem_join_uncontracted_of_mem_right_uncontracted {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin [φsΛ]ᵘᶜ.length)
|
||||
(ha : i ∈ φsucΛ.uncontracted) :
|
||||
uncontractedListEmd i ∈ (join φsΛ φsucΛ).uncontracted := by
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
intro p hp
|
||||
simp [join] at hp
|
||||
rcases hp with hp | hp
|
||||
· have hi : uncontractedListEmd i ∈ φsΛ.uncontracted := by
|
||||
exact uncontractedListEmd_mem_uncontracted i
|
||||
rw [mem_uncontracted_iff_not_contracted] at hi
|
||||
exact hi p hp
|
||||
· obtain ⟨p, hp, rfl⟩ := hp
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
simp
|
||||
rw [mem_uncontracted_iff_not_contracted] at ha
|
||||
exact ha p hp
|
||||
|
||||
lemma exists_mem_left_uncontracted_of_mem_join_uncontracted {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin φs.length)
|
||||
(ha : i ∈ (join φsΛ φsucΛ).uncontracted) :
|
||||
i ∈ φsΛ.uncontracted := by
|
||||
rw [@mem_uncontracted_iff_not_contracted]
|
||||
rw [@mem_uncontracted_iff_not_contracted] at ha
|
||||
simp [join] at ha
|
||||
intro p hp
|
||||
have hp' := ha p
|
||||
simp_all
|
||||
|
||||
lemma exists_mem_right_uncontracted_of_mem_join_uncontracted {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin φs.length)
|
||||
(hi : i ∈ (join φsΛ φsucΛ).uncontracted) :
|
||||
∃ a, uncontractedListEmd a = i ∧ a ∈ φsucΛ.uncontracted := by
|
||||
have hi' := exists_mem_left_uncontracted_of_mem_join_uncontracted _ _ i hi
|
||||
obtain ⟨j, rfl⟩ := uncontractedListEmd_surjective_mem_uncontracted i hi'
|
||||
use j
|
||||
simp
|
||||
rw [mem_uncontracted_iff_not_contracted] at hi
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
intro p hp
|
||||
have hip := hi (p.map uncontractedListEmd) (by
|
||||
simp [join]
|
||||
right
|
||||
use p
|
||||
simp [hp]
|
||||
rw [Finset.mapEmbedding_apply])
|
||||
simpa using hip
|
||||
|
||||
lemma join_uncontractedList {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).uncontractedList = List.map uncontractedListEmd φsucΛ.uncontractedList := by
|
||||
rw [uncontractedList_eq_sort]
|
||||
rw [uncontractedList_eq_sort]
|
||||
rw [fin_finset_sort_map_monotone]
|
||||
congr
|
||||
ext a
|
||||
simp
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
obtain ⟨a, rfl, ha⟩ := exists_mem_right_uncontracted_of_mem_join_uncontracted _ _ a h
|
||||
use a, ha
|
||||
· intro h
|
||||
obtain ⟨a, ha, rfl⟩ := h
|
||||
exact mem_join_uncontracted_of_mem_right_uncontracted φsΛ φsucΛ a ha
|
||||
· intro a b h
|
||||
exact uncontractedListEmd_strictMono h
|
||||
|
||||
lemma join_uncontractedList_get {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
((join φsΛ φsucΛ).uncontractedList).get =
|
||||
φsΛ.uncontractedListEmd ∘ (φsucΛ.uncontractedList).get ∘ (
|
||||
Fin.cast (by rw [join_uncontractedList]; simp) ):= by
|
||||
have h1 {n : ℕ} (l1 l2 : List (Fin n)) (h : l1 = l2) : l1.get = l2.get ∘ Fin.cast (by rw [h]):= by
|
||||
subst h
|
||||
rfl
|
||||
have hl := h1 _ _ (join_uncontractedList φsΛ φsucΛ)
|
||||
conv_lhs => rw [h1 _ _ (join_uncontractedList φsΛ φsucΛ)]
|
||||
ext i
|
||||
simp
|
||||
|
||||
lemma join_uncontractedListGet {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).uncontractedListGet = φsucΛ.uncontractedListGet := by
|
||||
simp [uncontractedListGet, join_uncontractedList]
|
||||
intro a ha
|
||||
simp [uncontractedListEmd, uncontractedIndexEquiv]
|
||||
rfl
|
||||
|
||||
lemma join_uncontractedListEmb {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).uncontractedListEmd =
|
||||
((finCongr (congrArg List.length (join_uncontractedListGet _ _))).toEmbedding.trans φsucΛ.uncontractedListEmd).trans φsΛ.uncontractedListEmd := by
|
||||
refine Function.Embedding.ext_iff.mpr (congrFun ?_)
|
||||
change uncontractedListEmd.toFun = _
|
||||
rw [uncontractedListEmd_toFun_eq_get]
|
||||
rw [join_uncontractedList_get]
|
||||
rfl
|
||||
|
||||
lemma join_assoc {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (φsucΛ' : WickContraction [φsΛ.join φsucΛ]ᵘᶜ.length) :
|
||||
join (join φsΛ φsucΛ) (φsucΛ') = join φsΛ (join φsucΛ (congr
|
||||
(congrArg List.length (join_uncontractedListGet _ _)) φsucΛ')) := by
|
||||
apply Subtype.ext
|
||||
ext a
|
||||
by_cases ha : a ∈ φsΛ.1
|
||||
· simp [ha, join]
|
||||
simp [ha, join]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
rcases h with h | h
|
||||
· obtain ⟨a, ha', rfl⟩ := h
|
||||
use a
|
||||
simp [ha']
|
||||
· obtain ⟨a, ha', rfl⟩ := h
|
||||
let a' := congrLift (congrArg List.length (join_uncontractedListGet _ _)) ⟨a, ha'⟩
|
||||
let a'' := joinLiftRight a'
|
||||
use a''
|
||||
apply And.intro
|
||||
· right
|
||||
use a'
|
||||
apply And.intro
|
||||
· exact a'.2
|
||||
· simp only [joinLiftRight, a'']
|
||||
rfl
|
||||
· simp only [a'']
|
||||
rw [Finset.mapEmbedding_apply, Finset.mapEmbedding_apply]
|
||||
simp only [a', joinLiftRight, congrLift]
|
||||
rw [join_uncontractedListEmb]
|
||||
simp [Finset.map_map]
|
||||
· intro h
|
||||
obtain ⟨a, ha', rfl⟩ := h
|
||||
rcases ha' with ha' | ha'
|
||||
· left
|
||||
use a
|
||||
· obtain ⟨a, ha', rfl⟩ := ha'
|
||||
right
|
||||
let a' := congrLiftInv _ ⟨a, ha'⟩
|
||||
use a'
|
||||
simp
|
||||
simp only [a']
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
rw [join_uncontractedListEmb]
|
||||
simp only [congrLiftInv, ← Finset.map_map]
|
||||
congr
|
||||
rw [Finset.map_map]
|
||||
change Finset.map (Equiv.refl _).toEmbedding a = _
|
||||
simp only [Equiv.refl_toEmbedding, Finset.map_refl]
|
||||
|
||||
@[simp]
|
||||
lemma join_getDual?_apply_uncontractedListEmb_eq_none_iff {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).getDual? (uncontractedListEmd i) = none
|
||||
↔ φsucΛ.getDual? i = none := by
|
||||
rw [getDual?_eq_none_iff_mem_uncontracted, getDual?_eq_none_iff_mem_uncontracted]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
obtain ⟨a, ha', ha⟩ := exists_mem_right_uncontracted_of_mem_join_uncontracted _ _ (uncontractedListEmd i) h
|
||||
simp at ha'
|
||||
subst ha'
|
||||
exact ha
|
||||
· intro h
|
||||
exact mem_join_uncontracted_of_mem_right_uncontracted φsΛ φsucΛ i h
|
||||
|
||||
@[simp]
|
||||
lemma join_getDual?_apply_uncontractedListEmb_isSome_iff {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin [φsΛ]ᵘᶜ.length) :
|
||||
((join φsΛ φsucΛ).getDual? (uncontractedListEmd i)).isSome
|
||||
↔ (φsucΛ.getDual? i).isSome := by
|
||||
rw [← Decidable.not_iff_not]
|
||||
simp
|
||||
|
||||
lemma join_getDual?_apply_uncontractedListEmb_some {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin [φsΛ]ᵘᶜ.length)
|
||||
(hi :((join φsΛ φsucΛ).getDual? (uncontractedListEmd i)).isSome) :
|
||||
((join φsΛ φsucΛ).getDual? (uncontractedListEmd i)) =
|
||||
some (uncontractedListEmd ((φsucΛ.getDual? i).get (by simpa using hi))) := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
simp [join]
|
||||
right
|
||||
use {i, (φsucΛ.getDual? i).get (by simpa using hi)}
|
||||
simp
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma join_getDual?_apply_uncontractedListEmb {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin [φsΛ]ᵘᶜ.length) :
|
||||
((join φsΛ φsucΛ).getDual? (uncontractedListEmd i)) = Option.map uncontractedListEmd (φsucΛ.getDual? i) := by
|
||||
by_cases h : (φsucΛ.getDual? i).isSome
|
||||
· rw [join_getDual?_apply_uncontractedListEmb_some]
|
||||
have h1 : (φsucΛ.getDual? i) =(φsucΛ.getDual? i).get (by simpa using h) :=
|
||||
Eq.symm (Option.some_get h)
|
||||
conv_rhs => rw [h1]
|
||||
simp [- Option.some_get]
|
||||
exact (join_getDual?_apply_uncontractedListEmb_isSome_iff φsΛ φsucΛ i).mpr h
|
||||
· simp only [Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none] at h
|
||||
rw [h]
|
||||
simp
|
||||
exact h
|
||||
|
||||
/-!
|
||||
|
||||
## Subcontractions and quotient contractions
|
||||
|
||||
-/
|
||||
|
||||
section
|
||||
|
||||
variable {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
|
||||
lemma join_sub_quot (S : Finset (Finset (Fin φs.length))) (ha : S ⊆ φsΛ.1) :
|
||||
join (subContraction S ha) (quotContraction S ha) = φsΛ := by
|
||||
apply Subtype.ext
|
||||
ext a
|
||||
simp [join]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
rcases h with h | h
|
||||
· exact mem_of_mem_subContraction h
|
||||
· obtain ⟨a, ha, rfl⟩ := h
|
||||
apply mem_of_mem_quotContraction ha
|
||||
· intro h
|
||||
have h1 := mem_subContraction_or_quotContraction (S := S) (a := a) (hs := ha) h
|
||||
rcases h1 with h1 | h1
|
||||
· simp [h1]
|
||||
· right
|
||||
obtain ⟨a, rfl, ha⟩ := h1
|
||||
use a
|
||||
simp [ha]
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
|
||||
lemma subContraction_card_plus_quotContraction_card_eq (S : Finset (Finset (Fin φs.length)))
|
||||
(ha : S ⊆ φsΛ.1) :
|
||||
(subContraction S ha).1.card + (quotContraction S ha).1.card = φsΛ.1.card := by
|
||||
rw [← join_card]
|
||||
simp [join_sub_quot]
|
||||
|
||||
end
|
||||
open FieldStatistic
|
||||
|
||||
lemma stat_signFinset_right {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i j : Fin [φsΛ]ᵘᶜ.length) :
|
||||
(𝓕 |>ₛ ⟨[φsΛ]ᵘᶜ.get, φsucΛ.signFinset i j⟩) =
|
||||
(𝓕 |>ₛ ⟨φs.get, (φsucΛ.signFinset i j).map uncontractedListEmd⟩) := by
|
||||
simp [ofFinset]
|
||||
congr 1
|
||||
rw [← fin_finset_sort_map_monotone]
|
||||
simp
|
||||
intro i j h
|
||||
exact uncontractedListEmd_strictMono h
|
||||
|
||||
lemma signFinset_right_map_uncontractedListEmd_eq_filter {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length) (φsucΛ : WickContraction [φsΛ]ᵘᶜ.length)
|
||||
(i j : Fin [φsΛ]ᵘᶜ.length) : (φsucΛ.signFinset i j).map uncontractedListEmd =
|
||||
((join φsΛ φsucΛ).signFinset (uncontractedListEmd i) (uncontractedListEmd j)).filter
|
||||
(fun c => c ∈ φsΛ.uncontracted) := by
|
||||
ext a
|
||||
simp
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
obtain ⟨a, ha, rfl⟩ := h
|
||||
apply And.intro
|
||||
· simp_all [signFinset]
|
||||
apply And.intro
|
||||
· exact uncontractedListEmd_strictMono ha.1
|
||||
· apply And.intro
|
||||
· exact uncontractedListEmd_strictMono ha.2.1
|
||||
· have ha2 := ha.2.2
|
||||
simp_all
|
||||
rcases ha2 with ha2 | ha2
|
||||
· simp [ha2]
|
||||
· right
|
||||
intro h
|
||||
have h1 : (φsucΛ.getDual? a) = some ((φsucΛ.getDual? a).get h) :=
|
||||
Eq.symm (Option.some_get h)
|
||||
apply lt_of_lt_of_eq (uncontractedListEmd_strictMono (ha2 h))
|
||||
rw [Option.get_map]
|
||||
· exact uncontractedListEmd_mem_uncontracted a
|
||||
· intro h
|
||||
have h2 := h.2
|
||||
have h2' := uncontractedListEmd_surjective_mem_uncontracted a h.2
|
||||
obtain ⟨a, rfl⟩ := h2'
|
||||
use a
|
||||
have h1 := h.1
|
||||
simp_all [signFinset]
|
||||
apply And.intro
|
||||
· have h1 := h.1
|
||||
rw [StrictMono.lt_iff_lt] at h1
|
||||
exact h1
|
||||
exact fun _ _ h => uncontractedListEmd_strictMono h
|
||||
· apply And.intro
|
||||
· have h1 := h.2.1
|
||||
rw [StrictMono.lt_iff_lt] at h1
|
||||
exact h1
|
||||
exact fun _ _ h => uncontractedListEmd_strictMono h
|
||||
· have h1 := h.2.2
|
||||
simp_all
|
||||
rcases h1 with h1 | h1
|
||||
· simp [h1]
|
||||
· right
|
||||
intro h
|
||||
have h1' := h1 h
|
||||
have hl : uncontractedListEmd i < uncontractedListEmd ((φsucΛ.getDual? a).get h) := by
|
||||
apply lt_of_lt_of_eq h1'
|
||||
simp [Option.get_map]
|
||||
rw [StrictMono.lt_iff_lt] at hl
|
||||
exact hl
|
||||
exact fun _ _ h => uncontractedListEmd_strictMono h
|
||||
|
||||
lemma sign_right_eq_prod_mul_prod {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
φsucΛ.sign = (∏ a, 𝓢(𝓕|>ₛ [φsΛ]ᵘᶜ[φsucΛ.sndFieldOfContract a], 𝓕|>ₛ ⟨φs.get ,
|
||||
((join φsΛ φsucΛ).signFinset (uncontractedListEmd (φsucΛ.fstFieldOfContract a)) (uncontractedListEmd (φsucΛ.sndFieldOfContract a))).filter
|
||||
(fun c => ¬ c ∈ φsΛ.uncontracted)⟩)) *
|
||||
(∏ a, 𝓢(𝓕|>ₛ [φsΛ]ᵘᶜ[φsucΛ.sndFieldOfContract a], 𝓕|>ₛ ⟨φs.get ,
|
||||
((join φsΛ φsucΛ).signFinset (uncontractedListEmd (φsucΛ.fstFieldOfContract a))
|
||||
(uncontractedListEmd (φsucΛ.sndFieldOfContract a)))⟩)) := by
|
||||
rw [← Finset.prod_mul_distrib, sign]
|
||||
congr
|
||||
funext a
|
||||
rw [← map_mul]
|
||||
congr
|
||||
rw [stat_signFinset_right, signFinset_right_map_uncontractedListEmd_eq_filter]
|
||||
rw [ofFinset_filter]
|
||||
|
||||
lemma join_singleton_signFinset_eq_filter {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(join (singleton h) φsucΛ).signFinset i j =
|
||||
((singleton h).signFinset i j).filter (fun c => ¬
|
||||
(((join (singleton h) φsucΛ).getDual? c).isSome ∧
|
||||
((h1 : ((join (singleton h) φsucΛ).getDual? c).isSome) →
|
||||
(((join (singleton h) φsucΛ).getDual? c).get h1) < i))) := by
|
||||
ext a
|
||||
simp [signFinset, and_assoc]
|
||||
intro h1 h2
|
||||
have h1 : (singleton h).getDual? a = none := by
|
||||
rw [singleton_getDual?_eq_none_iff_neq]
|
||||
omega
|
||||
simp [h1]
|
||||
apply Iff.intro
|
||||
· intro h1 h2
|
||||
rcases h1 with h1 | h1
|
||||
· simp [h1]
|
||||
have h2' : ¬ (((singleton h).join φsucΛ).getDual? a).isSome := by
|
||||
exact Option.not_isSome_iff_eq_none.mpr h1
|
||||
exact h2' h2
|
||||
use h2
|
||||
have h1 := h1 h2
|
||||
omega
|
||||
· intro h2
|
||||
by_cases h2' : (((singleton h).join φsucΛ).getDual? a).isSome = true
|
||||
· have h2 := h2 h2'
|
||||
obtain ⟨hb, h2⟩ := h2
|
||||
right
|
||||
intro hl
|
||||
apply lt_of_le_of_ne h2
|
||||
by_contra hn
|
||||
have hij : ((singleton h).join φsucΛ).getDual? i = j := by
|
||||
rw [@getDual?_eq_some_iff_mem]
|
||||
simp [join, singleton]
|
||||
simp [hn] at hij
|
||||
omega
|
||||
· simp at h2'
|
||||
simp [h2']
|
||||
|
||||
|
||||
lemma join_singleton_left_signFinset_eq_filter {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(𝓕 |>ₛ ⟨φs.get, (singleton h).signFinset i j⟩)
|
||||
= (𝓕 |>ₛ ⟨φs.get, (join (singleton h) φsucΛ).signFinset i j⟩) *
|
||||
(𝓕 |>ₛ ⟨φs.get, ((singleton h).signFinset i j).filter (fun c =>
|
||||
(((join (singleton h) φsucΛ).getDual? c).isSome ∧
|
||||
((h1 : ((join (singleton h) φsucΛ).getDual? c).isSome) →
|
||||
(((join (singleton h) φsucΛ).getDual? c).get h1) < i)))⟩) := by
|
||||
conv_rhs =>
|
||||
left
|
||||
rw [join_singleton_signFinset_eq_filter]
|
||||
rw [mul_comm]
|
||||
exact (ofFinset_filter_mul_neg 𝓕.statesStatistic _ _ _).symm
|
||||
|
||||
def joinSignRightExtra {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) : ℂ :=
|
||||
∏ a, 𝓢(𝓕|>ₛ [singleton h]ᵘᶜ[φsucΛ.sndFieldOfContract a], 𝓕|>ₛ ⟨φs.get ,
|
||||
((join (singleton h) φsucΛ).signFinset (uncontractedListEmd (φsucΛ.fstFieldOfContract a))
|
||||
(uncontractedListEmd (φsucΛ.sndFieldOfContract a))).filter
|
||||
(fun c => ¬ c ∈ (singleton h).uncontracted)⟩)
|
||||
|
||||
|
||||
def joinSignLeftExtra {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) : ℂ :=
|
||||
𝓢(𝓕 |>ₛ φs[j], (𝓕 |>ₛ ⟨φs.get, ((singleton h).signFinset i j).filter (fun c =>
|
||||
(((join (singleton h) φsucΛ).getDual? c).isSome ∧
|
||||
((h1 : ((join (singleton h) φsucΛ).getDual? c).isSome) →
|
||||
(((join (singleton h) φsucΛ).getDual? c).get h1) < i)))⟩))
|
||||
|
||||
lemma join_singleton_sign_left {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(singleton h).sign = 𝓢(𝓕 |>ₛ φs[j], (𝓕 |>ₛ ⟨φs.get, (join (singleton h) φsucΛ).signFinset i j⟩)) *
|
||||
(joinSignLeftExtra h φsucΛ) := by
|
||||
rw [singleton_sign_expand]
|
||||
rw [join_singleton_left_signFinset_eq_filter h φsucΛ]
|
||||
rw [map_mul]
|
||||
rfl
|
||||
|
||||
|
||||
lemma join_singleton_sign_right {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
φsucΛ.sign =
|
||||
(joinSignRightExtra h φsucΛ) *
|
||||
(∏ a, 𝓢(𝓕|>ₛ [singleton h]ᵘᶜ[φsucΛ.sndFieldOfContract a], 𝓕|>ₛ ⟨φs.get ,
|
||||
((join (singleton h) φsucΛ).signFinset (uncontractedListEmd (φsucΛ.fstFieldOfContract a))
|
||||
(uncontractedListEmd (φsucΛ.sndFieldOfContract a)))⟩)) := by
|
||||
rw [sign_right_eq_prod_mul_prod]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma join_singleton_getDual?_left {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(join (singleton h) φsucΛ).getDual? i = some j := by
|
||||
rw [@getDual?_eq_some_iff_mem]
|
||||
simp [singleton, join]
|
||||
|
||||
@[simp]
|
||||
lemma join_singleton_getDual?_right {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(join (singleton h) φsucΛ).getDual? j = some i := by
|
||||
rw [@getDual?_eq_some_iff_mem]
|
||||
simp [singleton, join]
|
||||
left
|
||||
exact Finset.pair_comm j i
|
||||
|
||||
|
||||
lemma joinSignRightExtra_eq_i_j_finset_eq_if {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
joinSignRightExtra h φsucΛ = ∏ a,
|
||||
𝓢((𝓕|>ₛ [singleton h]ᵘᶜ[φsucΛ.sndFieldOfContract a]),
|
||||
𝓕 |>ₛ ⟨φs.get, (if uncontractedListEmd (φsucΛ.fstFieldOfContract a) < j ∧
|
||||
j < uncontractedListEmd (φsucΛ.sndFieldOfContract a) ∧
|
||||
uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i then {j} else ∅)
|
||||
∪ (if uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i ∧
|
||||
i < uncontractedListEmd (φsucΛ.sndFieldOfContract a) then {i} else ∅)⟩) := by
|
||||
rw [joinSignRightExtra]
|
||||
congr
|
||||
funext a
|
||||
congr
|
||||
rw [signFinset]
|
||||
rw [Finset.filter_comm]
|
||||
have h11 : (Finset.filter (fun c => c ∉ (singleton h).uncontracted) Finset.univ) = {i, j} := by
|
||||
ext x
|
||||
simp
|
||||
rw [@mem_uncontracted_iff_not_contracted]
|
||||
simp [singleton]
|
||||
omega
|
||||
rw [h11]
|
||||
ext x
|
||||
simp
|
||||
have hjneqfst := singleton_uncontractedEmd_neq_right h (φsucΛ.fstFieldOfContract a)
|
||||
have hjneqsnd := singleton_uncontractedEmd_neq_right h (φsucΛ.sndFieldOfContract a)
|
||||
have hineqfst := singleton_uncontractedEmd_neq_left h (φsucΛ.fstFieldOfContract a)
|
||||
have hineqsnd := singleton_uncontractedEmd_neq_left h (φsucΛ.sndFieldOfContract a)
|
||||
by_cases hj1 : ¬ uncontractedListEmd (φsucΛ.fstFieldOfContract a) < j
|
||||
· simp [hj1]
|
||||
have hi1 : ¬ uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i := by omega
|
||||
simp [hi1]
|
||||
intro hxij h1 h2
|
||||
omega
|
||||
· have hj1 : uncontractedListEmd (φsucΛ.fstFieldOfContract a) < j := by
|
||||
omega
|
||||
by_cases hi1 : ¬ i < uncontractedListEmd (φsucΛ.sndFieldOfContract a)
|
||||
· simp [hi1]
|
||||
have hj2 : ¬ j < uncontractedListEmd (φsucΛ.sndFieldOfContract a) := by omega
|
||||
simp [hj2]
|
||||
intro hxij h1 h2
|
||||
omega
|
||||
· have hi1 : i < uncontractedListEmd (φsucΛ.sndFieldOfContract a) := by
|
||||
omega
|
||||
simp [hi1, hj1]
|
||||
by_cases hi2 : ¬ uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i
|
||||
· simp [hi2]
|
||||
by_cases hj3 : ¬ j < uncontractedListEmd (φsucΛ.sndFieldOfContract a)
|
||||
· omega
|
||||
· have hj4 : j < uncontractedListEmd (φsucΛ.sndFieldOfContract a) := by omega
|
||||
intro h
|
||||
rcases h with h | h
|
||||
· subst h
|
||||
omega
|
||||
· subst h
|
||||
simp
|
||||
omega
|
||||
· have hi2 : uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i := by omega
|
||||
simp [hi2]
|
||||
by_cases hj3 : ¬ j < uncontractedListEmd (φsucΛ.sndFieldOfContract a)
|
||||
· simp [hj3]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
omega
|
||||
· intro h
|
||||
subst h
|
||||
simp
|
||||
omega
|
||||
· have hj3 : j < uncontractedListEmd (φsucΛ.sndFieldOfContract a) := by omega
|
||||
simp [hj3]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
omega
|
||||
· intro h
|
||||
rcases h with h1 | h1
|
||||
· subst h1
|
||||
simp
|
||||
omega
|
||||
· subst h1
|
||||
simp
|
||||
omega
|
||||
|
||||
|
||||
lemma joinSignLeftExtra_eq_joinSignRightExtra {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j) (hs : (𝓕 |>ₛ φs[i]) = (𝓕 |>ₛ φs[j]))
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
joinSignLeftExtra h φsucΛ = joinSignRightExtra h φsucΛ := by
|
||||
/- Simplifying joinSignLeftExtra. -/
|
||||
rw [joinSignLeftExtra]
|
||||
rw [ofFinset_eq_prod]
|
||||
rw [map_prod]
|
||||
let e2 : Fin φs.length ≃ {x // (((singleton h).join φsucΛ).getDual? x).isSome} ⊕ {x // ¬ (((singleton h).join φsucΛ).getDual? x).isSome} := by
|
||||
exact (Equiv.sumCompl fun a => (((singleton h).join φsucΛ).getDual? a).isSome = true).symm
|
||||
rw [← e2.symm.prod_comp]
|
||||
simp only [Fin.getElem_fin, Fintype.prod_sum_type, instCommGroup]
|
||||
conv_lhs =>
|
||||
enter [2, 2, x]
|
||||
simp only [Equiv.symm_symm, Equiv.sumCompl_apply_inl, Equiv.sumCompl_apply_inr, e2]
|
||||
rw [if_neg (
|
||||
by
|
||||
simp
|
||||
intro h1 h2
|
||||
have hx := x.2
|
||||
simp_all)]
|
||||
simp
|
||||
rw [← ((singleton h).join φsucΛ).sigmaContractedEquiv.prod_comp]
|
||||
erw [Finset.prod_sigma]
|
||||
conv_lhs =>
|
||||
enter [2, a]
|
||||
rw [prod_finset_eq_mul_fst_snd]
|
||||
simp [e2, sigmaContractedEquiv]
|
||||
rw [prod_join]
|
||||
rw [singleton_prod]
|
||||
simp only [join_fstFieldOfContract_joinLiftLeft, singleton_fstFieldOfContract,
|
||||
join_sndFieldOfContract_joinLift, singleton_sndFieldOfContract, lt_self_iff_false, and_false,
|
||||
↓reduceIte, map_one, mul_one, join_fstFieldOfContract_joinLiftRight,
|
||||
join_sndFieldOfContract_joinLiftRight, getElem_uncontractedListEmd]
|
||||
rw [if_neg (by omega)]
|
||||
simp only [map_one, one_mul]
|
||||
/- Introducing joinSignRightExtra. -/
|
||||
rw [joinSignRightExtra_eq_i_j_finset_eq_if]
|
||||
congr
|
||||
funext a
|
||||
have hjneqfst := singleton_uncontractedEmd_neq_right h (φsucΛ.fstFieldOfContract a)
|
||||
have hjneqsnd := singleton_uncontractedEmd_neq_right h (φsucΛ.sndFieldOfContract a)
|
||||
have hineqfst := singleton_uncontractedEmd_neq_left h (φsucΛ.fstFieldOfContract a)
|
||||
have hineqsnd := singleton_uncontractedEmd_neq_left h (φsucΛ.sndFieldOfContract a)
|
||||
have hl : uncontractedListEmd (φsucΛ.fstFieldOfContract a) < uncontractedListEmd (φsucΛ.sndFieldOfContract a) := by
|
||||
apply uncontractedListEmd_strictMono
|
||||
exact fstFieldOfContract_lt_sndFieldOfContract φsucΛ a
|
||||
by_cases hj1 : ¬ uncontractedListEmd (φsucΛ.fstFieldOfContract a) < j
|
||||
· have hi1 : ¬ uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i := by omega
|
||||
simp [hj1, hi1]
|
||||
· have hj1 : uncontractedListEmd (φsucΛ.fstFieldOfContract a) < j := by omega
|
||||
simp [hj1]
|
||||
by_cases hi2 : ¬ i < uncontractedListEmd (φsucΛ.sndFieldOfContract a)
|
||||
· have hi1 : ¬ i < uncontractedListEmd (φsucΛ.fstFieldOfContract a) := by omega
|
||||
have hj2 : ¬ j < uncontractedListEmd (φsucΛ.sndFieldOfContract a) := by omega
|
||||
simp [hi2, hj2, hi1]
|
||||
· have hi2 : i < uncontractedListEmd (φsucΛ.sndFieldOfContract a) := by omega
|
||||
have hi2n : ¬ uncontractedListEmd (φsucΛ.sndFieldOfContract a) < i := by omega
|
||||
simp [hi2, hi2n]
|
||||
by_cases hj2 : ¬ j < uncontractedListEmd (φsucΛ.sndFieldOfContract a)
|
||||
· simp [hj2]
|
||||
have hj2 : uncontractedListEmd (φsucΛ.sndFieldOfContract a) < j:= by omega
|
||||
simp [hj2]
|
||||
by_cases hi1 : ¬ uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i
|
||||
· simp [hi1]
|
||||
· have hi1 : uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i := by omega
|
||||
simp [hi1, ofFinset_singleton]
|
||||
erw [hs]
|
||||
exact exchangeSign_symm (𝓕|>ₛφs[↑j]) (𝓕|>ₛ[singleton h]ᵘᶜ[↑(φsucΛ.sndFieldOfContract a)])
|
||||
· simp at hj2
|
||||
simp [hj2]
|
||||
by_cases hi1 : ¬ uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i
|
||||
· simp [hi1]
|
||||
· have hi1 : uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i := by omega
|
||||
simp [hi1]
|
||||
have hj2 : ¬ uncontractedListEmd (φsucΛ.sndFieldOfContract a) < j := by omega
|
||||
simp [hj2]
|
||||
rw [← ofFinset_union_disjoint]
|
||||
simp only [instCommGroup, ofFinset_singleton, List.get_eq_getElem, hs]
|
||||
erw [hs]
|
||||
simp
|
||||
simp
|
||||
omega
|
||||
|
||||
lemma join_sign_singleton {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j) (hs : (𝓕 |>ₛ φs[i]) = (𝓕 |>ₛ φs[j]))
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(join (singleton h) φsucΛ).sign = (singleton h).sign * φsucΛ.sign := by
|
||||
rw [join_singleton_sign_right]
|
||||
rw [join_singleton_sign_left h φsucΛ]
|
||||
rw [joinSignLeftExtra_eq_joinSignRightExtra h hs φsucΛ]
|
||||
rw [← mul_assoc]
|
||||
rw [mul_assoc _ _ (joinSignRightExtra h φsucΛ)]
|
||||
have h1 : (joinSignRightExtra h φsucΛ * joinSignRightExtra h φsucΛ) = 1 := by
|
||||
rw [← joinSignLeftExtra_eq_joinSignRightExtra h hs φsucΛ]
|
||||
simp [joinSignLeftExtra]
|
||||
simp only [instCommGroup, Fin.getElem_fin, h1, mul_one]
|
||||
rw [sign]
|
||||
rw [prod_join]
|
||||
congr
|
||||
· rw [singleton_prod]
|
||||
simp
|
||||
· funext a
|
||||
simp
|
||||
|
||||
lemma exists_contraction_pair_of_card_ge_zero {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(h : 0 < φsΛ.1.card) :
|
||||
∃ a, a ∈ φsΛ.1 := by
|
||||
simpa using h
|
||||
|
||||
lemma exists_join_singleton_of_card_ge_zero {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(h : 0 < φsΛ.1.card) (hc : φsΛ.GradingCompliant) :
|
||||
∃ (i j : Fin φs.length) (h : i < j) (φsucΛ : WickContraction [singleton h]ᵘᶜ.length),
|
||||
φsΛ = join (singleton h) φsucΛ ∧ (𝓕 |>ₛ φs[i]) = (𝓕 |>ₛ φs[j])
|
||||
∧ φsucΛ.GradingCompliant ∧ φsucΛ.1.card + 1 = φsΛ.1.card := by
|
||||
obtain ⟨a, ha⟩ := exists_contraction_pair_of_card_ge_zero φsΛ h
|
||||
use φsΛ.fstFieldOfContract ⟨a, ha⟩
|
||||
use φsΛ.sndFieldOfContract ⟨a, ha⟩
|
||||
use φsΛ.fstFieldOfContract_lt_sndFieldOfContract ⟨a, ha⟩
|
||||
let φsucΛ :
|
||||
WickContraction [singleton (φsΛ.fstFieldOfContract_lt_sndFieldOfContract ⟨a, ha⟩)]ᵘᶜ.length :=
|
||||
congr (by simp [← subContraction_singleton_eq_singleton]) (φsΛ.quotContraction {a} (by simpa using ha))
|
||||
use φsucΛ
|
||||
simp [← subContraction_singleton_eq_singleton]
|
||||
apply And.intro
|
||||
· have h1 := join_congr (subContraction_singleton_eq_singleton _ ⟨a, ha⟩).symm (φsucΛ := φsucΛ)
|
||||
simp [h1, φsucΛ]
|
||||
rw [join_sub_quot]
|
||||
· apply And.intro (hc ⟨a, ha⟩)
|
||||
apply And.intro
|
||||
· simp [φsucΛ]
|
||||
rw [gradingCompliant_congr (φs' := [(φsΛ.subContraction {a} (by simpa using ha))]ᵘᶜ)]
|
||||
simp
|
||||
exact quotContraction_gradingCompliant hc
|
||||
rw [← subContraction_singleton_eq_singleton]
|
||||
· simp [φsucΛ]
|
||||
have h1 := subContraction_card_plus_quotContraction_card_eq _ {a} (by simpa using ha)
|
||||
simp [subContraction] at h1
|
||||
omega
|
||||
|
||||
lemma join_sign_induction {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (hc : φsΛ.GradingCompliant) :
|
||||
(n : ℕ) → (hn : φsΛ.1.card = n) →
|
||||
(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign
|
||||
| 0, hn => by
|
||||
rw [@card_zero_iff_empty] at hn
|
||||
subst hn
|
||||
simp [sign_empty]
|
||||
apply sign_congr
|
||||
simp
|
||||
| Nat.succ n, hn => by
|
||||
obtain ⟨i, j, hij, φsucΛ', rfl, h1, h2, h3⟩ := exists_join_singleton_of_card_ge_zero φsΛ (by simp [hn]) hc
|
||||
rw [join_assoc]
|
||||
rw [join_sign_singleton hij h1 ]
|
||||
rw [join_sign_singleton hij h1 ]
|
||||
have hn : φsucΛ'.1.card = n := by
|
||||
omega
|
||||
rw [join_sign_induction φsucΛ' (congr (by simp [join_uncontractedListGet]) φsucΛ) h2
|
||||
n hn]
|
||||
rw [mul_assoc]
|
||||
simp [sign_congr]
|
||||
left
|
||||
left
|
||||
apply sign_congr
|
||||
exact join_uncontractedListGet (singleton hij) φsucΛ'
|
||||
|
||||
lemma join_sign {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (hc : φsΛ.GradingCompliant) :
|
||||
(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign := by
|
||||
exact join_sign_induction φsΛ φsucΛ hc (φsΛ).1.card rfl
|
||||
|
||||
end WickContraction
|
|
@ -324,6 +324,16 @@ def sign (φs : List 𝓕.States) (φsΛ : WickContraction φs.length) : ℂ :=
|
|||
∏ (a : φsΛ.1), 𝓢(𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a],
|
||||
𝓕 |>ₛ ⟨φs.get, φsΛ.signFinset (φsΛ.fstFieldOfContract a) (φsΛ.sndFieldOfContract a)⟩)
|
||||
|
||||
lemma sign_empty (φs : List 𝓕.States) :
|
||||
sign φs empty = 1 := by
|
||||
rw [sign]
|
||||
simp [empty]
|
||||
|
||||
lemma sign_congr {φs φs' : List 𝓕.States} (h : φs = φs') (φsΛ : WickContraction φs.length) :
|
||||
sign φs' (congr (by simp [h]) φsΛ) = sign φs φsΛ := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Sign insert
|
||||
|
|
125
HepLean/PerturbationTheory/WickContraction/Singleton.lean
Normal file
125
HepLean/PerturbationTheory/WickContraction/Singleton.lean
Normal file
|
@ -0,0 +1,125 @@
|
|||
/-
|
||||
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.TimeContract
|
||||
import HepLean.PerturbationTheory.WickContraction.StaticContract
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.TimeContraction
|
||||
import HepLean.PerturbationTheory.WickContraction.SubContraction
|
||||
/-!
|
||||
|
||||
# Singleton of contractions
|
||||
|
||||
-/
|
||||
|
||||
open FieldSpecification
|
||||
variable {𝓕 : FieldSpecification}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
open FieldOpAlgebra
|
||||
open FieldStatistic
|
||||
|
||||
def singleton {i j : Fin n} (hij : i < j) : WickContraction n :=
|
||||
⟨{{i, j}}, by
|
||||
intro i hi
|
||||
simp at hi
|
||||
subst hi
|
||||
rw [@Finset.card_eq_two]
|
||||
use i, j
|
||||
simp
|
||||
omega
|
||||
, by
|
||||
intro i hi j hj
|
||||
simp_all⟩
|
||||
|
||||
lemma mem_singleton {i j : Fin n} (hij : i < j) :
|
||||
{i, j} ∈ (singleton hij).1 := by
|
||||
simp [singleton]
|
||||
|
||||
lemma mem_singleton_iff {i j : Fin n} (hij : i < j) {a : Finset (Fin n)} :
|
||||
a ∈ (singleton hij).1 ↔ a = {i, j} := by
|
||||
simp [singleton]
|
||||
|
||||
@[simp]
|
||||
lemma of_singleton_eq {i j : Fin n} (hij : i < j) (a : (singleton hij).1):
|
||||
a = ⟨{i, j}, mem_singleton hij⟩ := by
|
||||
have ha2 := a.2
|
||||
rw [@mem_singleton_iff] at ha2
|
||||
exact Subtype.coe_eq_of_eq_mk ha2
|
||||
|
||||
lemma singleton_prod {φs : List 𝓕.States} {i j : Fin φs.length} (hij : i < j)
|
||||
(f : (singleton hij).1 → M) [CommMonoid M] :
|
||||
∏ a, f a = f ⟨{i,j}, mem_singleton hij⟩:= by
|
||||
simp [singleton]
|
||||
|
||||
@[simp]
|
||||
lemma singleton_fstFieldOfContract {i j : Fin n} (hij : i < j) :
|
||||
(singleton hij).fstFieldOfContract ⟨{i, j}, mem_singleton hij⟩ = i := by
|
||||
refine eq_fstFieldOfContract_of_mem (singleton hij) ⟨{i, j}, mem_singleton hij⟩ i j ?_ ?_ ?_
|
||||
· simp
|
||||
· simp
|
||||
· exact hij
|
||||
|
||||
@[simp]
|
||||
lemma singleton_sndFieldOfContract {i j : Fin n} (hij : i < j) :
|
||||
(singleton hij).sndFieldOfContract ⟨{i, j}, mem_singleton hij⟩ = j := by
|
||||
refine eq_sndFieldOfContract_of_mem (singleton hij) ⟨{i, j}, mem_singleton hij⟩ i j ?_ ?_ ?_
|
||||
· simp
|
||||
· simp
|
||||
· exact hij
|
||||
|
||||
lemma singleton_sign_expand {φs : List 𝓕.States} {i j : Fin φs.length} (hij : i < j) :
|
||||
(singleton hij).sign = 𝓢(𝓕 |>ₛ φs[j], 𝓕 |>ₛ ⟨φs.get, (singleton hij).signFinset i j⟩) := by
|
||||
rw [sign, singleton_prod]
|
||||
simp
|
||||
|
||||
lemma singleton_getDual?_eq_none_iff_neq {i j : Fin n} (hij : i < j) (a : Fin n) :
|
||||
(singleton hij).getDual? a = none ↔ (i ≠ a ∧ j ≠ a) := by
|
||||
rw [getDual?_eq_none_iff_mem_uncontracted]
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
simp [singleton]
|
||||
omega
|
||||
|
||||
lemma singleton_uncontractedEmd_neq_left {φs : List 𝓕.States} {i j : Fin φs.length} (hij : i < j)
|
||||
(a : Fin [singleton hij]ᵘᶜ.length ) :
|
||||
(singleton hij).uncontractedListEmd a ≠ i := by
|
||||
by_contra hn
|
||||
have h1 : (singleton hij).uncontractedListEmd a ∈ (singleton hij).uncontracted := by
|
||||
simp [uncontractedListEmd]
|
||||
have h2 : i ∉ (singleton hij).uncontracted := by
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
simp [singleton]
|
||||
simp_all
|
||||
|
||||
lemma singleton_uncontractedEmd_neq_right {φs : List 𝓕.States} {i j : Fin φs.length} (hij : i < j)
|
||||
(a : Fin [singleton hij]ᵘᶜ.length ) :
|
||||
(singleton hij).uncontractedListEmd a ≠ j := by
|
||||
by_contra hn
|
||||
have h1 : (singleton hij).uncontractedListEmd a ∈ (singleton hij).uncontracted := by
|
||||
simp [uncontractedListEmd]
|
||||
have h2 : j ∉ (singleton hij).uncontracted := by
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
simp [singleton]
|
||||
simp_all
|
||||
|
||||
@[simp]
|
||||
lemma mem_signFinset {i j : Fin n} (hij : i < j) (a : Fin n) :
|
||||
a ∈ (singleton hij).signFinset i j ↔ i < a ∧ a < j := by
|
||||
simp [signFinset]
|
||||
intro h1 h2
|
||||
rw [@singleton_getDual?_eq_none_iff_neq]
|
||||
apply Or.inl
|
||||
omega
|
||||
|
||||
lemma subContraction_singleton_eq_singleton {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(a : φsΛ.1) : φsΛ.subContraction {a.1} (by simp) =
|
||||
singleton (φsΛ.fstFieldOfContract_lt_sndFieldOfContract a) := by
|
||||
apply Subtype.ext
|
||||
simp [subContraction, singleton]
|
||||
exact finset_eq_fstFieldOfContract_sndFieldOfContract φsΛ a
|
||||
|
||||
end WickContraction
|
137
HepLean/PerturbationTheory/WickContraction/SubContraction.lean
Normal file
137
HepLean/PerturbationTheory/WickContraction/SubContraction.lean
Normal file
|
@ -0,0 +1,137 @@
|
|||
/-
|
||||
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.TimeContract
|
||||
import HepLean.PerturbationTheory.WickContraction.StaticContract
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.TimeContraction
|
||||
/-!
|
||||
|
||||
# Sub contractions
|
||||
|
||||
-/
|
||||
|
||||
open FieldSpecification
|
||||
variable {𝓕 : FieldSpecification}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
open HepLean.List
|
||||
open FieldOpAlgebra
|
||||
|
||||
def subContraction (S : Finset (Finset (Fin φs.length))) (ha : S ⊆ φsΛ.1) : WickContraction φs.length :=
|
||||
⟨S, by
|
||||
intro i hi
|
||||
exact φsΛ.2.1 i (ha hi),
|
||||
by
|
||||
intro i hi j hj
|
||||
exact φsΛ.2.2 i (ha hi) j (ha hj)⟩
|
||||
|
||||
lemma mem_of_mem_subContraction {S : Finset (Finset (Fin φs.length))} {hs : S ⊆ φsΛ.1}
|
||||
{a : Finset (Fin φs.length)} (ha : a ∈ (φsΛ.subContraction S hs).1) : a ∈ φsΛ.1 := by
|
||||
exact hs ha
|
||||
|
||||
def quotContraction (S : Finset (Finset (Fin φs.length))) (ha : S ⊆ φsΛ.1) :
|
||||
WickContraction [φsΛ.subContraction S ha]ᵘᶜ.length :=
|
||||
⟨Finset.filter (fun a => Finset.map uncontractedListEmd a ∈ φsΛ.1) Finset.univ,
|
||||
by
|
||||
intro a ha'
|
||||
simp at ha'
|
||||
simpa using φsΛ.2.1 (Finset.map uncontractedListEmd a) ha'
|
||||
, by
|
||||
intro a ha b hb
|
||||
simp at ha hb
|
||||
by_cases hab : a = b
|
||||
· simp [hab]
|
||||
· simp [hab]
|
||||
have hx := φsΛ.2.2 (Finset.map uncontractedListEmd a) ha (Finset.map uncontractedListEmd b) hb
|
||||
simp_all⟩
|
||||
|
||||
lemma mem_of_mem_quotContraction {S : Finset (Finset (Fin φs.length))} {hs : S ⊆ φsΛ.1}
|
||||
{a : Finset (Fin [φsΛ.subContraction S hs]ᵘᶜ.length)}
|
||||
(ha : a ∈ (quotContraction S hs).1) : Finset.map uncontractedListEmd a ∈ φsΛ.1 := by
|
||||
simp [quotContraction] at ha
|
||||
exact ha
|
||||
|
||||
lemma mem_subContraction_or_quotContraction {S : Finset (Finset (Fin φs.length))} {hs : S ⊆ φsΛ.1}
|
||||
{a : Finset (Fin φs.length)} (ha : a ∈ φsΛ.1) :
|
||||
a ∈ (φsΛ.subContraction S hs).1 ∨
|
||||
∃ a', Finset.map uncontractedListEmd a' = a ∧ a' ∈ (quotContraction S hs).1 := by
|
||||
by_cases h1 : a ∈ (φsΛ.subContraction S hs).1
|
||||
· simp [h1]
|
||||
simp [h1]
|
||||
simp [subContraction] at h1
|
||||
have h2 := φsΛ.2.1 a ha
|
||||
rw [@Finset.card_eq_two] at h2
|
||||
obtain ⟨i, j, hij, rfl⟩ := h2
|
||||
have hi : i ∈ (φsΛ.subContraction S hs).uncontracted := by
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
intro p hp
|
||||
have hp' : p ∈ φsΛ.1 := hs hp
|
||||
have hp2 := φsΛ.2.2 p hp' {i, j} ha
|
||||
simp [subContraction] at hp
|
||||
rcases hp2 with hp2 | hp2
|
||||
· simp_all
|
||||
simp at hp2
|
||||
exact hp2.1
|
||||
have hj : j ∈ (φsΛ.subContraction S hs).uncontracted := by
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
intro p hp
|
||||
have hp' : p ∈ φsΛ.1 := hs hp
|
||||
have hp2 := φsΛ.2.2 p hp' {i, j} ha
|
||||
simp [subContraction] at hp
|
||||
rcases hp2 with hp2 | hp2
|
||||
· simp_all
|
||||
simp at hp2
|
||||
exact hp2.2
|
||||
obtain ⟨i, rfl⟩ := uncontractedListEmd_surjective_mem_uncontracted i hi
|
||||
obtain ⟨j, rfl⟩ := uncontractedListEmd_surjective_mem_uncontracted j hj
|
||||
use {i, j}
|
||||
simp [quotContraction]
|
||||
exact ha
|
||||
|
||||
@[simp]
|
||||
lemma subContraction_uncontractedList_get {S : Finset (Finset (Fin φs.length))} {hs : S ⊆ φsΛ.1}
|
||||
{a : Fin [subContraction S hs]ᵘᶜ.length} :
|
||||
[subContraction S hs]ᵘᶜ[a] = φs[uncontractedListEmd a] := by
|
||||
erw [← getElem_uncontractedListEmd]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma quotContraction_fstFieldOfContract_uncontractedListEmd {S : Finset (Finset (Fin φs.length))}
|
||||
{hs : S ⊆ φsΛ.1} (a : (quotContraction S hs).1) :
|
||||
uncontractedListEmd ((quotContraction S hs).fstFieldOfContract a) =
|
||||
(φsΛ.fstFieldOfContract ⟨Finset.map uncontractedListEmd a.1, mem_of_mem_quotContraction a.2⟩) := by
|
||||
symm
|
||||
apply eq_fstFieldOfContract_of_mem _ _ _ (uncontractedListEmd ((quotContraction S hs).sndFieldOfContract a) )
|
||||
· simp only [Finset.mem_map', fstFieldOfContract_mem]
|
||||
· simp
|
||||
· apply uncontractedListEmd_strictMono
|
||||
exact fstFieldOfContract_lt_sndFieldOfContract (quotContraction S hs) a
|
||||
|
||||
@[simp]
|
||||
lemma quotContraction_sndFieldOfContract_uncontractedListEmd {S : Finset (Finset (Fin φs.length))}
|
||||
{hs : S ⊆ φsΛ.1} (a : (quotContraction S hs).1) :
|
||||
uncontractedListEmd ((quotContraction S hs).sndFieldOfContract a) =
|
||||
(φsΛ.sndFieldOfContract ⟨Finset.map uncontractedListEmd a.1, mem_of_mem_quotContraction a.2⟩) := by
|
||||
symm
|
||||
apply eq_sndFieldOfContract_of_mem _ _ (uncontractedListEmd ((quotContraction S hs).fstFieldOfContract a) )
|
||||
· simp only [Finset.mem_map', fstFieldOfContract_mem]
|
||||
· simp
|
||||
· apply uncontractedListEmd_strictMono
|
||||
exact fstFieldOfContract_lt_sndFieldOfContract (quotContraction S hs) a
|
||||
|
||||
lemma quotContraction_gradingCompliant {S : Finset (Finset (Fin φs.length))} {hs : S ⊆ φsΛ.1}
|
||||
(hsΛ : φsΛ.GradingCompliant) :
|
||||
GradingCompliant [φsΛ.subContraction S hs]ᵘᶜ (quotContraction S hs) := by
|
||||
simp [GradingCompliant]
|
||||
intro a ha
|
||||
have h1' := mem_of_mem_quotContraction ha
|
||||
erw [subContraction_uncontractedList_get]
|
||||
erw [subContraction_uncontractedList_get]
|
||||
simp
|
||||
apply hsΛ
|
||||
|
||||
|
||||
end WickContraction
|
|
@ -24,6 +24,11 @@ lemma congr_uncontracted {n m : ℕ} (c : WickContraction n) (h : n = m) :
|
|||
subst h
|
||||
simp
|
||||
|
||||
|
||||
lemma getDual?_eq_none_iff_mem_uncontracted (i : Fin n) :
|
||||
c.getDual? i = none ↔ i ∈ c.uncontracted := by
|
||||
simp [uncontracted]
|
||||
|
||||
/-- The equivalence of `Option c.uncontracted` for two propositionally equal Wick contractions. -/
|
||||
def uncontractedCongr {c c': WickContraction n} (h : c = c') :
|
||||
Option c.uncontracted ≃ Option c'.uncontracted :=
|
||||
|
@ -64,4 +69,18 @@ lemma mem_uncontracted_iff_not_contracted (i : Fin n) :
|
|||
apply h {i, j} hj
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma mem_uncontracted_empty (i : Fin n) : i ∈ empty.uncontracted := by
|
||||
rw [@mem_uncontracted_iff_not_contracted]
|
||||
intro p hp
|
||||
simp [empty] at hp
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_empty_eq_none (i : Fin n) : empty.getDual? i = none := by
|
||||
simpa [uncontracted] using mem_uncontracted_empty i
|
||||
|
||||
@[simp]
|
||||
lemma uncontracted_empty {n : ℕ} : (@empty n).uncontracted = Finset.univ := by
|
||||
simp [ uncontracted]
|
||||
|
||||
end WickContraction
|
||||
|
|
|
@ -46,6 +46,26 @@ lemma fin_list_sorted_succAboveEmb_sorted (l: List (Fin n)) (hl : l.Sorted (·
|
|||
simp only [Fin.coe_succAboveEmb]
|
||||
exact Fin.strictMono_succAbove i
|
||||
|
||||
lemma fin_finset_sort_map_monotone {n m : ℕ} (a : Finset (Fin n)) (f : Fin n ↪ Fin m)
|
||||
(hf : StrictMono f) : (Finset.sort (· ≤ ·) a).map f =
|
||||
(Finset.sort (· ≤ ·) (a.map f)) := by
|
||||
have h1 : ((Finset.sort (· ≤ ·) a).map f).Sorted (· ≤ ·) := by
|
||||
apply fin_list_sorted_monotone_sorted
|
||||
exact Finset.sort_sorted (fun x1 x2 => x1 ≤ x2) a
|
||||
exact hf
|
||||
have h2 : ((Finset.sort (· ≤ ·) a).map f).Nodup := by
|
||||
refine (List.nodup_map_iff_inj_on ?_).mpr ?_
|
||||
exact Finset.sort_nodup (fun x1 x2 => x1 ≤ x2) a
|
||||
intro a ha b hb hf
|
||||
exact f.2 hf
|
||||
have h3 : ((Finset.sort (· ≤ ·) a).map f).toFinset = (a.map f) := by
|
||||
ext a
|
||||
simp
|
||||
rw [← h3]
|
||||
exact ((List.toFinset_sort (· ≤ ·) h2).mpr h1).symm
|
||||
|
||||
|
||||
|
||||
lemma fin_list_sorted_split :
|
||||
(l : List (Fin n)) → (hl : l.Sorted (· ≤ ·)) → (i : ℕ) →
|
||||
l = l.filter (fun x => x.1 < i) ++ l.filter (fun x => i ≤ x.1)
|
||||
|
@ -177,6 +197,10 @@ lemma uncontractedList_mem_iff (i : Fin n) :
|
|||
i ∈ c.uncontractedList ↔ i ∈ c.uncontracted := by
|
||||
simp [uncontractedList]
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedList_empty : (empty (n := n)).uncontractedList = List.finRange n := by
|
||||
simp [uncontractedList]
|
||||
|
||||
@[simp]
|
||||
lemma nil_zero_uncontractedList : (empty (n := 0)).uncontractedList = [] := by
|
||||
simp [empty, uncontractedList]
|
||||
|
@ -197,6 +221,12 @@ lemma uncontractedList_sorted : List.Sorted (· ≤ ·) c.uncontractedList := by
|
|||
rw [← List.ofFn_id]
|
||||
exact Monotone.ofFn_sorted fun ⦃a b⦄ a => a
|
||||
|
||||
lemma uncontractedList_sorted_lt : List.Sorted (· < ·) c.uncontractedList := by
|
||||
rw [uncontractedList]
|
||||
apply List.Sorted.filter
|
||||
rw [← List.ofFn_id]
|
||||
exact List.sorted_lt_ofFn_iff.mpr fun ⦃a b⦄ a => a
|
||||
|
||||
lemma uncontractedList_nodup : c.uncontractedList.Nodup := by
|
||||
rw [uncontractedList]
|
||||
refine List.Nodup.filter (fun x => decide (x ∈ c.uncontracted)) ?_
|
||||
|
@ -294,6 +324,12 @@ def uncontractedListGet {φs : List 𝓕.States} (φsΛ : WickContraction φs.le
|
|||
|
||||
@[inherit_doc uncontractedListGet]
|
||||
scoped[WickContraction] notation "[" φsΛ "]ᵘᶜ" => uncontractedListGet φsΛ
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedListGet_empty {φs : List 𝓕.States} :
|
||||
(empty (n := φs.length)).uncontractedListGet = φs := by
|
||||
simp [uncontractedListGet]
|
||||
|
||||
/-!
|
||||
|
||||
## uncontractedStatesEquiv
|
||||
|
@ -321,6 +357,70 @@ lemma uncontractedStatesEquiv_list_sum [AddCommMonoid α] (φs : List 𝓕.State
|
|||
|
||||
/-!
|
||||
|
||||
## uncontractedListEmd
|
||||
|
||||
-/
|
||||
|
||||
/-- The embedding of `Fin [φsΛ]ᵘᶜ.length` into `Fin φs.length`. -/
|
||||
def uncontractedListEmd {φs : List 𝓕.States} {φsΛ : WickContraction φs.length} :
|
||||
Fin [φsΛ]ᵘᶜ.length ↪ Fin φs.length :=
|
||||
((finCongr (by simp [uncontractedListGet])).trans φsΛ.uncontractedIndexEquiv).toEmbedding.trans
|
||||
(Function.Embedding.subtype fun x => x ∈ φsΛ.uncontracted)
|
||||
|
||||
lemma uncontractedListEmd_toFun_eq_get (φs : List 𝓕.States) (φsΛ : WickContraction φs.length) :
|
||||
(uncontractedListEmd (φsΛ := φsΛ)).toFun =
|
||||
φsΛ.uncontractedList.get ∘ (finCongr (by simp [uncontractedListGet])):= by
|
||||
rfl
|
||||
|
||||
lemma uncontractedListEmd_strictMono {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
{i j : Fin [φsΛ]ᵘᶜ.length} (h : i < j) : uncontractedListEmd i < uncontractedListEmd j := by
|
||||
simp [uncontractedListEmd, uncontractedIndexEquiv]
|
||||
exact List.Sorted.get_strictMono φsΛ.uncontractedList_sorted_lt h
|
||||
|
||||
lemma uncontractedListEmd_mem_uncontracted {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
(i : Fin [φsΛ]ᵘᶜ.length) : uncontractedListEmd i ∈ φsΛ.uncontracted := by
|
||||
simp [uncontractedListEmd]
|
||||
|
||||
lemma uncontractedListEmd_surjective_mem_uncontracted {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
(i : Fin φs.length) (hi : i ∈ φsΛ.uncontracted) :
|
||||
∃ j, φsΛ.uncontractedListEmd j = i := by
|
||||
simp [uncontractedListEmd]
|
||||
have hj : ∃ j, φsΛ.uncontractedIndexEquiv j = ⟨i, hi⟩ := by
|
||||
exact φsΛ.uncontractedIndexEquiv.surjective ⟨i, hi⟩
|
||||
obtain ⟨j, hj⟩ := hj
|
||||
have hj' : ∃ j', Fin.cast uncontractedListEmd.proof_1 j' = j := by
|
||||
exact (finCongr uncontractedListEmd.proof_1).surjective j
|
||||
obtain ⟨j', rfl⟩ := hj'
|
||||
use j'
|
||||
rw [hj]
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedListEmd_finset_disjoint_left {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
(a : Finset (Fin [φsΛ]ᵘᶜ.length)) (b : Finset (Fin φs.length)) (hb : b ∈ φsΛ.1) : Disjoint (a.map uncontractedListEmd) b := by
|
||||
rw [Finset.disjoint_left]
|
||||
intro x hx
|
||||
simp at hx
|
||||
obtain ⟨x, hx, rfl⟩ := hx
|
||||
have h1 : uncontractedListEmd x ∈ φsΛ.uncontracted :=
|
||||
uncontractedListEmd_mem_uncontracted x
|
||||
rw [mem_uncontracted_iff_not_contracted] at h1
|
||||
exact h1 b hb
|
||||
|
||||
@[simp]
|
||||
lemma getElem_uncontractedListEmd {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
(k : Fin [φsΛ]ᵘᶜ.length) : φs[(uncontractedListEmd k).1] = [φsΛ]ᵘᶜ[k.1] := by
|
||||
simp [uncontractedListGet]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedListEmd_empty {φs : List 𝓕.States} :
|
||||
(empty (n := φs.length)).uncontractedListEmd = (finCongr (by simp)).toEmbedding := by
|
||||
ext x
|
||||
simp [uncontractedListEmd, uncontractedIndexEquiv]
|
||||
|
||||
|
||||
/-!
|
||||
|
||||
## Uncontracted List for extractEquiv symm none
|
||||
|
||||
-/
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue