refactor: Lint
This commit is contained in:
parent
fca3f02eca
commit
c8e9c285a3
8 changed files with 229 additions and 194 deletions
|
@ -21,11 +21,11 @@ open EqTimeOnly
|
|||
|
||||
lemma timeOrder_ofFieldOpList_eqTimeOnly (φs : List 𝓕.States) :
|
||||
timeOrder (ofFieldOpList φs) = ∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs)}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)):= by
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
|
||||
rw [static_wick_theorem φs]
|
||||
let e2 : WickContraction φs.length ≃
|
||||
{φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} ⊕
|
||||
{φsΛ : WickContraction φs.length // ¬ φsΛ.EqTimeOnly} :=
|
||||
{φsΛ : WickContraction φs.length // ¬ φsΛ.EqTimeOnly} :=
|
||||
(Equiv.sumCompl _).symm
|
||||
rw [← e2.symm.sum_comp]
|
||||
simp only [Equiv.symm_symm, Algebra.smul_mul_assoc, Fintype.sum_sum_type,
|
||||
|
@ -43,22 +43,23 @@ lemma timeOrder_ofFieldOpList_eqTimeOnly (φs : List 𝓕.States) :
|
|||
exact x.2
|
||||
|
||||
lemma timeOrder_ofFieldOpList_eq_eqTimeOnly_empty (φs : List 𝓕.States) :
|
||||
timeOrder (ofFieldOpList φs) = 𝓣(𝓝(ofFieldOpList φs)) +
|
||||
timeOrder (ofFieldOpList φs) = 𝓣(𝓝(ofFieldOpList φs)) +
|
||||
∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
|
||||
let e1 : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} ≃
|
||||
{φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} // φsΛ.1 = empty} ⊕
|
||||
{φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} // ¬ φsΛ.1 = empty} :=
|
||||
(Equiv.sumCompl _).symm
|
||||
(Equiv.sumCompl _).symm
|
||||
rw [timeOrder_ofFieldOpList_eqTimeOnly, ← e1.symm.sum_comp]
|
||||
simp only [Equiv.symm_symm, Algebra.smul_mul_assoc, Fintype.sum_sum_type,
|
||||
Equiv.sumCompl_apply_inl, Equiv.sumCompl_apply_inr, ne_eq, e1]
|
||||
congr 1
|
||||
· let e2 : { φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} // φsΛ.1 = empty } ≃ Unit := {
|
||||
· let e2 : {φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} // φsΛ.1 = empty } ≃
|
||||
Unit := {
|
||||
toFun := fun x => (), invFun := fun x => ⟨⟨empty, by simp⟩, rfl⟩,
|
||||
left_inv a := by
|
||||
ext
|
||||
simp [a.2], right_inv a := by simp }
|
||||
simp [a.2], right_inv a := by simp}
|
||||
rw [← e2.symm.sum_comp]
|
||||
simp [e2, sign_empty]
|
||||
· let e2 : { φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} // ¬ φsΛ.1 = empty } ≃
|
||||
|
@ -69,22 +70,22 @@ lemma timeOrder_ofFieldOpList_eq_eqTimeOnly_empty (φs : List 𝓕.States) :
|
|||
rfl
|
||||
|
||||
lemma normalOrder_timeOrder_ofFieldOpList_eq_eqTimeOnly_empty (φs : List 𝓕.States) :
|
||||
𝓣(𝓝(ofFieldOpList φs)) = 𝓣(ofFieldOpList φs) -
|
||||
𝓣(𝓝(ofFieldOpList φs)) = 𝓣(ofFieldOpList φs) -
|
||||
∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
|
||||
rw [timeOrder_ofFieldOpList_eq_eqTimeOnly_empty]
|
||||
simp
|
||||
|
||||
lemma normalOrder_timeOrder_ofFieldOpList_eq_haveEqTime_sum_not_haveEqTime (φs : List 𝓕.States) :
|
||||
𝓣(𝓝(ofFieldOpList φs)) = (∑ (φsΛ : {φsΛ : WickContraction φs.length // ¬ HaveEqTime φsΛ}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))
|
||||
+ (∑ (φsΛ : {φsΛ : WickContraction φs.length // HaveEqTime φsΛ}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))
|
||||
- ∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
|
||||
𝓣(𝓝(ofFieldOpList φs)) = (∑ (φsΛ : {φsΛ : WickContraction φs.length // ¬ HaveEqTime φsΛ}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))
|
||||
+ (∑ (φsΛ : {φsΛ : WickContraction φs.length // HaveEqTime φsΛ}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))
|
||||
- ∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
|
||||
rw [normalOrder_timeOrder_ofFieldOpList_eq_eqTimeOnly_empty]
|
||||
rw [wicks_theorem]
|
||||
let e1 : WickContraction φs.length ≃ {φsΛ // HaveEqTime φsΛ} ⊕ {φsΛ // ¬ HaveEqTime φsΛ} := by
|
||||
let e1 : WickContraction φs.length ≃ {φsΛ // HaveEqTime φsΛ} ⊕ {φsΛ // ¬ HaveEqTime φsΛ} := by
|
||||
exact (Equiv.sumCompl HaveEqTime).symm
|
||||
rw [← e1.symm.sum_comp]
|
||||
simp only [Equiv.symm_symm, Algebra.smul_mul_assoc, Fintype.sum_sum_type,
|
||||
|
@ -93,8 +94,9 @@ lemma normalOrder_timeOrder_ofFieldOpList_eq_haveEqTime_sum_not_haveEqTime (φs
|
|||
|
||||
lemma haveEqTime_wick_sum_eq_split (φs : List 𝓕.States) :
|
||||
(∑ (φsΛ : {φsΛ : WickContraction φs.length // HaveEqTime φsΛ}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) =
|
||||
∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}), (sign φs ↑φsΛ • (φsΛ.1).timeContract *
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) =
|
||||
∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
|
||||
(sign φs ↑φsΛ • (φsΛ.1).timeContract *
|
||||
∑ φssucΛ : { φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬φssucΛ.HaveEqTime },
|
||||
sign [φsΛ.1]ᵘᶜ φssucΛ •
|
||||
(φssucΛ.1).timeContract * normalOrder (ofFieldOpList [φssucΛ.1]ᵘᶜ)) := by
|
||||
|
@ -104,7 +106,7 @@ lemma haveEqTime_wick_sum_eq_split (φs : List 𝓕.States) :
|
|||
rw [sum_haveEqTime]
|
||||
congr
|
||||
funext φsΛ
|
||||
simp only [ f]
|
||||
simp only [f]
|
||||
conv_lhs =>
|
||||
enter [2, φsucΛ]
|
||||
enter [1]
|
||||
|
@ -119,15 +121,14 @@ lemma haveEqTime_wick_sum_eq_split (φs : List 𝓕.States) :
|
|||
congr 1
|
||||
rw [@join_uncontractedListGet]
|
||||
|
||||
|
||||
lemma normalOrder_timeOrder_ofFieldOpList_eq_not_haveEqTime_sub_inductive (φs : List 𝓕.States) :
|
||||
𝓣(𝓝(ofFieldOpList φs)) = (∑ (φsΛ : {φsΛ : WickContraction φs.length // ¬ HaveEqTime φsΛ}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))
|
||||
+ ∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
|
||||
sign φs ↑φsΛ • (φsΛ.1).timeContract *
|
||||
(∑ φssucΛ : { φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬ φssucΛ.HaveEqTime },
|
||||
sign [φsΛ.1]ᵘᶜ φssucΛ • (φssucΛ.1).timeContract * normalOrder (ofFieldOpList [φssucΛ.1]ᵘᶜ) -
|
||||
𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))) := by
|
||||
𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))) := by
|
||||
rw [normalOrder_timeOrder_ofFieldOpList_eq_haveEqTime_sum_not_haveEqTime]
|
||||
rw [add_sub_assoc]
|
||||
congr 1
|
||||
|
@ -139,8 +140,9 @@ lemma normalOrder_timeOrder_ofFieldOpList_eq_not_haveEqTime_sub_inductive (φs :
|
|||
simp only
|
||||
rw [← smul_sub, ← mul_sub]
|
||||
|
||||
lemma wicks_theorem_normal_order_empty : 𝓣(𝓝(ofFieldOpList [])) = ∑ (φsΛ : {φsΛ : WickContraction ([] : List 𝓕.States).length // ¬ HaveEqTime φsΛ}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ) := by
|
||||
lemma wicks_theorem_normal_order_empty : 𝓣(𝓝(ofFieldOpList [])) =
|
||||
∑ (φsΛ : {φsΛ : WickContraction ([] : List 𝓕.States).length // ¬ HaveEqTime φsΛ}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ) := by
|
||||
let e2 : {φsΛ : WickContraction ([] : List 𝓕.States).length // ¬ HaveEqTime φsΛ} ≃ Unit :=
|
||||
{
|
||||
toFun := fun x => (),
|
||||
|
@ -173,11 +175,11 @@ lemma wicks_theorem_normal_order_empty : 𝓣(𝓝(ofFieldOpList [])) = ∑ (φs
|
|||
|
||||
theorem wicks_theorem_normal_order : (φs : List 𝓕.States) →
|
||||
𝓣(𝓝(ofFieldOpList φs)) = ∑ (φsΛ : {φsΛ : WickContraction φs.length // ¬ HaveEqTime φsΛ}),
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)
|
||||
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)
|
||||
| [] => wicks_theorem_normal_order_empty
|
||||
| φ :: φs => by
|
||||
rw [normalOrder_timeOrder_ofFieldOpList_eq_not_haveEqTime_sub_inductive]
|
||||
simp only [ Algebra.smul_mul_assoc, ne_eq, add_right_eq_self]
|
||||
simp only [Algebra.smul_mul_assoc, ne_eq, add_right_eq_self]
|
||||
apply Finset.sum_eq_zero
|
||||
intro φsΛ hφsΛ
|
||||
simp only [smul_eq_zero]
|
||||
|
@ -195,6 +197,5 @@ decreasing_by
|
|||
simp_all only [Algebra.smul_mul_assoc, List.length_cons, Finset.mem_univ, gt_iff_lt]
|
||||
omega
|
||||
|
||||
|
||||
end FieldOpAlgebra
|
||||
end FieldSpecification
|
||||
|
|
|
@ -44,7 +44,6 @@ 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]
|
||||
|
|
|
@ -23,6 +23,9 @@ variable {n : ℕ} (c : WickContraction n)
|
|||
open HepLean.List
|
||||
open FieldOpAlgebra
|
||||
|
||||
/-- Given a Wick contraction `φsΛ` on `φs` and a Wick contraction `φsucΛ` on the uncontracted fields
|
||||
within `φsΛ`, a Wick contraction on `φs`consisting of the contractins in `φsΛ` and
|
||||
the contractions in `φsucΛ`. -/
|
||||
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
|
||||
|
@ -60,12 +63,13 @@ def join {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
|||
|
||||
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
|
||||
(h1 : φsΛ = φsΛ') :
|
||||
join φsΛ φsucΛ = join φsΛ' (congr (by simp [h1]) φsucΛ) := by
|
||||
subst h1
|
||||
rfl
|
||||
|
||||
|
||||
/-- Given a contracting pair within `φsΛ` the corresponding contracting pair within
|
||||
`(join φsΛ φsucΛ)`. -/
|
||||
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]⟩
|
||||
|
@ -78,6 +82,8 @@ lemma jointLiftLeft_injective {φs : List 𝓕.States} {φsΛ : WickContraction
|
|||
rw [Subtype.mk_eq_mk] at h
|
||||
refine Subtype.eq h
|
||||
|
||||
/-- Given a contracting pair within `φsucΛ` the corresponding contracting pair within
|
||||
`(join φsΛ φsucΛ)`. -/
|
||||
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
|
||||
|
@ -116,6 +122,8 @@ lemma jointLiftLeft_neq_joinLiftRight {φs : List 𝓕.States} {φsΛ : WickCont
|
|||
rw [h1] at hj
|
||||
simp at hj
|
||||
|
||||
/-- The map from contracted pairs of `φsΛ` and `φsucΛ` to contracted pairs in
|
||||
`(join φsΛ φsucΛ)`. -/
|
||||
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
|
||||
|
@ -163,7 +171,7 @@ lemma joinLift_bijective {φs : List 𝓕.States} {φsΛ : WickContraction φs.l
|
|||
· 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]:
|
||||
(φ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
|
||||
|
@ -171,9 +179,12 @@ lemma prod_join {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
|||
simp only [Fintype.prod_sum_type, Finset.univ_eq_attach]
|
||||
rfl
|
||||
|
||||
lemma joinLiftLeft_or_joinLiftRight_of_mem_join {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) {a : Finset (Fin φs.length)} (ha : a ∈ (join φsΛ φsucΛ).1) :
|
||||
(∃ b, a = (joinLiftLeft (φsucΛ := φsucΛ) b).1) ∨ (∃ b, a = (joinLiftRight (φsucΛ := φsucΛ) b).1):= by
|
||||
lemma joinLiftLeft_or_joinLiftRight_of_mem_join {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) {a : Finset (Fin φs.length)}
|
||||
(ha : a ∈ (join φsΛ φsucΛ).1) :
|
||||
(∃ b, a = (joinLiftLeft (φsucΛ := φsucΛ) b).1) ∨
|
||||
(∃ b, a = (joinLiftRight (φsucΛ := φsucΛ) b).1) := by
|
||||
simp only [join, Finset.le_eq_subset, Finset.mem_union, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding] at ha
|
||||
rcases ha with ha | ⟨a, ha, rfl⟩
|
||||
|
@ -184,7 +195,6 @@ lemma joinLiftLeft_or_joinLiftRight_of_mem_join {φs : List 𝓕.States} (φsΛ
|
|||
use ⟨a, ha⟩
|
||||
rfl
|
||||
|
||||
|
||||
@[simp]
|
||||
lemma join_fstFieldOfContract_joinLiftRight {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (a : φsucΛ.1) :
|
||||
|
@ -227,13 +237,12 @@ lemma join_sndFieldOfContract_joinLift {φs : List 𝓕.States} (φsΛ : WickCon
|
|||
· simp [joinLiftLeft]
|
||||
· exact fstFieldOfContract_lt_sndFieldOfContract φsΛ a
|
||||
|
||||
|
||||
lemma mem_join_right_iff {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (a : Finset (Fin [φsΛ]ᵘᶜ.length)) :
|
||||
a ∈ φsucΛ.1 ↔ a.map uncontractedListEmd ∈ (join φsΛ φsucΛ).1 := by
|
||||
a ∈ φsucΛ.1 ↔ a.map uncontractedListEmd ∈ (join φsΛ φsucΛ).1 := by
|
||||
simp only [join, Finset.le_eq_subset, Finset.mem_union, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding]
|
||||
have h1' : ¬ Finset.map uncontractedListEmd a ∈ φsΛ.1 :=
|
||||
have h1' : ¬ Finset.map uncontractedListEmd a ∈ φsΛ.1 :=
|
||||
uncontractedListEmd_finset_not_mem a
|
||||
simp only [h1', false_or]
|
||||
apply Iff.intro
|
||||
|
@ -259,7 +268,7 @@ lemma join_card {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
|||
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding, not_exists, not_and]
|
||||
intro x hx
|
||||
by_contra hn
|
||||
have hdis : Disjoint (Finset.map uncontractedListEmd x) a := by
|
||||
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
|
||||
|
@ -333,7 +342,7 @@ lemma mem_join_uncontracted_of_mem_right_uncontracted {φs : List 𝓕.States}
|
|||
simp only [join, Finset.le_eq_subset, Finset.mem_union, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding] at hp
|
||||
rcases hp with hp | hp
|
||||
· have hi : uncontractedListEmd i ∈ φsΛ.uncontracted := by
|
||||
· have hi : uncontractedListEmd i ∈ φsΛ.uncontracted := by
|
||||
exact uncontractedListEmd_mem_uncontracted i
|
||||
rw [mem_uncontracted_iff_not_contracted] at hi
|
||||
exact hi p hp
|
||||
|
@ -353,7 +362,6 @@ lemma exists_mem_left_uncontracted_of_mem_join_uncontracted {φs : List 𝓕.Sta
|
|||
simp only [join, Finset.le_eq_subset, Finset.mem_union, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding] at ha
|
||||
intro p hp
|
||||
have hp' := ha p
|
||||
simp_all
|
||||
|
||||
lemma exists_mem_right_uncontracted_of_mem_join_uncontracted {φs : List 𝓕.States}
|
||||
|
@ -399,12 +407,12 @@ lemma join_uncontractedList {φs : List 𝓕.States} (φsΛ : WickContraction φ
|
|||
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
|
||||
φ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
|
||||
|
@ -420,10 +428,11 @@ lemma join_uncontractedListGet {φs : List 𝓕.States} (φsΛ : WickContraction
|
|||
Function.Embedding.coe_subtype]
|
||||
rfl
|
||||
|
||||
lemma join_uncontractedListEmb {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
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
|
||||
((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]
|
||||
|
@ -447,8 +456,8 @@ lemma join_assoc {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
|||
use a
|
||||
simp [ha']
|
||||
· obtain ⟨a, ha', rfl⟩ := h
|
||||
let a' := congrLift (congrArg List.length (join_uncontractedListGet _ _)) ⟨a, ha'⟩
|
||||
let a'' := joinLiftRight a'
|
||||
let a' := congrLift (congrArg List.length (join_uncontractedListGet _ _)) ⟨a, ha'⟩
|
||||
let a'' := joinLiftRight a'
|
||||
use a''
|
||||
apply And.intro
|
||||
· right
|
||||
|
@ -481,39 +490,43 @@ lemma join_assoc {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
|||
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)
|
||||
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
|
||||
(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
|
||||
obtain ⟨a, ha', ha⟩ := exists_mem_right_uncontracted_of_mem_join_uncontracted _ _
|
||||
(uncontractedListEmd i) h
|
||||
simp only [EmbeddingLike.apply_eq_iff_eq] 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)
|
||||
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
|
||||
((join φsΛ φsucΛ).getDual? (uncontractedListEmd i)).isSome
|
||||
↔ (φsucΛ.getDual? i).isSome := by
|
||||
rw [← Decidable.not_iff_not]
|
||||
simp
|
||||
simp [join_getDual?_apply_uncontractedListEmb_eq_none_iff]
|
||||
|
||||
lemma join_getDual?_apply_uncontractedListEmb_some {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
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
|
||||
(hi :((join φsΛ φsucΛ).getDual? (uncontractedListEmd i)).isSome) :
|
||||
((join φsΛ φsucΛ).getDual? (uncontractedListEmd i)) =
|
||||
some (uncontractedListEmd ((φsucΛ.getDual? i).get (by
|
||||
simpa [join_getDual?_apply_uncontractedListEmb_isSome_iff]using hi))) := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
simp only [join, Finset.le_eq_subset, Finset.mem_union, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding]
|
||||
right
|
||||
use {i, (φsucΛ.getDual? i).get (by simpa using hi)}
|
||||
use {i, (φsucΛ.getDual? i).get (by
|
||||
simpa [join_getDual?_apply_uncontractedListEmb_isSome_iff] using hi)}
|
||||
simp only [self_getDual?_get_mem, true_and]
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
simp
|
||||
|
@ -521,10 +534,11 @@ lemma join_getDual?_apply_uncontractedListEmb_some {φs : List 𝓕.States} (φs
|
|||
@[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
|
||||
((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) :=
|
||||
have h1 : (φsucΛ.getDual? i) = (φsucΛ.getDual? i).get (by simpa using h) :=
|
||||
Eq.symm (Option.some_get h)
|
||||
conv_rhs => rw [h1]
|
||||
simp only [Option.map_some']
|
||||
|
@ -610,8 +624,6 @@ lemma signFinset_right_map_uncontractedListEmd_eq_filter {φs : List 𝓕.States
|
|||
· 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
|
||||
|
@ -620,7 +632,6 @@ lemma signFinset_right_map_uncontractedListEmd_eq_filter {φs : List 𝓕.States
|
|||
have h2' := uncontractedListEmd_surjective_mem_uncontracted a h.2
|
||||
obtain ⟨a, rfl⟩ := h2'
|
||||
use a
|
||||
have h1 := h.1
|
||||
simp_all only [signFinset, Finset.mem_filter, Finset.mem_univ,
|
||||
join_getDual?_apply_uncontractedListEmb, Option.map_eq_none', Option.isSome_map', true_and,
|
||||
and_true, and_self]
|
||||
|
@ -650,10 +661,11 @@ lemma signFinset_right_map_uncontractedListEmd_eq_filter {φs : List 𝓕.States
|
|||
|
||||
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 ,
|
||||
φ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]
|
||||
|
@ -666,7 +678,7 @@ lemma sign_right_eq_prod_mul_prod {φs : List 𝓕.States} (φsΛ : WickContract
|
|||
|
||||
lemma join_singleton_signFinset_eq_filter {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(φ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 ∧
|
||||
|
@ -676,7 +688,7 @@ lemma join_singleton_signFinset_eq_filter {φs : List 𝓕.States}
|
|||
simp only [signFinset, Finset.mem_filter, Finset.mem_univ, true_and, not_and, not_forall, not_lt,
|
||||
and_assoc, and_congr_right_iff]
|
||||
intro h1 h2
|
||||
have h1 : (singleton h).getDual? a = none := by
|
||||
have h1 : (singleton h).getDual? a = none := by
|
||||
rw [singleton_getDual?_eq_none_iff_neq]
|
||||
omega
|
||||
simp only [h1, Option.isSome_none, Bool.false_eq_true, IsEmpty.forall_iff, or_self, true_and]
|
||||
|
@ -684,14 +696,14 @@ lemma join_singleton_signFinset_eq_filter {φs : List 𝓕.States}
|
|||
· intro h1 h2
|
||||
rcases h1 with h1 | h1
|
||||
· simp only [h1, Option.isSome_none, Bool.false_eq_true, IsEmpty.exists_iff]
|
||||
have h2' : ¬ (((singleton h).join φsucΛ).getDual? a).isSome := by
|
||||
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
|
||||
by_cases h2' : (((singleton h).join φsucΛ).getDual? a).isSome = true
|
||||
· have h2 := h2 h2'
|
||||
obtain ⟨hb, h2⟩ := h2
|
||||
right
|
||||
|
@ -706,10 +718,9 @@ lemma join_singleton_signFinset_eq_filter {φs : List 𝓕.States}
|
|||
· simp only [Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none] 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) :
|
||||
(φ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 =>
|
||||
|
@ -722,18 +733,21 @@ lemma join_singleton_left_signFinset_eq_filter {φs : List 𝓕.States}
|
|||
rw [mul_comm]
|
||||
exact (ofFinset_filter_mul_neg 𝓕.statesStatistic _ _ _).symm
|
||||
|
||||
/-- The difference in sign between `φsucΛ.sign` and the direct contribution of `φsucΛ` to
|
||||
`(join (singleton h) φsucΛ)`. -/
|
||||
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 ,
|
||||
(φ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)⟩)
|
||||
|
||||
(fun c => ¬ c ∈ (singleton h).uncontracted)⟩)
|
||||
|
||||
/-- The difference in sign between `(singleton h).sign` and the direct contribution of
|
||||
`(singleton h)` to `(join (singleton h) φsucΛ)`. -/
|
||||
def joinSignLeftExtra {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) : ℂ :=
|
||||
(φ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) →
|
||||
|
@ -741,30 +755,29 @@ def joinSignLeftExtra {φs : List 𝓕.States}
|
|||
|
||||
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
|
||||
(φ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Λ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
φsucΛ.sign =
|
||||
(joinSignRightExtra h φsucΛ) *
|
||||
(∏ a, 𝓢(𝓕|>ₛ [singleton h]ᵘᶜ[φsucΛ.sndFieldOfContract a], 𝓕|>ₛ ⟨φs.get ,
|
||||
(∏ a, 𝓢(𝓕|>ₛ [singleton h]ᵘᶜ[φsucΛ.sndFieldOfContract a], 𝓕|>ₛ ⟨φs.get,
|
||||
((join (singleton h) φsucΛ).signFinset (uncontractedListEmd (φsucΛ.fstFieldOfContract a))
|
||||
(uncontractedListEmd (φsucΛ.sndFieldOfContract a)))⟩)) := by
|
||||
(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) :
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(join (singleton h) φsucΛ).getDual? i = some j := by
|
||||
rw [@getDual?_eq_some_iff_mem]
|
||||
simp [singleton, join]
|
||||
|
@ -772,7 +785,7 @@ lemma join_singleton_getDual?_left {φs : List 𝓕.States}
|
|||
@[simp]
|
||||
lemma join_singleton_getDual?_right {φs : List 𝓕.States}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(join (singleton h) φsucΛ).getDual? j = some i := by
|
||||
rw [@getDual?_eq_some_iff_mem]
|
||||
simp only [join, singleton, Finset.le_eq_subset, Finset.mem_union, Finset.mem_singleton,
|
||||
|
@ -780,17 +793,16 @@ lemma join_singleton_getDual?_right {φs : List 𝓕.States}
|
|||
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) :
|
||||
(φ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 ∅)
|
||||
uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i then {j} else ∅)
|
||||
∪ (if uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i ∧
|
||||
i < uncontractedListEmd (φsucΛ.sndFieldOfContract a) then {i} else ∅)⟩) := by
|
||||
i < uncontractedListEmd (φsucΛ.sndFieldOfContract a) then {i} else ∅)⟩) := by
|
||||
rw [joinSignRightExtra]
|
||||
congr
|
||||
funext a
|
||||
|
@ -831,7 +843,7 @@ lemma joinSignRightExtra_eq_i_j_finset_eq_if {φs : List 𝓕.States}
|
|||
· have hi1 : i < uncontractedListEmd (φsucΛ.sndFieldOfContract a) := by
|
||||
omega
|
||||
simp only [hj1, true_and, hi1, and_true]
|
||||
by_cases hi2 : ¬ uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i
|
||||
by_cases hi2 : ¬ uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i
|
||||
· simp only [hi2, and_false, ↓reduceIte, Finset.not_mem_empty, or_self, iff_false, not_and,
|
||||
not_or, not_forall, not_lt]
|
||||
by_cases hj3 : ¬ j < uncontractedListEmd (φsucΛ.sndFieldOfContract a)
|
||||
|
@ -873,7 +885,6 @@ lemma joinSignRightExtra_eq_i_j_finset_eq_if {φs : List 𝓕.States}
|
|||
Option.get_some, forall_const, false_or, true_and]
|
||||
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) :
|
||||
|
@ -882,7 +893,8 @@ lemma joinSignLeftExtra_eq_joinSignRightExtra {φs : List 𝓕.States}
|
|||
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
|
||||
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]
|
||||
|
@ -913,11 +925,9 @@ lemma joinSignLeftExtra_eq_joinSignRightExtra {φs : List 𝓕.States}
|
|||
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
|
||||
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
|
||||
|
@ -926,7 +936,7 @@ lemma joinSignLeftExtra_eq_joinSignRightExtra {φs : List 𝓕.States}
|
|||
· have hj1 : uncontractedListEmd (φsucΛ.fstFieldOfContract a) < j := by omega
|
||||
simp only [hj1, and_true, instCommGroup, Fin.getElem_fin, true_and]
|
||||
by_cases hi2 : ¬ i < uncontractedListEmd (φsucΛ.sndFieldOfContract a)
|
||||
· have hi1 : ¬ i < uncontractedListEmd (φsucΛ.fstFieldOfContract a) := by omega
|
||||
· 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
|
||||
|
@ -934,7 +944,7 @@ lemma joinSignLeftExtra_eq_joinSignRightExtra {φs : List 𝓕.States}
|
|||
simp only [hi2n, and_false, ↓reduceIte, map_one, hi2, true_and, one_mul, and_true]
|
||||
by_cases hj2 : ¬ j < uncontractedListEmd (φsucΛ.sndFieldOfContract a)
|
||||
· simp only [hj2, false_and, ↓reduceIte, Finset.empty_union]
|
||||
have hj2 : uncontractedListEmd (φsucΛ.sndFieldOfContract a) < j:= by omega
|
||||
have hj2 : uncontractedListEmd (φsucΛ.sndFieldOfContract a) < j:= by omega
|
||||
simp only [hj2, true_and]
|
||||
by_cases hi1 : ¬ uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i
|
||||
· simp [hi1]
|
||||
|
@ -948,7 +958,7 @@ lemma joinSignLeftExtra_eq_joinSignRightExtra {φs : List 𝓕.States}
|
|||
· simp [hi1]
|
||||
· have hi1 : uncontractedListEmd (φsucΛ.fstFieldOfContract a) < i := by omega
|
||||
simp only [hi1, and_true, ↓reduceIte]
|
||||
have hj2 : ¬ uncontractedListEmd (φsucΛ.sndFieldOfContract a) < j := by omega
|
||||
have hj2 : ¬ uncontractedListEmd (φsucΛ.sndFieldOfContract a) < j := by omega
|
||||
simp only [hj2, ↓reduceIte, map_one]
|
||||
rw [← ofFinset_union_disjoint]
|
||||
simp only [instCommGroup, ofFinset_singleton, List.get_eq_getElem, hs]
|
||||
|
@ -993,15 +1003,16 @@ lemma exists_join_singleton_of_card_ge_zero {φs : List 𝓕.States} (φsΛ : Wi
|
|||
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))
|
||||
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 only [Fin.getElem_fin]
|
||||
apply And.intro
|
||||
· have h1 := join_congr (subContraction_singleton_eq_singleton _ ⟨a, ha⟩).symm (φsucΛ := φsucΛ)
|
||||
simp only [id_eq, eq_mpr_eq_cast, h1, congr_trans_apply, congr_refl, φsucΛ]
|
||||
rw [join_sub_quot]
|
||||
· apply And.intro (hc ⟨a, ha⟩)
|
||||
· apply And.intro (hc ⟨a, ha⟩)
|
||||
apply And.intro
|
||||
· simp only [id_eq, eq_mpr_eq_cast, φsucΛ]
|
||||
rw [gradingCompliant_congr (φs' := [(φsΛ.subContraction {a} (by simpa using ha))]ᵘᶜ)]
|
||||
|
@ -1024,11 +1035,12 @@ lemma join_sign_induction {φs : List 𝓕.States} (φsΛ : WickContraction φs.
|
|||
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
|
||||
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
|
||||
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]
|
||||
|
@ -1058,7 +1070,7 @@ lemma join_not_gradingCompliant_of_left_not_gradingCompliant {φs : List 𝓕.St
|
|||
lemma join_sign_timeContract {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).sign • (join φsΛ φsucΛ).timeContract.1 =
|
||||
(φsΛ.sign • φsΛ.timeContract.1) * (φsucΛ.sign • φsucΛ.timeContract.1) := by
|
||||
(φsΛ.sign • φsΛ.timeContract.1) * (φsucΛ.sign • φsucΛ.timeContract.1) := by
|
||||
rw [join_timeContract]
|
||||
by_cases h : φsΛ.GradingCompliant
|
||||
· rw [join_sign _ _ h]
|
||||
|
|
|
@ -22,6 +22,7 @@ open HepLean.List
|
|||
open FieldOpAlgebra
|
||||
open FieldStatistic
|
||||
|
||||
/-- The Wick contraction formed from a single ordered pair. -/
|
||||
def singleton {i j : Fin n} (hij : i < j) : WickContraction n :=
|
||||
⟨{{i, j}}, by
|
||||
intro i hi
|
||||
|
@ -30,30 +31,28 @@ def singleton {i j : Fin n} (hij : i < j) : WickContraction n :=
|
|||
rw [@Finset.card_eq_two]
|
||||
use i, j
|
||||
simp only [ne_eq, and_true]
|
||||
omega
|
||||
, by
|
||||
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
|
||||
{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
|
||||
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):
|
||||
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] :
|
||||
(f : (singleton hij).1 → M) [CommMonoid M] :
|
||||
∏ a, f a = f ⟨{i,j}, mem_singleton hij⟩:= by
|
||||
simp [singleton]
|
||||
simp [singleton, of_singleton_eq]
|
||||
|
||||
@[simp]
|
||||
lemma singleton_fstFieldOfContract {i j : Fin n} (hij : i < j) :
|
||||
|
@ -84,23 +83,23 @@ lemma singleton_getDual?_eq_none_iff_neq {i j : Fin n} (hij : i < j) (a : Fin n)
|
|||
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
|
||||
(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
|
||||
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
|
||||
(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
|
||||
have h2 : j ∉ (singleton hij).uncontracted := by
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
simp [singleton]
|
||||
simp_all
|
||||
|
|
|
@ -8,7 +8,7 @@ import HepLean.PerturbationTheory.WickContraction.StaticContract
|
|||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.TimeContraction
|
||||
/-!
|
||||
|
||||
# Sub contractions
|
||||
# Sub contractions
|
||||
|
||||
-/
|
||||
|
||||
|
@ -20,7 +20,10 @@ 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 :=
|
||||
/-- Given a Wick contraction `φsΛ`, and a subset of `φsΛ.1`, the Wick contraction
|
||||
conisting of contracted pairs within that subset. -/
|
||||
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),
|
||||
|
@ -32,13 +35,16 @@ lemma mem_of_mem_subContraction {S : Finset (Finset (Fin φs.length))} {hs : S
|
|||
{a : Finset (Fin φs.length)} (ha : a ∈ (φsΛ.subContraction S hs).1) : a ∈ φsΛ.1 := by
|
||||
exact hs ha
|
||||
|
||||
/-- Given a Wick contraction `φsΛ`, and a subset `S` of `φsΛ.1`, the Wick contraction
|
||||
on the uncontracted list `[φsΛ.subContraction S ha]ᵘᶜ`
|
||||
consisting of the remaining contracted pairs of `φsΛ` not in `S`. -/
|
||||
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 only [Finset.mem_filter, Finset.mem_univ, true_and] at ha'
|
||||
simpa using φsΛ.2.1 (Finset.map uncontractedListEmd a) ha' , by
|
||||
simpa using φsΛ.2.1 (Finset.map uncontractedListEmd a) ha', by
|
||||
intro a ha b hb
|
||||
simp only [Finset.mem_filter, Finset.mem_univ, true_and] at ha hb
|
||||
by_cases hab : a = b
|
||||
|
@ -103,13 +109,16 @@ lemma subContraction_fstFieldOfContract {S : Finset (Finset (Fin φs.length))} {
|
|||
(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⟩
|
||||
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 only at ha
|
||||
conv_lhs =>
|
||||
rw [ha]
|
||||
simp
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _ ⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _
|
||||
⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
simp only at ha
|
||||
conv_lhs =>
|
||||
rw [ha]
|
||||
|
@ -121,27 +130,31 @@ lemma subContraction_sndFieldOfContract {S : Finset (Finset (Fin φs.length))} {
|
|||
(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⟩
|
||||
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 only at ha
|
||||
conv_lhs =>
|
||||
rw [ha]
|
||||
simp
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _ ⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _
|
||||
⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
simp only 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) :
|
||||
uncontractedListEmd ((quotContraction S hs).fstFieldOfContract a) =
|
||||
(φsΛ.fstFieldOfContract ⟨Finset.map uncontractedListEmd a.1, mem_of_mem_quotContraction a.2⟩) := by
|
||||
(φ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) )
|
||||
apply eq_fstFieldOfContract_of_mem _ _ _
|
||||
(uncontractedListEmd ((quotContraction S hs).sndFieldOfContract a))
|
||||
· simp only [Finset.mem_map', fstFieldOfContract_mem]
|
||||
· simp
|
||||
· apply uncontractedListEmd_strictMono
|
||||
|
@ -151,9 +164,11 @@ lemma quotContraction_fstFieldOfContract_uncontractedListEmd {S : Finset (Finset
|
|||
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
|
||||
(φ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) )
|
||||
apply eq_sndFieldOfContract_of_mem _ _
|
||||
(uncontractedListEmd ((quotContraction S hs).fstFieldOfContract a))
|
||||
· simp only [Finset.mem_map', fstFieldOfContract_mem]
|
||||
· simp
|
||||
· apply uncontractedListEmd_strictMono
|
||||
|
@ -164,7 +179,6 @@ lemma quotContraction_gradingCompliant {S : Finset (Finset (Fin φs.length))} {h
|
|||
GradingCompliant [φsΛ.subContraction S hs]ᵘᶜ (quotContraction S hs) := by
|
||||
simp only [GradingCompliant, Fin.getElem_fin, Subtype.forall]
|
||||
intro a ha
|
||||
have h1' := mem_of_mem_quotContraction ha
|
||||
erw [subContraction_uncontractedList_get]
|
||||
erw [subContraction_uncontractedList_get]
|
||||
simp only [quotContraction_fstFieldOfContract_uncontractedListEmd, Fin.getElem_fin,
|
||||
|
@ -173,7 +187,7 @@ lemma quotContraction_gradingCompliant {S : Finset (Finset (Fin φs.length))} {h
|
|||
|
||||
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 ∈ (quotContraction S hs).1 ↔ a.map uncontractedListEmd ∈ φsΛ.1
|
||||
∧ a.map uncontractedListEmd ∉ (subContraction S hs).1 := by
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
|
@ -184,5 +198,4 @@ lemma mem_quotContraction_iff {S : Finset (Finset (Fin φs.length))} {hs : S ⊆
|
|||
have h2 := mem_subContraction_or_quotContraction (S := S) (hs := hs) h.1
|
||||
simp_all
|
||||
|
||||
|
||||
end WickContraction
|
||||
|
|
|
@ -26,11 +26,10 @@ def EqTimeOnly {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) : P
|
|||
∀ (i j), {i, j} ∈ φsΛ.1 → timeOrderRel φs[i] φs[j]
|
||||
noncomputable section
|
||||
|
||||
instance {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
instance {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
Decidable (EqTimeOnly φsΛ) :=
|
||||
inferInstanceAs (Decidable (∀ (i j), {i, j} ∈ φsΛ.1 → timeOrderRel φs[i] φs[j]))
|
||||
|
||||
|
||||
namespace EqTimeOnly
|
||||
variable {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
|
||||
|
@ -44,7 +43,7 @@ lemma timeOrderRel_of_eqTimeOnly_pair {i j : Fin φs.length} (h : {i, j} ∈ φs
|
|||
|
||||
lemma timeOrderRel_both_of_eqTimeOnly {i j : Fin φs.length} (h : {i, j} ∈ φsΛ.1)
|
||||
(hc : EqTimeOnly φsΛ) :
|
||||
timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i] := by
|
||||
timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i] := by
|
||||
apply And.intro
|
||||
· exact timeOrderRel_of_eqTimeOnly_pair φsΛ h hc
|
||||
· apply timeOrderRel_of_eqTimeOnly_pair φsΛ _ hc
|
||||
|
@ -52,9 +51,9 @@ lemma timeOrderRel_both_of_eqTimeOnly {i j : Fin φs.length} (h : {i, j} ∈ φs
|
|||
exact h
|
||||
|
||||
lemma eqTimeOnly_iff_forall_finset {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
φsΛ.EqTimeOnly ↔ ∀ (a : φsΛ.1),
|
||||
φsΛ.EqTimeOnly ↔ ∀ (a : φsΛ.1),
|
||||
timeOrderRel (φs[φsΛ.fstFieldOfContract a]) (φs[φsΛ.sndFieldOfContract a])
|
||||
∧ timeOrderRel (φs[φsΛ.sndFieldOfContract a]) (φs[φsΛ.fstFieldOfContract a]) := by
|
||||
∧ timeOrderRel (φs[φsΛ.sndFieldOfContract a]) (φs[φsΛ.fstFieldOfContract a]) := by
|
||||
apply Iff.intro
|
||||
· intro h a
|
||||
apply timeOrderRel_both_of_eqTimeOnly φsΛ _ h
|
||||
|
@ -101,7 +100,7 @@ lemma empty_mem {φs : List 𝓕.States} : empty (n := φs.length).EqTimeOnly :=
|
|||
|
||||
lemma staticContract_eq_timeContract_of_eqTimeOnly (h : φsΛ.EqTimeOnly) :
|
||||
φsΛ.staticContract = φsΛ.timeContract := by
|
||||
simp only [staticContract, timeContract]
|
||||
simp only [staticContract, timeContract]
|
||||
apply congrArg
|
||||
funext a
|
||||
ext
|
||||
|
@ -112,14 +111,14 @@ lemma staticContract_eq_timeContract_of_eqTimeOnly (h : φsΛ.EqTimeOnly) :
|
|||
exact a.2
|
||||
exact h
|
||||
|
||||
lemma eqTimeOnly_congr {φs φs' : List 𝓕.States} (h : φs = φs') (φsΛ : WickContraction φs.length):
|
||||
lemma eqTimeOnly_congr {φs φs' : List 𝓕.States} (h : φs = φs') (φsΛ : WickContraction φs.length) :
|
||||
(congr (by simp [h]) φsΛ).EqTimeOnly (φs := φs') ↔ φsΛ.EqTimeOnly := by
|
||||
subst h
|
||||
simp
|
||||
|
||||
lemma quotContraction_eqTimeOnly {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
(h : φsΛ.EqTimeOnly) (S : Finset (Finset (Fin φs.length))) (ha : S ⊆ φsΛ.1) :
|
||||
(φsΛ.quotContraction S ha).EqTimeOnly := by
|
||||
(h : φsΛ.EqTimeOnly) (S : Finset (Finset (Fin φs.length))) (ha : S ⊆ φsΛ.1) :
|
||||
(φsΛ.quotContraction S ha).EqTimeOnly := by
|
||||
rw [eqTimeOnly_iff_forall_finset]
|
||||
intro a
|
||||
simp only [Fin.getElem_fin]
|
||||
|
@ -131,7 +130,7 @@ lemma quotContraction_eqTimeOnly {φs : List 𝓕.States} {φsΛ : WickContracti
|
|||
apply h
|
||||
|
||||
lemma exists_join_singleton_of_card_ge_zero {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
(h : 0 < φsΛ.1.card) (h1 : φsΛ.EqTimeOnly) :
|
||||
(h : 0 < φsΛ.1.card) (h1 : φsΛ.EqTimeOnly) :
|
||||
∃ (i j : Fin φs.length) (h : i < j) (φsucΛ : WickContraction [singleton h]ᵘᶜ.length),
|
||||
φsΛ = join (singleton h) φsucΛ ∧ (timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i])
|
||||
∧ φsucΛ.EqTimeOnly ∧ φsucΛ.1.card + 1 = φsΛ.1.card := by
|
||||
|
@ -140,8 +139,9 @@ lemma exists_join_singleton_of_card_ge_zero {φs : List 𝓕.States} (φsΛ : Wi
|
|||
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))
|
||||
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 only [Fin.getElem_fin]
|
||||
apply And.intro
|
||||
|
@ -163,7 +163,8 @@ lemma exists_join_singleton_of_card_ge_zero {φs : List 𝓕.States} (φsΛ : Wi
|
|||
simp only [subContraction, Finset.card_singleton, id_eq, eq_mpr_eq_cast] at h1
|
||||
omega
|
||||
|
||||
lemma timeOrder_timeContract_mul_of_eqTimeOnly_mid_induction {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma timeOrder_timeContract_mul_of_eqTimeOnly_mid_induction {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(hl : φsΛ.EqTimeOnly) (a b: 𝓕.FieldOpAlgebra) : (n : ℕ) → (hn : φsΛ.1.card = n) →
|
||||
𝓣(a * φsΛ.timeContract.1 * b) = φsΛ.timeContract.1 * 𝓣(a * b)
|
||||
| 0, hn => by
|
||||
|
@ -171,7 +172,8 @@ lemma timeOrder_timeContract_mul_of_eqTimeOnly_mid_induction {φs : List 𝓕.S
|
|||
subst hn
|
||||
simp
|
||||
| Nat.succ n, hn => by
|
||||
obtain ⟨i, j, hij, φsucΛ, rfl, h2, h3, h4⟩ := exists_join_singleton_of_card_ge_zero φsΛ (by simp [hn]) hl
|
||||
obtain ⟨i, j, hij, φsucΛ, rfl, h2, h3, h4⟩ :=
|
||||
exists_join_singleton_of_card_ge_zero φsΛ (by simp [hn]) hl
|
||||
rw [join_timeContract]
|
||||
rw [singleton_timeContract]
|
||||
simp only [Fin.getElem_fin, MulMemClass.coe_mul]
|
||||
|
@ -192,8 +194,8 @@ lemma timeOrder_timeContract_mul_of_eqTimeOnly_mid {φs : List 𝓕.States}
|
|||
|
||||
lemma timeOrder_timeContract_mul_of_eqTimeOnly_left {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(hl : φsΛ.EqTimeOnly) ( b : 𝓕.FieldOpAlgebra) :
|
||||
𝓣( φsΛ.timeContract.1 * b) = φsΛ.timeContract.1 * 𝓣( b) := by
|
||||
(hl : φsΛ.EqTimeOnly) (b : 𝓕.FieldOpAlgebra) :
|
||||
𝓣(φsΛ.timeContract.1 * b) = φsΛ.timeContract.1 * 𝓣(b) := by
|
||||
trans 𝓣(1 * φsΛ.timeContract.1 * b)
|
||||
simp only [one_mul]
|
||||
rw [timeOrder_timeContract_mul_of_eqTimeOnly_mid φsΛ hl]
|
||||
|
@ -210,8 +212,9 @@ lemma exists_join_singleton_of_not_eqTimeOnly {φs : List 𝓕.States} (φsΛ :
|
|||
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))
|
||||
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 only [Fin.getElem_fin]
|
||||
apply And.intro
|
||||
|
@ -223,7 +226,8 @@ lemma exists_join_singleton_of_not_eqTimeOnly {φs : List 𝓕.States} (φsΛ :
|
|||
· simp_all [h1]
|
||||
· simp_all [h1]
|
||||
|
||||
lemma timeOrder_timeContract_of_not_eqTimeOnly {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma timeOrder_timeContract_of_not_eqTimeOnly {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(hl : ¬ φsΛ.EqTimeOnly) : 𝓣(φsΛ.timeContract.1) = 0 := by
|
||||
obtain ⟨i, j, hij, φsucΛ, rfl, hr⟩ := exists_join_singleton_of_not_eqTimeOnly φsΛ hl
|
||||
rw [join_timeContract]
|
||||
|
@ -250,19 +254,21 @@ lemma timeOrder_staticContract_of_not_mem {φs : List 𝓕.States} (φsΛ : Wick
|
|||
|
||||
end EqTimeOnly
|
||||
|
||||
|
||||
/-- The condition on a Wick contraction which is true if it has at least one contraction
|
||||
which is between two equal time fields. -/
|
||||
def HaveEqTime {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) : Prop :=
|
||||
∃ (i j : Fin φs.length) (h : {i, j} ∈ φsΛ.1),
|
||||
timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i]
|
||||
|
||||
|
||||
noncomputable instance {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
noncomputable instance {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
Decidable (HaveEqTime φsΛ) :=
|
||||
inferInstanceAs (Decidable (∃ (i j : Fin φs.length) (h : ({i, j} : Finset (Fin φs.length)) ∈ φsΛ.1),
|
||||
timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i]))
|
||||
inferInstanceAs (Decidable (∃ (i j : Fin φs.length)
|
||||
(h : ({i, j} : Finset (Fin φs.length)) ∈ φsΛ.1),
|
||||
timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i]))
|
||||
|
||||
lemma haveEqTime_iff_finset {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
HaveEqTime φsΛ ↔ ∃ (a : Finset (Fin φs.length)) (h : a ∈ φsΛ.1), timeOrderRel φs[φsΛ.fstFieldOfContract ⟨a, h⟩] φs[φsΛ.sndFieldOfContract ⟨a, h⟩]
|
||||
HaveEqTime φsΛ ↔ ∃ (a : Finset (Fin φs.length)) (h : a ∈ φsΛ.1),
|
||||
timeOrderRel φs[φsΛ.fstFieldOfContract ⟨a, h⟩] φs[φsΛ.sndFieldOfContract ⟨a, h⟩]
|
||||
∧ timeOrderRel φs[φsΛ.sndFieldOfContract ⟨a, h⟩] φs[φsΛ.fstFieldOfContract ⟨a, h⟩] := by
|
||||
simp only [HaveEqTime, Fin.getElem_fin, exists_and_left, exists_prop]
|
||||
apply Iff.intro
|
||||
|
@ -293,16 +299,18 @@ lemma haveEqTime_iff_finset {φs : List 𝓕.States} (φsΛ : WickContraction φ
|
|||
exact h1
|
||||
|
||||
@[simp]
|
||||
lemma empty_not_haveEqTime {φs : List 𝓕.States} : ¬ HaveEqTime (empty : WickContraction φs.length) := by
|
||||
lemma empty_not_haveEqTime {φs : List 𝓕.States} :
|
||||
¬ HaveEqTime (empty : WickContraction φs.length) := by
|
||||
rw [haveEqTime_iff_finset]
|
||||
simp [empty]
|
||||
|
||||
/-- Given a Wick contraction the subset of contracted pairs between eqaul time fields. -/
|
||||
def eqTimeContractSet {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
Finset (Finset (Fin φs.length)) :=
|
||||
Finset.univ.filter (fun a =>
|
||||
a ∈ φsΛ.1 ∧ ∀ (h : a ∈ φsΛ.1),
|
||||
timeOrderRel φs[φsΛ.fstFieldOfContract ⟨a, h⟩] φs[φsΛ.sndFieldOfContract ⟨a, h⟩]
|
||||
∧ timeOrderRel φs[φsΛ.sndFieldOfContract ⟨a, h⟩] φs[φsΛ.fstFieldOfContract ⟨a, h⟩])
|
||||
timeOrderRel φs[φsΛ.fstFieldOfContract ⟨a, h⟩] φs[φsΛ.sndFieldOfContract ⟨a, h⟩]
|
||||
∧ timeOrderRel φs[φsΛ.sndFieldOfContract ⟨a, h⟩] φs[φsΛ.fstFieldOfContract ⟨a, h⟩])
|
||||
|
||||
lemma eqTimeContractSet_subset {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
eqTimeContractSet φsΛ ⊆ φsΛ.1 := by
|
||||
|
@ -366,14 +374,14 @@ lemma join_eqTimeContractSet {φs : List 𝓕.States} (φsΛ : WickContraction
|
|||
· simp [join, h1]
|
||||
· intro h'
|
||||
have h2 := h1.2 h1.1
|
||||
have hj : ⟨(Finset.mapEmbedding uncontractedListEmd) b, h'⟩ = joinLiftRight ⟨b, h1.1⟩ := by rfl
|
||||
have hj : ⟨(Finset.mapEmbedding uncontractedListEmd) b, h'⟩
|
||||
= joinLiftRight ⟨b, h1.1⟩ := by rfl
|
||||
simp only [hj, join_fstFieldOfContract_joinLiftRight, getElem_uncontractedListEmd,
|
||||
join_sndFieldOfContract_joinLiftRight]
|
||||
simpa using h2
|
||||
|
||||
|
||||
lemma eqTimeContractSet_of_not_haveEqTime {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
(h : ¬ HaveEqTime φsΛ) : eqTimeContractSet φsΛ = ∅ := by
|
||||
(h : ¬ HaveEqTime φsΛ) : eqTimeContractSet φsΛ = ∅ := by
|
||||
ext a
|
||||
simp only [Finset.not_mem_empty, iff_false]
|
||||
by_contra hn
|
||||
|
@ -381,7 +389,6 @@ lemma eqTimeContractSet_of_not_haveEqTime {φs : List 𝓕.States} {φsΛ : Wick
|
|||
simp only [Fin.getElem_fin, not_exists, not_and] at h
|
||||
simp only [eqTimeContractSet, Fin.getElem_fin, Finset.mem_filter, Finset.mem_univ, true_and] at hn
|
||||
have h2 := hn.2 hn.1
|
||||
have h3 := h a hn.1
|
||||
simp_all
|
||||
|
||||
lemma eqTimeContractSet_of_mem_eqTimeOnly {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
|
@ -392,18 +399,20 @@ lemma eqTimeContractSet_of_mem_eqTimeOnly {φs : List 𝓕.States} {φsΛ : Wick
|
|||
rw [EqTimeOnly.eqTimeOnly_iff_forall_finset] at h
|
||||
exact fun h_1 => h ⟨a, h_1⟩
|
||||
|
||||
lemma subContraction_eqTimeContractSet_eqTimeOnly {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
lemma subContraction_eqTimeContractSet_eqTimeOnly {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length) :
|
||||
(φsΛ.subContraction (eqTimeContractSet φsΛ) (eqTimeContractSet_subset φsΛ)).EqTimeOnly := by
|
||||
rw [EqTimeOnly.eqTimeOnly_iff_forall_finset]
|
||||
intro a
|
||||
have ha2 := a.2
|
||||
have ha2 := a.2
|
||||
simp only [subContraction, eqTimeContractSet, Fin.getElem_fin, Finset.mem_filter, Finset.mem_univ,
|
||||
true_and] at ha2
|
||||
simp only [subContraction_fstFieldOfContract, Fin.getElem_fin, subContraction_sndFieldOfContract]
|
||||
exact ha2.2 ha2.1
|
||||
|
||||
lemma pair_mem_eqTimeContractSet_iff {φs : List 𝓕.States} {i j : Fin φs.length} (φsΛ : WickContraction φs.length) (h : {i, j} ∈ φsΛ.1) :
|
||||
{i, j} ∈ φsΛ.eqTimeContractSet ↔ timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i] := by
|
||||
lemma pair_mem_eqTimeContractSet_iff {φs : List 𝓕.States} {i j : Fin φs.length}
|
||||
(φsΛ : WickContraction φs.length) (h : {i, j} ∈ φsΛ.1) :
|
||||
{i, j} ∈ φsΛ.eqTimeContractSet ↔ timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i] := by
|
||||
simp only [eqTimeContractSet, Fin.getElem_fin, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
by_cases hij : i < j
|
||||
· have h1 := eq_fstFieldOfContract_of_mem φsΛ ⟨{i,j}, h⟩ i j (by simp) (by simp) hij
|
||||
|
@ -441,8 +450,9 @@ lemma subContraction_eqTimeContractSet_not_empty_of_haveEqTime
|
|||
simp_all only [Fin.getElem_fin, and_self]
|
||||
exact h1
|
||||
|
||||
lemma quotContraction_eqTimeContractSet_not_haveEqTime {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
¬ HaveEqTime (φsΛ.quotContraction (eqTimeContractSet φsΛ) (eqTimeContractSet_subset φsΛ)) := by
|
||||
lemma quotContraction_eqTimeContractSet_not_haveEqTime {φs : List 𝓕.States}
|
||||
(φsΛ : WickContraction φs.length) :
|
||||
¬ HaveEqTime (φsΛ.quotContraction (eqTimeContractSet φsΛ) (eqTimeContractSet_subset φsΛ)) := by
|
||||
rw [haveEqTime_iff_finset]
|
||||
simp only [Fin.getElem_fin, not_exists, not_and]
|
||||
intro a ha
|
||||
|
@ -451,8 +461,8 @@ lemma quotContraction_eqTimeContractSet_not_haveEqTime {φs : List 𝓕.States}
|
|||
simp only [quotContraction_fstFieldOfContract_uncontractedListEmd, Fin.getElem_fin,
|
||||
quotContraction_sndFieldOfContract_uncontractedListEmd]
|
||||
simp only [quotContraction, Finset.mem_filter, Finset.mem_univ, true_and] at ha
|
||||
have hn' : Finset.map uncontractedListEmd a ∉
|
||||
(φsΛ.subContraction (eqTimeContractSet φsΛ) (eqTimeContractSet_subset φsΛ) ).1 := by
|
||||
have hn' : Finset.map uncontractedListEmd a ∉
|
||||
(φsΛ.subContraction (eqTimeContractSet φsΛ) (eqTimeContractSet_subset φsΛ)).1 := by
|
||||
exact uncontractedListEmd_finset_not_mem a
|
||||
simp only [subContraction, eqTimeContractSet, Fin.getElem_fin, Finset.mem_filter, Finset.mem_univ,
|
||||
true_and, not_and, not_forall] at hn'
|
||||
|
@ -470,13 +480,13 @@ lemma join_haveEqTime_of_eqTimeOnly_nonEmpty {φs : List 𝓕.States} (φsΛ : W
|
|||
true_and] at h1
|
||||
obtain ⟨i, j, h⟩ := exists_pair_of_not_eq_empty _ h2
|
||||
use i, j
|
||||
have h1 := h1 i j h
|
||||
simp_all only [ne_eq, true_or, true_and]
|
||||
apply h1 j i
|
||||
rw [Finset.pair_comm]
|
||||
exact h
|
||||
|
||||
lemma hasEqTimeEquiv_ext_sigma {φs : List 𝓕.States} {x1 x2 : Σ (φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly ∧ φsΛ ≠ empty}),
|
||||
lemma hasEqTimeEquiv_ext_sigma {φs : List 𝓕.States} {x1 x2 :
|
||||
Σ (φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly ∧ φsΛ ≠ empty}),
|
||||
{φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬ HaveEqTime φssucΛ}}
|
||||
(h1 : x1.1.1 = x2.1.1) (h2 : x1.2.1 = congr (by simp [h1]) x2.2.1) : x1 = x2 := by
|
||||
match x1, x2 with
|
||||
|
@ -486,16 +496,20 @@ lemma hasEqTimeEquiv_ext_sigma {φs : List 𝓕.States} {x1 x2 : Σ (φsΛ : {
|
|||
simp only [ne_eq, congr_refl] at h2
|
||||
simp [h2]
|
||||
|
||||
/-- The equivalence which seperates a Wick contraction which has an equal time contraction
|
||||
into a non-empty contraction only between equal-time fields and a Wick contraction which
|
||||
does not have equal time contractions. -/
|
||||
def hasEqTimeEquiv (φs : List 𝓕.States) :
|
||||
{φsΛ : WickContraction φs.length // HaveEqTime φsΛ} ≃
|
||||
Σ (φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly ∧ φsΛ ≠ empty}),
|
||||
{φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬ HaveEqTime φssucΛ} where
|
||||
{φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬ HaveEqTime φssucΛ} where
|
||||
toFun φsΛ := ⟨⟨φsΛ.1.subContraction (eqTimeContractSet φsΛ.1) (eqTimeContractSet_subset φsΛ.1),
|
||||
subContraction_eqTimeContractSet_eqTimeOnly φsΛ.1,
|
||||
subContraction_eqTimeContractSet_not_empty_of_haveEqTime φsΛ.1 φsΛ.2⟩,
|
||||
⟨φsΛ.1.quotContraction (eqTimeContractSet φsΛ.1) (eqTimeContractSet_subset φsΛ.1),
|
||||
quotContraction_eqTimeContractSet_not_haveEqTime φsΛ.1⟩⟩
|
||||
invFun φsΛ := ⟨join φsΛ.1.1 φsΛ.2.1 , join_haveEqTime_of_eqTimeOnly_nonEmpty φsΛ.1.1 φsΛ.1.2.1 φsΛ.1.2.2 φsΛ.2⟩
|
||||
invFun φsΛ := ⟨join φsΛ.1.1 φsΛ.2.1,
|
||||
join_haveEqTime_of_eqTimeOnly_nonEmpty φsΛ.1.1 φsΛ.1.2.1 φsΛ.1.2.2 φsΛ.2⟩
|
||||
left_inv φsΛ := by
|
||||
match φsΛ with
|
||||
| ⟨φsΛ, h1, h2⟩ =>
|
||||
|
@ -504,8 +518,8 @@ def hasEqTimeEquiv (φs : List 𝓕.States) :
|
|||
right_inv φsΛ' := by
|
||||
match φsΛ' with
|
||||
| ⟨⟨φsΛ, h1⟩, ⟨φsucΛ, h2⟩⟩ =>
|
||||
have hs : subContraction (φsΛ.join φsucΛ).eqTimeContractSet (
|
||||
eqTimeContractSet_subset (φsΛ.join φsucΛ)) = φsΛ := by
|
||||
have hs : subContraction (φsΛ.join φsucΛ).eqTimeContractSet
|
||||
(eqTimeContractSet_subset (φsΛ.join φsucΛ)) = φsΛ := by
|
||||
apply Subtype.ext
|
||||
ext a
|
||||
simp only [subContraction]
|
||||
|
@ -526,7 +540,6 @@ def hasEqTimeEquiv (φs : List 𝓕.States) :
|
|||
rw [uncontractedListEmd_congr hs]
|
||||
rw [Finset.map_map]
|
||||
|
||||
|
||||
lemma sum_haveEqTime (φs : List 𝓕.States)
|
||||
(f : WickContraction φs.length → M) [AddCommMonoid M]:
|
||||
∑ (φsΛ : {φsΛ : WickContraction φs.length // HaveEqTime φsΛ}), f φsΛ =
|
||||
|
|
|
@ -68,7 +68,6 @@ 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
|
||||
|
|
|
@ -199,7 +199,6 @@ lemma uncontractedList_mem_iff (i : Fin n) :
|
|||
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]
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue