feat: Time dependent Wick theorem. (#274)
feat: Proof of the time-dependent Wick's theorem
This commit is contained in:
parent
4d43698b3c
commit
17f84b7153
53 changed files with 8563 additions and 3329 deletions
47
HepLean.lean
47
HepLean.lean
|
@ -104,6 +104,8 @@ import HepLean.Mathematics.Fin
|
|||
import HepLean.Mathematics.Fin.Involutions
|
||||
import HepLean.Mathematics.LinearMaps
|
||||
import HepLean.Mathematics.List
|
||||
import HepLean.Mathematics.List.InsertIdx
|
||||
import HepLean.Mathematics.List.InsertionSort
|
||||
import HepLean.Mathematics.PiTensorProduct
|
||||
import HepLean.Mathematics.SO3.Basic
|
||||
import HepLean.Mathematics.SchurTriangulation
|
||||
|
@ -116,29 +118,42 @@ import HepLean.Meta.Notes.HTMLNote
|
|||
import HepLean.Meta.Notes.NoteFile
|
||||
import HepLean.Meta.Notes.ToHTML
|
||||
import HepLean.Meta.TransverseTactics
|
||||
import HepLean.PerturbationTheory.Contractions.Basic
|
||||
import HepLean.PerturbationTheory.Contractions.Card
|
||||
import HepLean.PerturbationTheory.Contractions.Involutions
|
||||
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.NormalOrder
|
||||
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.SuperCommute
|
||||
import HepLean.PerturbationTheory.Algebras.OperatorAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Algebras.OperatorAlgebra.NormalOrder
|
||||
import HepLean.PerturbationTheory.Algebras.OperatorAlgebra.TimeContraction
|
||||
import HepLean.PerturbationTheory.Algebras.StateAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Algebras.StateAlgebra.TimeOrder
|
||||
import HepLean.PerturbationTheory.CreateAnnihilate
|
||||
import HepLean.PerturbationTheory.FeynmanDiagrams.Basic
|
||||
import HepLean.PerturbationTheory.FeynmanDiagrams.Instances.ComplexScalar
|
||||
import HepLean.PerturbationTheory.FeynmanDiagrams.Instances.Phi4
|
||||
import HepLean.PerturbationTheory.FeynmanDiagrams.Momentum
|
||||
import HepLean.PerturbationTheory.FieldStatistics
|
||||
import HepLean.PerturbationTheory.FieldStatistics.Basic
|
||||
import HepLean.PerturbationTheory.FieldStatistics.ExchangeSign
|
||||
import HepLean.PerturbationTheory.FieldStatistics.OfFinset
|
||||
import HepLean.PerturbationTheory.FieldStruct.Basic
|
||||
import HepLean.PerturbationTheory.FieldStruct.CrAnSection
|
||||
import HepLean.PerturbationTheory.FieldStruct.CreateAnnihilate
|
||||
import HepLean.PerturbationTheory.FieldStruct.CreateAnnihilateSect
|
||||
import HepLean.PerturbationTheory.Wick.CreateAnnihilateSection
|
||||
import HepLean.PerturbationTheory.Wick.KoszulOrder
|
||||
import HepLean.PerturbationTheory.Wick.OfList
|
||||
import HepLean.PerturbationTheory.Wick.OperatorMap
|
||||
import HepLean.PerturbationTheory.Wick.Signs.InsertSign
|
||||
import HepLean.PerturbationTheory.Wick.Signs.KoszulSign
|
||||
import HepLean.PerturbationTheory.Wick.Signs.KoszulSignInsert
|
||||
import HepLean.PerturbationTheory.Wick.Signs.StaticWickCoef
|
||||
import HepLean.PerturbationTheory.Wick.Signs.SuperCommuteCoef
|
||||
import HepLean.PerturbationTheory.Wick.StaticTheorem
|
||||
import HepLean.PerturbationTheory.Wick.SuperCommute
|
||||
import HepLean.PerturbationTheory.FieldStruct.Filters
|
||||
import HepLean.PerturbationTheory.FieldStruct.NormalOrder
|
||||
import HepLean.PerturbationTheory.FieldStruct.TimeOrder
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSignInsert
|
||||
import HepLean.PerturbationTheory.WickContraction.Basic
|
||||
import HepLean.PerturbationTheory.WickContraction.Erase
|
||||
import HepLean.PerturbationTheory.WickContraction.ExtractEquiv
|
||||
import HepLean.PerturbationTheory.WickContraction.Insert
|
||||
import HepLean.PerturbationTheory.WickContraction.InsertList
|
||||
import HepLean.PerturbationTheory.WickContraction.Involutions
|
||||
import HepLean.PerturbationTheory.WickContraction.IsFull
|
||||
import HepLean.PerturbationTheory.WickContraction.Sign
|
||||
import HepLean.PerturbationTheory.WickContraction.TimeContract
|
||||
import HepLean.PerturbationTheory.WickContraction.Uncontracted
|
||||
import HepLean.PerturbationTheory.WickContraction.UncontractedList
|
||||
import HepLean.PerturbationTheory.WicksTheorem
|
||||
import HepLean.SpaceTime.Basic
|
||||
import HepLean.SpaceTime.CliffordAlgebra
|
||||
import HepLean.StandardModel.Basic
|
||||
|
|
|
@ -122,6 +122,11 @@ lemma dropWile_eraseIdx {I : Type} (P : I → Prop) [DecidablePred P] :
|
|||
simpa using h (Fin.succ i) (Fin.succ j) (by simpa using hij) (by simpa using hP)
|
||||
· simp [hPa]
|
||||
|
||||
lemma insertionSort_length {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (l : List I) :
|
||||
(List.insertionSort le1 l).length = l.length := by
|
||||
apply List.Perm.length_eq
|
||||
exact List.perm_insertionSort le1 l
|
||||
|
||||
/-- The position `r0` ends up in `r` on adding it via `List.orderedInsert _ r0 r`. -/
|
||||
def orderedInsertPos {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I) (r0 : I) :
|
||||
Fin (List.orderedInsert le1 r0 r).length :=
|
||||
|
@ -265,6 +270,20 @@ lemma lt_orderedInsertPos_rel {I : Type} (le1 : I → I → Prop) [DecidableRel
|
|||
have htake' := List.mem_takeWhile_imp htake
|
||||
simpa using htake'
|
||||
|
||||
lemma lt_orderedInsertPos_rel_fin {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
|
||||
(r0 : I) (r : List I) (n : Fin (List.orderedInsert le1 r0 r).length)
|
||||
(hn : n < (orderedInsertPos le1 r r0)) : ¬ le1 r0 ((List.orderedInsert le1 r0 r).get n) := by
|
||||
have htake : (List.orderedInsert le1 r0 r).get n ∈ List.take (orderedInsertPos le1 r r0) r := by
|
||||
rw [orderedInsertPos_take_eq_orderedInsert]
|
||||
rw [List.mem_take_iff_getElem]
|
||||
use n
|
||||
simp only [List.get_eq_getElem, Fin.is_le', inf_of_le_left, Fin.val_fin_lt, exists_prop,
|
||||
and_true]
|
||||
exact hn
|
||||
rw [orderedInsertPos_take] at htake
|
||||
have htake' := List.mem_takeWhile_imp htake
|
||||
simpa using htake'
|
||||
|
||||
lemma gt_orderedInsertPos_rel {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
|
||||
[IsTotal I le1] [IsTrans I le1] (r0 : I) (r : List I) (hs : List.Sorted le1 r)
|
||||
(n : Fin r.length)
|
||||
|
@ -384,6 +403,17 @@ lemma orderedInsertEquiv_fin_succ {I : Type} (le1 : I → I → Prop) [Decidable
|
|||
rfl
|
||||
exact ne_of_beq_false rfl
|
||||
|
||||
lemma orderedInsertEquiv_monotone_fin_succ {I : Type}
|
||||
(le1 : I → I → Prop) [DecidableRel le1] (r : List I)
|
||||
(r0 : I) (n m : Fin r.length)
|
||||
(hx : orderedInsertEquiv le1 r r0 n.succ < orderedInsertEquiv le1 r r0 m.succ) :
|
||||
n < m := by
|
||||
rw [orderedInsertEquiv_fin_succ, orderedInsertEquiv_fin_succ] at hx
|
||||
rw [Fin.lt_def] at hx
|
||||
simp only [Fin.eta, Fin.coe_cast, Fin.val_fin_lt] at hx
|
||||
rw [Fin.succAbove_lt_succAbove_iff] at hx
|
||||
exact hx
|
||||
|
||||
lemma orderedInsertEquiv_congr {α : Type} {r : α → α → Prop} [DecidableRel r] (a : α)
|
||||
(l l' : List α) (h : l = l') :
|
||||
orderedInsertEquiv r l a = (Fin.castOrderIso (by simp [h])).toEquiv.trans
|
||||
|
@ -612,6 +642,38 @@ lemma insertionSort_eq_ofFn {α : Type} {r : α → α → Prop} [DecidableRel r
|
|||
rw [insertionSortEquiv_get (r := r)]
|
||||
exact Eq.symm (List.ofFn_get (List.insertionSort r l))
|
||||
|
||||
lemma insertionSortEquiv_order {α : Type} {r : α → α → Prop} [DecidableRel r] :
|
||||
(l : List α) → (i : Fin l.length) → (j : Fin l.length) → (hij : i < j)
|
||||
→ (hij' : insertionSortEquiv r l j < insertionSortEquiv r l i) →
|
||||
¬ r l[i] l[j]
|
||||
| [], i, _, _, _ => Fin.elim0 i
|
||||
| a :: as, ⟨0, hi⟩, ⟨j + 1, hj⟩, hij, hij' => by
|
||||
simp only [List.length_cons, Fin.zero_eta, Fin.getElem_fin, Fin.val_zero,
|
||||
List.getElem_cons_zero, List.getElem_cons_succ]
|
||||
nth_rewrite 2 [insertionSortEquiv] at hij'
|
||||
simp only [List.insertionSort.eq_2, List.length_cons, Nat.succ_eq_add_one, Fin.zero_eta,
|
||||
Equiv.trans_apply, equivCons_zero] at hij'
|
||||
have hx := orderedInsertEquiv_zero r (List.insertionSort r as) a
|
||||
simp only [List.length_cons, Fin.zero_eta] at hx
|
||||
rw [hx] at hij'
|
||||
have h1 := lt_orderedInsertPos_rel_fin r a (List.insertionSort r as) _ hij'
|
||||
rw [insertionSortEquiv] at h1
|
||||
simp only [Nat.succ_eq_add_one, List.insertionSort.eq_2, Equiv.trans_apply,
|
||||
equivCons_succ] at h1
|
||||
rw [← orderedInsertEquiv_get] at h1
|
||||
simp only [List.length_cons, Function.comp_apply, Equiv.symm_apply_apply, List.get_eq_getElem,
|
||||
Fin.val_succ, List.getElem_cons_succ] at h1
|
||||
rw [← List.get_eq_getElem] at h1
|
||||
rw [← insertionSortEquiv_get] at h1
|
||||
simpa using h1
|
||||
| a :: as, ⟨i + 1, hi⟩, ⟨j + 1, hj⟩, hij, hij' => by
|
||||
simp only [List.insertionSort.eq_2, List.length_cons, insertionSortEquiv, Nat.succ_eq_add_one,
|
||||
Equiv.trans_apply, equivCons_succ] at hij'
|
||||
have h1 := orderedInsertEquiv_monotone_fin_succ _ _ _ _ _ hij'
|
||||
have h2 := insertionSortEquiv_order as ⟨i, Nat.succ_lt_succ_iff.mp hi⟩
|
||||
⟨j, Nat.succ_lt_succ_iff.mp hj⟩ (by simpa using hij) h1
|
||||
simpa using h2
|
||||
|
||||
/-- Optional erase of an element in a list. For `none` returns the list, for `some i` returns
|
||||
the list with the `i`'th element erased. -/
|
||||
def optionErase {I : Type} (l : List I) (i : Option (Fin l.length)) : List I :=
|
||||
|
@ -619,12 +681,22 @@ def optionErase {I : Type} (l : List I) (i : Option (Fin l.length)) : List I :=
|
|||
| none => l
|
||||
| some i => List.eraseIdx l i
|
||||
|
||||
lemma eraseIdx_length' {I : Type} (l : List I) (i : Fin l.length) :
|
||||
(List.eraseIdx l i).length = l.length - 1 := by
|
||||
simp only [List.length_eraseIdx, Fin.is_lt, ↓reduceIte]
|
||||
|
||||
lemma eraseIdx_length {I : Type} (l : List I) (i : Fin l.length) :
|
||||
(List.eraseIdx l i).length + 1 = l.length := by
|
||||
simp only [List.length_eraseIdx, Fin.is_lt, ↓reduceIte]
|
||||
have hi := i.prop
|
||||
omega
|
||||
|
||||
lemma eraseIdx_length_succ {I : Type} (l : List I) (i : Fin l.length) :
|
||||
(List.eraseIdx l i).length.succ = l.length := by
|
||||
simp only [List.length_eraseIdx, Fin.is_lt, ↓reduceIte]
|
||||
have hi := i.prop
|
||||
omega
|
||||
|
||||
lemma eraseIdx_cons_length {I : Type} (a : I) (l : List I) (i : Fin (a :: l).length) :
|
||||
(List.eraseIdx (a :: l) i).length= l.length := by
|
||||
simp [List.length_eraseIdx]
|
||||
|
@ -650,7 +722,7 @@ lemma eraseIdx_get {I : Type} (l : List I) (i : Fin l.length) :
|
|||
lemma eraseIdx_insertionSort {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
|
||||
[IsTotal I le1] [IsTrans I le1] :
|
||||
(n : ℕ) → (r : List I) → (hn : n < r.length) →
|
||||
(List.insertionSort le1 r).eraseIdx ↑((HepLean.List.insertionSortEquiv le1 r) ⟨n, hn⟩)
|
||||
(List.insertionSort le1 r).eraseIdx ↑((insertionSortEquiv le1 r) ⟨n, hn⟩)
|
||||
= List.insertionSort le1 (r.eraseIdx n)
|
||||
| 0, [], _ => by rfl
|
||||
| 0, (r0 :: r), hn => by
|
||||
|
@ -683,6 +755,53 @@ lemma eraseIdx_insertionSort_fin {I : Type} (le1 : I → I → Prop) [DecidableR
|
|||
= List.insertionSort le1 (r.eraseIdx n) :=
|
||||
eraseIdx_insertionSort le1 n.val r (Fin.prop n)
|
||||
|
||||
/-- Given a list `i :: l` the left-most minimial position `a` of `i :: l` wrt `r`.
|
||||
That is the first position
|
||||
of `l` such that for every element `(i :: l)[b]` before that position
|
||||
`r ((i :: l)[b]) ((i :: l)[a])` is not true. The use of `i :: l` here
|
||||
rather then just `l` is to ensure that such a position exists. . -/
|
||||
def insertionSortMinPos {α : Type} (r : α → α → Prop) [DecidableRel r] (i : α) (l : List α) :
|
||||
Fin (i :: l).length := (insertionSortEquiv r (i :: l)).symm ⟨0, by
|
||||
rw [insertionSort_length]
|
||||
exact Nat.zero_lt_succ l.length⟩
|
||||
|
||||
/-- The element of `i :: l` at `insertionSortMinPos`. -/
|
||||
def insertionSortMin {α : Type} (r : α → α → Prop) [DecidableRel r] (i : α) (l : List α) :
|
||||
α := (i :: l).get (insertionSortMinPos r i l)
|
||||
|
||||
lemma insertionSortMin_eq_insertionSort_head {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
(i : α) (l : List α) :
|
||||
insertionSortMin r i l = (List.insertionSort r (i :: l)).head (by
|
||||
refine List.ne_nil_of_length_pos ?_
|
||||
rw [insertionSort_length]
|
||||
exact Nat.zero_lt_succ l.length) := by
|
||||
trans (List.insertionSort r (i :: l)).get (⟨0, by
|
||||
rw [insertionSort_length]; exact Nat.zero_lt_succ l.length⟩)
|
||||
· rw [← insertionSortEquiv_get]
|
||||
rfl
|
||||
· exact List.get_mk_zero
|
||||
(Eq.mpr (id (congrArg (fun _a => 0 < _a) (insertionSort_length r (i :: l))))
|
||||
(Nat.zero_lt_succ l.length))
|
||||
|
||||
/-- The list remaining after dropping the element at the position determined by
|
||||
`insertionSortMinPos`. -/
|
||||
def insertionSortDropMinPos {α : Type} (r : α → α → Prop) [DecidableRel r] (i : α) (l : List α) :
|
||||
List α := (i :: l).eraseIdx (insertionSortMinPos r i l)
|
||||
|
||||
lemma insertionSort_eq_insertionSortMin_cons {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] (i : α) (l : List α) :
|
||||
List.insertionSort r (i :: l) =
|
||||
(insertionSortMin r i l) :: List.insertionSort r (insertionSortDropMinPos r i l) := by
|
||||
rw [insertionSortDropMinPos, ← eraseIdx_insertionSort_fin]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rhs
|
||||
rw [insertionSortMinPos]
|
||||
rw [Equiv.apply_symm_apply]
|
||||
simp only [List.insertionSort, List.eraseIdx_zero]
|
||||
rw [insertionSortMin_eq_insertionSort_head]
|
||||
exact (List.head_cons_tail _ _).symm
|
||||
|
||||
/-- Optional erase of an element in a list, with addition for `none`. For `none` adds `a` to the
|
||||
front of the list, for `some i` removes the `i`th element of the list (does not add `a`).
|
||||
E.g. `optionEraseZ [0, 1, 2] 4 none = [4, 0, 1, 2]` and
|
||||
|
@ -707,4 +826,28 @@ lemma optionEraseZ_ext {I : Type} {l l' : List I} {a a' : I} {i : Option (Fin l.
|
|||
congr
|
||||
simp
|
||||
|
||||
lemma mem_take_finrange : (n m : ℕ) → (a : Fin n) → a ∈ List.take m (List.finRange n) ↔ a.val < m
|
||||
| 0, m, a => Fin.elim0 a
|
||||
| n+1, 0, a => by
|
||||
simp
|
||||
| n +1, m + 1, ⟨0, h⟩ => by
|
||||
simp [List.finRange_succ]
|
||||
| n +1, m + 1, ⟨i + 1, h⟩ => by
|
||||
simp only [List.finRange_succ, List.take_succ_cons, List.mem_cons, Fin.ext_iff, Fin.val_zero,
|
||||
AddLeftCancelMonoid.add_eq_zero, one_ne_zero, and_false, false_or, add_lt_add_iff_right]
|
||||
rw [← List.map_take]
|
||||
rw [@List.mem_map]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
obtain ⟨a, ha⟩ := h
|
||||
rw [mem_take_finrange n m a] at ha
|
||||
rw [Fin.ext_iff] at ha
|
||||
simp_all only [Fin.val_succ, add_left_inj]
|
||||
omega
|
||||
· intro h1
|
||||
use ⟨i, Nat.succ_lt_succ_iff.mp h⟩
|
||||
simp only [Fin.succ_mk, and_true]
|
||||
rw [mem_take_finrange n m ⟨i, Nat.succ_lt_succ_iff.mp h⟩]
|
||||
exact h1
|
||||
|
||||
end HepLean.List
|
||||
|
|
182
HepLean/Mathematics/List/InsertIdx.lean
Normal file
182
HepLean/Mathematics/List/InsertIdx.lean
Normal file
|
@ -0,0 +1,182 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.Mathematics.List
|
||||
/-!
|
||||
# List lemmas
|
||||
|
||||
-/
|
||||
namespace HepLean.List
|
||||
|
||||
open Fin
|
||||
open HepLean
|
||||
variable {n : Nat}
|
||||
|
||||
lemma insertIdx_map {I J : Type} (f : I → J) : (i : ℕ) → (r : List I) → (r0 : I) →
|
||||
(List.insertIdx i r0 r).map f = List.insertIdx i (f r0) (r.map f)
|
||||
| 0, [], r0 => by simp
|
||||
| n+1, [], r0 => by simp
|
||||
| 0, a::as, r0 => by simp
|
||||
| n+1, a::as, r0 => by
|
||||
simp only [List.insertIdx_succ_cons, List.map_cons, List.cons.injEq, true_and]
|
||||
exact insertIdx_map f n as r0
|
||||
|
||||
lemma eraseIdx_sorted {I : Type} (le : I → I → Prop) :
|
||||
(r : List I) → (n : ℕ) →
|
||||
List.Sorted le r → List.Sorted le (r.eraseIdx n)
|
||||
| [], _, _ => by simp
|
||||
| a::as, 0, h => by
|
||||
simp only [List.eraseIdx]
|
||||
simp only [List.sorted_cons] at h
|
||||
exact h.2
|
||||
| a::as, n+1, h => by
|
||||
simp only [List.eraseIdx_cons_succ, List.sorted_cons]
|
||||
simp only [List.sorted_cons] at h
|
||||
refine And.intro ?_ (eraseIdx_sorted le as n h.2)
|
||||
intro b hb
|
||||
refine h.1 _ ?_
|
||||
exact List.mem_of_mem_eraseIdx hb
|
||||
|
||||
lemma mem_eraseIdx_nodup {I : Type} (i : I) :
|
||||
(l : List I) → (n : ℕ) → (hn : n < l.length) → (h : List.Nodup l) →
|
||||
i ∈ l.eraseIdx n ↔ i ∈ l ∧ i ≠ l[n]
|
||||
| [], _, _, _ => by simp
|
||||
| a1 :: as, 0, _, h => by
|
||||
simp only [List.eraseIdx_zero, List.tail_cons, List.mem_cons, List.getElem_cons_zero, ne_eq]
|
||||
by_cases hi : i = a1
|
||||
· subst hi
|
||||
simp only [List.nodup_cons] at h
|
||||
simp [h]
|
||||
· simp [hi]
|
||||
| a1 :: as, n+1, hn, h => by
|
||||
simp only [List.eraseIdx_cons_succ, List.mem_cons, List.getElem_cons_succ, ne_eq]
|
||||
simp only [List.nodup_cons] at h
|
||||
rw [mem_eraseIdx_nodup i as n (Nat.succ_lt_succ_iff.mp hn) h.2]
|
||||
simp_all only [ne_eq]
|
||||
obtain ⟨left, right⟩ := h
|
||||
apply Iff.intro
|
||||
· intro a
|
||||
cases a with
|
||||
| inl h =>
|
||||
subst h
|
||||
simp_all only [or_false, true_and]
|
||||
apply Aesop.BuiltinRules.not_intro
|
||||
intro a
|
||||
simp_all only [List.getElem_mem, not_true_eq_false]
|
||||
| inr h_1 => simp_all only [or_true, not_false_eq_true, and_self]
|
||||
· intro a
|
||||
simp_all only [not_false_eq_true, and_true]
|
||||
|
||||
lemma insertIdx_eq_take_drop {I : Type} (i : I) : (r : List I) → (n : Fin r.length.succ) →
|
||||
List.insertIdx n i r = List.take n r ++ i :: r.drop n
|
||||
| [], 0 => by simp
|
||||
| a :: as, 0 => by
|
||||
simp
|
||||
| a :: as, ⟨n + 1, h⟩ => by
|
||||
simp only [List.insertIdx_succ_cons, List.take_succ_cons, List.drop_succ_cons, List.cons_append,
|
||||
List.cons.injEq, true_and]
|
||||
exact insertIdx_eq_take_drop i as ⟨n, Nat.succ_lt_succ_iff.mp h⟩
|
||||
|
||||
@[simp]
|
||||
lemma insertIdx_length_fin {I : Type} (i : I) :
|
||||
(r : List I) → (n : Fin r.length.succ) →
|
||||
(List.insertIdx n i r).length = r.length.succ
|
||||
| [], 0 => by simp
|
||||
| a :: as, 0 => by simp
|
||||
| a :: as, ⟨n + 1, h⟩ => by
|
||||
simp only [List.insertIdx_succ_cons, List.length_cons, Nat.succ_eq_add_one, add_left_inj]
|
||||
exact insertIdx_length_fin i as ⟨n, Nat.succ_lt_succ_iff.mp h⟩
|
||||
|
||||
@[simp]
|
||||
lemma insertIdx_getElem_fin {I : Type} (i : I) :
|
||||
(r : List I) → (k : Fin r.length.succ) → (m : Fin r.length) →
|
||||
(List.insertIdx k i r)[(k.succAbove m).val] = r[m.val]
|
||||
| [], 0, m => by exact Fin.elim0 m
|
||||
| a :: as, 0, m => by simp
|
||||
| a :: as, ⟨n + 1, h⟩, ⟨0, h0⟩ => by
|
||||
simp [Fin.succAbove, Fin.lt_def]
|
||||
| a :: as, ⟨n + 1, h⟩, ⟨m+1, hm⟩ => by
|
||||
simp only [List.insertIdx_succ_cons, List.length_cons, Nat.succ_eq_add_one,
|
||||
List.getElem_cons_succ]
|
||||
conv_rhs => rw [← insertIdx_getElem_fin i as ⟨n, Nat.succ_lt_succ_iff.mp h⟩
|
||||
⟨m, Nat.lt_of_succ_lt_succ hm⟩]
|
||||
simp only [Fin.succAbove, Fin.castSucc_mk, Fin.lt_def, add_lt_add_iff_right, Fin.succ_mk,
|
||||
Nat.succ_eq_add_one]
|
||||
split
|
||||
· simp_all only [List.getElem_cons_succ]
|
||||
· simp_all only [List.getElem_cons_succ]
|
||||
|
||||
lemma insertIdx_eraseIdx_fin {I : Type} :
|
||||
(r : List I) → (k : Fin r.length) →
|
||||
(List.eraseIdx r k).insertIdx k r[k] = r
|
||||
| [], k => by exact Fin.elim0 k
|
||||
| a :: as, ⟨0, h⟩ => by simp
|
||||
| a :: as, ⟨n + 1, h⟩ => by
|
||||
simp only [List.length_cons, Fin.getElem_fin, List.getElem_cons_succ, List.eraseIdx_cons_succ,
|
||||
List.insertIdx_succ_cons, List.cons.injEq, true_and]
|
||||
exact insertIdx_eraseIdx_fin as ⟨n, Nat.lt_of_succ_lt_succ h⟩
|
||||
|
||||
lemma get_eq_insertIdx_succAbove {I : Type} (i : I) (r : List I) (k : Fin r.length.succ) :
|
||||
r.get = (List.insertIdx k i r).get ∘
|
||||
(finCongr (insertIdx_length_fin i r k).symm) ∘ k.succAbove := by
|
||||
funext i
|
||||
simp
|
||||
|
||||
lemma take_insert_same {I : Type} (i : I) :
|
||||
(n : ℕ) → (r : List I) →
|
||||
List.take n (List.insertIdx n i r) = List.take n r
|
||||
| 0, _ => by simp
|
||||
| _+1, [] => by simp
|
||||
| n+1, a::as => by
|
||||
simp only [List.insertIdx_succ_cons, List.take_succ_cons, List.cons.injEq, true_and]
|
||||
exact take_insert_same i n as
|
||||
|
||||
lemma take_eraseIdx_same {I : Type} :
|
||||
(n : ℕ) → (r : List I) →
|
||||
List.take n (List.eraseIdx r n) = List.take n r
|
||||
| 0, _ => by simp
|
||||
| _+1, [] => by simp
|
||||
| n+1, a::as => by
|
||||
simp only [List.eraseIdx_cons_succ, List.take_succ_cons, List.cons.injEq, true_and]
|
||||
exact take_eraseIdx_same n as
|
||||
|
||||
lemma drop_eraseIdx_succ {I : Type} :
|
||||
(n : ℕ) → (r : List I) → (hn : n < r.length) →
|
||||
r[n] :: List.drop n (List.eraseIdx r n) = List.drop n r
|
||||
| 0, _, _=> by
|
||||
simp only [List.eraseIdx_zero, List.drop_tail, zero_add, List.drop_one, List.drop_zero]
|
||||
rw [@List.getElem_zero]
|
||||
exact List.head_cons_tail _ _
|
||||
| n+1, [], hn => by simp at hn
|
||||
| n+1, a::as, hn => by
|
||||
simp only [List.getElem_cons_succ, List.eraseIdx_cons_succ, List.drop_succ_cons]
|
||||
refine drop_eraseIdx_succ n as _
|
||||
|
||||
lemma take_insert_gt {I : Type} (i : I) :
|
||||
(n m : ℕ) → (h : n < m) → (r : List I) →
|
||||
List.take n (List.insertIdx m i r) = List.take n r
|
||||
| 0, 0, _, _ => by simp
|
||||
| 0, m + 1, _, _ => by simp
|
||||
| n+1, m + 1, _, [] => by simp
|
||||
| n+1, m + 1, h, a::as => by
|
||||
simp only [List.insertIdx_succ_cons, List.take_succ_cons, List.cons.injEq, true_and]
|
||||
refine take_insert_gt i n m (Nat.succ_lt_succ_iff.mp h) as
|
||||
|
||||
lemma take_insert_let {I : Type} (i : I) :
|
||||
(n m : ℕ) → (h : m ≤ n) → (r : List I) → (hm : m ≤ r.length) →
|
||||
(List.take (n + 1) (List.insertIdx m i r)).Perm (i :: List.take n r)
|
||||
| 0, 0, h, _, _ => by simp
|
||||
| m + 1, 0, h, r, _ => by simp
|
||||
| n + 1, m + 1, h, [], hm => by
|
||||
simp at hm
|
||||
| n + 1, m + 1, h, a::as, hm => by
|
||||
simp only [List.insertIdx_succ_cons, List.take_succ_cons]
|
||||
have hp : (i :: a :: List.take n as).Perm (a :: i :: List.take n as) := by
|
||||
exact List.Perm.swap a i (List.take n as)
|
||||
refine List.Perm.trans ?_ hp.symm
|
||||
refine List.Perm.cons a ?_
|
||||
exact take_insert_let i n m (Nat.le_of_succ_le_succ h) as (Nat.le_of_succ_le_succ hm)
|
||||
|
||||
end HepLean.List
|
94
HepLean/Mathematics/List/InsertionSort.lean
Normal file
94
HepLean/Mathematics/List/InsertionSort.lean
Normal file
|
@ -0,0 +1,94 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.Mathematics.List.InsertIdx
|
||||
/-!
|
||||
# List lemmas
|
||||
|
||||
-/
|
||||
namespace HepLean.List
|
||||
|
||||
open Fin
|
||||
open HepLean
|
||||
variable {n : Nat}
|
||||
|
||||
lemma insertionSortMin_lt_length_succ {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
(i : α) (l : List α) :
|
||||
insertionSortMinPos r i l < (insertionSortDropMinPos r i l).length.succ := by
|
||||
rw [insertionSortMinPos]
|
||||
simp only [List.length_cons, List.insertionSort.eq_2, insertionSortDropMinPos,
|
||||
Nat.succ_eq_add_one]
|
||||
rw [eraseIdx_length']
|
||||
simp
|
||||
|
||||
/-- Given a list `i :: l` the left-most minimial position `a` of `i :: l` wrt `r`
|
||||
as an element of `Fin (insertionSortDropMinPos r i l).length.succ`. -/
|
||||
def insertionSortMinPosFin {α : Type} (r : α → α → Prop) [DecidableRel r] (i : α) (l : List α) :
|
||||
Fin (insertionSortDropMinPos r i l).length.succ :=
|
||||
⟨insertionSortMinPos r i l, insertionSortMin_lt_length_succ r i l⟩
|
||||
|
||||
lemma insertionSortMin_lt_mem_insertionSortDropMinPos {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] (a : α) (l : List α)
|
||||
(i : Fin (insertionSortDropMinPos r a l).length) :
|
||||
r (insertionSortMin r a l) ((insertionSortDropMinPos r a l)[i]) := by
|
||||
let l1 := List.insertionSort r (a :: l)
|
||||
have hl1 : l1.Sorted r := by exact List.sorted_insertionSort r (a :: l)
|
||||
simp only [l1] at hl1
|
||||
rw [insertionSort_eq_insertionSortMin_cons r a l] at hl1
|
||||
simp only [List.sorted_cons, List.mem_insertionSort] at hl1
|
||||
apply hl1.1 ((insertionSortDropMinPos r a l)[i])
|
||||
simp
|
||||
|
||||
lemma insertionSortMinPos_insertionSortEquiv {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
(a : α) (l : List α) :
|
||||
insertionSortEquiv r (a ::l) (insertionSortMinPos r a l) =
|
||||
⟨0, by simp [List.orderedInsert_length]⟩ := by
|
||||
rw [insertionSortMinPos]
|
||||
exact
|
||||
Equiv.apply_symm_apply (insertionSortEquiv r (a :: l)) ⟨0, insertionSortMinPos.proof_1 r a l⟩
|
||||
|
||||
lemma insertionSortEquiv_gt_zero_of_ne_insertionSortMinPos {α : Type} (r : α → α → Prop)
|
||||
[DecidableRel r] (a : α) (l : List α) (k : Fin (a :: l).length)
|
||||
(hk : k ≠ insertionSortMinPos r a l) :
|
||||
⟨0, by simp [List.orderedInsert_length]⟩ < insertionSortEquiv r (a :: l) k := by
|
||||
by_contra hn
|
||||
simp only [List.insertionSort.eq_2, List.length_cons, not_lt] at hn
|
||||
refine hk ((Equiv.apply_eq_iff_eq_symm_apply (insertionSortEquiv r (a :: l))).mp ?_)
|
||||
simp_all only [List.length_cons, ne_eq, Fin.le_def, nonpos_iff_eq_zero, List.insertionSort.eq_2]
|
||||
simp only [Fin.ext_iff]
|
||||
omega
|
||||
|
||||
lemma insertionSortMin_lt_mem_insertionSortDropMinPos_of_lt {α : Type} (r : α → α → Prop)
|
||||
[DecidableRel r] (a : α) (l : List α)
|
||||
(i : Fin (insertionSortDropMinPos r a l).length)
|
||||
(h : (insertionSortMinPosFin r a l).succAbove i < insertionSortMinPosFin r a l) :
|
||||
¬ r ((insertionSortDropMinPos r a l)[i]) (insertionSortMin r a l) := by
|
||||
simp only [Fin.getElem_fin, insertionSortMin, List.get_eq_getElem, List.length_cons]
|
||||
have h1 : (insertionSortDropMinPos r a l)[i] =
|
||||
(a :: l).get (finCongr (eraseIdx_length_succ (a :: l) (insertionSortMinPos r a l))
|
||||
((insertionSortMinPosFin r a l).succAbove i)) := by
|
||||
trans (insertionSortDropMinPos r a l).get i
|
||||
simp only [Fin.getElem_fin, List.get_eq_getElem]
|
||||
simp only [insertionSortDropMinPos, List.length_cons, Nat.succ_eq_add_one,
|
||||
finCongr_apply, Fin.coe_cast]
|
||||
rw [eraseIdx_get]
|
||||
simp only [List.length_cons, Function.comp_apply, List.get_eq_getElem, Fin.coe_cast]
|
||||
rfl
|
||||
erw [h1]
|
||||
simp only [List.length_cons, Nat.succ_eq_add_one, List.get_eq_getElem,
|
||||
Fin.coe_cast]
|
||||
apply insertionSortEquiv_order
|
||||
simpa using h
|
||||
simp only [List.insertionSort.eq_2, List.length_cons, finCongr_apply]
|
||||
apply lt_of_eq_of_lt (insertionSortMinPos_insertionSortEquiv r a l)
|
||||
simp only [List.insertionSort.eq_2]
|
||||
apply insertionSortEquiv_gt_zero_of_ne_insertionSortMinPos r a l
|
||||
simp only [List.length_cons, ne_eq, Fin.ext_iff, Fin.coe_cast]
|
||||
have hl : (insertionSortMinPos r a l).val = (insertionSortMinPosFin r a l).val := by
|
||||
rfl
|
||||
simp only [hl, Nat.succ_eq_add_one, Fin.val_eq_val, ne_eq]
|
||||
exact Fin.succAbove_ne (insertionSortMinPosFin r a l) i
|
||||
|
||||
end HepLean.List
|
266
HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Basic.lean
Normal file
266
HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Basic.lean
Normal file
|
@ -0,0 +1,266 @@
|
|||
/-
|
||||
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.Algebras.StateAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.FieldStruct.CrAnSection
|
||||
/-!
|
||||
|
||||
# Creation and annihlation free-algebra
|
||||
|
||||
This module defines the creation and annihilation algebra for a field structure.
|
||||
|
||||
The creation and annihilation algebra extends from the state algebra by adding information about
|
||||
whether a state is a creation or annihilation operator.
|
||||
|
||||
The algebra is spanned by lists of creation/annihilation states.
|
||||
|
||||
The main structures defined in this module are:
|
||||
|
||||
* `CrAnAlgebra` - The creation and annihilation algebra
|
||||
* `ofCrAnState` - Maps a creation/annihilation state to the algebra
|
||||
* `ofCrAnList` - Maps a list of creation/annihilation states to the algebra
|
||||
* `ofState` - Maps a state to a sum of creation and annihilation operators
|
||||
* `crPart` - The creation part of a state in the algebra
|
||||
* `anPart` - The annihilation part of a state in the algebra
|
||||
* `superCommute` - The super commutator on the algebra
|
||||
|
||||
The key lemmas show how these operators interact, particularly focusing on the
|
||||
super commutation relations between creation and annihilation operators.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
/-- The creation and annihlation free-algebra.
|
||||
The free algebra generated by `CrAnStates`,
|
||||
that is a position based states or assymptotic states with a specification of
|
||||
whether the state is a creation or annihlation state.
|
||||
As a module `CrAnAlgebra` is spanned by lists of `CrAnStates`. -/
|
||||
abbrev CrAnAlgebra (𝓕 : FieldStruct) : Type := FreeAlgebra ℂ 𝓕.CrAnStates
|
||||
|
||||
namespace CrAnAlgebra
|
||||
|
||||
open StateAlgebra
|
||||
|
||||
/-- Maps a creation and annihlation state to the creation and annihlation free-algebra. -/
|
||||
def ofCrAnState (φ : 𝓕.CrAnStates) : CrAnAlgebra 𝓕 :=
|
||||
FreeAlgebra.ι ℂ φ
|
||||
|
||||
/-- Maps a list creation and annihlation state to the creation and annihlation free-algebra
|
||||
by taking their product. -/
|
||||
def ofCrAnList (φs : List 𝓕.CrAnStates) : CrAnAlgebra 𝓕 := (List.map ofCrAnState φs).prod
|
||||
|
||||
@[simp]
|
||||
lemma ofCrAnList_nil : ofCrAnList ([] : List 𝓕.CrAnStates) = 1 := rfl
|
||||
|
||||
lemma ofCrAnList_cons (φ : 𝓕.CrAnStates) (φs : List 𝓕.CrAnStates) :
|
||||
ofCrAnList (φ :: φs) = ofCrAnState φ * ofCrAnList φs := rfl
|
||||
|
||||
lemma ofCrAnList_append (φs φs' : List 𝓕.CrAnStates) :
|
||||
ofCrAnList (φs ++ φs') = ofCrAnList φs * ofCrAnList φs' := by
|
||||
dsimp only [ofCrAnList]
|
||||
rw [List.map_append, List.prod_append]
|
||||
|
||||
lemma ofCrAnList_singleton (φ : 𝓕.CrAnStates) :
|
||||
ofCrAnList [φ] = ofCrAnState φ := by
|
||||
simp [ofCrAnList]
|
||||
|
||||
/-- Maps a state to the sum of creation and annihilation operators in
|
||||
creation and annihilation free-algebra. -/
|
||||
def ofState (φ : 𝓕.States) : CrAnAlgebra 𝓕 :=
|
||||
∑ (i : 𝓕.statesToCrAnType φ), ofCrAnState ⟨φ, i⟩
|
||||
|
||||
/-- The algebra map from the state free-algebra to the creation and annihlation free-algebra. -/
|
||||
def ofStateAlgebra : StateAlgebra 𝓕 →ₐ[ℂ] CrAnAlgebra 𝓕 :=
|
||||
FreeAlgebra.lift ℂ ofState
|
||||
|
||||
@[simp]
|
||||
lemma ofStateAlgebra_ofState (φ : 𝓕.States) :
|
||||
ofStateAlgebra (StateAlgebra.ofState φ) = ofState φ := by
|
||||
simp [ofStateAlgebra, StateAlgebra.ofState]
|
||||
|
||||
/-- Maps a list of states to the creation and annihilation free-algebra by taking
|
||||
the product of their sums of creation and annihlation operators.
|
||||
Roughly `[φ1, φ2]` gets sent to `(φ1ᶜ+ φ1ᵃ) * (φ2ᶜ+ φ2ᵃ)` etc. -/
|
||||
def ofStateList (φs : List 𝓕.States) : CrAnAlgebra 𝓕 := (List.map ofState φs).prod
|
||||
|
||||
@[simp]
|
||||
lemma ofStateList_nil : ofStateList ([] : List 𝓕.States) = 1 := rfl
|
||||
|
||||
lemma ofStateList_cons (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
ofStateList (φ :: φs) = ofState φ * ofStateList φs := rfl
|
||||
|
||||
lemma ofStateList_singleton (φ : 𝓕.States) :
|
||||
ofStateList [φ] = ofState φ := by
|
||||
simp [ofStateList]
|
||||
|
||||
lemma ofStateList_append (φs φs' : List 𝓕.States) :
|
||||
ofStateList (φs ++ φs') = ofStateList φs * ofStateList φs' := by
|
||||
dsimp only [ofStateList]
|
||||
rw [List.map_append, List.prod_append]
|
||||
|
||||
lemma ofStateAlgebra_ofList_eq_ofStateList : (φs : List 𝓕.States) →
|
||||
ofStateAlgebra (ofList φs) = ofStateList φs
|
||||
| [] => by
|
||||
simp [ofStateList]
|
||||
| φ :: φs => by
|
||||
rw [ofStateList_cons, StateAlgebra.ofList_cons]
|
||||
rw [map_mul]
|
||||
simp only [ofStateAlgebra_ofState, mul_eq_mul_left_iff]
|
||||
apply Or.inl (ofStateAlgebra_ofList_eq_ofStateList φs)
|
||||
|
||||
lemma ofStateList_sum (φs : List 𝓕.States) :
|
||||
ofStateList φs = ∑ (s : CrAnSection φs), ofCrAnList s.1 := by
|
||||
induction φs with
|
||||
| nil => simp
|
||||
| cons φ φs ih =>
|
||||
rw [CrAnSection.sum_cons]
|
||||
dsimp only [CrAnSection.cons, ofCrAnList_cons]
|
||||
conv_rhs =>
|
||||
enter [2, x]
|
||||
rw [← Finset.mul_sum]
|
||||
rw [← Finset.sum_mul, ofStateList_cons, ← ih]
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Creation and annihilation parts of a state
|
||||
|
||||
-/
|
||||
|
||||
/-- The algebra map taking an element of the free-state algbra to
|
||||
the part of it in the creation and annihlation free algebra
|
||||
spanned by creation operators. -/
|
||||
def crPart : 𝓕.StateAlgebra →ₐ[ℂ] 𝓕.CrAnAlgebra :=
|
||||
FreeAlgebra.lift ℂ fun φ =>
|
||||
match φ with
|
||||
| States.negAsymp φ => ofCrAnState ⟨States.negAsymp φ, ()⟩
|
||||
| States.position φ => ofCrAnState ⟨States.position φ, CreateAnnihilate.create⟩
|
||||
| States.posAsymp _ => 0
|
||||
|
||||
@[simp]
|
||||
lemma crPart_negAsymp (φ : 𝓕.AsymptoticNegTime) :
|
||||
crPart (StateAlgebra.ofState (States.negAsymp φ)) = ofCrAnState ⟨States.negAsymp φ, ()⟩ := by
|
||||
dsimp only [crPart, StateAlgebra.ofState]
|
||||
rw [FreeAlgebra.lift_ι_apply]
|
||||
|
||||
@[simp]
|
||||
lemma crPart_position (φ : 𝓕.PositionStates) :
|
||||
crPart (StateAlgebra.ofState (States.position φ)) =
|
||||
ofCrAnState ⟨States.position φ, CreateAnnihilate.create⟩ := by
|
||||
dsimp only [crPart, StateAlgebra.ofState]
|
||||
rw [FreeAlgebra.lift_ι_apply]
|
||||
|
||||
@[simp]
|
||||
lemma crPart_posAsymp (φ : 𝓕.AsymptoticPosTime) :
|
||||
crPart (StateAlgebra.ofState (States.posAsymp φ)) = 0 := by
|
||||
dsimp only [crPart, StateAlgebra.ofState]
|
||||
rw [FreeAlgebra.lift_ι_apply]
|
||||
|
||||
/-- The algebra map taking an element of the free-state algbra to
|
||||
the part of it in the creation and annihilation free algebra
|
||||
spanned by annihilation operators. -/
|
||||
def anPart : 𝓕.StateAlgebra →ₐ[ℂ] 𝓕.CrAnAlgebra :=
|
||||
FreeAlgebra.lift ℂ fun φ =>
|
||||
match φ with
|
||||
| States.negAsymp _ => 0
|
||||
| States.position φ => ofCrAnState ⟨States.position φ, CreateAnnihilate.annihilate⟩
|
||||
| States.posAsymp φ => ofCrAnState ⟨States.posAsymp φ, ()⟩
|
||||
|
||||
@[simp]
|
||||
lemma anPart_negAsymp (φ : 𝓕.AsymptoticNegTime) :
|
||||
anPart (StateAlgebra.ofState (States.negAsymp φ)) = 0 := by
|
||||
dsimp only [anPart, StateAlgebra.ofState]
|
||||
rw [FreeAlgebra.lift_ι_apply]
|
||||
|
||||
@[simp]
|
||||
lemma anPart_position (φ : 𝓕.PositionStates) :
|
||||
anPart (StateAlgebra.ofState (States.position φ)) =
|
||||
ofCrAnState ⟨States.position φ, CreateAnnihilate.annihilate⟩ := by
|
||||
dsimp only [anPart, StateAlgebra.ofState]
|
||||
rw [FreeAlgebra.lift_ι_apply]
|
||||
|
||||
@[simp]
|
||||
lemma anPart_posAsymp (φ : 𝓕.AsymptoticPosTime) :
|
||||
anPart (StateAlgebra.ofState (States.posAsymp φ)) = ofCrAnState ⟨States.posAsymp φ, ()⟩ := by
|
||||
dsimp only [anPart, StateAlgebra.ofState]
|
||||
rw [FreeAlgebra.lift_ι_apply]
|
||||
|
||||
lemma ofState_eq_crPart_add_anPart (φ : 𝓕.States) :
|
||||
ofState φ = crPart (StateAlgebra.ofState φ) + anPart (StateAlgebra.ofState φ) := by
|
||||
rw [ofState]
|
||||
cases φ with
|
||||
| negAsymp φ =>
|
||||
dsimp only [statesToCrAnType]
|
||||
simp
|
||||
| position φ =>
|
||||
dsimp only [statesToCrAnType]
|
||||
rw [CreateAnnihilate.sum_eq]
|
||||
simp
|
||||
| posAsymp φ =>
|
||||
dsimp only [statesToCrAnType]
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## The basis of the creation and annihlation free-algebra.
|
||||
|
||||
-/
|
||||
|
||||
/-- The basis of the free creation and annihilation algebra formed by lists of CrAnStates. -/
|
||||
noncomputable def ofCrAnListBasis : Basis (List 𝓕.CrAnStates) ℂ (CrAnAlgebra 𝓕) where
|
||||
repr := FreeAlgebra.equivMonoidAlgebraFreeMonoid.toLinearEquiv
|
||||
|
||||
@[simp]
|
||||
lemma ofListBasis_eq_ofList (φs : List 𝓕.CrAnStates) :
|
||||
ofCrAnListBasis φs = ofCrAnList φs := by
|
||||
simp only [ofCrAnListBasis, FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
Basis.coe_ofRepr, AlgEquiv.toLinearEquiv_symm, AlgEquiv.toLinearEquiv_apply,
|
||||
AlgEquiv.ofAlgHom_symm_apply, ofCrAnList]
|
||||
erw [MonoidAlgebra.lift_apply]
|
||||
simp only [zero_smul, Finsupp.sum_single_index, one_smul]
|
||||
rw [@FreeMonoid.lift_apply]
|
||||
simp only [List.prod]
|
||||
match φs with
|
||||
| [] => rfl
|
||||
| φ :: φs =>
|
||||
erw [List.map_cons]
|
||||
|
||||
/-!
|
||||
|
||||
## Some useful multi-linear maps.
|
||||
|
||||
-/
|
||||
|
||||
/-- The bi-linear map associated with multiplication in `CrAnAlgebra`. -/
|
||||
noncomputable def mulLinearMap : CrAnAlgebra 𝓕 →ₗ[ℂ] CrAnAlgebra 𝓕 →ₗ[ℂ] CrAnAlgebra 𝓕 where
|
||||
toFun a := {
|
||||
toFun := fun b => a * b,
|
||||
map_add' := fun c d => by
|
||||
simp [mul_add]
|
||||
map_smul' := by simp}
|
||||
map_add' := fun a b => by
|
||||
ext c
|
||||
simp [add_mul]
|
||||
map_smul' := by
|
||||
intros
|
||||
ext c
|
||||
simp [smul_mul']
|
||||
|
||||
lemma mulLinearMap_apply (a b : CrAnAlgebra 𝓕) :
|
||||
mulLinearMap a b = a * b := by rfl
|
||||
|
||||
/-- The linear map associated with scalar-multiplication in `CrAnAlgebra`. -/
|
||||
noncomputable def smulLinearMap (c : ℂ) : CrAnAlgebra 𝓕 →ₗ[ℂ] CrAnAlgebra 𝓕 where
|
||||
toFun a := c • a
|
||||
map_add' := by simp
|
||||
map_smul' m x := by
|
||||
simp only [smul_smul, RingHom.id_apply]
|
||||
rw [NonUnitalNormedCommRing.mul_comm]
|
||||
|
||||
end CrAnAlgebra
|
||||
|
||||
end FieldStruct
|
555
HepLean/PerturbationTheory/Algebras/CrAnAlgebra/NormalOrder.lean
Normal file
555
HepLean/PerturbationTheory/Algebras/CrAnAlgebra/NormalOrder.lean
Normal file
|
@ -0,0 +1,555 @@
|
|||
/-
|
||||
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.FieldStruct.NormalOrder
|
||||
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.SuperCommute
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
/-!
|
||||
|
||||
# Normal Ordering
|
||||
|
||||
The normal ordering puts all creation operators to the left and all annihilation operators to the
|
||||
right. It acts on `CrAnStates` and defines a linear map from the `CrAnAlgebra` to itself.
|
||||
|
||||
The normal ordering satisfies a number of nice properties with relation to the operator
|
||||
algebra 𝓞.A.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
open FieldStatistic
|
||||
/-!
|
||||
|
||||
## Normal order on the CrAnAlgebra
|
||||
|
||||
-/
|
||||
namespace CrAnAlgebra
|
||||
|
||||
noncomputable section
|
||||
|
||||
/-- The linear map on the free creation and annihlation
|
||||
algebra defined as the map taking
|
||||
a list of CrAnStates to the normal-ordered list of states multiplied by
|
||||
the sign corresponding to the number of fermionic-fermionic
|
||||
exchanges done in ordering. -/
|
||||
def normalOrder : CrAnAlgebra 𝓕 →ₗ[ℂ] CrAnAlgebra 𝓕 :=
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs =>
|
||||
normalOrderSign φs • ofCrAnList (normalOrderList φs)
|
||||
|
||||
lemma normalOrder_ofCrAnList (φs : List 𝓕.CrAnStates) :
|
||||
normalOrder (ofCrAnList φs) = normalOrderSign φs • ofCrAnList (normalOrderList φs) := by
|
||||
rw [← ofListBasis_eq_ofList]
|
||||
simp only [normalOrder, Basis.constr_basis]
|
||||
|
||||
lemma ofCrAnList_eq_normalOrder (φs : List 𝓕.CrAnStates) :
|
||||
ofCrAnList (normalOrderList φs) = normalOrderSign φs • normalOrder (ofCrAnList φs) := by
|
||||
rw [normalOrder_ofCrAnList, normalOrderList]
|
||||
rw [smul_smul]
|
||||
simp only [normalOrderSign]
|
||||
rw [Wick.koszulSign_mul_self]
|
||||
simp
|
||||
|
||||
lemma normalOrder_one : normalOrder (𝓕 := 𝓕) 1 = 1 := by
|
||||
rw [← ofCrAnList_nil, normalOrder_ofCrAnList]
|
||||
simp
|
||||
|
||||
lemma normalOrder_ofCrAnList_cons_create (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create) (φs : List 𝓕.CrAnStates) :
|
||||
normalOrder (ofCrAnList (φ :: φs)) =
|
||||
ofCrAnState φ * normalOrder (ofCrAnList φs) := by
|
||||
rw [normalOrder_ofCrAnList]
|
||||
rw [normalOrderSign_cons_create φ hφ, normalOrderList_cons_create φ hφ φs]
|
||||
rw [ofCrAnList_cons, normalOrder_ofCrAnList, mul_smul_comm]
|
||||
|
||||
lemma normalOrder_create_mul (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create)
|
||||
(a : CrAnAlgebra 𝓕) :
|
||||
normalOrder (ofCrAnState φ * a) = ofCrAnState φ * normalOrder a := by
|
||||
change (normalOrder ∘ₗ mulLinearMap (ofCrAnState φ)) a =
|
||||
(mulLinearMap (ofCrAnState φ) ∘ₗ normalOrder) a
|
||||
refine LinearMap.congr_fun ?h a
|
||||
apply ofCrAnListBasis.ext
|
||||
intro l
|
||||
simp only [mulLinearMap, LinearMap.coe_mk, AddHom.coe_mk, ofListBasis_eq_ofList,
|
||||
LinearMap.coe_comp, Function.comp_apply]
|
||||
rw [← ofCrAnList_cons]
|
||||
rw [normalOrder_ofCrAnList_cons_create φ hφ]
|
||||
|
||||
lemma normalOrder_ofCrAnList_append_annihilate (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) (φs : List 𝓕.CrAnStates) :
|
||||
normalOrder (ofCrAnList (φs ++ [φ])) =
|
||||
normalOrder (ofCrAnList φs) * ofCrAnState φ := by
|
||||
rw [normalOrder_ofCrAnList]
|
||||
rw [normalOrderSign_append_annihlate φ hφ φs, normalOrderList_append_annihilate φ hφ φs]
|
||||
rw [ofCrAnList_append, ofCrAnList_singleton, normalOrder_ofCrAnList, smul_mul_assoc]
|
||||
|
||||
lemma normalOrder_mul_annihilate (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate)
|
||||
(a : CrAnAlgebra 𝓕) :
|
||||
normalOrder (a * ofCrAnState φ) = normalOrder a * ofCrAnState φ := by
|
||||
change (normalOrder ∘ₗ mulLinearMap.flip (ofCrAnState φ)) a =
|
||||
(mulLinearMap.flip (ofCrAnState φ) ∘ₗ normalOrder) a
|
||||
refine LinearMap.congr_fun ?h a
|
||||
apply ofCrAnListBasis.ext
|
||||
intro l
|
||||
simp only [mulLinearMap, ofListBasis_eq_ofList, LinearMap.coe_comp, Function.comp_apply,
|
||||
LinearMap.flip_apply, LinearMap.coe_mk, AddHom.coe_mk]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_append, ofCrAnList_singleton]
|
||||
rw [normalOrder_ofCrAnList_append_annihilate φ hφ]
|
||||
|
||||
lemma normalOrder_swap_create_annihlate_ofCrAnList_ofCrAnList (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(φs φs' : List 𝓕.CrAnStates) :
|
||||
normalOrder (ofCrAnList φs' * ofCrAnState φc * ofCrAnState φa * ofCrAnList φs) =
|
||||
𝓢(𝓕 |>ₛ φc, 𝓕 |>ₛ φa) •
|
||||
normalOrder (ofCrAnList φs' * ofCrAnState φa * ofCrAnState φc * ofCrAnList φs) := by
|
||||
rw [mul_assoc, mul_assoc, ← ofCrAnList_cons, ← ofCrAnList_cons, ← ofCrAnList_append]
|
||||
rw [normalOrder_ofCrAnList, normalOrderSign_swap_create_annihlate φc φa hφc hφa]
|
||||
rw [normalOrderList_swap_create_annihlate φc φa hφc hφa]
|
||||
rw [← smul_smul, ← normalOrder_ofCrAnList]
|
||||
congr
|
||||
rw [ofCrAnList_append, ofCrAnList_cons, ofCrAnList_cons]
|
||||
noncomm_ring
|
||||
|
||||
lemma normalOrder_swap_create_annihlate_ofCrAnList (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(φs : List 𝓕.CrAnStates) (a : 𝓕.CrAnAlgebra) :
|
||||
normalOrder (ofCrAnList φs * ofCrAnState φc * ofCrAnState φa * a) =
|
||||
𝓢(𝓕 |>ₛ φc, 𝓕 |>ₛ φa) •
|
||||
normalOrder (ofCrAnList φs * ofCrAnState φa * ofCrAnState φc * a) := by
|
||||
change (normalOrder ∘ₗ mulLinearMap (ofCrAnList φs * ofCrAnState φc * ofCrAnState φa)) a =
|
||||
(smulLinearMap _ ∘ₗ normalOrder ∘ₗ
|
||||
mulLinearMap (ofCrAnList φs * ofCrAnState φa * ofCrAnState φc)) a
|
||||
refine LinearMap.congr_fun ?h a
|
||||
apply ofCrAnListBasis.ext
|
||||
intro l
|
||||
simp only [mulLinearMap, LinearMap.coe_mk, AddHom.coe_mk, ofListBasis_eq_ofList,
|
||||
LinearMap.coe_comp, Function.comp_apply, instCommGroup.eq_1]
|
||||
rw [normalOrder_swap_create_annihlate_ofCrAnList_ofCrAnList φc φa hφc hφa]
|
||||
rfl
|
||||
|
||||
lemma normalOrder_swap_create_annihlate (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(a b : 𝓕.CrAnAlgebra) :
|
||||
normalOrder (a * ofCrAnState φc * ofCrAnState φa * b) =
|
||||
𝓢(𝓕 |>ₛ φc, 𝓕 |>ₛ φa) •
|
||||
normalOrder (a * ofCrAnState φa * ofCrAnState φc * b) := by
|
||||
rw [mul_assoc, mul_assoc, mul_assoc, mul_assoc]
|
||||
change (normalOrder ∘ₗ mulLinearMap.flip (ofCrAnState φc * (ofCrAnState φa * b))) a =
|
||||
(smulLinearMap (𝓢(𝓕 |>ₛ φc, 𝓕 |>ₛ φa)) ∘ₗ
|
||||
normalOrder ∘ₗ mulLinearMap.flip (ofCrAnState φa * (ofCrAnState φc * b))) a
|
||||
apply LinearMap.congr_fun
|
||||
apply ofCrAnListBasis.ext
|
||||
intro l
|
||||
simp only [mulLinearMap, ofListBasis_eq_ofList, LinearMap.coe_comp, Function.comp_apply,
|
||||
LinearMap.flip_apply, LinearMap.coe_mk, AddHom.coe_mk, instCommGroup.eq_1]
|
||||
repeat rw [← mul_assoc]
|
||||
rw [normalOrder_swap_create_annihlate_ofCrAnList φc φa hφc hφa]
|
||||
rfl
|
||||
|
||||
lemma normalOrder_superCommute_create_annihilate (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(a b : 𝓕.CrAnAlgebra) :
|
||||
normalOrder (a * superCommute (ofCrAnState φc) (ofCrAnState φa) * b) = 0 := by
|
||||
rw [superCommute_ofCrAnState]
|
||||
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [mul_sub, sub_mul, map_sub, ← smul_mul_assoc]
|
||||
rw [← mul_assoc, ← mul_assoc]
|
||||
rw [normalOrder_swap_create_annihlate φc φa hφc hφa]
|
||||
simp only [FieldStatistic.instCommGroup.eq_1, Algebra.mul_smul_comm, Algebra.smul_mul_assoc,
|
||||
map_smul, sub_self]
|
||||
|
||||
lemma normalOrder_superCommute_annihilate_create (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(a b : 𝓕.CrAnAlgebra) :
|
||||
normalOrder (a * superCommute (ofCrAnState φa) (ofCrAnState φc) * b) = 0 := by
|
||||
rw [superCommute_ofCrAnState_symm]
|
||||
simp only [instCommGroup.eq_1, neg_smul, mul_neg, Algebra.mul_smul_comm, neg_mul,
|
||||
Algebra.smul_mul_assoc, map_neg, map_smul, neg_eq_zero, smul_eq_zero]
|
||||
apply Or.inr
|
||||
exact normalOrder_superCommute_create_annihilate φc φa hφc hφa a b
|
||||
|
||||
lemma normalOrder_crPart_mul (φ : 𝓕.States) (a : CrAnAlgebra 𝓕) :
|
||||
normalOrder (crPart (StateAlgebra.ofState φ) * a) =
|
||||
crPart (StateAlgebra.ofState φ) * normalOrder a := by
|
||||
match φ with
|
||||
| .negAsymp φ =>
|
||||
dsimp only [crPart, StateAlgebra.ofState]
|
||||
simp only [FreeAlgebra.lift_ι_apply]
|
||||
exact normalOrder_create_mul ⟨States.negAsymp φ, ()⟩ rfl a
|
||||
| .position φ =>
|
||||
dsimp only [crPart, StateAlgebra.ofState]
|
||||
simp only [FreeAlgebra.lift_ι_apply]
|
||||
refine normalOrder_create_mul _ ?_ _
|
||||
simp [crAnStatesToCreateAnnihilate]
|
||||
| .posAsymp φ =>
|
||||
simp
|
||||
|
||||
lemma normalOrder_mul_anPart (φ : 𝓕.States) (a : CrAnAlgebra 𝓕) :
|
||||
normalOrder (a * anPart (StateAlgebra.ofState φ)) =
|
||||
normalOrder a * anPart (StateAlgebra.ofState φ) := by
|
||||
match φ with
|
||||
| .negAsymp φ =>
|
||||
simp
|
||||
| .position φ =>
|
||||
dsimp only [anPart, StateAlgebra.ofState]
|
||||
simp only [FreeAlgebra.lift_ι_apply]
|
||||
refine normalOrder_mul_annihilate _ ?_ _
|
||||
simp [crAnStatesToCreateAnnihilate]
|
||||
| .posAsymp φ =>
|
||||
dsimp only [anPart, StateAlgebra.ofState]
|
||||
simp only [FreeAlgebra.lift_ι_apply]
|
||||
refine normalOrder_mul_annihilate _ ?_ _
|
||||
simp [crAnStatesToCreateAnnihilate]
|
||||
|
||||
lemma normalOrder_swap_crPart_anPart (φ φ' : 𝓕.States) (a b : CrAnAlgebra 𝓕) :
|
||||
normalOrder (a * (crPart (StateAlgebra.ofState φ)) * (anPart (StateAlgebra.ofState φ')) * b) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
normalOrder (a * (anPart (StateAlgebra.ofState φ')) *
|
||||
(crPart (StateAlgebra.ofState φ)) * b) := by
|
||||
match φ, φ' with
|
||||
| _, .negAsymp φ' =>
|
||||
simp
|
||||
| .posAsymp φ, _ =>
|
||||
simp
|
||||
| .position φ, .position φ' =>
|
||||
simp only [crPart_position, anPart_position, instCommGroup.eq_1]
|
||||
rw [normalOrder_swap_create_annihlate]
|
||||
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod]
|
||||
rfl
|
||||
rfl
|
||||
| .negAsymp φ, .posAsymp φ' =>
|
||||
simp only [crPart_negAsymp, anPart_posAsymp, instCommGroup.eq_1]
|
||||
rw [normalOrder_swap_create_annihlate]
|
||||
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod]
|
||||
rfl
|
||||
rfl
|
||||
| .negAsymp φ, .position φ' =>
|
||||
simp only [crPart_negAsymp, anPart_position, instCommGroup.eq_1]
|
||||
rw [normalOrder_swap_create_annihlate]
|
||||
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod]
|
||||
rfl
|
||||
rfl
|
||||
| .position φ, .posAsymp φ' =>
|
||||
simp only [crPart_position, anPart_posAsymp, instCommGroup.eq_1]
|
||||
rw [normalOrder_swap_create_annihlate]
|
||||
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnStatesToStates_prod]
|
||||
rfl
|
||||
rfl
|
||||
|
||||
lemma normalOrder_swap_anPart_crPart (φ φ' : 𝓕.States) (a b : CrAnAlgebra 𝓕) :
|
||||
normalOrder (a * (anPart (StateAlgebra.ofState φ)) * (crPart (StateAlgebra.ofState φ')) * b) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • normalOrder (a * (crPart (StateAlgebra.ofState φ')) *
|
||||
(anPart (StateAlgebra.ofState φ)) * b) := by
|
||||
rw [normalOrder_swap_crPart_anPart]
|
||||
rw [smul_smul, FieldStatistic.exchangeSign_symm, FieldStatistic.exchangeSign_mul_self]
|
||||
simp
|
||||
|
||||
lemma normalOrder_superCommute_crPart_anPart (φ φ' : 𝓕.States) (a b : CrAnAlgebra 𝓕) :
|
||||
normalOrder (a * superCommute
|
||||
(crPart (StateAlgebra.ofState φ)) (anPart (StateAlgebra.ofState φ')) * b) = 0 := by
|
||||
match φ, φ' with
|
||||
| _, .negAsymp φ' =>
|
||||
simp
|
||||
| .posAsymp φ', _ =>
|
||||
simp
|
||||
| .position φ, .position φ' =>
|
||||
simp only [crPart_position, anPart_position]
|
||||
refine normalOrder_superCommute_create_annihilate _ _ (by rfl) (by rfl) _ _
|
||||
| .negAsymp φ, .posAsymp φ' =>
|
||||
simp only [crPart_negAsymp, anPart_posAsymp]
|
||||
refine normalOrder_superCommute_create_annihilate _ _ (by rfl) (by rfl) _ _
|
||||
| .negAsymp φ, .position φ' =>
|
||||
simp only [crPart_negAsymp, anPart_position]
|
||||
refine normalOrder_superCommute_create_annihilate _ _ (by rfl) (by rfl) _ _
|
||||
| .position φ, .posAsymp φ' =>
|
||||
simp only [crPart_position, anPart_posAsymp]
|
||||
refine normalOrder_superCommute_create_annihilate _ _ (by rfl) (by rfl) _ _
|
||||
|
||||
lemma normalOrder_superCommute_anPart_crPart (φ φ' : 𝓕.States) (a b : CrAnAlgebra 𝓕) :
|
||||
normalOrder (a * superCommute
|
||||
(anPart (StateAlgebra.ofState φ)) (crPart (StateAlgebra.ofState φ')) * b) = 0 := by
|
||||
match φ, φ' with
|
||||
| .negAsymp φ', _ =>
|
||||
simp
|
||||
| _, .posAsymp φ' =>
|
||||
simp
|
||||
| .position φ, .position φ' =>
|
||||
simp only [anPart_position, crPart_position]
|
||||
refine normalOrder_superCommute_annihilate_create _ _ (by rfl) (by rfl) _ _
|
||||
| .posAsymp φ', .negAsymp φ =>
|
||||
simp only [anPart_posAsymp, crPart_negAsymp]
|
||||
refine normalOrder_superCommute_annihilate_create _ _ (by rfl) (by rfl) _ _
|
||||
| .position φ', .negAsymp φ =>
|
||||
simp only [anPart_position, crPart_negAsymp]
|
||||
refine normalOrder_superCommute_annihilate_create _ _ (by rfl) (by rfl) _ _
|
||||
| .posAsymp φ, .position φ' =>
|
||||
simp only [anPart_posAsymp, crPart_position]
|
||||
refine normalOrder_superCommute_annihilate_create _ _ (by rfl) (by rfl) _ _
|
||||
|
||||
lemma normalOrder_superCommute_ofCrAnList_create_create_ofCrAnList
|
||||
(φc φc' : 𝓕.CrAnStates) (hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφc' : 𝓕 |>ᶜ φc' = CreateAnnihilate.create) (φs φs' : List 𝓕.CrAnStates) :
|
||||
(normalOrder (ofCrAnList φs *
|
||||
superCommute (ofCrAnState φc) (ofCrAnState φc') * ofCrAnList φs')) =
|
||||
normalOrderSign (φs ++ φc' :: φc :: φs') •
|
||||
(ofCrAnList (createFilter φs) * superCommute (ofCrAnState φc) (ofCrAnState φc') *
|
||||
ofCrAnList (createFilter φs') * ofCrAnList (annihilateFilter (φs ++ φs'))) := by
|
||||
rw [superCommute_ofCrAnState]
|
||||
rw [mul_sub, sub_mul, map_sub]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
rhs
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_append, ← ofCrAnList_append,
|
||||
← ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
rw [normalOrder_ofCrAnList]
|
||||
rw [normalOrderList_eq_createFilter_append_annihilateFilter]
|
||||
rw [createFilter_append, createFilter_append, createFilter_append,
|
||||
createFilter_singleton_create _ hφc, createFilter_singleton_create _ hφc']
|
||||
rw [annihilateFilter_append, annihilateFilter_append, annihilateFilter_append,
|
||||
annihilateFilter_singleton_create _ hφc, annihilateFilter_singleton_create _ hφc']
|
||||
enter [2, 1, 2]
|
||||
simp only [List.singleton_append, List.append_assoc, List.cons_append, List.append_nil,
|
||||
instCommGroup.eq_1, Algebra.smul_mul_assoc, Algebra.mul_smul_comm, map_smul]
|
||||
rw [← annihilateFilter_append]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
rw [smul_mul_assoc]
|
||||
rw [Algebra.mul_smul_comm, smul_mul_assoc]
|
||||
rhs
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_append, ← ofCrAnList_append,
|
||||
← ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rw [map_smul]
|
||||
rhs
|
||||
rw [normalOrder_ofCrAnList]
|
||||
rw [normalOrderList_eq_createFilter_append_annihilateFilter]
|
||||
rw [createFilter_append, createFilter_append, createFilter_append,
|
||||
createFilter_singleton_create _ hφc, createFilter_singleton_create _ hφc']
|
||||
rw [annihilateFilter_append, annihilateFilter_append, annihilateFilter_append,
|
||||
annihilateFilter_singleton_create _ hφc, annihilateFilter_singleton_create _ hφc']
|
||||
enter [2, 1, 2]
|
||||
simp only [List.singleton_append, List.append_assoc, List.cons_append, instCommGroup.eq_1,
|
||||
List.append_nil, Algebra.smul_mul_assoc]
|
||||
rw [← annihilateFilter_append]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
lhs
|
||||
simp
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
lhs
|
||||
simp
|
||||
rw [normalOrderSign_swap_create_create φc φc' hφc hφc']
|
||||
rw [smul_smul, mul_comm, ← smul_smul]
|
||||
rw [← smul_sub, ofCrAnList_append, ofCrAnList_append, ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
rw [ofCrAnList_append, ofCrAnList_append, ofCrAnList_append]
|
||||
rw [← smul_mul_assoc, ← smul_mul_assoc, ← Algebra.mul_smul_comm]
|
||||
rw [← sub_mul, ← sub_mul, ← mul_sub]
|
||||
congr
|
||||
rw [ofCrAnList_append, ofCrAnList_singleton, ofCrAnList_singleton]
|
||||
rw [ofCrAnList_append, ofCrAnList_singleton, ofCrAnList_singleton, smul_mul_assoc]
|
||||
|
||||
lemma normalOrder_superCommute_ofCrAnList_annihilate_annihilate_ofCrAnList
|
||||
(φa φa' : 𝓕.CrAnStates)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(hφa' : 𝓕 |>ᶜ φa' = CreateAnnihilate.annihilate)
|
||||
(φs φs' : List 𝓕.CrAnStates) :
|
||||
(normalOrder (ofCrAnList φs *
|
||||
superCommute (ofCrAnState φa) (ofCrAnState φa') * ofCrAnList φs')) =
|
||||
normalOrderSign (φs ++ φa' :: φa :: φs') •
|
||||
(ofCrAnList (createFilter (φs ++ φs'))
|
||||
* ofCrAnList (annihilateFilter φs) * superCommute (ofCrAnState φa) (ofCrAnState φa')
|
||||
* ofCrAnList (annihilateFilter φs')) := by
|
||||
rw [superCommute_ofCrAnState]
|
||||
rw [mul_sub, sub_mul, map_sub]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
rhs
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_append, ← ofCrAnList_append,
|
||||
← ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
rw [normalOrder_ofCrAnList]
|
||||
rw [normalOrderList_eq_createFilter_append_annihilateFilter]
|
||||
rw [createFilter_append, createFilter_append, createFilter_append,
|
||||
createFilter_singleton_annihilate _ hφa, createFilter_singleton_annihilate _ hφa']
|
||||
rw [annihilateFilter_append, annihilateFilter_append, annihilateFilter_append,
|
||||
annihilateFilter_singleton_annihilate _ hφa, annihilateFilter_singleton_annihilate _ hφa']
|
||||
enter [2, 1, 1]
|
||||
simp only [List.singleton_append, List.append_assoc, List.cons_append, List.append_nil,
|
||||
instCommGroup.eq_1, Algebra.smul_mul_assoc, Algebra.mul_smul_comm, map_smul]
|
||||
rw [← createFilter_append]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
rw [smul_mul_assoc]
|
||||
rw [Algebra.mul_smul_comm, smul_mul_assoc]
|
||||
rhs
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_append, ← ofCrAnList_append,
|
||||
← ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rw [map_smul]
|
||||
rhs
|
||||
rw [normalOrder_ofCrAnList]
|
||||
rw [normalOrderList_eq_createFilter_append_annihilateFilter]
|
||||
rw [createFilter_append, createFilter_append, createFilter_append,
|
||||
createFilter_singleton_annihilate _ hφa, createFilter_singleton_annihilate _ hφa']
|
||||
rw [annihilateFilter_append, annihilateFilter_append, annihilateFilter_append,
|
||||
annihilateFilter_singleton_annihilate _ hφa, annihilateFilter_singleton_annihilate _ hφa']
|
||||
enter [2, 1, 1]
|
||||
simp only [List.singleton_append, List.append_assoc, List.cons_append, instCommGroup.eq_1,
|
||||
List.append_nil, Algebra.smul_mul_assoc]
|
||||
rw [← createFilter_append]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
lhs
|
||||
simp
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
lhs
|
||||
simp
|
||||
rw [normalOrderSign_swap_annihilate_annihilate φa φa' hφa hφa']
|
||||
rw [smul_smul, mul_comm, ← smul_smul]
|
||||
rw [← smul_sub, ofCrAnList_append, ofCrAnList_append, ofCrAnList_append]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
rw [ofCrAnList_append, ofCrAnList_append, ofCrAnList_append]
|
||||
rw [← Algebra.mul_smul_comm, ← smul_mul_assoc, ← Algebra.mul_smul_comm]
|
||||
rw [← mul_sub, ← sub_mul, ← mul_sub]
|
||||
apply congrArg
|
||||
conv_rhs => rw [mul_assoc, mul_assoc]
|
||||
apply congrArg
|
||||
rw [mul_assoc]
|
||||
apply congrArg
|
||||
congr
|
||||
rw [ofCrAnList_append, ofCrAnList_singleton, ofCrAnList_singleton]
|
||||
rw [ofCrAnList_append, ofCrAnList_singleton, ofCrAnList_singleton, smul_mul_assoc]
|
||||
|
||||
@[simp]
|
||||
lemma normalOrder_crPart_mul_crPart (φ φ' : 𝓕.States) :
|
||||
normalOrder (crPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ')) =
|
||||
crPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ') := by
|
||||
rw [normalOrder_crPart_mul]
|
||||
conv_lhs => rw [← mul_one (crPart (StateAlgebra.ofState φ'))]
|
||||
rw [normalOrder_crPart_mul, normalOrder_one]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma normalOrder_anPart_mul_anPart (φ φ' : 𝓕.States) :
|
||||
normalOrder (anPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ')) =
|
||||
anPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') := by
|
||||
rw [normalOrder_mul_anPart]
|
||||
conv_lhs => rw [← one_mul (anPart (StateAlgebra.ofState φ))]
|
||||
rw [normalOrder_mul_anPart, normalOrder_one]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma normalOrder_crPart_mul_anPart (φ φ' : 𝓕.States) :
|
||||
normalOrder (crPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ')) =
|
||||
crPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') := by
|
||||
rw [normalOrder_crPart_mul]
|
||||
conv_lhs => rw [← one_mul (anPart (StateAlgebra.ofState φ'))]
|
||||
rw [normalOrder_mul_anPart, normalOrder_one]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma normalOrder_anPart_mul_crPart (φ φ' : 𝓕.States) :
|
||||
normalOrder (anPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ')) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
(crPart (StateAlgebra.ofState φ') * anPart (StateAlgebra.ofState φ)) := by
|
||||
conv_lhs => rw [← one_mul (anPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ'))]
|
||||
conv_lhs => rw [← mul_one (1 * (anPart (StateAlgebra.ofState φ) *
|
||||
crPart (StateAlgebra.ofState φ')))]
|
||||
rw [← mul_assoc, normalOrder_swap_anPart_crPart]
|
||||
simp
|
||||
|
||||
lemma normalOrder_ofState_mul_ofState (φ φ' : 𝓕.States) :
|
||||
normalOrder (ofState φ * ofState φ') =
|
||||
crPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ') +
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
(crPart (StateAlgebra.ofState φ') * anPart (StateAlgebra.ofState φ)) +
|
||||
crPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') +
|
||||
anPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') := by
|
||||
rw [ofState_eq_crPart_add_anPart, ofState_eq_crPart_add_anPart]
|
||||
rw [mul_add, add_mul, add_mul]
|
||||
simp only [map_add, normalOrder_crPart_mul_crPart, normalOrder_anPart_mul_crPart,
|
||||
instCommGroup.eq_1, normalOrder_crPart_mul_anPart, normalOrder_anPart_mul_anPart]
|
||||
abel
|
||||
|
||||
lemma ofCrAnList_superCommute_normalOrder_ofCrAnList (φs φs' : List 𝓕.CrAnStates) :
|
||||
⟨ofCrAnList φs, normalOrder (ofCrAnList φs')⟩ₛca =
|
||||
ofCrAnList φs * normalOrder (ofCrAnList φs') -
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • normalOrder (ofCrAnList φs') * ofCrAnList φs := by
|
||||
simp [normalOrder_ofCrAnList, map_smul, superCommute_ofCrAnList, ofCrAnList_append,
|
||||
smul_sub, smul_smul, mul_comm]
|
||||
|
||||
lemma ofCrAnList_superCommute_normalOrder_ofStateList (φs : List 𝓕.CrAnStates)
|
||||
(φs' : List 𝓕.States) : ⟨ofCrAnList φs, normalOrder (ofStateList φs')⟩ₛca =
|
||||
ofCrAnList φs * normalOrder (ofStateList φs') -
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • normalOrder (ofStateList φs') * ofCrAnList φs := by
|
||||
rw [ofStateList_sum, map_sum, Finset.mul_sum, Finset.smul_sum, Finset.sum_mul,
|
||||
← Finset.sum_sub_distrib, map_sum]
|
||||
congr
|
||||
funext n
|
||||
rw [ofCrAnList_superCommute_normalOrder_ofCrAnList,
|
||||
CrAnSection.statistics_eq_state_statistics]
|
||||
|
||||
lemma ofCrAnList_mul_normalOrder_ofStateList_eq_superCommute (φs : List 𝓕.CrAnStates)
|
||||
(φs' : List 𝓕.States) :
|
||||
ofCrAnList φs * normalOrder (ofStateList φs') =
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • normalOrder (ofStateList φs') * ofCrAnList φs
|
||||
+ ⟨ofCrAnList φs, normalOrder (ofStateList φs')⟩ₛca := by
|
||||
rw [ofCrAnList_superCommute_normalOrder_ofStateList]
|
||||
simp
|
||||
|
||||
lemma ofCrAnState_mul_normalOrder_ofStateList_eq_superCommute (φ : 𝓕.CrAnStates)
|
||||
(φs' : List 𝓕.States) :
|
||||
ofCrAnState φ * normalOrder (ofStateList φs') =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • normalOrder (ofStateList φs') * ofCrAnState φ
|
||||
+ ⟨ofCrAnState φ, normalOrder (ofStateList φs')⟩ₛca := by
|
||||
rw [← ofCrAnList_singleton, ofCrAnList_mul_normalOrder_ofStateList_eq_superCommute]
|
||||
simp [ofCrAnList_singleton]
|
||||
|
||||
lemma anPart_mul_normalOrder_ofStateList_eq_superCommute (φ : 𝓕.States)
|
||||
(φs' : List 𝓕.States) :
|
||||
anPart (StateAlgebra.ofState φ) * normalOrder (ofStateList φs') =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • normalOrder (ofStateList φs' * anPart (StateAlgebra.ofState φ))
|
||||
+ ⟨anPart (StateAlgebra.ofState φ), normalOrder (ofStateList φs')⟩ₛca := by
|
||||
rw [normalOrder_mul_anPart]
|
||||
match φ with
|
||||
| .negAsymp φ =>
|
||||
simp
|
||||
| .position φ =>
|
||||
simp only [anPart_position, instCommGroup.eq_1]
|
||||
rw [ofCrAnState_mul_normalOrder_ofStateList_eq_superCommute]
|
||||
simp [crAnStatistics]
|
||||
| .posAsymp φ =>
|
||||
simp only [anPart_posAsymp, instCommGroup.eq_1]
|
||||
rw [ofCrAnState_mul_normalOrder_ofStateList_eq_superCommute]
|
||||
simp [crAnStatistics]
|
||||
|
||||
end
|
||||
|
||||
end CrAnAlgebra
|
||||
|
||||
end FieldStruct
|
|
@ -0,0 +1,424 @@
|
|||
/-
|
||||
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.Algebras.CrAnAlgebra.Basic
|
||||
/-!
|
||||
|
||||
# Super Commute
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
namespace CrAnAlgebra
|
||||
|
||||
open StateAlgebra
|
||||
|
||||
/-!
|
||||
|
||||
## The super commutor on the creation and annihlation algebra.
|
||||
|
||||
-/
|
||||
|
||||
open FieldStatistic
|
||||
|
||||
/-- The super commutor on the creation and annihlation algebra. For two bosonic operators
|
||||
or a bosonic and fermionic operator this corresponds to the usual commutator
|
||||
whilst for two fermionic operators this corresponds to the anti-commutator. -/
|
||||
noncomputable def superCommute : 𝓕.CrAnAlgebra →ₗ[ℂ] 𝓕.CrAnAlgebra →ₗ[ℂ] 𝓕.CrAnAlgebra :=
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs =>
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs' =>
|
||||
ofCrAnList (φs ++ φs') - 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofCrAnList (φs' ++ φs)
|
||||
|
||||
/-- The super commutor on the creation and annihlation algebra. For two bosonic operators
|
||||
or a bosonic and fermionic operator this corresponds to the usual commutator
|
||||
whilst for two fermionic operators this corresponds to the anti-commutator. -/
|
||||
scoped[FieldStruct.CrAnAlgebra] notation "⟨" φs "," φs' "⟩ₛca" => superCommute φs φs'
|
||||
|
||||
lemma superCommute_ofCrAnList (φs φs' : List 𝓕.CrAnStates) : ⟨ofCrAnList φs, ofCrAnList φs'⟩ₛca =
|
||||
ofCrAnList (φs ++ φs') - 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofCrAnList (φs' ++ φs) := by
|
||||
rw [← ofListBasis_eq_ofList, ← ofListBasis_eq_ofList]
|
||||
simp only [superCommute, Basis.constr_basis]
|
||||
|
||||
lemma superCommute_ofCrAnList_ofStatesList (φcas : List 𝓕.CrAnStates) (φs : List 𝓕.States) :
|
||||
⟨ofCrAnList φcas, ofStateList φs⟩ₛca = ofCrAnList φcas * ofStateList φs -
|
||||
𝓢(𝓕 |>ₛ φcas, 𝓕 |>ₛ φs) • ofStateList φs * ofCrAnList φcas := by
|
||||
conv_lhs => rw [ofStateList_sum]
|
||||
rw [map_sum]
|
||||
conv_lhs =>
|
||||
enter [2, x]
|
||||
rw [superCommute_ofCrAnList, CrAnSection.statistics_eq_state_statistics,
|
||||
ofCrAnList_append, ofCrAnList_append]
|
||||
rw [Finset.sum_sub_distrib, ← Finset.mul_sum, ← Finset.smul_sum,
|
||||
← Finset.sum_mul, ← ofStateList_sum]
|
||||
simp
|
||||
|
||||
lemma superCommute_ofStateList_ofStatesList (φ : List 𝓕.States) (φs : List 𝓕.States) :
|
||||
⟨ofStateList φ, ofStateList φs⟩ₛca = ofStateList φ * ofStateList φs -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) • ofStateList φs * ofStateList φ := by
|
||||
conv_lhs => rw [ofStateList_sum]
|
||||
simp only [map_sum, LinearMap.coeFn_sum, Finset.sum_apply, instCommGroup.eq_1,
|
||||
Algebra.smul_mul_assoc]
|
||||
conv_lhs =>
|
||||
enter [2, x]
|
||||
rw [superCommute_ofCrAnList_ofStatesList]
|
||||
simp only [instCommGroup.eq_1, CrAnSection.statistics_eq_state_statistics,
|
||||
Algebra.smul_mul_assoc, Finset.sum_sub_distrib]
|
||||
rw [← Finset.sum_mul, ← Finset.smul_sum, ← Finset.mul_sum, ← ofStateList_sum]
|
||||
|
||||
lemma superCommute_ofState_ofStatesList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
⟨ofState φ, ofStateList φs⟩ₛca = ofState φ * ofStateList φs -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) • ofStateList φs * ofState φ := by
|
||||
rw [← ofStateList_singleton, superCommute_ofStateList_ofStatesList, ofStateList_singleton]
|
||||
simp
|
||||
|
||||
lemma superCommute_ofStateList_ofStates (φs : List 𝓕.States) (φ : 𝓕.States) :
|
||||
⟨ofStateList φs, ofState φ⟩ₛca = ofStateList φs * ofState φ -
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ) • ofState φ * ofStateList φs := by
|
||||
rw [← ofStateList_singleton, superCommute_ofStateList_ofStatesList, ofStateList_singleton]
|
||||
simp
|
||||
|
||||
lemma ofCrAnList_mul_ofCrAnList_eq_superCommute (φs φs' : List 𝓕.CrAnStates) :
|
||||
ofCrAnList φs * ofCrAnList φs' = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofCrAnList φs' * ofCrAnList φs
|
||||
+ ⟨ofCrAnList φs, ofCrAnList φs'⟩ₛca := by
|
||||
rw [superCommute_ofCrAnList]
|
||||
simp [ofCrAnList_append]
|
||||
|
||||
lemma ofCrAnState_mul_ofCrAnList_eq_superCommute (φ : 𝓕.CrAnStates) (φs' : List 𝓕.CrAnStates) :
|
||||
ofCrAnState φ * ofCrAnList φs' = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • ofCrAnList φs' * ofCrAnState φ
|
||||
+ ⟨ofCrAnState φ, ofCrAnList φs'⟩ₛca := by
|
||||
rw [← ofCrAnList_singleton, ofCrAnList_mul_ofCrAnList_eq_superCommute]
|
||||
simp
|
||||
|
||||
lemma ofStateList_mul_ofStateList_eq_superCommute (φs φs' : List 𝓕.States) :
|
||||
ofStateList φs * ofStateList φs' = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofStateList φs' * ofStateList φs
|
||||
+ ⟨ofStateList φs, ofStateList φs'⟩ₛca := by
|
||||
rw [superCommute_ofStateList_ofStatesList]
|
||||
simp
|
||||
|
||||
lemma ofState_mul_ofStateList_eq_superCommute (φ : 𝓕.States) (φs' : List 𝓕.States) :
|
||||
ofState φ * ofStateList φs' = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • ofStateList φs' * ofState φ
|
||||
+ ⟨ofState φ, ofStateList φs'⟩ₛca := by
|
||||
rw [superCommute_ofState_ofStatesList]
|
||||
simp
|
||||
|
||||
lemma ofStateList_mul_ofState_eq_superCommute (φs : List 𝓕.States) (φ : 𝓕.States) :
|
||||
ofStateList φs * ofState φ = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ) • ofState φ * ofStateList φs
|
||||
+ ⟨ofStateList φs, ofState φ⟩ₛca := by
|
||||
rw [superCommute_ofStateList_ofStates]
|
||||
simp
|
||||
|
||||
lemma superCommute_anPart_crPart (φ φ' : 𝓕.States) :
|
||||
⟨anPart (StateAlgebra.ofState φ), crPart (StateAlgebra.ofState φ')⟩ₛca =
|
||||
anPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ') -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
crPart (StateAlgebra.ofState φ') * anPart (StateAlgebra.ofState φ) := by
|
||||
match φ, φ' with
|
||||
| States.negAsymp φ, _ =>
|
||||
simp
|
||||
| _, States.posAsymp φ =>
|
||||
simp only [crPart_posAsymp, map_zero, mul_zero, instCommGroup.eq_1, smul_zero, zero_mul,
|
||||
sub_self]
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [anPart_position, crPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.posAsymp φ, States.position φ' =>
|
||||
simp only [anPart_posAsymp, crPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.negAsymp φ' =>
|
||||
simp only [anPart_position, crPart_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp only [List.singleton_append, instCommGroup.eq_1, crAnStatistics,
|
||||
FieldStatistic.ofList_singleton, Function.comp_apply, crAnStatesToStates_prod, ←
|
||||
ofCrAnList_append]
|
||||
| States.posAsymp φ, States.negAsymp φ' =>
|
||||
simp only [anPart_posAsymp, crPart_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma superCommute_crPart_anPart (φ φ' : 𝓕.States) :
|
||||
⟨crPart (StateAlgebra.ofState φ), anPart (StateAlgebra.ofState φ')⟩ₛca =
|
||||
crPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
anPart (StateAlgebra.ofState φ') * crPart (StateAlgebra.ofState φ) := by
|
||||
match φ, φ' with
|
||||
| States.posAsymp φ, _ =>
|
||||
simp only [crPart_posAsymp, map_zero, LinearMap.zero_apply, zero_mul, instCommGroup.eq_1,
|
||||
mul_zero, sub_self]
|
||||
| _, States.negAsymp φ =>
|
||||
simp only [anPart_negAsymp, map_zero, mul_zero, instCommGroup.eq_1, smul_zero, zero_mul,
|
||||
sub_self]
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [crPart_position, anPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.posAsymp φ' =>
|
||||
simp only [crPart_position, anPart_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.negAsymp φ, States.position φ' =>
|
||||
simp only [crPart_negAsymp, anPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.negAsymp φ, States.posAsymp φ' =>
|
||||
simp only [crPart_negAsymp, anPart_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma superCommute_crPart_crPart (φ φ' : 𝓕.States) :
|
||||
⟨crPart (StateAlgebra.ofState φ), crPart (StateAlgebra.ofState φ')⟩ₛca =
|
||||
crPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ') -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
crPart (StateAlgebra.ofState φ') * crPart (StateAlgebra.ofState φ) := by
|
||||
match φ, φ' with
|
||||
| States.posAsymp φ, _ =>
|
||||
simp only [crPart_posAsymp, map_zero, LinearMap.zero_apply, zero_mul, instCommGroup.eq_1,
|
||||
mul_zero, sub_self]
|
||||
| _, States.posAsymp φ =>
|
||||
simp only [crPart_posAsymp, map_zero, mul_zero, instCommGroup.eq_1, smul_zero, zero_mul, sub_self]
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [crPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.negAsymp φ' =>
|
||||
simp only [crPart_position, crPart_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.negAsymp φ, States.position φ' =>
|
||||
simp only [crPart_negAsymp, crPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.negAsymp φ, States.negAsymp φ' =>
|
||||
simp only [crPart_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma superCommute_anPart_anPart (φ φ' : 𝓕.States) :
|
||||
⟨anPart (StateAlgebra.ofState φ), anPart (StateAlgebra.ofState φ')⟩ₛca =
|
||||
anPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
anPart (StateAlgebra.ofState φ') * anPart (StateAlgebra.ofState φ) := by
|
||||
match φ, φ' with
|
||||
| States.negAsymp φ, _ =>
|
||||
simp
|
||||
| _, States.negAsymp φ =>
|
||||
simp
|
||||
| States.position φ, States.position φ' =>
|
||||
simp only [anPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.position φ, States.posAsymp φ' =>
|
||||
simp only [anPart_position, anPart_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.posAsymp φ, States.position φ' =>
|
||||
simp only [anPart_posAsymp, anPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
| States.posAsymp φ, States.posAsymp φ' =>
|
||||
simp only [anPart_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList]
|
||||
simp [crAnStatistics, ← ofCrAnList_append]
|
||||
|
||||
lemma crPart_anPart_eq_superCommute (φ φ' : 𝓕.States) :
|
||||
crPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
anPart (StateAlgebra.ofState φ') * crPart (StateAlgebra.ofState φ) +
|
||||
⟨crPart (StateAlgebra.ofState φ), anPart (StateAlgebra.ofState φ')⟩ₛca := by
|
||||
rw [superCommute_crPart_anPart]
|
||||
simp
|
||||
|
||||
lemma anPart_crPart_eq_superCommute (φ φ' : 𝓕.States) :
|
||||
anPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ') =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
crPart (StateAlgebra.ofState φ') * anPart (StateAlgebra.ofState φ) +
|
||||
⟨anPart (StateAlgebra.ofState φ), crPart (StateAlgebra.ofState φ')⟩ₛca := by
|
||||
rw [superCommute_anPart_crPart]
|
||||
simp
|
||||
|
||||
lemma crPart_crPart_eq_superCommute (φ φ' : 𝓕.States) :
|
||||
crPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ') =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
crPart (StateAlgebra.ofState φ') * crPart (StateAlgebra.ofState φ) +
|
||||
⟨crPart (StateAlgebra.ofState φ), crPart (StateAlgebra.ofState φ')⟩ₛca := by
|
||||
rw [superCommute_crPart_crPart]
|
||||
simp
|
||||
|
||||
lemma anPart_anPart_eq_superCommute (φ φ' : 𝓕.States) :
|
||||
anPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
anPart (StateAlgebra.ofState φ') * anPart (StateAlgebra.ofState φ) +
|
||||
⟨anPart (StateAlgebra.ofState φ), anPart (StateAlgebra.ofState φ')⟩ₛca := by
|
||||
rw [superCommute_anPart_anPart]
|
||||
simp
|
||||
|
||||
lemma superCommute_crPart_ofStatesList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
⟨crPart (StateAlgebra.ofState φ), ofStateList φs⟩ₛca =
|
||||
crPart (StateAlgebra.ofState φ) * ofStateList φs - 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) • ofStateList φs *
|
||||
crPart (StateAlgebra.ofState φ) := by
|
||||
match φ with
|
||||
| States.negAsymp φ =>
|
||||
simp only [crPart_negAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
| States.position φ =>
|
||||
simp only [crPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
| States.posAsymp φ =>
|
||||
simp
|
||||
|
||||
lemma superCommute_anPart_ofStatesList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
⟨anPart (StateAlgebra.ofState φ), ofStateList φs⟩ₛca =
|
||||
anPart (StateAlgebra.ofState φ) * ofStateList φs - 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) •
|
||||
ofStateList φs * anPart (StateAlgebra.ofState φ) := by
|
||||
match φ with
|
||||
| States.negAsymp φ =>
|
||||
simp
|
||||
| States.position φ =>
|
||||
simp only [anPart_position, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
| States.posAsymp φ =>
|
||||
simp only [anPart_posAsymp, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofStatesList]
|
||||
simp [crAnStatistics]
|
||||
|
||||
lemma superCommute_crPart_ofState (φ φ' : 𝓕.States) :
|
||||
⟨crPart (StateAlgebra.ofState φ), ofState φ'⟩ₛca =
|
||||
crPart (StateAlgebra.ofState φ) * ofState φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
ofState φ' * crPart (StateAlgebra.ofState φ) := by
|
||||
rw [← ofStateList_singleton, superCommute_crPart_ofStatesList]
|
||||
simp
|
||||
|
||||
lemma superCommute_anPart_ofState (φ φ' : 𝓕.States) :
|
||||
⟨anPart (StateAlgebra.ofState φ), ofState φ'⟩ₛca =
|
||||
anPart (StateAlgebra.ofState φ) * ofState φ' -
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
ofState φ' * anPart (StateAlgebra.ofState φ) := by
|
||||
rw [← ofStateList_singleton, superCommute_anPart_ofStatesList]
|
||||
simp
|
||||
|
||||
lemma superCommute_ofCrAnState (φ φ' : 𝓕.CrAnStates) : ⟨ofCrAnState φ, ofCrAnState φ'⟩ₛca =
|
||||
ofCrAnState φ * ofCrAnState φ' - 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • ofCrAnState φ' * ofCrAnState φ := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rw [superCommute_ofCrAnList, ofCrAnList_append]
|
||||
congr
|
||||
rw [ofCrAnList_append]
|
||||
rw [FieldStatistic.ofList_singleton, FieldStatistic.ofList_singleton, smul_mul_assoc]
|
||||
|
||||
lemma superCommute_ofCrAnList_symm (φs φs' : List 𝓕.CrAnStates) :
|
||||
⟨ofCrAnList φs, ofCrAnList φs'⟩ₛca =
|
||||
(- 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs')) •
|
||||
⟨ofCrAnList φs', ofCrAnList φs⟩ₛca := by
|
||||
rw [superCommute_ofCrAnList, superCommute_ofCrAnList, smul_sub]
|
||||
simp only [instCommGroup.eq_1, neg_smul, sub_neg_eq_add]
|
||||
rw [smul_smul]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [exchangeSign_symm, exchangeSign_mul_self]
|
||||
simp only [one_smul]
|
||||
abel
|
||||
|
||||
lemma superCommute_ofCrAnState_symm (φ φ' : 𝓕.CrAnStates) :
|
||||
⟨ofCrAnState φ, ofCrAnState φ'⟩ₛca =
|
||||
(- 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ')) • ⟨ofCrAnState φ', ofCrAnState φ⟩ₛca := by
|
||||
rw [superCommute_ofCrAnState, superCommute_ofCrAnState]
|
||||
rw [smul_sub]
|
||||
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc, neg_smul, sub_neg_eq_add]
|
||||
rw [smul_smul]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [exchangeSign_symm, exchangeSign_mul_self]
|
||||
simp only [one_smul]
|
||||
abel
|
||||
|
||||
lemma superCommute_ofCrAnList_ofCrAnList_cons (φ : 𝓕.CrAnStates) (φs φs' : List 𝓕.CrAnStates) :
|
||||
⟨ofCrAnList φs, ofCrAnList (φ :: φs')⟩ₛca =
|
||||
⟨ofCrAnList φs, ofCrAnState φ⟩ₛca * ofCrAnList φs' +
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ)
|
||||
• ofCrAnState φ * ⟨ofCrAnList φs, ofCrAnList φs'⟩ₛca := by
|
||||
rw [superCommute_ofCrAnList]
|
||||
conv_rhs =>
|
||||
lhs
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList, sub_mul, ← ofCrAnList_append]
|
||||
rhs
|
||||
rw [FieldStatistic.ofList_singleton, ofCrAnList_append, ofCrAnList_singleton, smul_mul_assoc,
|
||||
mul_assoc, ← ofCrAnList_append]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [superCommute_ofCrAnList, mul_sub, smul_mul_assoc]
|
||||
simp only [instCommGroup.eq_1, List.cons_append, List.append_assoc, List.nil_append,
|
||||
Algebra.mul_smul_comm, Algebra.smul_mul_assoc, sub_add_sub_cancel, sub_right_inj]
|
||||
rw [← ofCrAnList_cons, smul_smul, FieldStatistic.ofList_cons_eq_mul]
|
||||
simp only [instCommGroup, map_mul, mul_comm]
|
||||
|
||||
lemma superCommute_ofCrAnList_ofStateList_cons (φ : 𝓕.States) (φs : List 𝓕.CrAnStates)
|
||||
(φs' : List 𝓕.States) : ⟨ofCrAnList φs, ofStateList (φ :: φs')⟩ₛca =
|
||||
⟨ofCrAnList φs, ofState φ⟩ₛca * ofStateList φs' +
|
||||
𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φ) • ofState φ * ⟨ofCrAnList φs, ofStateList φs'⟩ₛca := by
|
||||
rw [superCommute_ofCrAnList_ofStatesList]
|
||||
conv_rhs =>
|
||||
lhs
|
||||
rw [← ofStateList_singleton, superCommute_ofCrAnList_ofStatesList, sub_mul, mul_assoc,
|
||||
← ofStateList_append]
|
||||
rhs
|
||||
rw [FieldStatistic.ofList_singleton, ofStateList_singleton, smul_mul_assoc,
|
||||
smul_mul_assoc, mul_assoc]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [superCommute_ofCrAnList_ofStatesList, mul_sub, smul_mul_assoc]
|
||||
simp only [instCommGroup, Algebra.smul_mul_assoc, List.singleton_append, Algebra.mul_smul_comm,
|
||||
sub_add_sub_cancel, sub_right_inj]
|
||||
rw [ofStateList_cons, mul_assoc, smul_smul, FieldStatistic.ofList_cons_eq_mul]
|
||||
simp [mul_comm]
|
||||
|
||||
lemma ofCrAnList_mul_ofStateList_eq_superCommute (φs : List 𝓕.CrAnStates) (φs' : List 𝓕.States) :
|
||||
ofCrAnList φs * ofStateList φs' = 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofStateList φs' * ofCrAnList φs
|
||||
+ ⟨ofCrAnList φs, ofStateList φs'⟩ₛca := by
|
||||
rw [superCommute_ofCrAnList_ofStatesList]
|
||||
simp
|
||||
|
||||
lemma superCommute_ofCrAnList_ofCrAnList_eq_sum (φs : List 𝓕.CrAnStates) :
|
||||
(φs' : List 𝓕.CrAnStates) →
|
||||
⟨ofCrAnList φs, ofCrAnList φs'⟩ₛca =
|
||||
∑ (n : Fin φs'.length), 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ (List.take n φs')) •
|
||||
ofCrAnList (φs'.take n) * ⟨ofCrAnList φs, ofCrAnState (φs'.get n)⟩ₛca *
|
||||
ofCrAnList (φs'.drop (n + 1))
|
||||
| [] => by
|
||||
simp [← ofCrAnList_nil, superCommute_ofCrAnList]
|
||||
| φ :: φs' => by
|
||||
rw [superCommute_ofCrAnList_ofCrAnList_cons, superCommute_ofCrAnList_ofCrAnList_eq_sum φs φs']
|
||||
conv_rhs => erw [Fin.sum_univ_succ]
|
||||
congr 1
|
||||
· simp
|
||||
· simp [Finset.mul_sum, smul_smul, ofCrAnList_cons, mul_assoc,
|
||||
FieldStatistic.ofList_cons_eq_mul, mul_comm]
|
||||
|
||||
lemma superCommute_ofCrAnList_ofStateList_eq_sum (φs : List 𝓕.CrAnStates) :
|
||||
(φs' : List 𝓕.States) →
|
||||
⟨ofCrAnList φs, ofStateList φs'⟩ₛca =
|
||||
∑ (n : Fin φs'.length), 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ List.take n φs') •
|
||||
ofStateList (φs'.take n) * ⟨ofCrAnList φs, ofState (φs'.get n)⟩ₛca *
|
||||
ofStateList (φs'.drop (n + 1))
|
||||
| [] => by
|
||||
simp only [superCommute_ofCrAnList_ofStatesList, instCommGroup, ofList_empty,
|
||||
exchangeSign_bosonic, one_smul, List.length_nil, Finset.univ_eq_empty, List.take_nil,
|
||||
List.get_eq_getElem, List.drop_nil, Finset.sum_empty]
|
||||
simp
|
||||
| φ :: φs' => by
|
||||
rw [superCommute_ofCrAnList_ofStateList_cons, superCommute_ofCrAnList_ofStateList_eq_sum φs φs']
|
||||
conv_rhs => erw [Fin.sum_univ_succ]
|
||||
congr 1
|
||||
· simp
|
||||
· simp [Finset.mul_sum, smul_smul, ofStateList_cons, mul_assoc,
|
||||
FieldStatistic.ofList_cons_eq_mul, mul_comm]
|
||||
|
||||
end CrAnAlgebra
|
||||
|
||||
end FieldStruct
|
202
HepLean/PerturbationTheory/Algebras/OperatorAlgebra/Basic.lean
Normal file
202
HepLean/PerturbationTheory/Algebras/OperatorAlgebra/Basic.lean
Normal file
|
@ -0,0 +1,202 @@
|
|||
/-
|
||||
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.Algebras.CrAnAlgebra.SuperCommute
|
||||
/-!
|
||||
|
||||
# The operator algebras
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable (𝓕 : FieldStruct)
|
||||
open CrAnAlgebra
|
||||
|
||||
/-- The structure of an algebra with properties necessary for that algebra
|
||||
to be isomorphic to the actual operator algebra of a field theory.
|
||||
These properties are sufficent to prove certain theorems about the Operator algebra
|
||||
in particular Wick's theorem. -/
|
||||
structure OperatorAlgebra where
|
||||
/-- The algebra representing the operator algebra. -/
|
||||
A : Type
|
||||
/-- The instance of the type `A` as a semi-ring. -/
|
||||
[A_semiRing : Semiring A]
|
||||
/-- The instance of the type `A` as an algebra. -/
|
||||
[A_algebra : Algebra ℂ A]
|
||||
/-- An algebra map from the creation and annihilation free algebra to the
|
||||
algebra A. -/
|
||||
crAnF : 𝓕.CrAnAlgebra →ₐ[ℂ] A
|
||||
superCommute_crAn_center : ∀ (φ ψ : 𝓕.CrAnStates),
|
||||
crAnF (superCommute (ofCrAnState φ) (ofCrAnState ψ))
|
||||
∈ Subalgebra.center ℂ A
|
||||
superCommute_create_create : ∀ (φc φc' : 𝓕.CrAnStates)
|
||||
(_ : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(_ : 𝓕 |>ᶜ φc' = CreateAnnihilate.create),
|
||||
crAnF (superCommute (ofCrAnState φc) (ofCrAnState φc')) = 0
|
||||
superCommute_annihilate_annihilate : ∀ (φa φa' : 𝓕.CrAnStates)
|
||||
(_ : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(_ : 𝓕 |>ᶜ φa' = CreateAnnihilate.annihilate),
|
||||
crAnF (superCommute (ofCrAnState φa) (ofCrAnState φa')) = 0
|
||||
superCommute_different_statistics : ∀ (φ φ' : 𝓕.CrAnStates)
|
||||
(_ : ¬ (𝓕 |>ₛ φ) = (𝓕 |>ₛ φ')),
|
||||
crAnF (superCommute (ofCrAnState φ) (ofCrAnState φ')) = 0
|
||||
|
||||
namespace OperatorAlgebra
|
||||
open FieldStatistic
|
||||
variable {𝓕 : FieldStruct} (𝓞 : 𝓕.OperatorAlgebra)
|
||||
|
||||
/-- The algebra `𝓞.A` carries the instance of a semi-ring induced via `A_seimRing`. -/
|
||||
instance : Semiring 𝓞.A := 𝓞.A_semiRing
|
||||
|
||||
/-- The algebra `𝓞.A` carries the instance of aan algebra over `ℂ` induced via `A_algebra`. -/
|
||||
instance : Algebra ℂ 𝓞.A := 𝓞.A_algebra
|
||||
|
||||
lemma crAnF_superCommute_ofCrAnState_ofState_mem_center (φ : 𝓕.CrAnStates) (ψ : 𝓕.States) :
|
||||
𝓞.crAnF (superCommute (ofCrAnState φ) (ofState ψ)) ∈ Subalgebra.center ℂ 𝓞.A := by
|
||||
rw [ofState, map_sum, map_sum]
|
||||
refine Subalgebra.sum_mem (Subalgebra.center ℂ 𝓞.A) ?h
|
||||
intro x _
|
||||
exact 𝓞.superCommute_crAn_center φ ⟨ψ, x⟩
|
||||
|
||||
lemma crAnF_superCommute_anPart_ofState_mem_center (φ ψ : 𝓕.States) :
|
||||
𝓞.crAnF ⟨anPart (StateAlgebra.ofState φ), ofState ψ⟩ₛca ∈ Subalgebra.center ℂ 𝓞.A := by
|
||||
match φ with
|
||||
| States.negAsymp _ =>
|
||||
simp only [anPart_negAsymp, map_zero, LinearMap.zero_apply]
|
||||
exact Subalgebra.zero_mem (Subalgebra.center ℂ 𝓞.A)
|
||||
| States.position φ =>
|
||||
simp only [anPart_position]
|
||||
exact 𝓞.crAnF_superCommute_ofCrAnState_ofState_mem_center _ _
|
||||
| States.posAsymp _ =>
|
||||
simp only [anPart_posAsymp]
|
||||
exact 𝓞.crAnF_superCommute_ofCrAnState_ofState_mem_center _ _
|
||||
|
||||
lemma crAnF_superCommute_ofCrAnState_ofState_diff_grade_zero (φ : 𝓕.CrAnStates) (ψ : 𝓕.States)
|
||||
(h : (𝓕 |>ₛ φ) ≠ (𝓕 |>ₛ ψ)) :
|
||||
𝓞.crAnF (superCommute (ofCrAnState φ) (ofState ψ)) = 0 := by
|
||||
rw [ofState, map_sum, map_sum]
|
||||
rw [Finset.sum_eq_zero]
|
||||
intro x hx
|
||||
apply 𝓞.superCommute_different_statistics
|
||||
simpa [crAnStatistics] using h
|
||||
|
||||
lemma crAnF_superCommute_anPart_ofState_diff_grade_zero (φ ψ : 𝓕.States)
|
||||
(h : (𝓕 |>ₛ φ) ≠ (𝓕 |>ₛ ψ)) :
|
||||
𝓞.crAnF (superCommute (anPart (StateAlgebra.ofState φ)) (ofState ψ)) = 0 := by
|
||||
match φ with
|
||||
| States.negAsymp _ =>
|
||||
simp
|
||||
| States.position φ =>
|
||||
simp only [anPart_position]
|
||||
apply 𝓞.crAnF_superCommute_ofCrAnState_ofState_diff_grade_zero _ _ _
|
||||
simpa [crAnStatistics] using h
|
||||
| States.posAsymp _ =>
|
||||
simp only [anPart_posAsymp]
|
||||
apply 𝓞.crAnF_superCommute_ofCrAnState_ofState_diff_grade_zero _ _
|
||||
simpa [crAnStatistics] using h
|
||||
|
||||
lemma crAnF_superCommute_ofState_ofState_mem_center (φ ψ : 𝓕.States) :
|
||||
𝓞.crAnF (superCommute (ofState φ) (ofState ψ)) ∈ Subalgebra.center ℂ 𝓞.A := by
|
||||
rw [ofState, map_sum]
|
||||
simp only [LinearMap.coeFn_sum, Finset.sum_apply, map_sum]
|
||||
refine Subalgebra.sum_mem (Subalgebra.center ℂ 𝓞.A) ?h
|
||||
intro x _
|
||||
exact crAnF_superCommute_ofCrAnState_ofState_mem_center 𝓞 ⟨φ, x⟩ ψ
|
||||
|
||||
lemma crAnF_superCommute_anPart_anPart (φ ψ : 𝓕.States) :
|
||||
𝓞.crAnF ⟨anPart (StateAlgebra.ofState φ), anPart (StateAlgebra.ofState ψ)⟩ₛca = 0 := by
|
||||
match φ, ψ with
|
||||
| _, States.negAsymp _ =>
|
||||
simp
|
||||
| States.negAsymp _, _ =>
|
||||
simp
|
||||
| States.position φ, States.position ψ =>
|
||||
simp only [anPart_position]
|
||||
rw [𝓞.superCommute_annihilate_annihilate]
|
||||
rfl
|
||||
rfl
|
||||
| States.position φ, States.posAsymp _ =>
|
||||
simp only [anPart_position, anPart_posAsymp]
|
||||
rw [𝓞.superCommute_annihilate_annihilate]
|
||||
rfl
|
||||
rfl
|
||||
| States.posAsymp _, States.posAsymp _ =>
|
||||
simp only [anPart_posAsymp]
|
||||
rw [𝓞.superCommute_annihilate_annihilate]
|
||||
rfl
|
||||
rfl
|
||||
| States.posAsymp _, States.position _ =>
|
||||
simp only [anPart_posAsymp, anPart_position]
|
||||
rw [𝓞.superCommute_annihilate_annihilate]
|
||||
rfl
|
||||
rfl
|
||||
|
||||
lemma crAnF_superCommute_crPart_crPart (φ ψ : 𝓕.States) :
|
||||
𝓞.crAnF ⟨crPart (StateAlgebra.ofState φ), crPart (StateAlgebra.ofState ψ)⟩ₛca = 0 := by
|
||||
match φ, ψ with
|
||||
| _, States.posAsymp _ =>
|
||||
simp
|
||||
| States.posAsymp _, _ =>
|
||||
simp
|
||||
| States.position φ, States.position ψ =>
|
||||
simp only [crPart_position]
|
||||
rw [𝓞.superCommute_create_create]
|
||||
rfl
|
||||
rfl
|
||||
| States.position φ, States.negAsymp _ =>
|
||||
simp only [crPart_position, crPart_negAsymp]
|
||||
rw [𝓞.superCommute_create_create]
|
||||
rfl
|
||||
rfl
|
||||
| States.negAsymp _, States.negAsymp _ =>
|
||||
simp only [crPart_negAsymp]
|
||||
rw [𝓞.superCommute_create_create]
|
||||
rfl
|
||||
rfl
|
||||
| States.negAsymp _, States.position _ =>
|
||||
simp only [crPart_negAsymp, crPart_position]
|
||||
rw [𝓞.superCommute_create_create]
|
||||
rfl
|
||||
rfl
|
||||
|
||||
lemma crAnF_superCommute_ofCrAnState_ofCrAnList_eq_sum (φ : 𝓕.CrAnStates) (φs : List 𝓕.CrAnStates) :
|
||||
𝓞.crAnF ⟨ofCrAnState φ, ofCrAnList φs⟩ₛca
|
||||
= 𝓞.crAnF (∑ (n : Fin φs.length), 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ (List.take n φs)) •
|
||||
⟨ofCrAnState φ, ofCrAnState (φs.get n)⟩ₛca * ofCrAnList (φs.eraseIdx n)) := by
|
||||
conv_lhs =>
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList_eq_sum]
|
||||
rw [map_sum, map_sum]
|
||||
congr
|
||||
funext x
|
||||
repeat rw [map_mul]
|
||||
rw [map_smul, map_smul, ofCrAnList_singleton]
|
||||
have h := Subalgebra.mem_center_iff.mp (𝓞.superCommute_crAn_center φ (φs.get x))
|
||||
rw [h, mul_smul_comm, smul_mul_assoc, smul_mul_assoc, mul_assoc]
|
||||
congr 1
|
||||
· simp
|
||||
· congr
|
||||
rw [← map_mul, ← ofCrAnList_append, ← List.eraseIdx_eq_take_drop_succ]
|
||||
|
||||
lemma crAnF_superCommute_ofCrAnState_ofStateList_eq_sum (φ : 𝓕.CrAnStates) (φs : List 𝓕.States) :
|
||||
𝓞.crAnF ⟨ofCrAnState φ, ofStateList φs⟩ₛca
|
||||
= 𝓞.crAnF (∑ (n : Fin φs.length), 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ (List.take n φs)) •
|
||||
⟨ofCrAnState φ, ofState (φs.get n)⟩ₛca * ofStateList (φs.eraseIdx n)) := by
|
||||
conv_lhs =>
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofStateList_eq_sum]
|
||||
rw [map_sum, map_sum]
|
||||
congr
|
||||
funext x
|
||||
repeat rw [map_mul]
|
||||
rw [map_smul, map_smul, ofCrAnList_singleton]
|
||||
have h := Subalgebra.mem_center_iff.mp
|
||||
(crAnF_superCommute_ofCrAnState_ofState_mem_center 𝓞 φ (φs.get x))
|
||||
rw [h, mul_smul_comm, smul_mul_assoc, smul_mul_assoc, mul_assoc]
|
||||
congr 1
|
||||
· simp
|
||||
· congr
|
||||
rw [← map_mul, ← ofStateList_append, ← List.eraseIdx_eq_take_drop_succ]
|
||||
|
||||
end OperatorAlgebra
|
||||
end FieldStruct
|
|
@ -0,0 +1,386 @@
|
|||
/-
|
||||
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.Algebras.CrAnAlgebra.NormalOrder
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
/-!
|
||||
|
||||
# Normal Ordering
|
||||
|
||||
The normal ordering puts all creation operators to the left and all annihilation operators to the
|
||||
right. It acts on `CrAnStates` and defines a linear map from the `CrAnAlgebra` to itself.
|
||||
|
||||
The normal ordering satisfies a number of nice properties with relation to the operator
|
||||
algebra 𝓞.A.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
namespace OperatorAlgebra
|
||||
variable {𝓞 : OperatorAlgebra 𝓕}
|
||||
open CrAnAlgebra
|
||||
open FieldStatistic
|
||||
|
||||
lemma crAnF_normalOrder_superCommute_ofCrAnList_create_create_ofCrAnList
|
||||
(φc φc' : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφc' : 𝓕 |>ᶜ φc' = CreateAnnihilate.create)
|
||||
(φs φs' : List 𝓕.CrAnStates) :
|
||||
𝓞.crAnF (normalOrder
|
||||
(ofCrAnList φs * superCommute (ofCrAnState φc) (ofCrAnState φc') * ofCrAnList φs')) = 0 := by
|
||||
rw [normalOrder_superCommute_ofCrAnList_create_create_ofCrAnList φc φc' hφc hφc' φs φs']
|
||||
rw [map_smul, map_mul, map_mul, map_mul, 𝓞.superCommute_create_create φc φc' hφc hφc']
|
||||
simp
|
||||
|
||||
lemma crAnF_normalOrder_superCommute_ofCrAnList_annihilate_annihilate_ofCrAnList
|
||||
(φa φa' : 𝓕.CrAnStates)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(hφa' : 𝓕 |>ᶜ φa' = CreateAnnihilate.annihilate)
|
||||
(φs φs' : List 𝓕.CrAnStates) :
|
||||
𝓞.crAnF (normalOrder
|
||||
(ofCrAnList φs * superCommute (ofCrAnState φa) (ofCrAnState φa') * ofCrAnList φs')) = 0 := by
|
||||
rw [normalOrder_superCommute_ofCrAnList_annihilate_annihilate_ofCrAnList φa φa' hφa hφa' φs φs']
|
||||
rw [map_smul, map_mul, map_mul, map_mul, 𝓞.superCommute_annihilate_annihilate φa φa' hφa hφa']
|
||||
simp
|
||||
|
||||
lemma crAnF_normalOrder_superCommute_ofCrAnList_ofCrAnList_eq_zero
|
||||
(φa φa' : 𝓕.CrAnStates) (φs φs' : List 𝓕.CrAnStates) :
|
||||
𝓞.crAnF (normalOrder
|
||||
(ofCrAnList φs * superCommute (ofCrAnState φa) (ofCrAnState φa') * ofCrAnList φs')) = 0 := by
|
||||
rcases CreateAnnihilate.eq_create_or_annihilate (𝓕 |>ᶜ φa) with hφa | hφa
|
||||
<;> rcases CreateAnnihilate.eq_create_or_annihilate (𝓕 |>ᶜ φa') with hφa' | hφa'
|
||||
· rw [normalOrder_superCommute_ofCrAnList_create_create_ofCrAnList φa φa' hφa hφa' φs φs']
|
||||
rw [map_smul, map_mul, map_mul, map_mul, 𝓞.superCommute_create_create φa φa' hφa hφa']
|
||||
simp
|
||||
· rw [normalOrder_superCommute_create_annihilate φa φa' hφa hφa' (ofCrAnList φs)
|
||||
(ofCrAnList φs')]
|
||||
simp
|
||||
· rw [normalOrder_superCommute_annihilate_create φa' φa hφa' hφa (ofCrAnList φs)
|
||||
(ofCrAnList φs')]
|
||||
simp
|
||||
· rw [normalOrder_superCommute_ofCrAnList_annihilate_annihilate_ofCrAnList φa φa' hφa hφa' φs φs']
|
||||
rw [map_smul, map_mul, map_mul, map_mul, 𝓞.superCommute_annihilate_annihilate φa φa' hφa hφa']
|
||||
simp
|
||||
|
||||
lemma crAnF_normalOrder_superCommute_ofCrAnList_eq_zero
|
||||
(φa φa' : 𝓕.CrAnStates) (φs : List 𝓕.CrAnStates)
|
||||
(a : 𝓕.CrAnAlgebra) : 𝓞.crAnF (normalOrder (ofCrAnList φs *
|
||||
superCommute (ofCrAnState φa) (ofCrAnState φa') * a)) = 0 := by
|
||||
change (𝓞.crAnF.toLinearMap ∘ₗ normalOrder ∘ₗ
|
||||
mulLinearMap ((ofCrAnList φs * superCommute (ofCrAnState φa) (ofCrAnState φa')))) a = 0
|
||||
have hf : 𝓞.crAnF.toLinearMap ∘ₗ normalOrder ∘ₗ
|
||||
mulLinearMap ((ofCrAnList φs * superCommute (ofCrAnState φa) (ofCrAnState φa'))) = 0 := by
|
||||
apply ofCrAnListBasis.ext
|
||||
intro l
|
||||
simp only [ofListBasis_eq_ofList, LinearMap.coe_comp, Function.comp_apply,
|
||||
AlgHom.toLinearMap_apply, LinearMap.zero_apply]
|
||||
exact crAnF_normalOrder_superCommute_ofCrAnList_ofCrAnList_eq_zero φa φa' φs l
|
||||
rw [hf]
|
||||
simp
|
||||
|
||||
lemma crAnF_normalOrder_superCommute_ofCrAnState_eq_zero_mul (φa φa' : 𝓕.CrAnStates)
|
||||
(a b : 𝓕.CrAnAlgebra) :
|
||||
𝓞.crAnF (normalOrder (a * superCommute (ofCrAnState φa) (ofCrAnState φa') * b)) = 0 := by
|
||||
rw [mul_assoc]
|
||||
change (𝓞.crAnF.toLinearMap ∘ₗ normalOrder ∘ₗ mulLinearMap.flip
|
||||
((superCommute (ofCrAnState φa) (ofCrAnState φa') * b))) a = 0
|
||||
have hf : (𝓞.crAnF.toLinearMap ∘ₗ normalOrder ∘ₗ mulLinearMap.flip
|
||||
((superCommute (ofCrAnState φa) (ofCrAnState φa') * b))) = 0 := by
|
||||
apply ofCrAnListBasis.ext
|
||||
intro l
|
||||
simp only [mulLinearMap, ofListBasis_eq_ofList, LinearMap.coe_comp, Function.comp_apply,
|
||||
LinearMap.flip_apply, LinearMap.coe_mk, AddHom.coe_mk, AlgHom.toLinearMap_apply,
|
||||
LinearMap.zero_apply]
|
||||
rw [← mul_assoc]
|
||||
exact crAnF_normalOrder_superCommute_ofCrAnList_eq_zero φa φa' _ _
|
||||
rw [hf]
|
||||
simp
|
||||
|
||||
lemma crAnF_normalOrder_superCommute_ofCrAnState_ofCrAnList_eq_zero_mul (φa : 𝓕.CrAnStates)
|
||||
(φs : List 𝓕.CrAnStates)
|
||||
(a b : 𝓕.CrAnAlgebra) :
|
||||
𝓞.crAnF (normalOrder (a * superCommute (ofCrAnState φa) (ofCrAnList φs) * b)) = 0 := by
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList_eq_sum]
|
||||
rw [Finset.mul_sum, Finset.sum_mul]
|
||||
rw [map_sum, map_sum]
|
||||
apply Fintype.sum_eq_zero
|
||||
intro n
|
||||
rw [← mul_assoc, ← mul_assoc]
|
||||
rw [mul_assoc _ _ b, ofCrAnList_singleton]
|
||||
rw [crAnF_normalOrder_superCommute_ofCrAnState_eq_zero_mul]
|
||||
|
||||
lemma crAnF_normalOrder_superCommute_ofCrAnList_ofCrAnState_eq_zero_mul (φa : 𝓕.CrAnStates)
|
||||
(φs : List 𝓕.CrAnStates)
|
||||
(a b : 𝓕.CrAnAlgebra) :
|
||||
𝓞.crAnF (normalOrder (a * superCommute (ofCrAnList φs) (ofCrAnState φa) * b)) = 0 := by
|
||||
rw [← ofCrAnList_singleton, superCommute_ofCrAnList_symm, ofCrAnList_singleton]
|
||||
simp only [FieldStatistic.instCommGroup.eq_1, FieldStatistic.ofList_singleton, mul_neg,
|
||||
Algebra.mul_smul_comm, neg_mul, Algebra.smul_mul_assoc, map_neg, map_smul]
|
||||
rw [crAnF_normalOrder_superCommute_ofCrAnState_ofCrAnList_eq_zero_mul]
|
||||
simp
|
||||
|
||||
lemma crAnF_normalOrder_superCommute_ofCrAnList_ofCrAnList_eq_zero_mul
|
||||
(φs φs' : List 𝓕.CrAnStates)
|
||||
(a b : 𝓕.CrAnAlgebra) :
|
||||
𝓞.crAnF (normalOrder (a * superCommute (ofCrAnList φs) (ofCrAnList φs') * b)) = 0 := by
|
||||
rw [superCommute_ofCrAnList_ofCrAnList_eq_sum, Finset.mul_sum, Finset.sum_mul]
|
||||
rw [map_sum, map_sum]
|
||||
apply Fintype.sum_eq_zero
|
||||
intro n
|
||||
rw [← mul_assoc, ← mul_assoc]
|
||||
rw [mul_assoc _ _ b]
|
||||
rw [crAnF_normalOrder_superCommute_ofCrAnList_ofCrAnState_eq_zero_mul]
|
||||
|
||||
lemma crAnF_normalOrder_superCommute_ofCrAnList_eq_zero_mul
|
||||
(φs : List 𝓕.CrAnStates)
|
||||
(a b c : 𝓕.CrAnAlgebra) :
|
||||
𝓞.crAnF (normalOrder (a * superCommute (ofCrAnList φs) c * b)) = 0 := by
|
||||
change (𝓞.crAnF.toLinearMap ∘ₗ normalOrder ∘ₗ
|
||||
mulLinearMap.flip b ∘ₗ mulLinearMap a ∘ₗ superCommute (ofCrAnList φs)) c = 0
|
||||
have hf : (𝓞.crAnF.toLinearMap ∘ₗ normalOrder ∘ₗ
|
||||
mulLinearMap.flip b ∘ₗ mulLinearMap a ∘ₗ superCommute (ofCrAnList φs)) = 0 := by
|
||||
apply ofCrAnListBasis.ext
|
||||
intro φs'
|
||||
simp only [mulLinearMap, LinearMap.coe_mk, AddHom.coe_mk, ofListBasis_eq_ofList,
|
||||
LinearMap.coe_comp, Function.comp_apply, LinearMap.flip_apply, AlgHom.toLinearMap_apply,
|
||||
LinearMap.zero_apply]
|
||||
rw [crAnF_normalOrder_superCommute_ofCrAnList_ofCrAnList_eq_zero_mul]
|
||||
rw [hf]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma crAnF_normalOrder_superCommute_eq_zero_mul
|
||||
(a b c d : 𝓕.CrAnAlgebra) : 𝓞.crAnF (normalOrder (a * superCommute d c * b)) = 0 := by
|
||||
change (𝓞.crAnF.toLinearMap ∘ₗ normalOrder ∘ₗ
|
||||
mulLinearMap.flip b ∘ₗ mulLinearMap a ∘ₗ superCommute.flip c) d = 0
|
||||
have hf : (𝓞.crAnF.toLinearMap ∘ₗ normalOrder ∘ₗ
|
||||
mulLinearMap.flip b ∘ₗ mulLinearMap a ∘ₗ superCommute.flip c) = 0 := by
|
||||
apply ofCrAnListBasis.ext
|
||||
intro φs
|
||||
simp only [mulLinearMap, LinearMap.coe_mk, AddHom.coe_mk, ofListBasis_eq_ofList,
|
||||
LinearMap.coe_comp, Function.comp_apply, LinearMap.flip_apply, AlgHom.toLinearMap_apply,
|
||||
LinearMap.zero_apply]
|
||||
rw [crAnF_normalOrder_superCommute_ofCrAnList_eq_zero_mul]
|
||||
rw [hf]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma crAnF_normalOrder_superCommute_eq_zero_mul_right
|
||||
(b c d : 𝓕.CrAnAlgebra) : 𝓞.crAnF (normalOrder (superCommute d c * b)) = 0 := by
|
||||
rw [← crAnF_normalOrder_superCommute_eq_zero_mul 1 b c d]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma crAnF_normalOrder_superCommute_eq_zero_mul_left
|
||||
(a c d : 𝓕.CrAnAlgebra) : 𝓞.crAnF (normalOrder (a * superCommute d c)) = 0 := by
|
||||
rw [← crAnF_normalOrder_superCommute_eq_zero_mul a 1 c d]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma crAnF_normalOrder_superCommute_eq_zero
|
||||
(c d : 𝓕.CrAnAlgebra) : 𝓞.crAnF (normalOrder (superCommute d c)) = 0 := by
|
||||
rw [← crAnF_normalOrder_superCommute_eq_zero_mul 1 1 c d]
|
||||
simp
|
||||
|
||||
lemma crAnF_normalOrder_ofState_ofState_swap (φ φ' : 𝓕.States) :
|
||||
𝓞.crAnF (normalOrder (ofState φ * ofState φ')) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
𝓞.crAnF (normalOrder (ofState φ' * ofState φ)) := by
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
rw [ofState_eq_crPart_add_anPart, ofState_eq_crPart_add_anPart]
|
||||
rw [mul_add, add_mul, add_mul]
|
||||
rw [crPart_crPart_eq_superCommute, crPart_anPart_eq_superCommute,
|
||||
anPart_anPart_eq_superCommute, anPart_crPart_eq_superCommute]
|
||||
simp only [FieldStatistic.instCommGroup.eq_1, Algebra.smul_mul_assoc, map_add, map_smul,
|
||||
normalOrder_crPart_mul_crPart, normalOrder_crPart_mul_anPart, normalOrder_anPart_mul_crPart,
|
||||
normalOrder_anPart_mul_anPart, map_mul, crAnF_normalOrder_superCommute_eq_zero, add_zero]
|
||||
rw [normalOrder_ofState_mul_ofState]
|
||||
simp only [FieldStatistic.instCommGroup.eq_1, map_add, map_mul, map_smul, smul_add]
|
||||
rw [smul_smul]
|
||||
simp only [FieldStatistic.exchangeSign_mul_self_swap, one_smul]
|
||||
abel
|
||||
|
||||
open FieldStatistic
|
||||
|
||||
lemma crAnF_normalOrder_ofCrAnState_ofCrAnList_swap (φ : 𝓕.CrAnStates)
|
||||
(φs : List 𝓕.CrAnStates) :
|
||||
𝓞.crAnF (normalOrder (ofCrAnState φ * ofCrAnList φs)) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs) • 𝓞.crAnF (normalOrder (ofCrAnList φs * ofCrAnState φ)) := by
|
||||
rw [← ofCrAnList_singleton, ofCrAnList_mul_ofCrAnList_eq_superCommute]
|
||||
simp
|
||||
|
||||
lemma crAnF_normalOrder_ofCrAnState_ofStatesList_swap (φ : 𝓕.CrAnStates)
|
||||
(φ' : List 𝓕.States) :
|
||||
𝓞.crAnF (normalOrder (ofCrAnState φ * ofStateList φ')) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
𝓞.crAnF (normalOrder (ofStateList φ' * ofCrAnState φ)) := by
|
||||
rw [← ofCrAnList_singleton, ofCrAnList_mul_ofStateList_eq_superCommute]
|
||||
simp
|
||||
|
||||
lemma crAnF_normalOrder_anPart_ofStatesList_swap (φ : 𝓕.States)
|
||||
(φ' : List 𝓕.States) :
|
||||
𝓞.crAnF (normalOrder (anPart (StateAlgebra.ofState φ) * ofStateList φ')) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
𝓞.crAnF (normalOrder (ofStateList φ' * anPart (StateAlgebra.ofState φ))) := by
|
||||
match φ with
|
||||
| .negAsymp φ =>
|
||||
simp
|
||||
| .position φ =>
|
||||
simp only [anPart_position, instCommGroup.eq_1]
|
||||
rw [crAnF_normalOrder_ofCrAnState_ofStatesList_swap]
|
||||
rfl
|
||||
| .posAsymp φ =>
|
||||
simp only [anPart_posAsymp, instCommGroup.eq_1]
|
||||
rw [crAnF_normalOrder_ofCrAnState_ofStatesList_swap]
|
||||
rfl
|
||||
|
||||
lemma crAnF_normalOrder_ofStatesList_anPart_swap (φ : 𝓕.States) (φ' : List 𝓕.States) :
|
||||
𝓞.crAnF (normalOrder (ofStateList φ' * anPart (StateAlgebra.ofState φ)))
|
||||
= 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
𝓞.crAnF (normalOrder (anPart (StateAlgebra.ofState φ) * ofStateList φ')) := by
|
||||
rw [crAnF_normalOrder_anPart_ofStatesList_swap]
|
||||
simp [smul_smul, FieldStatistic.exchangeSign_mul_self]
|
||||
|
||||
lemma crAnF_normalOrder_ofStatesList_mul_anPart_swap (φ : 𝓕.States)
|
||||
(φ' : List 𝓕.States) :
|
||||
𝓞.crAnF (normalOrder (ofStateList φ') * anPart (StateAlgebra.ofState φ)) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
|
||||
𝓞.crAnF (normalOrder (anPart (StateAlgebra.ofState φ) * ofStateList φ')) := by
|
||||
rw [← normalOrder_mul_anPart]
|
||||
rw [crAnF_normalOrder_ofStatesList_anPart_swap]
|
||||
|
||||
lemma crAnF_ofCrAnState_superCommute_normalOrder_ofCrAnList_eq_sum (φ : 𝓕.CrAnStates)
|
||||
(φs : List 𝓕.CrAnStates) : 𝓞.crAnF (⟨ofCrAnState φ, normalOrder (ofCrAnList φs)⟩ₛca) =
|
||||
∑ n : Fin φs.length, 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ (φs.take n)) •
|
||||
𝓞.crAnF (⟨ofCrAnState φ, ofCrAnState φs[n]⟩ₛca)
|
||||
* 𝓞.crAnF (normalOrder (ofCrAnList (φs.eraseIdx n))) := by
|
||||
rw [normalOrder_ofCrAnList, map_smul, map_smul]
|
||||
rw [crAnF_superCommute_ofCrAnState_ofCrAnList_eq_sum, sum_normalOrderList_length]
|
||||
simp only [instCommGroup.eq_1, List.get_eq_getElem, normalOrderList_get_normalOrderEquiv,
|
||||
normalOrderList_eraseIdx_normalOrderEquiv, Algebra.smul_mul_assoc, map_sum, map_smul, map_mul,
|
||||
Finset.smul_sum, Fin.getElem_fin]
|
||||
congr
|
||||
funext n
|
||||
rw [ofCrAnList_eq_normalOrder, map_smul, mul_smul_comm, smul_smul, smul_smul]
|
||||
by_cases hs : (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[n])
|
||||
· congr
|
||||
erw [normalOrderSign_eraseIdx, ← hs]
|
||||
trans (normalOrderSign φs * normalOrderSign φs) *
|
||||
(𝓢(𝓕 |>ₛ (φs.get n), 𝓕 |>ₛ ((normalOrderList φs).take (normalOrderEquiv n))) *
|
||||
𝓢(𝓕 |>ₛ (φs.get n), 𝓕 |>ₛ ((normalOrderList φs).take (normalOrderEquiv n))))
|
||||
* 𝓢(𝓕 |>ₛ (φs.get n), 𝓕 |>ₛ (φs.take n))
|
||||
· ring_nf
|
||||
rw [hs]
|
||||
rfl
|
||||
· simp [hs]
|
||||
· erw [𝓞.superCommute_different_statistics _ _ hs]
|
||||
simp
|
||||
|
||||
lemma crAnF_ofCrAnState_superCommute_normalOrder_ofStateList_eq_sum (φ : 𝓕.CrAnStates)
|
||||
(φs : List 𝓕.States) : 𝓞.crAnF (⟨ofCrAnState φ, normalOrder (ofStateList φs)⟩ₛca) =
|
||||
∑ n : Fin φs.length, 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ (φs.take n)) •
|
||||
𝓞.crAnF (⟨ofCrAnState φ, ofState φs[n]⟩ₛca)
|
||||
* 𝓞.crAnF (normalOrder (ofStateList (φs.eraseIdx n))) := by
|
||||
conv_lhs =>
|
||||
rw [ofStateList_sum, map_sum, map_sum, map_sum]
|
||||
enter [2, s]
|
||||
rw [crAnF_ofCrAnState_superCommute_normalOrder_ofCrAnList_eq_sum,
|
||||
CrAnSection.sum_over_length]
|
||||
enter [2, n]
|
||||
rw [CrAnSection.take_statistics_eq_take_state_statistics, smul_mul_assoc]
|
||||
rw [Finset.sum_comm]
|
||||
refine Finset.sum_congr rfl (fun n _ => ?_)
|
||||
simp only [instCommGroup.eq_1, Fin.coe_cast, Fin.getElem_fin,
|
||||
CrAnSection.sum_eraseIdxEquiv n _ n.prop,
|
||||
CrAnSection.eraseIdxEquiv_symm_getElem,
|
||||
CrAnSection.eraseIdxEquiv_symm_eraseIdx, ← Finset.smul_sum, Algebra.smul_mul_assoc]
|
||||
conv_lhs =>
|
||||
enter [2, 2, n]
|
||||
rw [← Finset.mul_sum]
|
||||
rw [← Finset.sum_mul, ← map_sum, ← map_sum, ← ofState, ← map_sum, ← map_sum, ← ofStateList_sum]
|
||||
|
||||
lemma crAnF_anPart_superCommute_normalOrder_ofStateList_eq_sum (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
𝓞.crAnF (⟨anPart (StateAlgebra.ofState φ), normalOrder (ofStateList φs)⟩ₛca) =
|
||||
∑ n : Fin φs.length, 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ (φs.take n)) •
|
||||
𝓞.crAnF (⟨anPart (StateAlgebra.ofState φ), ofState φs[n]⟩ₛca)
|
||||
* 𝓞.crAnF (normalOrder (ofStateList (φs.eraseIdx n))) := by
|
||||
match φ with
|
||||
| .negAsymp φ =>
|
||||
simp
|
||||
| .position φ =>
|
||||
simp only [anPart_position, instCommGroup.eq_1, Fin.getElem_fin, Algebra.smul_mul_assoc]
|
||||
rw [crAnF_ofCrAnState_superCommute_normalOrder_ofStateList_eq_sum]
|
||||
simp [crAnStatistics]
|
||||
| .posAsymp φ =>
|
||||
simp only [anPart_posAsymp, instCommGroup.eq_1, Fin.getElem_fin, Algebra.smul_mul_assoc]
|
||||
rw [crAnF_ofCrAnState_superCommute_normalOrder_ofStateList_eq_sum]
|
||||
simp [crAnStatistics]
|
||||
|
||||
lemma crAnF_anPart_mul_normalOrder_ofStatesList_eq_superCommute (φ : 𝓕.States)
|
||||
(φ' : List 𝓕.States) :
|
||||
𝓞.crAnF (anPart (StateAlgebra.ofState φ) * normalOrder (ofStateList φ')) =
|
||||
𝓞.crAnF (normalOrder (anPart (StateAlgebra.ofState φ) * ofStateList φ')) +
|
||||
𝓞.crAnF (⟨anPart (StateAlgebra.ofState φ), normalOrder (ofStateList φ')⟩ₛca) := by
|
||||
rw [anPart_mul_normalOrder_ofStateList_eq_superCommute]
|
||||
simp only [instCommGroup.eq_1, map_add, map_smul]
|
||||
congr
|
||||
rw [crAnF_normalOrder_anPart_ofStatesList_swap]
|
||||
|
||||
lemma crAnF_ofState_mul_normalOrder_ofStatesList_eq_superCommute (φ : 𝓕.States)
|
||||
(φ' : List 𝓕.States) :
|
||||
𝓞.crAnF (ofState φ * normalOrder (ofStateList φ')) =
|
||||
𝓞.crAnF (normalOrder (ofState φ * ofStateList φ')) +
|
||||
𝓞.crAnF (⟨anPart (StateAlgebra.ofState φ), normalOrder (ofStateList φ')⟩ₛca) := by
|
||||
conv_lhs => rw [ofState_eq_crPart_add_anPart]
|
||||
rw [add_mul, map_add, crAnF_anPart_mul_normalOrder_ofStatesList_eq_superCommute, ← add_assoc,
|
||||
← normalOrder_crPart_mul, ← map_add]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
rw [← map_add, ← add_mul, ← ofState_eq_crPart_add_anPart]
|
||||
|
||||
/-- In the expansion of `ofState φ * normalOrder (ofStateList φs)` the element
|
||||
of `𝓞.A` associated with contracting `φ` with the (optional) `n`th element of `φs`. -/
|
||||
noncomputable def contractStateAtIndex (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(n : Option (Fin φs.length)) : 𝓞.A :=
|
||||
match n with
|
||||
| none => 1
|
||||
| some n => 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ (φs.take n)) •
|
||||
𝓞.crAnF (⟨anPart (StateAlgebra.ofState φ), ofState φs[n]⟩ₛca)
|
||||
|
||||
lemma crAnF_ofState_mul_normalOrder_ofStatesList_eq_sum (φ : 𝓕.States)
|
||||
(φs : List 𝓕.States) :
|
||||
𝓞.crAnF (ofState φ * normalOrder (ofStateList φs)) =
|
||||
∑ n : Option (Fin φs.length),
|
||||
contractStateAtIndex φ φs n *
|
||||
𝓞.crAnF (normalOrder (ofStateList (HepLean.List.optionEraseZ φs φ n))) := by
|
||||
rw [crAnF_ofState_mul_normalOrder_ofStatesList_eq_superCommute]
|
||||
rw [crAnF_anPart_superCommute_normalOrder_ofStateList_eq_sum]
|
||||
simp only [instCommGroup.eq_1, Fin.getElem_fin, Algebra.smul_mul_assoc, contractStateAtIndex,
|
||||
Fintype.sum_option, one_mul]
|
||||
rfl
|
||||
|
||||
lemma crAnF_ofState_normalOrder_insert (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(k : Fin φs.length.succ) :
|
||||
𝓞.crAnF (normalOrder (ofStateList (φ :: φs))) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs.take k) • 𝓞.crAnF (normalOrder (ofStateList (φs.insertIdx k φ))) := by
|
||||
have hl : φs.insertIdx k φ = φs.take k ++ [φ] ++ φs.drop k := by
|
||||
rw [HepLean.List.insertIdx_eq_take_drop]
|
||||
simp
|
||||
rw [hl]
|
||||
rw [ofStateList_append, ofStateList_append]
|
||||
rw [ofStateList_mul_ofStateList_eq_superCommute, add_mul]
|
||||
simp only [instCommGroup.eq_1, Nat.succ_eq_add_one, ofList_singleton, Algebra.smul_mul_assoc,
|
||||
map_add, map_smul, crAnF_normalOrder_superCommute_eq_zero_mul_right, add_zero, smul_smul,
|
||||
exchangeSign_mul_self_swap, one_smul]
|
||||
rw [← ofStateList_append, ← ofStateList_append]
|
||||
simp
|
||||
|
||||
end OperatorAlgebra
|
||||
|
||||
end FieldStruct
|
|
@ -0,0 +1,92 @@
|
|||
/-
|
||||
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.Algebras.OperatorAlgebra.NormalOrder
|
||||
import HepLean.PerturbationTheory.Algebras.StateAlgebra.TimeOrder
|
||||
/-!
|
||||
|
||||
# Time contractions
|
||||
|
||||
We define the state algebra of a field structure to be the free algebra
|
||||
generated by the states.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
open CrAnAlgebra
|
||||
noncomputable section
|
||||
|
||||
namespace OperatorAlgebra
|
||||
|
||||
variable (𝓞 : 𝓕.OperatorAlgebra)
|
||||
open FieldStatistic
|
||||
|
||||
/-- The time contraction of two States as an element of `𝓞.A` defined
|
||||
as their time ordering in the state algebra minus their normal ordering in the
|
||||
creation and annihlation algebra, both mapped to `𝓞.A`.. -/
|
||||
def timeContract (φ ψ : 𝓕.States) : 𝓞.A :=
|
||||
𝓞.crAnF (ofStateAlgebra (StateAlgebra.timeOrder (StateAlgebra.ofState φ * StateAlgebra.ofState ψ))
|
||||
- normalOrder (ofState φ * ofState ψ))
|
||||
|
||||
lemma timeContract_eq_smul (φ ψ : 𝓕.States) : 𝓞.timeContract φ ψ =
|
||||
𝓞.crAnF (ofStateAlgebra (StateAlgebra.timeOrder
|
||||
(StateAlgebra.ofState φ * StateAlgebra.ofState ψ))
|
||||
+ (-1 : ℂ) • normalOrder (ofState φ * ofState ψ)) := by rfl
|
||||
|
||||
lemma timeContract_of_timeOrderRel (φ ψ : 𝓕.States) (h : timeOrderRel φ ψ) :
|
||||
𝓞.timeContract φ ψ = 𝓞.crAnF (⟨anPart (StateAlgebra.ofState φ), ofState ψ⟩ₛca) := by
|
||||
conv_rhs =>
|
||||
rw [ofState_eq_crPart_add_anPart]
|
||||
rw [map_add, map_add, crAnF_superCommute_anPart_anPart, superCommute_anPart_crPart]
|
||||
simp only [timeContract, instCommGroup.eq_1, Algebra.smul_mul_assoc, add_zero]
|
||||
rw [StateAlgebra.timeOrder_ofState_ofState_ordered h]
|
||||
rw [normalOrder_ofState_mul_ofState]
|
||||
rw [map_mul]
|
||||
simp only [ofStateAlgebra_ofState, instCommGroup.eq_1]
|
||||
rw [ofState_eq_crPart_add_anPart, ofState_eq_crPart_add_anPart]
|
||||
simp only [mul_add, add_mul]
|
||||
abel_nf
|
||||
|
||||
lemma timeContract_of_not_timeOrderRel (φ ψ : 𝓕.States) (h : ¬ timeOrderRel φ ψ) :
|
||||
𝓞.timeContract φ ψ = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ψ) • 𝓞.timeContract ψ φ := by
|
||||
rw [timeContract_eq_smul]
|
||||
simp only [Int.reduceNeg, one_smul, map_add]
|
||||
rw [map_smul]
|
||||
rw [crAnF_normalOrder_ofState_ofState_swap]
|
||||
rw [StateAlgebra.timeOrder_ofState_ofState_not_ordered_eq_timeOrder h]
|
||||
rw [timeContract_eq_smul]
|
||||
simp only [FieldStatistic.instCommGroup.eq_1, map_smul, one_smul, map_add, smul_add]
|
||||
rw [smul_smul, smul_smul, mul_comm]
|
||||
|
||||
lemma timeContract_mem_center (φ ψ : 𝓕.States) : 𝓞.timeContract φ ψ ∈ Subalgebra.center ℂ 𝓞.A := by
|
||||
by_cases h : timeOrderRel φ ψ
|
||||
· rw [timeContract_of_timeOrderRel _ _ _ h]
|
||||
exact 𝓞.crAnF_superCommute_anPart_ofState_mem_center _ _
|
||||
· rw [timeContract_of_not_timeOrderRel _ _ _ h]
|
||||
refine Subalgebra.smul_mem (Subalgebra.center ℂ 𝓞.A) ?_ 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ψ)
|
||||
rw [timeContract_of_timeOrderRel]
|
||||
exact 𝓞.crAnF_superCommute_anPart_ofState_mem_center _ _
|
||||
have h1 := IsTotal.total (r := 𝓕.timeOrderRel) φ ψ
|
||||
simp_all
|
||||
|
||||
lemma timeContract_zero_of_diff_grade (φ ψ : 𝓕.States) (h : (𝓕 |>ₛ φ) ≠ (𝓕 |>ₛ ψ)) :
|
||||
𝓞.timeContract φ ψ = 0 := by
|
||||
by_cases h1 : timeOrderRel φ ψ
|
||||
· rw [timeContract_of_timeOrderRel _ _ _ h1]
|
||||
rw [crAnF_superCommute_anPart_ofState_diff_grade_zero]
|
||||
exact h
|
||||
· rw [timeContract_of_not_timeOrderRel _ _ _ h1]
|
||||
rw [timeContract_of_timeOrderRel _ _ _]
|
||||
rw [crAnF_superCommute_anPart_ofState_diff_grade_zero]
|
||||
simp only [instCommGroup.eq_1, smul_zero]
|
||||
exact h.symm
|
||||
have ht := IsTotal.total (r := 𝓕.timeOrderRel) φ ψ
|
||||
simp_all
|
||||
|
||||
end OperatorAlgebra
|
||||
|
||||
end
|
||||
end FieldStruct
|
93
HepLean/PerturbationTheory/Algebras/StateAlgebra/Basic.lean
Normal file
93
HepLean/PerturbationTheory/Algebras/StateAlgebra/Basic.lean
Normal file
|
@ -0,0 +1,93 @@
|
|||
/-
|
||||
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.FieldStruct.CreateAnnihilate
|
||||
/-!
|
||||
|
||||
# State algebra
|
||||
|
||||
We define the state algebra of a field structure to be the free algebra
|
||||
generated by the states.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
/-- The state free-algebra.
|
||||
The free algebra generated by `States`,
|
||||
that is a position based states or assymptotic states.
|
||||
As a module `StateAlgebra` is spanned by lists of `States`. -/
|
||||
abbrev StateAlgebra (𝓕 : FieldStruct) : Type := FreeAlgebra ℂ 𝓕.States
|
||||
|
||||
namespace StateAlgebra
|
||||
|
||||
open FieldStatistic
|
||||
|
||||
/-- The element of the states free-algebra generated by a single state. -/
|
||||
def ofState (φ : 𝓕.States) : StateAlgebra 𝓕 :=
|
||||
FreeAlgebra.ι ℂ φ
|
||||
|
||||
/-- The element of the states free-algebra generated by a list of states. -/
|
||||
def ofList (φs : List 𝓕.States) : StateAlgebra 𝓕 :=
|
||||
(List.map ofState φs).prod
|
||||
|
||||
@[simp]
|
||||
lemma ofList_nil : ofList ([] : List 𝓕.States) = 1 := rfl
|
||||
|
||||
lemma ofList_singleton (φ : 𝓕.States) : ofList [φ] = ofState φ := by
|
||||
simp [ofList]
|
||||
|
||||
lemma ofList_append (φs ψs : List 𝓕.States) :
|
||||
ofList (φs ++ ψs) = ofList φs * ofList ψs := by
|
||||
rw [ofList, List.map_append, List.prod_append]
|
||||
rfl
|
||||
|
||||
lemma ofList_cons (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
ofList (φ :: φs) = ofState φ * ofList φs := rfl
|
||||
|
||||
/-- The basis of the free state algebra formed by lists of states. -/
|
||||
noncomputable def ofListBasis : Basis (List 𝓕.States) ℂ 𝓕.StateAlgebra where
|
||||
repr := FreeAlgebra.equivMonoidAlgebraFreeMonoid.toLinearEquiv
|
||||
|
||||
@[simp]
|
||||
lemma ofListBasis_eq_ofList (φs : List 𝓕.States) :
|
||||
ofListBasis φs = ofList φs := by
|
||||
simp only [ofListBasis, FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
Basis.coe_ofRepr, AlgEquiv.toLinearEquiv_symm, AlgEquiv.toLinearEquiv_apply,
|
||||
AlgEquiv.ofAlgHom_symm_apply, ofList]
|
||||
erw [MonoidAlgebra.lift_apply]
|
||||
simp only [zero_smul, Finsupp.sum_single_index, one_smul]
|
||||
rw [@FreeMonoid.lift_apply]
|
||||
simp only [List.prod]
|
||||
match φs with
|
||||
| [] => rfl
|
||||
| φ :: φs =>
|
||||
erw [List.map_cons]
|
||||
|
||||
/-!
|
||||
|
||||
## The super commutor on the state algebra.
|
||||
|
||||
-/
|
||||
|
||||
/-- The super commutor on the free state algebra. For two bosonic operators
|
||||
or a bosonic and fermionic operator this corresponds to the usual commutator
|
||||
whilst for two fermionic operators this corresponds to the anti-commutator. -/
|
||||
noncomputable def superCommute : 𝓕.StateAlgebra →ₗ[ℂ] 𝓕.StateAlgebra →ₗ[ℂ] 𝓕.StateAlgebra :=
|
||||
Basis.constr ofListBasis ℂ fun φs =>
|
||||
Basis.constr ofListBasis ℂ fun φs' =>
|
||||
ofList (φs ++ φs') - 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofList (φs' ++ φs)
|
||||
|
||||
local notation "⟨" φs "," φs' "⟩ₛ" => superCommute φs φs'
|
||||
|
||||
lemma superCommute_ofList (φs φs' : List 𝓕.States) : ⟨ofList φs, ofList φs'⟩ₛ =
|
||||
ofList (φs ++ φs') - 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofList (φs' ++ φs) := by
|
||||
rw [← ofListBasis_eq_ofList, ← ofListBasis_eq_ofList]
|
||||
simp only [superCommute, Basis.constr_basis]
|
||||
|
||||
end StateAlgebra
|
||||
|
||||
end FieldStruct
|
|
@ -0,0 +1,83 @@
|
|||
/-
|
||||
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.FieldStruct.TimeOrder
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
/-!
|
||||
|
||||
# State algebra
|
||||
|
||||
We define the state algebra of a field structure to be the free algebra
|
||||
generated by the states.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
noncomputable section
|
||||
|
||||
namespace StateAlgebra
|
||||
|
||||
open FieldStatistic
|
||||
|
||||
/-- The linear map on the free state algebra defined as the map taking
|
||||
a list of states to the time-ordered list of states multiplied by
|
||||
the sign corresponding to the number of fermionic-fermionic
|
||||
exchanges done in ordering. -/
|
||||
def timeOrder : StateAlgebra 𝓕 →ₗ[ℂ] StateAlgebra 𝓕 :=
|
||||
Basis.constr ofListBasis ℂ fun φs =>
|
||||
timeOrderSign φs • ofList (timeOrderList φs)
|
||||
|
||||
lemma timeOrder_ofList (φs : List 𝓕.States) :
|
||||
timeOrder (ofList φs) = timeOrderSign φs • ofList (timeOrderList φs) := by
|
||||
rw [← ofListBasis_eq_ofList]
|
||||
simp only [timeOrder, Basis.constr_basis]
|
||||
|
||||
lemma timeOrder_ofList_nil : timeOrder (𝓕 := 𝓕) (ofList []) = 1 := by
|
||||
rw [timeOrder_ofList]
|
||||
simp [timeOrderSign, Wick.koszulSign, timeOrderList]
|
||||
|
||||
@[simp]
|
||||
lemma timeOrder_ofList_singleton (φ : 𝓕.States) : timeOrder (ofList [φ]) = ofList [φ] := by
|
||||
simp [timeOrder_ofList, timeOrderSign, timeOrderList]
|
||||
|
||||
lemma timeOrder_ofState_ofState_ordered {φ ψ : 𝓕.States} (h : timeOrderRel φ ψ) :
|
||||
timeOrder (ofState φ * ofState ψ) = ofState φ * ofState ψ := by
|
||||
rw [← ofList_singleton, ← ofList_singleton, ← ofList_append, timeOrder_ofList]
|
||||
simp only [List.singleton_append]
|
||||
rw [timeOrderSign_pair_ordered h, timeOrderList_pair_ordered h]
|
||||
simp
|
||||
|
||||
lemma timeOrder_ofState_ofState_not_ordered {φ ψ : 𝓕.States} (h :¬ timeOrderRel φ ψ) :
|
||||
timeOrder (ofState φ * ofState ψ) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ψ) • ofState ψ * ofState φ := by
|
||||
rw [← ofList_singleton, ← ofList_singleton, ← ofList_append, timeOrder_ofList]
|
||||
simp only [List.singleton_append, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [timeOrderSign_pair_not_ordered h, timeOrderList_pair_not_ordered h]
|
||||
simp [← ofList_append]
|
||||
|
||||
lemma timeOrder_ofState_ofState_not_ordered_eq_timeOrder {φ ψ : 𝓕.States} (h :¬ timeOrderRel φ ψ) :
|
||||
timeOrder (ofState φ * ofState ψ) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ψ) • timeOrder (ofState ψ * ofState φ) := by
|
||||
rw [timeOrder_ofState_ofState_not_ordered h]
|
||||
rw [timeOrder_ofState_ofState_ordered]
|
||||
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
have hx := IsTotal.total (r := timeOrderRel) ψ φ
|
||||
simp_all
|
||||
|
||||
lemma timeOrder_eq_maxTimeField_mul (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
timeOrder (ofList (φ :: φs)) =
|
||||
𝓢(𝓕 |>ₛ maxTimeField φ φs, 𝓕 |>ₛ (φ :: φs).take (maxTimeFieldPos φ φs)) •
|
||||
ofState (maxTimeField φ φs) * timeOrder (ofList (eraseMaxTimeField φ φs)) := by
|
||||
rw [timeOrder_ofList, timeOrderList_eq_maxTimeField_timeOrderList]
|
||||
rw [ofList_cons, timeOrder_ofList]
|
||||
simp only [instCommGroup.eq_1, Algebra.mul_smul_comm, Algebra.smul_mul_assoc, smul_smul]
|
||||
congr
|
||||
rw [timerOrderSign_of_eraseMaxTimeField, mul_assoc]
|
||||
simp
|
||||
|
||||
end StateAlgebra
|
||||
end
|
||||
end FieldStruct
|
|
@ -1,292 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.Wick.OperatorMap
|
||||
import HepLean.Mathematics.Fin.Involutions
|
||||
/-!
|
||||
|
||||
# Contractions of a list of fields
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
|
||||
open HepLean.List
|
||||
open HepLean.Fin
|
||||
open FieldStatistic
|
||||
|
||||
variable {𝓕 : Type}
|
||||
|
||||
/-- Given a list of fields `φs`, the type of pairwise-contractions associated with `φs`
|
||||
which have the list `φsᵤₙ` uncontracted. -/
|
||||
inductive ContractionsAux : (φs : List 𝓕) → (φsᵤₙ : List 𝓕) → Type
|
||||
| nil : ContractionsAux [] []
|
||||
| cons {φs : List 𝓕} {φsᵤₙ : List 𝓕} {φ : 𝓕} (i : Option (Fin φsᵤₙ.length)) :
|
||||
ContractionsAux φs φsᵤₙ → ContractionsAux (φ :: φs) (optionEraseZ φsᵤₙ φ i)
|
||||
|
||||
/-- Given a list of fields `l`, the type of pairwise-contractions associated with `l`. -/
|
||||
def Contractions (φs : List 𝓕) : Type := Σ aux, ContractionsAux φs aux
|
||||
|
||||
namespace Contractions
|
||||
|
||||
variable {l : List 𝓕} (c : Contractions l)
|
||||
|
||||
/-- The equivalence between `ContractionsAux` based on the propositionally equivalent
|
||||
uncontracted lists. -/
|
||||
def auxCongr : {φs : List 𝓕} → {φsᵤₙ φsᵤₙ' : List 𝓕} → (h : φsᵤₙ = φsᵤₙ') →
|
||||
ContractionsAux φs φsᵤₙ ≃ ContractionsAux φs φsᵤₙ'
|
||||
| _, _, _, Eq.refl _ => Equiv.refl _
|
||||
|
||||
lemma auxCongr_ext {φs : List 𝓕} {c c2 : Contractions φs} (h : c.1 = c2.1)
|
||||
(hx : c.2 = auxCongr h.symm c2.2) : c = c2 := by
|
||||
cases c
|
||||
cases c2
|
||||
simp only at h
|
||||
subst h
|
||||
simp only [auxCongr, Equiv.refl_apply] at hx
|
||||
subst hx
|
||||
rfl
|
||||
|
||||
/-- The list of uncontracted fields. -/
|
||||
def uncontracted : List 𝓕 := c.1
|
||||
|
||||
lemma uncontracted_length_even_iff : {l : List 𝓕} → (c : Contractions l) →
|
||||
Even l.length ↔ Even c.uncontracted.length
|
||||
| [], ⟨[], ContractionsAux.nil⟩ => by
|
||||
simp [uncontracted]
|
||||
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => by
|
||||
simp only [List.length_cons, uncontracted, optionEraseZ]
|
||||
rw [Nat.even_add_one, Nat.even_add_one]
|
||||
rw [uncontracted_length_even_iff ⟨aux, c⟩]
|
||||
rfl
|
||||
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩=> by
|
||||
simp only [List.length_cons, uncontracted, optionEraseZ_some_length]
|
||||
rw [Nat.even_sub, Nat.even_add_one]
|
||||
· simp only [Nat.not_even_iff_odd, Nat.not_even_one, iff_false]
|
||||
rw [← Nat.not_even_iff_odd, ← Nat.not_even_iff_odd]
|
||||
rw [uncontracted_length_even_iff ⟨aux, c⟩]
|
||||
rfl
|
||||
· refine Nat.one_le_iff_ne_zero.mpr (fun hn => ?_)
|
||||
rw [hn] at n
|
||||
exact Fin.elim0 n
|
||||
|
||||
lemma contractions_nil (a : Contractions ([] : List 𝓕)) : a = ⟨[], ContractionsAux.nil⟩ := by
|
||||
cases a
|
||||
rename_i aux c
|
||||
cases c
|
||||
rfl
|
||||
|
||||
/-- The embedding of the uncontracted fields into all fields. -/
|
||||
def embedUncontracted {l : List 𝓕} (c : Contractions l) :
|
||||
Fin c.uncontracted.length → Fin l.length :=
|
||||
match l, c with
|
||||
| [], ⟨[], ContractionsAux.nil⟩ => Fin.elim0
|
||||
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ =>
|
||||
Fin.cons ⟨0, Nat.zero_lt_succ φs.length⟩ (Fin.succ ∘ (embedUncontracted ⟨aux, c⟩))
|
||||
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => by
|
||||
let lc := embedUncontracted ⟨aux, c⟩
|
||||
refine Fin.succ ∘ lc ∘ Fin.cast ?_ ∘ Fin.succAbove ⟨n, by
|
||||
rw [uncontracted, optionEraseZ_some_length]
|
||||
omega⟩
|
||||
simp only [uncontracted, optionEraseZ_some_length]
|
||||
have hq : aux.length ≠ 0 := by
|
||||
by_contra hn
|
||||
rw [hn] at n
|
||||
exact Fin.elim0 n
|
||||
omega
|
||||
|
||||
lemma embedUncontracted_injective {l : List 𝓕} (c : Contractions l) :
|
||||
Function.Injective c.embedUncontracted := by
|
||||
match l, c with
|
||||
| [], ⟨[], ContractionsAux.nil⟩ =>
|
||||
dsimp only [List.length_nil, embedUncontracted]
|
||||
intro i
|
||||
exact Fin.elim0 i
|
||||
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ =>
|
||||
dsimp only [List.length_cons, embedUncontracted, Fin.zero_eta]
|
||||
refine Fin.cons_injective_iff.mpr ?_
|
||||
apply And.intro
|
||||
· simp only [Set.mem_range, Function.comp_apply, not_exists]
|
||||
exact fun x => Fin.succ_ne_zero (embedUncontracted ⟨aux, c⟩ x)
|
||||
· exact Function.Injective.comp (Fin.succ_injective φs.length)
|
||||
(embedUncontracted_injective ⟨aux, c⟩)
|
||||
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some i) c⟩ =>
|
||||
dsimp only [List.length_cons, embedUncontracted]
|
||||
refine Function.Injective.comp (Fin.succ_injective φs.length) ?hf
|
||||
refine Function.Injective.comp (embedUncontracted_injective ⟨aux, c⟩) ?hf.hf
|
||||
refine Function.Injective.comp (Fin.cast_injective (embedUncontracted.proof_5 φ φs aux i c))
|
||||
Fin.succAbove_right_injective
|
||||
|
||||
/-- Establishes uniqueness of contractions for a single field, showing that any contraction
|
||||
of a single field must be equivalent to the trivial contraction with no pairing. -/
|
||||
lemma contractions_single {i : 𝓕} (a : Contractions [i]) : a =
|
||||
⟨[i], ContractionsAux.cons none ContractionsAux.nil⟩ := by
|
||||
cases a
|
||||
rename_i aux c
|
||||
cases c
|
||||
rename_i aux' c'
|
||||
cases c'
|
||||
cases aux'
|
||||
simp only [List.length_nil, optionEraseZ]
|
||||
rename_i x
|
||||
exact Fin.elim0 x
|
||||
|
||||
/--
|
||||
This function provides an equivalence between the type of contractions for an empty list of fields
|
||||
and the unit type, indicating that there is only one possible contraction for an empty list.
|
||||
-/
|
||||
def nilEquiv : Contractions ([] : List 𝓕) ≃ Unit where
|
||||
toFun _ := ()
|
||||
invFun _ := ⟨[], ContractionsAux.nil⟩
|
||||
left_inv a := Eq.symm (contractions_nil a)
|
||||
right_inv _ := rfl
|
||||
|
||||
/-- The equivalence between contractions of `a :: l` and contractions of
|
||||
`Contractions l` paired with an optional element of `Fin (c.uncontracted).length` specifying
|
||||
what `a` contracts with if any. -/
|
||||
def consEquiv {φ : 𝓕} {φs : List 𝓕} :
|
||||
Contractions (φ :: φs) ≃ (c : Contractions φs) × Option (Fin c.uncontracted.length) where
|
||||
toFun c :=
|
||||
match c with
|
||||
| ⟨aux, c⟩ =>
|
||||
match c with
|
||||
| ContractionsAux.cons (φsᵤₙ := aux') i c => ⟨⟨aux', c⟩, i⟩
|
||||
invFun ci :=
|
||||
⟨(optionEraseZ (ci.fst.uncontracted) φ ci.2), ContractionsAux.cons (φ := φ) ci.2 ci.1.2⟩
|
||||
left_inv c := by
|
||||
match c with
|
||||
| ⟨aux, c⟩ =>
|
||||
match c with
|
||||
| ContractionsAux.cons (φsᵤₙ := aux') i c => rfl
|
||||
right_inv ci := by rfl
|
||||
|
||||
lemma consEquiv_ext {φs : List 𝓕} {c1 c2 : Contractions φs}
|
||||
{n1 : Option (Fin c1.uncontracted.length)} {n2 : Option (Fin c2.uncontracted.length)}
|
||||
(hc : c1 = c2) (hn : Option.map (finCongr (by rw [hc])) n1 = n2) :
|
||||
(⟨c1, n1⟩ : (c : Contractions φs) × Option (Fin c.uncontracted.length)) = ⟨c2, n2⟩ := by
|
||||
subst hc
|
||||
subst hn
|
||||
simp
|
||||
|
||||
/-- The type of contractions is decidable. -/
|
||||
instance decidable : (φs : List 𝓕) → DecidableEq (Contractions φs)
|
||||
| [] => fun a b =>
|
||||
match a, b with
|
||||
| ⟨_, a⟩, ⟨_, b⟩ =>
|
||||
match a, b with
|
||||
| ContractionsAux.nil, ContractionsAux.nil => isTrue rfl
|
||||
| _ :: φs =>
|
||||
haveI : DecidableEq (Contractions φs) := decidable φs
|
||||
haveI : DecidableEq ((c : Contractions φs) × Option (Fin (c.uncontracted).length)) :=
|
||||
Sigma.instDecidableEqSigma
|
||||
Equiv.decidableEq consEquiv
|
||||
|
||||
/-- The type of contractions is finite. -/
|
||||
instance fintype : (φs : List 𝓕) → Fintype (Contractions φs)
|
||||
| [] => {
|
||||
elems := {⟨[], ContractionsAux.nil⟩}
|
||||
complete := by
|
||||
intro a
|
||||
rw [Finset.mem_singleton]
|
||||
exact contractions_nil a}
|
||||
| φ :: φs =>
|
||||
haveI : Fintype (Contractions φs) := fintype φs
|
||||
haveI : Fintype ((c : Contractions φs) × Option (Fin (c.uncontracted).length)) :=
|
||||
Sigma.instFintype
|
||||
Fintype.ofEquiv _ consEquiv.symm
|
||||
|
||||
/-- A contraction is a full contraction if there normalizing list of fields is empty. -/
|
||||
def IsFull : Prop := c.uncontracted = []
|
||||
|
||||
/-- Provides a decidable instance for determining if a contraction is full
|
||||
(i.e., all fields are paired). -/
|
||||
instance isFull_decidable : Decidable c.IsFull := by
|
||||
have hn : c.IsFull ↔ c.uncontracted.length = 0 := by
|
||||
simp [IsFull]
|
||||
apply decidable_of_decidable_of_iff hn.symm
|
||||
|
||||
/-- A structure specifying when a type `I` and a map `f :I → Type` corresponds to
|
||||
the splitting of a fields `φ` into a creation `n` and annihilation part `p`. -/
|
||||
structure Splitting (f : 𝓕 → Type) [∀ i, Fintype (f i)]
|
||||
(le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le] where
|
||||
/-- The creation part of the fields. The label `n` corresponds to the fact that
|
||||
in normal ordering these fields get pushed to the negative (left) direction. -/
|
||||
𝓑n : 𝓕 → (Σ i, f i)
|
||||
/-- The annihilation part of the fields. The label `p` corresponds to the fact that
|
||||
in normal ordering these fields get pushed to the positive (right) direction. -/
|
||||
𝓑p : 𝓕 → (Σ i, f i)
|
||||
/-- The complex coefficient of creation part of a field `i`. This is usually `0` or `1`. -/
|
||||
𝓧n : 𝓕 → ℂ
|
||||
/-- The complex coefficient of annihilation part of a field `i`. This is usually `0` or `1`. -/
|
||||
𝓧p : 𝓕 → ℂ
|
||||
h𝓑 : ∀ i, ofListLift f [i] 1 = ofList [𝓑n i] (𝓧n i) + ofList [𝓑p i] (𝓧p i)
|
||||
h𝓑n : ∀ i j, le (𝓑n i) j
|
||||
h𝓑p : ∀ i j, le j (𝓑p i)
|
||||
|
||||
/-- In the static wick's theorem, the term associated with a contraction `c` containing
|
||||
the contractions. -/
|
||||
noncomputable def toCenterTerm (f : 𝓕 → Type) [∀ i, Fintype (f i)]
|
||||
(q : 𝓕 → FieldStatistic)
|
||||
(le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le]
|
||||
{A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(F : FreeAlgebra ℂ (Σ i, f i) →ₐ[ℂ] A) :
|
||||
{φs : List 𝓕} → (c : Contractions φs) → (S : Splitting f le) → A
|
||||
| [], ⟨[], .nil⟩, _ => 1
|
||||
| _ :: _, ⟨_, .cons (φsᵤₙ := aux') none c⟩, S => toCenterTerm f q le F ⟨aux', c⟩ S
|
||||
| a :: _, ⟨_, .cons (φsᵤₙ := aux') (some n) c⟩, S => toCenterTerm f q le F ⟨aux', c⟩ S *
|
||||
superCommuteCoef q [aux'.get n] (List.take (↑n) aux') •
|
||||
F (((superCommute fun i => q i.fst) (ofList [S.𝓑p a] (S.𝓧p a)))
|
||||
(ofListLift f [aux'.get n] 1))
|
||||
|
||||
/-- Shows that adding a field with no contractions (none) to an existing set of contractions
|
||||
results in the same center term as the original contractions.
|
||||
|
||||
Physically, this represents that an uncontracted (free) field does not affect
|
||||
the contraction structure of other fields in Wick's theorem. -/
|
||||
lemma toCenterTerm_none (f : 𝓕 → Type) [∀ i, Fintype (f i)]
|
||||
(q : 𝓕 → FieldStatistic) {φs : List 𝓕}
|
||||
(le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le]
|
||||
{A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(F : FreeAlgebra ℂ (Σ i, f i) →ₐ A)
|
||||
(S : Splitting f le) (φ : 𝓕) (c : Contractions φs) :
|
||||
toCenterTerm (φs := φ :: φs) f q le F (Contractions.consEquiv.symm ⟨c, none⟩) S =
|
||||
toCenterTerm f q le F c S := by
|
||||
rw [consEquiv]
|
||||
simp only [Equiv.coe_fn_symm_mk]
|
||||
dsimp only [toCenterTerm]
|
||||
rfl
|
||||
|
||||
/-- Proves that the part of the term gained from Wick contractions is in
|
||||
the center of the algebra. -/
|
||||
lemma toCenterTerm_center (f : 𝓕 → Type) [∀ i, Fintype (f i)]
|
||||
(q : 𝓕 → FieldStatistic)
|
||||
(le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le]
|
||||
{A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(F : FreeAlgebra ℂ (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le F] :
|
||||
{φs : List 𝓕} → (c : Contractions φs) → (S : Splitting f le) →
|
||||
(c.toCenterTerm f q le F S) ∈ Subalgebra.center ℂ A
|
||||
| [], ⟨[], .nil⟩, _ => by
|
||||
dsimp only [toCenterTerm]
|
||||
exact Subalgebra.one_mem (Subalgebra.center ℂ A)
|
||||
| _ :: _, ⟨_, .cons (φsᵤₙ := aux') none c⟩, S => by
|
||||
dsimp only [toCenterTerm]
|
||||
exact toCenterTerm_center f q le F ⟨aux', c⟩ S
|
||||
| a :: _, ⟨_, .cons (φsᵤₙ := aux') (some n) c⟩, S => by
|
||||
dsimp only [toCenterTerm, List.get_eq_getElem]
|
||||
refine Subalgebra.mul_mem (Subalgebra.center ℂ A) ?hx ?hy
|
||||
exact toCenterTerm_center f q le F ⟨aux', c⟩ S
|
||||
apply Subalgebra.smul_mem
|
||||
rw [ofListLift_expand]
|
||||
rw [map_sum, map_sum]
|
||||
refine Subalgebra.sum_mem (Subalgebra.center ℂ A) ?hy.hx.h
|
||||
intro x _
|
||||
simp only [CreateAnnihilateSect.toList]
|
||||
rw [ofList_singleton]
|
||||
exact OperatorMap.superCommute_ofList_singleton_ι_center (q := fun i => q i.1)
|
||||
(le := le) F (S.𝓑p a) ⟨aux'[↑n], x.head⟩
|
||||
|
||||
end Contractions
|
||||
|
||||
end Wick
|
|
@ -1,33 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.Contractions.Involutions
|
||||
/-!
|
||||
|
||||
# Cardinality of full contractions
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
namespace Contractions
|
||||
|
||||
open HepLean.Fin
|
||||
open Nat
|
||||
|
||||
/-- There are `(φs.length - 1)‼` full contractions of a list `φs` with an even number of fields. -/
|
||||
lemma card_of_full_contractions_even {φs : List 𝓕} (he : Even φs.length) :
|
||||
Fintype.card {c : Contractions φs // IsFull c} = (φs.length - 1)‼ := by
|
||||
rw [Fintype.card_congr (isFullInvolutionEquiv (φs := φs))]
|
||||
exact involutionNoFixed_card_even φs.length he
|
||||
|
||||
/-- There are no full contractions of a list with an odd number of fields. -/
|
||||
lemma card_of_full_contractions_odd {φs : List 𝓕} (ho : Odd φs.length) :
|
||||
Fintype.card {c : Contractions φs // IsFull c} = 0 := by
|
||||
rw [Fintype.card_congr (isFullInvolutionEquiv (φs := φs))]
|
||||
exact involutionNoFixed_card_odd φs.length ho
|
||||
|
||||
end Contractions
|
||||
|
||||
end Wick
|
|
@ -1,393 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.Contractions.Basic
|
||||
/-!
|
||||
|
||||
# Involutions
|
||||
|
||||
There is an isomorphism between the type of contractions of a list `l` and
|
||||
the type of involutions from `Fin l.length` to `Fin l.length`.
|
||||
|
||||
Likewise, there is an isomorphism from the type of full contractions of a list `l`
|
||||
and the type of fixed-point free involutions from `Fin l.length` to `Fin l.length`.
|
||||
|
||||
Given this, the number of full contractions of a list `l` is
|
||||
is given by the OEIS sequence A000085.
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
|
||||
open HepLean.List
|
||||
open HepLean.Fin
|
||||
open FieldStatistic
|
||||
|
||||
variable {𝓕 : Type}
|
||||
namespace Contractions
|
||||
|
||||
variable {l : List 𝓕}
|
||||
|
||||
/-!
|
||||
|
||||
## From Involution.
|
||||
|
||||
-/
|
||||
|
||||
/-- Given an involution the uncontracted fields associated with it (corresponding
|
||||
to the fixed points of that involution). -/
|
||||
def uncontractedFromInvolution : {φs : List 𝓕} →
|
||||
(f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) →
|
||||
{l : List 𝓕 // l.length = (Finset.univ.filter fun i => f.1 i = i).card}
|
||||
| [], _ => ⟨[], by simp⟩
|
||||
| φ :: φs, f =>
|
||||
let luc := uncontractedFromInvolution (involutionCons φs.length f).fst
|
||||
let n' := involutionAddEquiv (involutionCons φs.length f).1 (involutionCons φs.length f).2
|
||||
let np : Option (Fin luc.1.length) := Option.map (finCongr luc.2.symm) n'
|
||||
if hn : n' = none then
|
||||
have hn' := involutionAddEquiv_none_image_zero (n := φs.length) (f := f) hn
|
||||
⟨optionEraseZ luc φ none, by
|
||||
simp only [optionEraseZ, Nat.succ_eq_add_one, List.length_cons, List.Vector.length_val]
|
||||
rw [← luc.2]
|
||||
conv_rhs => rw [Finset.card_filter]
|
||||
rw [Fin.sum_univ_succ]
|
||||
conv_rhs => erw [if_pos hn']
|
||||
ring_nf
|
||||
simp only [Nat.succ_eq_add_one, List.Vector.length_val, Nat.cast_id,
|
||||
add_right_inj]
|
||||
rw [Finset.card_filter]
|
||||
apply congrArg
|
||||
funext i
|
||||
refine ite_congr ?h.h.h₁ (congrFun rfl) (congrFun rfl)
|
||||
rw [involutionAddEquiv_none_succ hn]⟩
|
||||
else
|
||||
let n := n'.get (Option.isSome_iff_ne_none.mpr hn)
|
||||
let np : Fin luc.1.length := Fin.cast luc.2.symm n
|
||||
⟨optionEraseZ luc φ (some np), by
|
||||
let k' := (involutionCons φs.length f).2
|
||||
have hkIsSome : (k'.1).isSome := by
|
||||
simp only [Nat.succ_eq_add_one, involutionAddEquiv, Option.isSome_some, Option.get_some,
|
||||
Option.isSome_none, Equiv.trans_apply, Equiv.coe_fn_mk, Equiv.optionCongr_apply,
|
||||
Equiv.coe_trans, RelIso.coe_fn_toEquiv, Option.map_eq_none', n'] at hn
|
||||
split at hn
|
||||
· simp_all only [reduceCtorEq, not_false_eq_true, Nat.succ_eq_add_one,
|
||||
Option.isSome_some, k']
|
||||
· simp_all only [not_true_eq_false]
|
||||
let k := k'.1.get hkIsSome
|
||||
rw [optionEraseZ_some_length]
|
||||
have hksucc : k.succ = f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ := by
|
||||
simp [k, k', involutionCons]
|
||||
have hzero : ⟨0, Nat.zero_lt_succ φs.length⟩ = f.1 k.succ := by
|
||||
rw [hksucc, f.2]
|
||||
have hksuccNe : f.1 k.succ ≠ k.succ := by
|
||||
conv_rhs => rw [hksucc]
|
||||
exact fun hn => Fin.succ_ne_zero k (Function.Involutive.injective f.2 hn)
|
||||
have hluc : 1 ≤ luc.1.length := by
|
||||
simp only [Nat.succ_eq_add_one, List.Vector.length_val, Finset.one_le_card]
|
||||
use k
|
||||
simp only [involutionCons, Nat.succ_eq_add_one, Fin.cons_update, Equiv.coe_fn_mk,
|
||||
dite_eq_left_iff, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
rw [hksucc, f.2]
|
||||
simp
|
||||
rw [propext (Nat.sub_eq_iff_eq_add' hluc)]
|
||||
conv_rhs =>
|
||||
rw [Finset.card_filter]
|
||||
erw [Fin.sum_univ_succ, if_neg (Option.isSome_dite'.mp hkIsSome)]
|
||||
simp only [Nat.succ_eq_add_one, List.Vector.length_val, List.length_cons,
|
||||
Nat.cast_id, zero_add]
|
||||
conv_rhs => lhs; rw [Eq.symm (Fintype.sum_ite_eq' k fun j => 1)]
|
||||
rw [← Finset.sum_add_distrib, Finset.card_filter]
|
||||
apply congrArg
|
||||
funext i
|
||||
by_cases hik : i = k
|
||||
· subst hik
|
||||
simp only [k'.2 hkIsSome, Nat.succ_eq_add_one, ↓reduceIte, hksuccNe, add_zero]
|
||||
· simp only [hik, ↓reduceIte, zero_add]
|
||||
refine ite_congr ?_ (congrFun rfl) (congrFun rfl)
|
||||
simp only [involutionCons, Nat.succ_eq_add_one, Fin.cons_update, Equiv.coe_fn_mk,
|
||||
dite_eq_left_iff, eq_iff_iff]
|
||||
have hfi : f.1 i.succ ≠ ⟨0, Nat.zero_lt_succ φs.length⟩ := by
|
||||
rw [hzero]
|
||||
by_contra hn
|
||||
have hik' := (Function.Involutive.injective f.2 hn)
|
||||
simp only [List.length_cons, Fin.succ_inj] at hik'
|
||||
exact hik hik'
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
have h' := h hfi
|
||||
conv_rhs => rw [← h']
|
||||
simp
|
||||
· intro h hfi
|
||||
simp only [Fin.ext_iff, Fin.coe_pred]
|
||||
rw [h]
|
||||
simp⟩
|
||||
|
||||
lemma uncontractedFromInvolution_cons {φs : List 𝓕} {φ : 𝓕}
|
||||
(f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}) :
|
||||
uncontractedFromInvolution f =
|
||||
optionEraseZ (uncontractedFromInvolution (involutionCons φs.length f).fst) φ
|
||||
(Option.map (finCongr ((uncontractedFromInvolution (involutionCons φs.length f).fst).2.symm))
|
||||
(involutionAddEquiv (involutionCons φs.length f).1 (involutionCons φs.length f).2)) := by
|
||||
let luc := uncontractedFromInvolution (involutionCons φs.length f).fst
|
||||
let n' := involutionAddEquiv (involutionCons φs.length f).1 (involutionCons φs.length f).2
|
||||
change _ = optionEraseZ luc φ (Option.map
|
||||
(finCongr ((uncontractedFromInvolution (involutionCons φs.length f).fst).2.symm)) n')
|
||||
dsimp only [List.length_cons, uncontractedFromInvolution, Nat.succ_eq_add_one, Fin.zero_eta]
|
||||
by_cases hn : n' = none
|
||||
· have hn' := hn
|
||||
simp only [Nat.succ_eq_add_one, n'] at hn'
|
||||
simp only [hn', ↓reduceDIte, hn]
|
||||
rfl
|
||||
· have hn' := hn
|
||||
simp only [Nat.succ_eq_add_one, n'] at hn'
|
||||
simp only [hn', ↓reduceDIte]
|
||||
congr
|
||||
simp only [Nat.succ_eq_add_one, n']
|
||||
simp_all only [Nat.succ_eq_add_one, not_false_eq_true, n', luc]
|
||||
obtain ⟨val, property⟩ := f
|
||||
obtain ⟨val_1, property_1⟩ := luc
|
||||
simp_all only [Nat.succ_eq_add_one, List.length_cons]
|
||||
ext a : 1
|
||||
simp_all only [Option.mem_def, Option.some.injEq, Option.map_eq_some', finCongr_apply]
|
||||
apply Iff.intro
|
||||
· intro a_1
|
||||
subst a_1
|
||||
apply Exists.intro
|
||||
· apply And.intro
|
||||
on_goal 2 => {rfl
|
||||
}
|
||||
· simp_all only [Option.some_get]
|
||||
· intro a_1
|
||||
obtain ⟨w, h⟩ := a_1
|
||||
obtain ⟨left, right⟩ := h
|
||||
subst right
|
||||
simp_all only [Option.get_some]
|
||||
|
||||
/-- The `ContractionsAux` associated to an involution. -/
|
||||
def fromInvolutionAux : {l : List 𝓕} →
|
||||
(f : {f : Fin l.length → Fin l.length // Function.Involutive f}) →
|
||||
ContractionsAux l (uncontractedFromInvolution f)
|
||||
| [] => fun _ => ContractionsAux.nil
|
||||
| _ :: φs => fun f =>
|
||||
let f' := involutionCons φs.length f
|
||||
let c' := fromInvolutionAux f'.1
|
||||
let n' := Option.map (finCongr ((uncontractedFromInvolution f'.fst).2.symm))
|
||||
(involutionAddEquiv f'.1 f'.2)
|
||||
auxCongr (uncontractedFromInvolution_cons f).symm (ContractionsAux.cons n' c')
|
||||
|
||||
/-- The contraction associated with an involution. -/
|
||||
def fromInvolution {φs : List 𝓕}
|
||||
(f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) : Contractions φs :=
|
||||
⟨uncontractedFromInvolution f, fromInvolutionAux f⟩
|
||||
|
||||
lemma fromInvolution_cons {φs : List 𝓕} {φ : 𝓕}
|
||||
(f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}) :
|
||||
let f' := involutionCons φs.length f
|
||||
fromInvolution f = consEquiv.symm
|
||||
⟨fromInvolution f'.1, Option.map (finCongr ((uncontractedFromInvolution f'.fst).2.symm))
|
||||
(involutionAddEquiv f'.1 f'.2)⟩ := by
|
||||
refine auxCongr_ext ?_ ?_
|
||||
· dsimp only [fromInvolution, List.length_cons, Nat.succ_eq_add_one]
|
||||
rw [uncontractedFromInvolution_cons]
|
||||
rfl
|
||||
· dsimp only [fromInvolution, List.length_cons, fromInvolutionAux, Nat.succ_eq_add_one, id_eq,
|
||||
eq_mpr_eq_cast]
|
||||
rfl
|
||||
|
||||
lemma fromInvolution_of_involutionCons {φs : List 𝓕} {φ : 𝓕}
|
||||
(f : {f : Fin φs.length → Fin φs.length // Function.Involutive f})
|
||||
(n : { i : Option (Fin φs.length) // ∀ (h : i.isSome = true), f.1 (i.get h) = i.get h }) :
|
||||
fromInvolution (φs := φ :: φs) ((involutionCons φs.length).symm ⟨f, n⟩) =
|
||||
consEquiv.symm ⟨fromInvolution f, Option.map (finCongr ((uncontractedFromInvolution f).2.symm))
|
||||
(involutionAddEquiv f n)⟩ := by
|
||||
rw [fromInvolution_cons]
|
||||
congr 1
|
||||
simp only [Nat.succ_eq_add_one, Sigma.mk.inj_iff, Equiv.apply_symm_apply, true_and]
|
||||
rw [Equiv.apply_symm_apply]
|
||||
|
||||
/-!
|
||||
|
||||
## To Involution.
|
||||
|
||||
-/
|
||||
|
||||
/-- The involution associated with a contraction. -/
|
||||
def toInvolution : {φs : List 𝓕} → (c : Contractions φs) →
|
||||
{f : {f : Fin φs.length → Fin φs.length // Function.Involutive f} //
|
||||
uncontractedFromInvolution f = c.1}
|
||||
| [], ⟨[], ContractionsAux.nil⟩ => ⟨⟨fun i => i, by
|
||||
intro i
|
||||
simp⟩, by rfl⟩
|
||||
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) n c⟩ => by
|
||||
let ⟨⟨f', hf1⟩, hf2⟩ := toInvolution ⟨aux, c⟩
|
||||
let n' : Option (Fin (uncontractedFromInvolution ⟨f', hf1⟩).1.length) :=
|
||||
Option.map (finCongr (by rw [hf2])) n
|
||||
let F := (involutionCons φs.length).symm ⟨⟨f', hf1⟩,
|
||||
(involutionAddEquiv ⟨f', hf1⟩).symm
|
||||
(Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩
|
||||
refine ⟨F, ?_⟩
|
||||
have hF0 : ((involutionCons φs.length) F) = ⟨⟨f', hf1⟩,
|
||||
(involutionAddEquiv ⟨f', hf1⟩).symm
|
||||
(Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩ := by
|
||||
simp [F]
|
||||
have hF1 : ((involutionCons φs.length) F).fst = ⟨f', hf1⟩ := by
|
||||
rw [hF0]
|
||||
have hF2L : ((uncontractedFromInvolution ⟨f', hf1⟩)).1.length =
|
||||
(Finset.filter (fun i => ((involutionCons φs.length) F).1.1 i = i) Finset.univ).card := by
|
||||
apply Eq.trans ((uncontractedFromInvolution ⟨f', hf1⟩)).2
|
||||
congr
|
||||
rw [hF1]
|
||||
have hF2 : ((involutionCons φs.length) F).snd =
|
||||
(involutionAddEquiv ((involutionCons φs.length) F).fst).symm
|
||||
(Option.map (finCongr hF2L) n') := by
|
||||
rw [@Sigma.subtype_ext_iff] at hF0
|
||||
ext1
|
||||
rw [hF0.2]
|
||||
simp only [Nat.succ_eq_add_one]
|
||||
congr 1
|
||||
· rw [hF1]
|
||||
· refine involutionAddEquiv_cast' ?_ n' _ _
|
||||
rw [hF1]
|
||||
rw [uncontractedFromInvolution_cons]
|
||||
have hx := (toInvolution ⟨aux, c⟩).2
|
||||
simp only at hx
|
||||
simp only [Nat.succ_eq_add_one]
|
||||
refine optionEraseZ_ext ?_ ?_ ?_
|
||||
· dsimp only [F]
|
||||
rw [Equiv.apply_symm_apply]
|
||||
simp only
|
||||
rw [← hx]
|
||||
simp_all only
|
||||
· rfl
|
||||
· simp only [hF2, Nat.succ_eq_add_one, Equiv.apply_symm_apply, Option.map_map]
|
||||
dsimp only [id_eq, eq_mpr_eq_cast, Nat.succ_eq_add_one, n']
|
||||
simp only [finCongr, Equiv.coe_fn_mk, Option.map_map]
|
||||
simp only [Nat.succ_eq_add_one, id_eq, eq_mpr_eq_cast, F, n']
|
||||
ext a : 1
|
||||
simp only [Option.mem_def, Option.map_eq_some', Function.comp_apply, Fin.cast_trans,
|
||||
Fin.cast_eq_self, exists_eq_right]
|
||||
|
||||
lemma toInvolution_length_uncontracted {φs φsᵤₙ : List 𝓕} {c : ContractionsAux φs φsᵤₙ} :
|
||||
φsᵤₙ.length =
|
||||
(Finset.filter (fun i => (toInvolution ⟨φsᵤₙ, c⟩).1.1 i = i) Finset.univ).card := by
|
||||
have h2 := (toInvolution ⟨φsᵤₙ, c⟩).2
|
||||
simp only at h2
|
||||
conv_lhs => rw [← h2]
|
||||
exact List.Vector.length_val (uncontractedFromInvolution (toInvolution ⟨φsᵤₙ, c⟩).1)
|
||||
|
||||
lemma toInvolution_cons {φs φsᵤₙ : List 𝓕} {φ : 𝓕}
|
||||
(c : ContractionsAux φs φsᵤₙ) (n : Option (Fin (φsᵤₙ.length))) :
|
||||
(toInvolution ⟨optionEraseZ φsᵤₙ φ n, ContractionsAux.cons n c⟩).1
|
||||
= (involutionCons φs.length).symm ⟨(toInvolution ⟨φsᵤₙ, c⟩).1,
|
||||
(involutionAddEquiv (toInvolution ⟨φsᵤₙ, c⟩).1).symm
|
||||
(Option.map (finCongr (toInvolution_length_uncontracted)) n)⟩ := by
|
||||
dsimp only [List.length_cons, toInvolution, Nat.succ_eq_add_one, subset_refl, Set.coe_inclusion]
|
||||
congr 3
|
||||
rw [Option.map_map]
|
||||
simp only [finCongr, Equiv.coe_fn_mk]
|
||||
rfl
|
||||
|
||||
lemma toInvolution_consEquiv {φs : List 𝓕} {φ : 𝓕}
|
||||
(c : Contractions φs) (n : Option (Fin (c.uncontracted.length))) :
|
||||
(toInvolution ((consEquiv (φ := φ)).symm ⟨c, n⟩)).1 =
|
||||
(involutionCons φs.length).symm ⟨(toInvolution c).1,
|
||||
(involutionAddEquiv (toInvolution c).1).symm
|
||||
(Option.map (finCongr (toInvolution_length_uncontracted)) n)⟩ := by
|
||||
erw [toInvolution_cons]
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Involution equiv.
|
||||
|
||||
-/
|
||||
|
||||
lemma toInvolution_fromInvolution : {φs : List 𝓕} → (c : Contractions φs) →
|
||||
fromInvolution (toInvolution c) = c
|
||||
| [], ⟨[], ContractionsAux.nil⟩ => rfl
|
||||
| φ :: φs, ⟨_, .cons (φsᵤₙ := φsᵤₙ) n c⟩ => by
|
||||
rw [toInvolution_cons, fromInvolution_of_involutionCons, Equiv.symm_apply_eq]
|
||||
dsimp only [consEquiv, Equiv.coe_fn_mk]
|
||||
refine consEquiv_ext ?_ ?_
|
||||
· exact toInvolution_fromInvolution ⟨φsᵤₙ, c⟩
|
||||
· simp only [finCongr, Equiv.coe_fn_mk, Equiv.apply_symm_apply, Option.map_map]
|
||||
ext a : 1
|
||||
simp only [Option.mem_def, Option.map_eq_some', Function.comp_apply, Fin.cast_trans,
|
||||
Fin.cast_eq_self, exists_eq_right]
|
||||
|
||||
lemma fromInvolution_toInvolution : {φs : List 𝓕} →
|
||||
(f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) →
|
||||
toInvolution (fromInvolution f) = f
|
||||
| [], _ => by
|
||||
ext x
|
||||
exact Fin.elim0 x
|
||||
| φ :: φs, f => by
|
||||
rw [fromInvolution_cons]
|
||||
rw [toInvolution_consEquiv]
|
||||
erw [Equiv.symm_apply_eq]
|
||||
have hx := fromInvolution_toInvolution ((involutionCons φs.length) f).fst
|
||||
apply involutionCons_ext ?_ ?_
|
||||
· simp only [Nat.succ_eq_add_one, List.length_cons]
|
||||
exact hx
|
||||
· simp only [Nat.succ_eq_add_one, Option.map_map, List.length_cons]
|
||||
rw [Equiv.symm_apply_eq]
|
||||
conv_rhs =>
|
||||
lhs
|
||||
rw [involutionAddEquiv_cast hx]
|
||||
simp only [Nat.succ_eq_add_one, Equiv.trans_apply]
|
||||
rfl
|
||||
|
||||
/-- The equivalence between contractions and involutions.
|
||||
Note: This shows that the type of contractions only depends on the length of the list `φs`. -/
|
||||
def equivInvolutions {φs : List 𝓕} :
|
||||
Contractions φs ≃ {f : Fin φs.length → Fin φs.length // Function.Involutive f} where
|
||||
toFun := fun c => toInvolution c
|
||||
invFun := fromInvolution
|
||||
left_inv := toInvolution_fromInvolution
|
||||
right_inv := fromInvolution_toInvolution
|
||||
|
||||
/-!
|
||||
|
||||
## Full contractions and involutions.
|
||||
-/
|
||||
|
||||
lemma isFull_iff_uncontractedFromInvolution_empty {φs : List 𝓕} (c : Contractions φs) :
|
||||
IsFull c ↔ (uncontractedFromInvolution (equivInvolutions c)).1 = [] := by
|
||||
let l := toInvolution c
|
||||
erw [l.2]
|
||||
rfl
|
||||
|
||||
lemma isFull_iff_filter_card_involution_zero {φs : List 𝓕} (c : Contractions φs) :
|
||||
IsFull c ↔ (Finset.univ.filter fun i => (equivInvolutions c).1 i = i).card = 0 := by
|
||||
rw [isFull_iff_uncontractedFromInvolution_empty, List.ext_get_iff]
|
||||
simp
|
||||
|
||||
lemma isFull_iff_involution_no_fixed_points {φs : List 𝓕} (c : Contractions φs) :
|
||||
IsFull c ↔ ∀ (i : Fin φs.length), (equivInvolutions c).1 i ≠ i := by
|
||||
rw [isFull_iff_filter_card_involution_zero]
|
||||
simp only [Finset.card_eq_zero, ne_eq]
|
||||
rw [Finset.filter_eq_empty_iff]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
intro i
|
||||
refine h (Finset.mem_univ i)
|
||||
· intro i h
|
||||
exact fun a => i h
|
||||
|
||||
/-- The equivalence between full contractions and fixed-point free involutions. -/
|
||||
def isFullInvolutionEquiv {φs : List 𝓕} : {c : Contractions φs // IsFull c} ≃
|
||||
{f : Fin φs.length → Fin φs.length // Function.Involutive f ∧ (∀ i, f i ≠ i)} where
|
||||
toFun c := ⟨equivInvolutions c.1, by
|
||||
apply And.intro (equivInvolutions c.1).2
|
||||
rw [← isFull_iff_involution_no_fixed_points]
|
||||
exact c.2⟩
|
||||
invFun f := ⟨equivInvolutions.symm ⟨f.1, f.2.1⟩, by
|
||||
rw [isFull_iff_involution_no_fixed_points]
|
||||
simpa using f.2.2⟩
|
||||
left_inv c := by simp
|
||||
right_inv f := by simp
|
||||
|
||||
end Contractions
|
||||
end Wick
|
|
@ -3,15 +3,14 @@ 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 Mathlib.Order.Defs.Unbundled
|
||||
import Mathlib.Data.Fintype.Basic
|
||||
import Mathlib.Algebra.BigOperators.Group.Finset
|
||||
/-!
|
||||
|
||||
# Creation and annihilation parts of fields
|
||||
# Creation and annihlation parts of fields
|
||||
|
||||
-/
|
||||
|
||||
/-- The type specifying whether an operator is a creation or annihilation operator. -/
|
||||
/-- The type specifing whether an operator is a creation or annihilation operator. -/
|
||||
inductive CreateAnnihilate where
|
||||
| create : CreateAnnihilate
|
||||
| annihilate : CreateAnnihilate
|
||||
|
@ -29,14 +28,23 @@ instance : Fintype CreateAnnihilate where
|
|||
· refine Finset.insert_eq_self.mp ?_
|
||||
exact rfl
|
||||
|
||||
/-- The normal ordering on creation and annihilation operators.
|
||||
Creation operators are put to the left. -/
|
||||
lemma eq_create_or_annihilate (φ : CreateAnnihilate) : φ = create ∨ φ = annihilate := by
|
||||
cases φ <;> simp
|
||||
|
||||
/-- The normal ordering on creation and annihlation operators.
|
||||
Under this relation, `normalOrder a b` is false only if `a` is annihlate and `b` is create. -/
|
||||
def normalOrder : CreateAnnihilate → CreateAnnihilate → Prop
|
||||
| create, create => True
|
||||
| create, _ => True
|
||||
| annihilate, annihilate => True
|
||||
| create, annihilate => True
|
||||
| annihilate, create => False
|
||||
|
||||
/-- The normal ordering on `CreateAnnihilate` is decidable. -/
|
||||
instance : (φ φ' : CreateAnnihilate) → Decidable (normalOrder φ φ')
|
||||
| create, create => isTrue True.intro
|
||||
| annihilate, annihilate => isTrue True.intro
|
||||
| create, annihilate => isTrue True.intro
|
||||
| annihilate, create => isFalse False.elim
|
||||
|
||||
/-- Normal ordering is total. -/
|
||||
instance : IsTotal CreateAnnihilate normalOrder where
|
||||
total a b := by
|
||||
|
@ -47,4 +55,16 @@ instance : IsTrans CreateAnnihilate normalOrder where
|
|||
trans a b c := by
|
||||
cases a <;> cases b <;> cases c <;> simp [normalOrder]
|
||||
|
||||
@[simp]
|
||||
lemma not_normalOrder_annihilate_iff_false (a : CreateAnnihilate) :
|
||||
(¬ normalOrder a annihilate) ↔ False := by
|
||||
cases a
|
||||
· simp [normalOrder]
|
||||
· simp [normalOrder]
|
||||
|
||||
lemma sum_eq {M : Type} [AddCommMonoid M] (f : CreateAnnihilate → M) :
|
||||
∑ i, f i = f create + f annihilate := by
|
||||
change ∑ i in {create, annihilate}, f i = f create + f annihilate
|
||||
simp
|
||||
|
||||
end CreateAnnihilate
|
||||
|
|
|
@ -1,153 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import Mathlib.Algebra.FreeAlgebra
|
||||
import Mathlib.Algebra.Lie.OfAssociative
|
||||
import Mathlib.Analysis.Complex.Basic
|
||||
/-!
|
||||
|
||||
# Field statistics
|
||||
|
||||
Basic properties related to whether a field, or list of fields, is bosonic or fermionic.
|
||||
|
||||
-/
|
||||
|
||||
/-- A field can either be bosonic or fermionic in nature.
|
||||
That is to say, they can either have Bose-Einstein statistics or
|
||||
Fermi-Dirac statistics. -/
|
||||
inductive FieldStatistic : Type where
|
||||
| bosonic : FieldStatistic
|
||||
| fermionic : FieldStatistic
|
||||
deriving DecidableEq
|
||||
|
||||
namespace FieldStatistic
|
||||
|
||||
variable {𝓕 : Type}
|
||||
|
||||
/-- Field statics form a finite type. -/
|
||||
instance : Fintype FieldStatistic where
|
||||
elems := {bosonic, fermionic}
|
||||
complete := by
|
||||
intro c
|
||||
cases c
|
||||
· exact Finset.mem_insert_self bosonic {fermionic}
|
||||
· refine Finset.insert_eq_self.mp ?_
|
||||
exact rfl
|
||||
|
||||
@[simp]
|
||||
lemma fermionic_not_eq_bonsic : ¬ fermionic = bosonic := by
|
||||
intro h
|
||||
exact FieldStatistic.noConfusion h
|
||||
|
||||
lemma bonsic_eq_fermionic_false : bosonic = fermionic ↔ false := by
|
||||
simp only [reduceCtorEq, Bool.false_eq_true]
|
||||
|
||||
@[simp]
|
||||
lemma neq_fermionic_iff_eq_bosonic (a : FieldStatistic) : ¬ a = fermionic ↔ a = bosonic := by
|
||||
fin_cases a
|
||||
· simp
|
||||
· simp
|
||||
|
||||
@[simp]
|
||||
lemma bosonic_neq_iff_fermionic_eq (a : FieldStatistic) : ¬ bosonic = a ↔ fermionic = a := by
|
||||
fin_cases a
|
||||
· simp
|
||||
· simp
|
||||
|
||||
@[simp]
|
||||
lemma fermionic_neq_iff_bosonic_eq (a : FieldStatistic) : ¬ fermionic = a ↔ bosonic = a := by
|
||||
fin_cases a
|
||||
· simp
|
||||
· simp
|
||||
|
||||
lemma eq_self_if_eq_bosonic {a : FieldStatistic} :
|
||||
(if a = bosonic then bosonic else fermionic) = a := by
|
||||
fin_cases a <;> rfl
|
||||
|
||||
lemma eq_self_if_bosonic_eq {a : FieldStatistic} :
|
||||
(if bosonic = a then bosonic else fermionic) = a := by
|
||||
fin_cases a <;> rfl
|
||||
|
||||
/-- The field statistics of a list of fields is fermionic if ther is an odd number of fermions,
|
||||
otherwise it is bosonic. -/
|
||||
def ofList (s : 𝓕 → FieldStatistic) : (φs : List 𝓕) → FieldStatistic
|
||||
| [] => bosonic
|
||||
| φ :: φs => if s φ = ofList s φs then bosonic else fermionic
|
||||
|
||||
@[simp]
|
||||
lemma ofList_singleton (s : 𝓕 → FieldStatistic) (φ : 𝓕) : ofList s [φ] = s φ := by
|
||||
simp only [ofList, Fin.isValue]
|
||||
rw [eq_self_if_eq_bosonic]
|
||||
|
||||
@[simp]
|
||||
lemma ofList_freeMonoid (s : 𝓕 → FieldStatistic) (φ : 𝓕) : ofList s (FreeMonoid.of φ) = s φ :=
|
||||
ofList_singleton s φ
|
||||
|
||||
@[simp]
|
||||
lemma ofList_empty (s : 𝓕 → FieldStatistic) : ofList s [] = bosonic := rfl
|
||||
|
||||
@[simp]
|
||||
lemma ofList_append (s : 𝓕 → FieldStatistic) (φs φs' : List 𝓕) :
|
||||
ofList s (φs ++ φs') = if ofList s φs = ofList s φs' then bosonic else fermionic := by
|
||||
induction φs with
|
||||
| nil =>
|
||||
simp only [List.nil_append, ofList_empty, Fin.isValue, eq_self_if_bosonic_eq]
|
||||
| cons a l ih =>
|
||||
have hab (a b c : FieldStatistic) :
|
||||
(if a = (if b = c then bosonic else fermionic) then bosonic else fermionic) =
|
||||
if (if a = b then bosonic else fermionic) = c then bosonic else fermionic := by
|
||||
fin_cases a <;> fin_cases b <;> fin_cases c <;> rfl
|
||||
simp only [ofList, List.append_eq, Fin.isValue, ih, hab]
|
||||
|
||||
lemma ofList_eq_countP (s : 𝓕 → FieldStatistic) (φs : List 𝓕) :
|
||||
ofList s φs = if Nat.mod (List.countP (fun i => decide (s i = fermionic)) φs) 2 = 0 then
|
||||
bosonic else fermionic := by
|
||||
induction φs with
|
||||
| nil => simp
|
||||
| cons r0 r ih =>
|
||||
simp only [ofList]
|
||||
rw [List.countP_cons]
|
||||
simp only [decide_eq_true_eq]
|
||||
by_cases hr : s r0 = fermionic
|
||||
· simp only [hr, ↓reduceIte]
|
||||
simp_all only
|
||||
split
|
||||
next h =>
|
||||
simp_all only [↓reduceIte, fermionic_not_eq_bonsic]
|
||||
split
|
||||
next h_1 =>
|
||||
simp_all only [fermionic_not_eq_bonsic]
|
||||
have ha (a : ℕ) (ha : a % 2 = 0) : (a + 1) % 2 ≠ 0 := by
|
||||
omega
|
||||
exact ha (List.countP (fun i => decide (s i = fermionic)) r) h h_1
|
||||
next h_1 => simp_all only
|
||||
next h =>
|
||||
simp_all only [↓reduceIte]
|
||||
split
|
||||
next h_1 => rfl
|
||||
next h_1 =>
|
||||
simp_all only [reduceCtorEq]
|
||||
have ha (a : ℕ) (ha : ¬ a % 2 = 0) : (a + 1) % 2 = 0 := by
|
||||
omega
|
||||
exact h_1 (ha (List.countP (fun i => decide (s i = fermionic)) r) h)
|
||||
· simp only [neq_fermionic_iff_eq_bosonic] at hr
|
||||
by_cases hx : (List.countP (fun i => decide (s i = fermionic)) r).mod 2 = 0
|
||||
· simpa [hx, hr] using ih.symm
|
||||
· simpa [hx, hr] using ih.symm
|
||||
|
||||
lemma ofList_perm (s : 𝓕 → FieldStatistic) {l l' : List 𝓕} (h : l.Perm l') :
|
||||
ofList s l = ofList s l' := by
|
||||
rw [ofList_eq_countP, ofList_eq_countP, h.countP_eq]
|
||||
|
||||
lemma ofList_orderedInsert (s : 𝓕 → FieldStatistic) (le1 : 𝓕 → 𝓕 → Prop) [DecidableRel le1]
|
||||
(φs : List 𝓕) (φ : 𝓕) : ofList s (List.orderedInsert le1 φ φs) = ofList s (φ :: φs) :=
|
||||
ofList_perm s (List.perm_orderedInsert le1 φ φs)
|
||||
|
||||
@[simp]
|
||||
lemma ofList_insertionSort (s : 𝓕 → FieldStatistic) (le1 : 𝓕 → 𝓕 → Prop) [DecidableRel le1]
|
||||
(φs : List 𝓕) : ofList s (List.insertionSort le1 φs) = ofList s φs :=
|
||||
ofList_perm s (List.perm_insertionSort le1 φs)
|
||||
|
||||
end FieldStatistic
|
276
HepLean/PerturbationTheory/FieldStatistics/Basic.lean
Normal file
276
HepLean/PerturbationTheory/FieldStatistics/Basic.lean
Normal file
|
@ -0,0 +1,276 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import Mathlib.Algebra.FreeAlgebra
|
||||
import Mathlib.Algebra.Lie.OfAssociative
|
||||
import Mathlib.Analysis.Complex.Basic
|
||||
import HepLean.Mathematics.List.InsertIdx
|
||||
/-!
|
||||
|
||||
# Field statistics
|
||||
|
||||
Basic properties related to whether a field, or list of fields, is bosonic or fermionic.
|
||||
|
||||
-/
|
||||
|
||||
/-- A field can either be bosonic or fermionic in nature.
|
||||
That is to say, they can either have Bose-Einstein statistics or
|
||||
Fermi-Dirac statistics. -/
|
||||
inductive FieldStatistic : Type where
|
||||
| bosonic : FieldStatistic
|
||||
| fermionic : FieldStatistic
|
||||
deriving DecidableEq
|
||||
|
||||
namespace FieldStatistic
|
||||
|
||||
variable {𝓕 : Type}
|
||||
|
||||
/-- Field statistics form a commuative group equivalent to `ℤ₂`. -/
|
||||
@[simp]
|
||||
instance : CommGroup FieldStatistic where
|
||||
one := bosonic
|
||||
mul a b :=
|
||||
match a, b with
|
||||
| bosonic, bosonic => bosonic
|
||||
| bosonic, fermionic => fermionic
|
||||
| fermionic, bosonic => fermionic
|
||||
| fermionic, fermionic => bosonic
|
||||
inv a := a
|
||||
mul_assoc a b c := by
|
||||
cases a <;> cases b <;> cases c <;>
|
||||
dsimp [HMul.hMul]
|
||||
one_mul a := by
|
||||
cases a <;> dsimp [HMul.hMul]
|
||||
mul_one a := by
|
||||
cases a <;> dsimp [HMul.hMul]
|
||||
inv_mul_cancel a := by
|
||||
cases a <;> dsimp only [HMul.hMul, Nat.succ_eq_add_one] <;> rfl
|
||||
mul_comm a b := by
|
||||
cases a <;> cases b <;> rfl
|
||||
|
||||
@[simp]
|
||||
lemma bosonic_mul_bosonic : bosonic * bosonic = bosonic := rfl
|
||||
|
||||
@[simp]
|
||||
lemma bosonic_mul_fermionic : bosonic * fermionic = fermionic := rfl
|
||||
|
||||
@[simp]
|
||||
lemma fermionic_mul_bosonic : fermionic * bosonic = fermionic := rfl
|
||||
|
||||
@[simp]
|
||||
lemma fermionic_mul_fermionic : fermionic * fermionic = bosonic := rfl
|
||||
|
||||
@[simp]
|
||||
lemma mul_self (a : FieldStatistic) : a * a = 1 := by
|
||||
cases a <;> rfl
|
||||
|
||||
/-- Field statics form a finite type. -/
|
||||
instance : Fintype FieldStatistic where
|
||||
elems := {bosonic, fermionic}
|
||||
complete := by
|
||||
intro c
|
||||
cases c
|
||||
· exact Finset.mem_insert_self bosonic {fermionic}
|
||||
· refine Finset.insert_eq_self.mp ?_
|
||||
exact rfl
|
||||
|
||||
@[simp]
|
||||
lemma fermionic_not_eq_bonsic : ¬ fermionic = bosonic := by
|
||||
intro h
|
||||
exact FieldStatistic.noConfusion h
|
||||
|
||||
lemma bonsic_eq_fermionic_false : bosonic = fermionic ↔ false := by
|
||||
simp only [reduceCtorEq, Bool.false_eq_true]
|
||||
|
||||
@[simp]
|
||||
lemma neq_fermionic_iff_eq_bosonic (a : FieldStatistic) : ¬ a = fermionic ↔ a = bosonic := by
|
||||
fin_cases a
|
||||
· simp
|
||||
· simp
|
||||
|
||||
@[simp]
|
||||
lemma bosonic_neq_iff_fermionic_eq (a : FieldStatistic) : ¬ bosonic = a ↔ fermionic = a := by
|
||||
fin_cases a
|
||||
· simp
|
||||
· simp
|
||||
|
||||
@[simp]
|
||||
lemma fermionic_neq_iff_bosonic_eq (a : FieldStatistic) : ¬ fermionic = a ↔ bosonic = a := by
|
||||
fin_cases a
|
||||
· simp
|
||||
· simp
|
||||
|
||||
lemma eq_self_if_eq_bosonic {a : FieldStatistic} :
|
||||
(if a = bosonic then bosonic else fermionic) = a := by
|
||||
fin_cases a <;> rfl
|
||||
|
||||
lemma eq_self_if_bosonic_eq {a : FieldStatistic} :
|
||||
(if bosonic = a then bosonic else fermionic) = a := by
|
||||
fin_cases a <;> rfl
|
||||
|
||||
lemma mul_eq_one_iff (a b : FieldStatistic) : a * b = 1 ↔ a = b := by
|
||||
fin_cases a <;> fin_cases b <;> simp
|
||||
|
||||
lemma one_eq_mul_iff (a b : FieldStatistic) : 1 = a * b ↔ a = b := by
|
||||
fin_cases a <;> fin_cases b <;> simp
|
||||
|
||||
lemma mul_eq_iff_eq_mul (a b c : FieldStatistic) : a * b = c ↔ a = b * c := by
|
||||
fin_cases a <;> fin_cases b <;> fin_cases c <;>
|
||||
simp only [bosonic_mul_fermionic, fermionic_not_eq_bonsic, mul_self,
|
||||
reduceCtorEq, fermionic_mul_bosonic, true_iff, iff_true]
|
||||
all_goals rfl
|
||||
|
||||
lemma mul_eq_iff_eq_mul' (a b c : FieldStatistic) : a * b = c ↔ b = a * c := by
|
||||
fin_cases a <;> fin_cases b <;> fin_cases c <;>
|
||||
simp only [bosonic_mul_fermionic, fermionic_not_eq_bonsic, mul_self,
|
||||
reduceCtorEq, fermionic_mul_bosonic, true_iff, iff_true]
|
||||
all_goals rfl
|
||||
|
||||
/-- The field statistics of a list of fields is fermionic if ther is an odd number of fermions,
|
||||
otherwise it is bosonic. -/
|
||||
def ofList (s : 𝓕 → FieldStatistic) : (φs : List 𝓕) → FieldStatistic
|
||||
| [] => bosonic
|
||||
| φ :: φs => if s φ = ofList s φs then bosonic else fermionic
|
||||
|
||||
lemma ofList_cons_eq_mul (s : 𝓕 → FieldStatistic) (φ : 𝓕) (φs : List 𝓕) :
|
||||
ofList s (φ :: φs) = s φ * ofList s φs := by
|
||||
have ha (a b : FieldStatistic) : (if a = b then bosonic else fermionic) = a * b := by
|
||||
fin_cases a <;> fin_cases b <;> rfl
|
||||
exact ha (s φ) (ofList s φs)
|
||||
|
||||
lemma ofList_eq_prod (s : 𝓕 → FieldStatistic) : (φs : List 𝓕) →
|
||||
ofList s φs = (List.map s φs).prod
|
||||
| [] => rfl
|
||||
| φ :: φs => by
|
||||
rw [ofList_cons_eq_mul, List.map_cons, List.prod_cons, ofList_eq_prod]
|
||||
|
||||
@[simp]
|
||||
lemma ofList_singleton (s : 𝓕 → FieldStatistic) (φ : 𝓕) : ofList s [φ] = s φ := by
|
||||
simp only [ofList, Fin.isValue]
|
||||
rw [eq_self_if_eq_bosonic]
|
||||
|
||||
@[simp]
|
||||
lemma ofList_freeMonoid (s : 𝓕 → FieldStatistic) (φ : 𝓕) : ofList s (FreeMonoid.of φ) = s φ :=
|
||||
ofList_singleton s φ
|
||||
|
||||
@[simp]
|
||||
lemma ofList_empty (s : 𝓕 → FieldStatistic) : ofList s [] = bosonic := rfl
|
||||
|
||||
@[simp]
|
||||
lemma ofList_append (s : 𝓕 → FieldStatistic) (φs φs' : List 𝓕) :
|
||||
ofList s (φs ++ φs') = if ofList s φs = ofList s φs' then bosonic else fermionic := by
|
||||
induction φs with
|
||||
| nil =>
|
||||
simp only [List.nil_append, ofList_empty, Fin.isValue, eq_self_if_bosonic_eq]
|
||||
| cons a l ih =>
|
||||
have hab (a b c : FieldStatistic) :
|
||||
(if a = (if b = c then bosonic else fermionic) then bosonic else fermionic) =
|
||||
if (if a = b then bosonic else fermionic) = c then bosonic else fermionic := by
|
||||
fin_cases a <;> fin_cases b <;> fin_cases c <;> rfl
|
||||
simp only [ofList, List.append_eq, Fin.isValue, ih, hab]
|
||||
|
||||
lemma ofList_append_eq_mul (s : 𝓕 → FieldStatistic) (φs φs' : List 𝓕) :
|
||||
ofList s (φs ++ φs') = ofList s φs * ofList s φs' := by
|
||||
rw [ofList_append]
|
||||
have ha (a b : FieldStatistic) : (if a = b then bosonic else fermionic) = a * b := by
|
||||
fin_cases a <;> fin_cases b <;> rfl
|
||||
exact ha _ _
|
||||
|
||||
lemma ofList_perm (s : 𝓕 → FieldStatistic) {l l' : List 𝓕} (h : l.Perm l') :
|
||||
ofList s l = ofList s l' := by
|
||||
rw [ofList_eq_prod, ofList_eq_prod]
|
||||
exact List.Perm.prod_eq (List.Perm.map s h)
|
||||
|
||||
lemma ofList_orderedInsert (s : 𝓕 → FieldStatistic) (le1 : 𝓕 → 𝓕 → Prop) [DecidableRel le1]
|
||||
(φs : List 𝓕) (φ : 𝓕) : ofList s (List.orderedInsert le1 φ φs) = ofList s (φ :: φs) :=
|
||||
ofList_perm s (List.perm_orderedInsert le1 φ φs)
|
||||
|
||||
@[simp]
|
||||
lemma ofList_insertionSort (s : 𝓕 → FieldStatistic) (le1 : 𝓕 → 𝓕 → Prop) [DecidableRel le1]
|
||||
(φs : List 𝓕) : ofList s (List.insertionSort le1 φs) = ofList s φs :=
|
||||
ofList_perm s (List.perm_insertionSort le1 φs)
|
||||
|
||||
lemma ofList_map_eq_finset_prod (s : 𝓕 → FieldStatistic) :
|
||||
(φs : List 𝓕) → (l : List (Fin φs.length)) → (hl : l.Nodup) →
|
||||
ofList s (l.map φs.get) = ∏ (i : Fin φs.length), if i ∈ l then s φs[i] else 1
|
||||
| [], [], _ => rfl
|
||||
| [], i :: l, hl => Fin.elim0 i
|
||||
| φ :: φs, [], hl => by
|
||||
simp only [List.length_cons, List.map_nil, ofList_empty, List.not_mem_nil, ↓reduceIte,
|
||||
Finset.prod_const_one]
|
||||
rfl
|
||||
| φ :: φs, i :: l, hl => by
|
||||
simp only [List.length_cons, List.map_cons, List.get_eq_getElem, List.mem_cons, instCommGroup,
|
||||
Fin.getElem_fin]
|
||||
rw [ofList_cons_eq_mul]
|
||||
rw [ofList_map_eq_finset_prod s (φ :: φs) l]
|
||||
have h1 : s (φ :: φs)[↑i] = ∏ (j : Fin (φ :: φs).length),
|
||||
if j = i then s (φ :: φs)[↑i] else 1 := by
|
||||
rw [Fintype.prod_ite_eq']
|
||||
erw [h1]
|
||||
rw [← Finset.prod_mul_distrib]
|
||||
congr
|
||||
funext a
|
||||
simp only [instCommGroup, List.length_cons, mul_ite, ite_mul, one_mul, mul_one]
|
||||
by_cases ha : a = i
|
||||
· simp only [ha, ↓reduceIte, mul_self, true_or]
|
||||
rw [if_neg]
|
||||
rfl
|
||||
simp only [List.length_cons, List.nodup_cons] at hl
|
||||
exact hl.1
|
||||
· simp only [ha, ↓reduceIte, false_or]
|
||||
rfl
|
||||
simp only [List.length_cons, List.nodup_cons] at hl
|
||||
exact hl.2
|
||||
|
||||
/-!
|
||||
|
||||
## ofList and take
|
||||
|
||||
-/
|
||||
|
||||
section ofListTake
|
||||
open HepLean.List
|
||||
variable (q : 𝓕 → FieldStatistic)
|
||||
lemma ofList_take_insert (n : ℕ) (φ : 𝓕) (φs : List 𝓕) :
|
||||
ofList q (List.take n φs) = ofList q (List.take n (List.insertIdx n φ φs)) := by
|
||||
congr 1
|
||||
rw [take_insert_same]
|
||||
|
||||
lemma ofList_take_eraseIdx (n : ℕ) (φs : List 𝓕) :
|
||||
ofList q (List.take n (φs.eraseIdx n)) = ofList q (List.take n φs) := by
|
||||
congr 1
|
||||
rw [take_eraseIdx_same]
|
||||
|
||||
lemma ofList_take_zero (φs : List 𝓕) :
|
||||
ofList q (List.take 0 φs) = 1 := by
|
||||
simp only [List.take_zero, ofList_empty]
|
||||
rfl
|
||||
|
||||
lemma ofList_take_succ_cons (n : ℕ) (φ1 : 𝓕) (φs : List 𝓕) :
|
||||
ofList q ((φ1 :: φs).take (n + 1)) = q φ1 * ofList q (φs.take n) := by
|
||||
simp only [List.take_succ_cons, instCommGroup]
|
||||
rw [ofList_cons_eq_mul]
|
||||
|
||||
lemma ofList_take_insertIdx_gt (n m : ℕ) (φ1 : 𝓕) (φs : List 𝓕) (hn : n < m) :
|
||||
ofList q ((List.insertIdx m φ1 φs).take n) = ofList q (φs.take n) := by
|
||||
rw [take_insert_gt φ1 n m hn φs]
|
||||
|
||||
lemma ofList_insert_lt_eq (n m : ℕ) (φ1 : 𝓕) (φs : List 𝓕) (hn : m ≤ n)
|
||||
(hm : m ≤ φs.length) :
|
||||
ofList q ((List.insertIdx m φ1 φs).take (n + 1)) =
|
||||
ofList q ((φ1 :: φs).take (n + 1)) := by
|
||||
apply ofList_perm
|
||||
simp only [List.take_succ_cons]
|
||||
refine take_insert_let φ1 n m hn φs hm
|
||||
|
||||
lemma ofList_take_insertIdx_le (n m : ℕ) (φ1 : 𝓕) (φs : List 𝓕) (hn : m ≤ n) (hm : m ≤ φs.length) :
|
||||
ofList q ((List.insertIdx m φ1 φs).take (n + 1)) = q φ1 * ofList q (φs.take n) := by
|
||||
rw [ofList_insert_lt_eq, ofList_take_succ_cons]
|
||||
· exact hn
|
||||
· exact hm
|
||||
|
||||
end ofListTake
|
||||
end FieldStatistic
|
82
HepLean/PerturbationTheory/FieldStatistics/ExchangeSign.lean
Normal file
82
HepLean/PerturbationTheory/FieldStatistics/ExchangeSign.lean
Normal file
|
@ -0,0 +1,82 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.FieldStatistics.Basic
|
||||
/-!
|
||||
|
||||
# Exchange sign
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStatistic
|
||||
|
||||
variable {𝓕 : Type}
|
||||
|
||||
/-- The echange sign of two field statistics.
|
||||
Defined to be `-1` if both field statistics are `fermionic` and `1` otherwise. -/
|
||||
def exchangeSign : FieldStatistic →* FieldStatistic →* ℂ where
|
||||
toFun a :=
|
||||
{
|
||||
toFun := fun b =>
|
||||
match a, b with
|
||||
| bosonic, _ => 1
|
||||
| _, bosonic => 1
|
||||
| fermionic, fermionic => -1
|
||||
map_one' := by
|
||||
fin_cases a
|
||||
<;> simp
|
||||
map_mul' := fun c b => by
|
||||
fin_cases a <;>
|
||||
fin_cases b <;>
|
||||
fin_cases c <;>
|
||||
simp
|
||||
}
|
||||
map_one' := by
|
||||
ext b
|
||||
fin_cases b
|
||||
<;> simp
|
||||
map_mul' c b := by
|
||||
ext a
|
||||
fin_cases a
|
||||
<;> fin_cases b <;> fin_cases c
|
||||
<;> simp
|
||||
|
||||
/-- The echange sign of two field statistics.
|
||||
Defined to be `-1` if both field statistics are `fermionic` and `1` otherwise. -/
|
||||
scoped[FieldStatistic] notation "𝓢(" a "," b ")" => exchangeSign a b
|
||||
|
||||
@[simp]
|
||||
lemma exchangeSign_bosonic (a : FieldStatistic) : 𝓢(a, bosonic) = 1 := by
|
||||
fin_cases a <;> rfl
|
||||
|
||||
@[simp]
|
||||
lemma bosonic_exchangeSign (a : FieldStatistic) : 𝓢(bosonic, a) = 1 := by
|
||||
fin_cases a <;> rfl
|
||||
|
||||
lemma exchangeSign_symm (a b : FieldStatistic) : 𝓢(a, b) = 𝓢(b, a) := by
|
||||
fin_cases a <;> fin_cases b <;> rfl
|
||||
|
||||
lemma exchangeSign_eq_if (a b : FieldStatistic) :
|
||||
𝓢(a, b) = if a = fermionic ∧ b = fermionic then - 1 else 1 := by
|
||||
fin_cases a <;> fin_cases b <;> rfl
|
||||
|
||||
@[simp]
|
||||
lemma exchangeSign_mul_self (a b : FieldStatistic) : 𝓢(a, b) * 𝓢(a, b) = 1 := by
|
||||
fin_cases a <;> fin_cases b <;> simp [exchangeSign]
|
||||
|
||||
@[simp]
|
||||
lemma exchangeSign_mul_self_swap (a b : FieldStatistic) : 𝓢(a, b) * 𝓢(b, a) = 1 := by
|
||||
fin_cases a <;> fin_cases b <;> simp [exchangeSign]
|
||||
|
||||
lemma exchangeSign_ofList_cons (a : FieldStatistic)
|
||||
(s : 𝓕 → FieldStatistic) (φ : 𝓕) (φs : List 𝓕) :
|
||||
𝓢(a, ofList s (φ :: φs)) = 𝓢(a, s φ) * 𝓢(a, ofList s φs) := by
|
||||
rw [ofList_cons_eq_mul, map_mul]
|
||||
|
||||
lemma exchangeSign_cocycle (a b c : FieldStatistic) :
|
||||
𝓢(a, b * c) * 𝓢(b, c) = 𝓢(a, b) * 𝓢(a * b, c) := by
|
||||
fin_cases a <;> fin_cases b <;> fin_cases c <;> simp
|
||||
|
||||
end FieldStatistic
|
109
HepLean/PerturbationTheory/FieldStatistics/OfFinset.lean
Normal file
109
HepLean/PerturbationTheory/FieldStatistics/OfFinset.lean
Normal file
|
@ -0,0 +1,109 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.FieldStatistics.ExchangeSign
|
||||
/-!
|
||||
|
||||
# Field statistics of a finite set.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStatistic
|
||||
|
||||
variable {𝓕 : Type}
|
||||
|
||||
/-- The field statistic associated with a map `f : Fin n → 𝓕` (usually `.get` of a list)
|
||||
and a finite set of elements of `Fin n`. -/
|
||||
def ofFinset {n : ℕ} (q : 𝓕 → FieldStatistic) (f : Fin n → 𝓕) (a : Finset (Fin n)) :
|
||||
FieldStatistic :=
|
||||
ofList q ((a.sort (· ≤ ·)).map f)
|
||||
|
||||
@[simp]
|
||||
lemma ofFinset_empty (q : 𝓕 → FieldStatistic) (f : Fin n → 𝓕) :
|
||||
ofFinset q f ∅ = 1 := by
|
||||
simp only [ofFinset, Finset.sort_empty, List.map_nil, ofList_empty]
|
||||
rfl
|
||||
|
||||
lemma ofFinset_singleton {n : ℕ} (q : 𝓕 → FieldStatistic) (f : Fin n → 𝓕) (i : Fin n) :
|
||||
ofFinset q f {i} = q (f i) := by
|
||||
simp [ofFinset]
|
||||
|
||||
lemma ofFinset_finset_map {n m : ℕ}
|
||||
(q : 𝓕 → FieldStatistic) (i : Fin m → Fin n) (hi : Function.Injective i)
|
||||
(f : Fin n → 𝓕) (a : Finset (Fin m)) :
|
||||
ofFinset q (f ∘ i) a = ofFinset q f (a.map ⟨i, hi⟩) := by
|
||||
simp only [ofFinset]
|
||||
apply FieldStatistic.ofList_perm
|
||||
rw [← List.map_map]
|
||||
refine List.Perm.map f ?_
|
||||
apply List.perm_of_nodup_nodup_toFinset_eq
|
||||
· refine (List.nodup_map_iff_inj_on ?_).mpr ?_
|
||||
exact Finset.sort_nodup (fun x1 x2 => x1 ≤ x2) a
|
||||
simp only [Finset.mem_sort]
|
||||
intro x hx y hy
|
||||
exact fun a => hi a
|
||||
· exact Finset.sort_nodup (fun x1 x2 => x1 ≤ x2) (Finset.map { toFun := i, inj' := hi } a)
|
||||
· ext a
|
||||
simp
|
||||
|
||||
lemma ofFinset_insert (q : 𝓕 → FieldStatistic) (φs : List 𝓕) (a : Finset (Fin φs.length))
|
||||
(i : Fin φs.length) (h : i ∉ a) :
|
||||
ofFinset q φs.get (Insert.insert i a) = (q φs[i]) * ofFinset q φs.get a := by
|
||||
simp only [ofFinset, instCommGroup, Fin.getElem_fin]
|
||||
rw [← ofList_cons_eq_mul]
|
||||
have h1 : (φs[↑i] :: List.map φs.get (Finset.sort (fun x1 x2 => x1 ≤ x2) a))
|
||||
= List.map φs.get (i :: Finset.sort (fun x1 x2 => x1 ≤ x2) a) := by
|
||||
simp
|
||||
erw [h1]
|
||||
apply ofList_perm
|
||||
refine List.Perm.map φs.get ?_
|
||||
refine (List.perm_ext_iff_of_nodup ?_ ?_).mpr ?_
|
||||
· exact Finset.sort_nodup (fun x1 x2 => x1 ≤ x2) (Insert.insert i a)
|
||||
· simp only [List.nodup_cons, Finset.mem_sort, Finset.sort_nodup, and_true]
|
||||
exact h
|
||||
intro a
|
||||
simp
|
||||
|
||||
lemma ofFinset_erase (q : 𝓕 → FieldStatistic) (φs : List 𝓕) (a : Finset (Fin φs.length))
|
||||
(i : Fin φs.length) (h : i ∈ a) :
|
||||
ofFinset q φs.get (a.erase i) = (q φs[i]) * ofFinset q φs.get a := by
|
||||
have ha : a = Insert.insert i (a.erase i) := by
|
||||
exact Eq.symm (Finset.insert_erase h)
|
||||
conv_rhs => rw [ha]
|
||||
rw [ofFinset_insert]
|
||||
rw [← mul_assoc]
|
||||
simp only [instCommGroup, Fin.getElem_fin, mul_self, one_mul]
|
||||
simp
|
||||
|
||||
lemma ofFinset_eq_prod (q : 𝓕 → FieldStatistic) (φs : List 𝓕) (a : Finset (Fin φs.length)) :
|
||||
ofFinset q φs.get a = ∏ (i : Fin φs.length), if i ∈ a then (q φs[i]) else 1 := by
|
||||
rw [ofFinset]
|
||||
rw [ofList_map_eq_finset_prod]
|
||||
congr
|
||||
funext i
|
||||
simp only [Finset.mem_sort, Fin.getElem_fin]
|
||||
exact Finset.sort_nodup (fun x1 x2 => x1 ≤ x2) a
|
||||
|
||||
lemma ofFinset_union (q : 𝓕 → FieldStatistic) (φs : List 𝓕) (a b : Finset (Fin φs.length)) :
|
||||
ofFinset q φs.get a * ofFinset q φs.get b = ofFinset q φs.get ((a ∪ b) \ (a ∩ b)) := by
|
||||
rw [ofFinset_eq_prod, ofFinset_eq_prod, ofFinset_eq_prod]
|
||||
rw [← Finset.prod_mul_distrib]
|
||||
congr
|
||||
funext x
|
||||
simp only [instCommGroup, Fin.getElem_fin, mul_ite, ite_mul, mul_self, one_mul, mul_one,
|
||||
Finset.mem_sdiff, Finset.mem_union, Finset.mem_inter, not_and]
|
||||
split
|
||||
· rename_i h
|
||||
simp [h]
|
||||
· rename_i h
|
||||
simp [h]
|
||||
|
||||
lemma ofFinset_union_disjoint (q : 𝓕 → FieldStatistic) (φs : List 𝓕) (a b : Finset (Fin φs.length))
|
||||
(h : Disjoint a b) :
|
||||
ofFinset q φs.get a * ofFinset q φs.get b = ofFinset q φs.get (a ∪ b) := by
|
||||
rw [ofFinset_union, Finset.disjoint_iff_inter_eq_empty.mp h]
|
||||
simp
|
||||
|
||||
end FieldStatistic
|
|
@ -3,10 +3,10 @@ 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.Wick.OperatorMap
|
||||
import HepLean.Mathematics.Fin.Involutions
|
||||
import HepLean.Lorentz.RealVector.Basic
|
||||
import HepLean.PerturbationTheory.FieldStatistics.ExchangeSign
|
||||
import HepLean.SpaceTime.Basic
|
||||
import HepLean.PerturbationTheory.FieldStatistics.OfFinset
|
||||
/-!
|
||||
|
||||
# Field structures
|
||||
|
@ -16,7 +16,7 @@ import HepLean.SpaceTime.Basic
|
|||
/-- A field structure is a type of fields plus a specification of the
|
||||
statistics (fermionic or bosonic) of each field. -/
|
||||
structure FieldStruct where
|
||||
/-- The type of fields. -/
|
||||
/-- The type of fields. This also includes anti-states. -/
|
||||
Fields : Type
|
||||
/-- The specification if a field is bosonic or fermionic. -/
|
||||
statistics : 𝓕 → FieldStatistic
|
||||
|
@ -48,27 +48,15 @@ def statesToField : 𝓕.States → 𝓕.Fields
|
|||
/-- The statistics associated to a state. -/
|
||||
def statesStatistic : 𝓕.States → FieldStatistic := 𝓕.statistics ∘ 𝓕.statesToField
|
||||
|
||||
/-- Returns true if `timeOrder a b` is true if `a` has time greater then or equal to `b`.
|
||||
This will put the stats at the greatest time to left. -/
|
||||
def timeOrder : 𝓕.States → 𝓕.States → Prop
|
||||
| States.posAsymp _, _ => True
|
||||
| States.position φ0, States.position φ1 => φ1.2 0 ≤ φ0.2 0
|
||||
| States.position _, States.negAsymp _ => True
|
||||
| States.position _, States.posAsymp _ => False
|
||||
| States.negAsymp _, States.posAsymp _ => False
|
||||
| States.negAsymp _, States.position _ => False
|
||||
| States.negAsymp _, States.negAsymp _ => True
|
||||
/-- The field statistics associated with a state. -/
|
||||
scoped[FieldStruct] notation 𝓕 "|>ₛ" φ => statesStatistic 𝓕 φ
|
||||
|
||||
/-- Time ordering is total. -/
|
||||
instance : IsTotal 𝓕.States 𝓕.timeOrder where
|
||||
total a b := by
|
||||
cases a <;> cases b <;> simp [timeOrder]
|
||||
exact LinearOrder.le_total _ _
|
||||
/-- The field statistics associated with a list states. -/
|
||||
scoped[FieldStruct] notation 𝓕 "|>ₛ" φ => FieldStatistic.ofList
|
||||
(statesStatistic 𝓕) φ
|
||||
|
||||
/-- Time ordering is transitive. -/
|
||||
instance : IsTrans 𝓕.States 𝓕.timeOrder where
|
||||
trans a b c := by
|
||||
cases a <;> cases b <;> cases c <;> simp [timeOrder]
|
||||
exact fun h1 h2 => Preorder.le_trans _ _ _ h2 h1
|
||||
/-- The field statistic associated with a finite set-/
|
||||
scoped[FieldStruct] notation 𝓕 "|>ₛ" "⟨" f ","a "⟩"=> FieldStatistic.ofFinset
|
||||
(statesStatistic 𝓕) f a
|
||||
|
||||
end FieldStruct
|
||||
|
|
395
HepLean/PerturbationTheory/FieldStruct/CrAnSection.lean
Normal file
395
HepLean/PerturbationTheory/FieldStruct/CrAnSection.lean
Normal file
|
@ -0,0 +1,395 @@
|
|||
/-
|
||||
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.FieldStruct.CreateAnnihilate
|
||||
/-!
|
||||
|
||||
# Creation and annihlation sections
|
||||
|
||||
This module defines creation and annihilation sections, which represent different ways to associate
|
||||
fields with creation or annihilation operators.
|
||||
|
||||
## Main definitions
|
||||
|
||||
- `CrAnSection φs` : Represents sections in `𝓕.CrAnStates` over a list of states `φs`.
|
||||
Given fields `φ₁...φₙ`, this captures all possible ways to assign each field as either a creation
|
||||
or annihilation operator.
|
||||
|
||||
- `head`, `tail` : Access the first element and remaining elements of a section
|
||||
|
||||
- `cons` : Constructs a new section by prepending a creation/annihilation state
|
||||
|
||||
- Various equivalences:
|
||||
* `nilEquiv` : Empty sections
|
||||
* `singletonEquiv` : Sections over single elements
|
||||
* `consEquiv` : Separates head and tail
|
||||
* `appendEquiv` : Splits sections at a given point
|
||||
|
||||
All sections form finite types and support operations like taking/dropping elements and
|
||||
concatenation while preserving the relationship between states and their operator assignments.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
/-- The sections in `𝓕.CrAnStates` over a list `φs : List 𝓕.States`.
|
||||
In terms of physics, given some fields `φ₁...φₙ`, the different ways one can associate
|
||||
each field as a `creation` or an `annilation` operator. E.g. the number of terms
|
||||
`φ₁⁰φ₂¹...φₙ⁰` `φ₁¹φ₂¹...φₙ⁰` etc. If some fields are exclusively creation or annhilation
|
||||
operators at this point (e.g. ansymptotic states) this is accounted for. -/
|
||||
def CrAnSection (φs : List 𝓕.States) : Type :=
|
||||
{ψs : List 𝓕.CrAnStates //
|
||||
List.map 𝓕.crAnStatesToStates ψs = φs}
|
||||
-- Π i, 𝓕.statesToCreateAnnihilateType (φs.get i)
|
||||
|
||||
namespace CrAnSection
|
||||
open FieldStatistic
|
||||
variable {𝓕 : FieldStruct} {φs : List 𝓕.States}
|
||||
|
||||
@[simp]
|
||||
lemma length_eq (ψs : CrAnSection φs) : ψs.1.length = φs.length := by
|
||||
simpa using congrArg List.length ψs.2
|
||||
|
||||
/-- The tail of a section for `φs`. -/
|
||||
def tail : {φs : List 𝓕.States} → (ψs : CrAnSection φs) → CrAnSection φs.tail
|
||||
| [], ⟨[], h⟩ => ⟨[], h⟩
|
||||
| φ :: φs, ⟨[], h⟩ => False.elim (by simp at h)
|
||||
| φ :: φs, ⟨ψ :: ψs, h⟩ => ⟨ψs, by rw [List.map_cons, List.cons.injEq] at h; exact h.2⟩
|
||||
|
||||
lemma head_state_eq {φ : 𝓕.States} : (ψs : CrAnSection (φ :: φs)) →
|
||||
(ψs.1.head (by simp [← List.length_pos_iff_ne_nil])).1 = φ
|
||||
| ⟨[], h⟩ => False.elim (by simp at h)
|
||||
| ⟨ψ :: ψs, h⟩ => by
|
||||
simp only [List.map_cons, List.cons.injEq] at h
|
||||
exact h.1
|
||||
|
||||
lemma statistics_eq_state_statistics (ψs : CrAnSection φs) :
|
||||
(𝓕 |>ₛ ψs.1) = 𝓕 |>ₛ φs := by
|
||||
erw [FieldStatistic.ofList_eq_prod, FieldStatistic.ofList_eq_prod, crAnStatistics]
|
||||
rw [← List.map_comp_map, Function.comp_apply, ψs.2]
|
||||
|
||||
lemma take_statistics_eq_take_state_statistics (ψs : CrAnSection φs) n :
|
||||
(𝓕 |>ₛ (ψs.1.take n)) = 𝓕 |>ₛ (φs.take n) := by
|
||||
erw [FieldStatistic.ofList_eq_prod, FieldStatistic.ofList_eq_prod, crAnStatistics]
|
||||
simp only [instCommGroup, List.map_take]
|
||||
rw [← List.map_comp_map, Function.comp_apply, ψs.2]
|
||||
|
||||
/-- The head of a section for `φ :: φs` as an element in `𝓕.statesToCreateAnnihilateType φ`. -/
|
||||
def head : {φ : 𝓕.States} → (ψs : CrAnSection (φ :: φs)) →
|
||||
𝓕.statesToCrAnType φ
|
||||
| φ, ⟨[], h⟩ => False.elim (by simp at h)
|
||||
| φ, ⟨ψ :: ψs, h⟩ => 𝓕.statesToCreateAnnihilateTypeCongr (by
|
||||
simpa using head_state_eq ⟨ψ :: ψs, h⟩) ψ.2
|
||||
|
||||
lemma eq_head_cons_tail {φ : 𝓕.States} {ψs : CrAnSection (φ :: φs)} :
|
||||
ψs.1 = ⟨φ, head ψs⟩ :: ψs.tail.1 := by
|
||||
match ψs with
|
||||
| ⟨[], h⟩ => exact False.elim (by simp at h)
|
||||
| ⟨ψ :: ψs, h⟩ =>
|
||||
have h2 := head_state_eq ⟨ψ :: ψs, h⟩
|
||||
simp only [List.head_cons] at h2
|
||||
subst h2
|
||||
rfl
|
||||
|
||||
/-- The creation of a section from for `φ : φs` from a section for `φs` and a
|
||||
element of `𝓕.statesToCreateAnnihilateType φ`. -/
|
||||
def cons {φ : 𝓕.States} (ψ : 𝓕.statesToCrAnType φ) (ψs : CrAnSection φs) :
|
||||
CrAnSection (φ :: φs) := ⟨⟨φ, ψ⟩ :: ψs.1, by
|
||||
simp [List.map_cons, ψs.2]⟩
|
||||
|
||||
/-- For the empty list of states there is only one `CrAnSection`. Corresponding to the
|
||||
empty list of `CrAnStates`. -/
|
||||
def nilEquiv : CrAnSection (𝓕 := 𝓕) [] ≃ Unit where
|
||||
toFun _ := ()
|
||||
invFun _ := ⟨[], rfl⟩
|
||||
left_inv ψs := Subtype.ext <| by
|
||||
have h2 := ψs.2
|
||||
simp only [List.map_eq_nil_iff] at h2
|
||||
simp [h2]
|
||||
right_inv _ := by
|
||||
simp
|
||||
|
||||
/-- The creation and annihlation sections for a singleton list is given by
|
||||
a choice of `𝓕.statesToCreateAnnihilateType φ`. If `φ` is a asymptotic state
|
||||
there is no choice here, else there are two choices. -/
|
||||
def singletonEquiv {φ : 𝓕.States} : CrAnSection [φ] ≃
|
||||
𝓕.statesToCrAnType φ where
|
||||
toFun ψs := ψs.head
|
||||
invFun ψ := ⟨[⟨φ, ψ⟩], by simp⟩
|
||||
left_inv ψs := by
|
||||
apply Subtype.ext
|
||||
simp only
|
||||
have h1 := eq_head_cons_tail (ψs := ψs)
|
||||
rw [h1]
|
||||
have h2 := ψs.tail.2
|
||||
simp only [List.tail_cons, List.map_eq_nil_iff] at h2
|
||||
simp [h2]
|
||||
right_inv ψ := by
|
||||
simp only [head]
|
||||
rfl
|
||||
|
||||
/-- An equivalence seperating the head of a creation and annhilation section
|
||||
from the tail. -/
|
||||
def consEquiv {φ : 𝓕.States} {φs : List 𝓕.States} : CrAnSection (φ :: φs) ≃
|
||||
𝓕.statesToCrAnType φ × CrAnSection φs where
|
||||
toFun ψs := ⟨ψs.head, ψs.tail⟩
|
||||
invFun ψψs :=
|
||||
match ψψs with
|
||||
| (ψ, ψs) => cons ψ ψs
|
||||
left_inv ψs := by
|
||||
apply Subtype.ext
|
||||
exact Eq.symm eq_head_cons_tail
|
||||
right_inv ψψs := by
|
||||
match ψψs with
|
||||
| (ψ, ψs) => rfl
|
||||
|
||||
/-- The instance of a finite type on `CrAnSection`s defined recursively through
|
||||
`consEquiv`. -/
|
||||
instance fintype : (φs : List 𝓕.States) → Fintype (CrAnSection φs)
|
||||
| [] => Fintype.ofEquiv _ nilEquiv.symm
|
||||
| _ :: φs =>
|
||||
haveI : Fintype (CrAnSection φs) := fintype φs
|
||||
Fintype.ofEquiv _ consEquiv.symm
|
||||
|
||||
@[simp]
|
||||
lemma sum_nil (f : CrAnSection (𝓕 := 𝓕) [] → M) [AddCommMonoid M] :
|
||||
∑ (s : CrAnSection []), f s = f ⟨[], rfl⟩ := by
|
||||
rw [← nilEquiv.symm.sum_comp]
|
||||
simp only [Finset.univ_unique, PUnit.default_eq_unit, Finset.sum_singleton]
|
||||
rfl
|
||||
|
||||
lemma sum_cons (f : CrAnSection (φ :: φs) → M) [AddCommMonoid M] :
|
||||
∑ (s : CrAnSection (φ :: φs)), f s = ∑ (a : 𝓕.statesToCrAnType φ),
|
||||
∑ (s : CrAnSection φs), f (cons a s) := by
|
||||
rw [← consEquiv.symm.sum_comp, Fintype.sum_prod_type]
|
||||
rfl
|
||||
|
||||
lemma sum_over_length {s : CrAnSection φs} (f : Fin s.1.length → M)
|
||||
[AddCommMonoid M] : ∑ (n : Fin s.1.length), f n =
|
||||
∑ (n : Fin φs.length), f (Fin.cast (length_eq s).symm n) := by
|
||||
rw [← (finCongr (length_eq s)).sum_comp]
|
||||
rfl
|
||||
|
||||
/-- The equivalence between `CrAnSection φs` and
|
||||
`CrAnSection φs'` induced by an equality `φs = φs'`. -/
|
||||
def congr : {φs φs' : List 𝓕.States} → (h : φs = φs') →
|
||||
CrAnSection φs ≃ CrAnSection φs'
|
||||
| _, _, rfl => Equiv.refl _
|
||||
|
||||
@[simp]
|
||||
lemma congr_fst {φs φs' : List 𝓕.States} (h : φs = φs') (ψs : CrAnSection φs) :
|
||||
(congr h ψs).1 = ψs.1 := by
|
||||
cases h
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma congr_symm {φs φs' : List 𝓕.States} (h : φs = φs') :
|
||||
(congr h).symm = congr h.symm := by
|
||||
cases h
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma congr_trans_apply {φs φs' φs'' : List 𝓕.States} (h1 : φs = φs') (h2 : φs' = φs'')
|
||||
(ψs : CrAnSection φs) :
|
||||
(congr h2 (congr h1 ψs)) = congr (by rw [h1, h2]) ψs := by
|
||||
subst h1 h2
|
||||
rfl
|
||||
|
||||
/-- Returns the first `n` elements of a section and its underlying list. -/
|
||||
def take (n : ℕ) (ψs : CrAnSection φs) : CrAnSection (φs.take n) :=
|
||||
⟨ψs.1.take n, by simp [ψs.2]⟩
|
||||
|
||||
@[simp]
|
||||
lemma take_congr {φs φs' : List 𝓕.States} (h : φs = φs') (n : ℕ)
|
||||
(ψs : CrAnSection φs) :
|
||||
(take n (congr h ψs)) = congr (by rw [h]) (take n ψs) := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
/-- Removes the first `n` elements of a section and its underlying list. -/
|
||||
def drop (n : ℕ) (ψs : CrAnSection φs) : CrAnSection (φs.drop n) :=
|
||||
⟨ψs.1.drop n, by simp [ψs.2]⟩
|
||||
|
||||
@[simp]
|
||||
lemma drop_congr {φs φs' : List 𝓕.States} (h : φs = φs') (n : ℕ)
|
||||
(ψs : CrAnSection φs) :
|
||||
(drop n (congr h ψs)) = congr (by rw [h]) (drop n ψs) := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
/-- Appends two sections and their underlying lists. -/
|
||||
def append {φs φs' : List 𝓕.States} (ψs : CrAnSection φs)
|
||||
(ψs' : CrAnSection φs') : CrAnSection (φs ++ φs') :=
|
||||
⟨ψs.1 ++ ψs'.1, by simp [ψs.2, ψs'.2]⟩
|
||||
|
||||
lemma append_assoc {φs φs' φs'' : List 𝓕.States} (ψs : CrAnSection φs)
|
||||
(ψs' : CrAnSection φs') (ψs'' : CrAnSection φs'') :
|
||||
append ψs (append ψs' ψs'') = congr (by simp) (append (append ψs ψs') ψs'') := by
|
||||
apply Subtype.ext
|
||||
simp [append]
|
||||
|
||||
lemma append_assoc' {φs φs' φs'' : List 𝓕.States} (ψs : CrAnSection φs)
|
||||
(ψs' : CrAnSection φs') (ψs'' : CrAnSection φs'') :
|
||||
(append (append ψs ψs') ψs'') = congr (by simp) (append ψs (append ψs' ψs'')) := by
|
||||
apply Subtype.ext
|
||||
simp [append]
|
||||
|
||||
lemma singletonEquiv_append_eq_cons {φs : List 𝓕.States} {φ : 𝓕.States}
|
||||
(ψs : CrAnSection φs) (ψ : 𝓕.statesToCrAnType φ) :
|
||||
append (singletonEquiv.symm ψ) ψs = cons ψ ψs := by
|
||||
apply Subtype.ext
|
||||
simp [append, cons, singletonEquiv]
|
||||
|
||||
@[simp]
|
||||
lemma take_append_drop {n : ℕ} (ψs : CrAnSection φs) :
|
||||
append (take n ψs) (drop n ψs) = congr (List.take_append_drop n φs).symm ψs := by
|
||||
apply Subtype.ext
|
||||
simp [take, drop, append]
|
||||
|
||||
lemma congr_append {φs1 φs1' φs2 φs2' : List 𝓕.States} (h1 : φs1 = φs1') (h2 : φs2 = φs2')
|
||||
(ψs1 : CrAnSection φs1) (ψs2 : CrAnSection φs2) :
|
||||
(append (congr h1 ψs1) (congr h2 ψs2)) = congr (by rw [h1, h2]) (append ψs1 ψs2) := by
|
||||
subst h1 h2
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma congr_fst_append {φs1 φs1' φs2 : List 𝓕.States} (h1 : φs1 = φs1')
|
||||
(ψs1 : CrAnSection φs1) (ψs2 : CrAnSection φs2) :
|
||||
(append (congr h1 ψs1) (ψs2)) = congr (by rw [h1]) (append ψs1 ψs2) := by
|
||||
subst h1
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma congr_snd_append {φs1 φs2 φs2' : List 𝓕.States} (h2 : φs2 = φs2')
|
||||
(ψs1 : CrAnSection φs1) (ψs2 : CrAnSection φs2) :
|
||||
(append ψs1 (congr h2 ψs2)) = congr (by rw [h2]) (append ψs1 ψs2) := by
|
||||
subst h2
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma take_left {φs φs' : List 𝓕.States} (ψs : CrAnSection φs)
|
||||
(ψs' : CrAnSection φs') :
|
||||
take φs.length (ψs.append ψs') = congr (by simp) ψs := by
|
||||
apply Subtype.ext
|
||||
simp [take, append]
|
||||
|
||||
@[simp]
|
||||
lemma drop_left {φs φs' : List 𝓕.States} (ψs : CrAnSection φs)
|
||||
(ψs' : CrAnSection φs') :
|
||||
drop φs.length (ψs.append ψs') = congr (by simp) ψs' := by
|
||||
apply Subtype.ext
|
||||
simp [drop, append]
|
||||
|
||||
/-- The equivalence between `CrAnSection (φs ++ φs')` and
|
||||
`CrAnSection φs × CrAnSection φs` formed by `append`, `take`
|
||||
and `drop` and their interrelationship. -/
|
||||
def appendEquiv {φs φs' : List 𝓕.States} : CrAnSection (φs ++ φs') ≃
|
||||
CrAnSection φs × CrAnSection φs' where
|
||||
toFun ψs := (congr (List.take_left φs φs') (take φs.length ψs),
|
||||
congr (List.drop_left φs φs') (drop φs.length ψs))
|
||||
invFun ψsψs' := append ψsψs'.1 ψsψs'.2
|
||||
left_inv ψs := by
|
||||
apply Subtype.ext
|
||||
simp
|
||||
right_inv ψsψs' := by
|
||||
match ψsψs' with
|
||||
| (ψs, ψs') =>
|
||||
simp only [take_left, drop_left, Prod.mk.injEq]
|
||||
refine And.intro (Subtype.ext ?_) (Subtype.ext ?_)
|
||||
· simp
|
||||
· simp
|
||||
|
||||
@[simp]
|
||||
lemma _root_.List.map_eraseIdx {α β : Type} (f : α → β) : (l : List α) → (n : ℕ) →
|
||||
List.map f (l.eraseIdx n) = (List.map f l).eraseIdx n
|
||||
| [], _ => rfl
|
||||
| a :: l, 0 => rfl
|
||||
| a :: l, n+1 => by
|
||||
simp only [List.eraseIdx, List.map_cons, List.cons.injEq, true_and]
|
||||
exact List.map_eraseIdx f l n
|
||||
|
||||
/-- Erasing an element from a section and it's underlying list. -/
|
||||
def eraseIdx (n : ℕ) (ψs : CrAnSection φs) : CrAnSection (φs.eraseIdx n) :=
|
||||
⟨ψs.1.eraseIdx n, by simp [ψs.2]⟩
|
||||
|
||||
/-- The equivalence formed by extracting an element from a section. -/
|
||||
def eraseIdxEquiv (n : ℕ) (φs : List 𝓕.States) (hn : n < φs.length) :
|
||||
CrAnSection φs ≃ 𝓕.statesToCrAnType φs[n] ×
|
||||
CrAnSection (φs.eraseIdx n) :=
|
||||
(congr (by simp only [List.take_concat_get', List.take_append_drop])).trans <|
|
||||
appendEquiv.trans <|
|
||||
(Equiv.prodCongr (appendEquiv.trans (Equiv.prodComm _ _)) (Equiv.refl _)).trans <|
|
||||
(Equiv.prodAssoc _ _ _).trans <|
|
||||
Equiv.prodCongr singletonEquiv <|
|
||||
appendEquiv.symm.trans <|
|
||||
congr (List.eraseIdx_eq_take_drop_succ φs n).symm
|
||||
|
||||
@[simp]
|
||||
lemma eraseIdxEquiv_apply_snd {n : ℕ} (ψs : CrAnSection φs) (hn : n < φs.length) :
|
||||
(eraseIdxEquiv n φs hn ψs).snd = eraseIdx n ψs := by
|
||||
apply Subtype.ext
|
||||
simp only [eraseIdxEquiv, appendEquiv, take, List.take_concat_get', List.length_take, drop,
|
||||
append, Equiv.trans_apply, Equiv.coe_fn_mk, congr_fst, Equiv.prodCongr_apply, Equiv.coe_trans,
|
||||
Equiv.coe_prodComm, Equiv.coe_refl, Prod.map_apply, Function.comp_apply, Prod.swap_prod_mk,
|
||||
id_eq, Equiv.prodAssoc_apply, Equiv.coe_fn_symm_mk, eraseIdx]
|
||||
rw [Nat.min_eq_left (Nat.le_of_succ_le hn), Nat.min_eq_left hn, List.take_take]
|
||||
simp only [Nat.succ_eq_add_one, le_add_iff_nonneg_right, zero_le, inf_of_le_left]
|
||||
exact Eq.symm (List.eraseIdx_eq_take_drop_succ ψs.1 n)
|
||||
|
||||
lemma eraseIdxEquiv_symm_eq_take_cons_drop {n : ℕ} (φs : List 𝓕.States) (hn : n < φs.length)
|
||||
(a : 𝓕.statesToCrAnType φs[n]) (s : CrAnSection (φs.eraseIdx n)) :
|
||||
(eraseIdxEquiv n φs hn).symm ⟨a, s⟩ =
|
||||
congr (by
|
||||
rw [HepLean.List.take_eraseIdx_same, HepLean.List.drop_eraseIdx_succ]
|
||||
conv_rhs => rw [← List.take_append_drop n φs]) (append (take n s) (cons a (drop n s))) := by
|
||||
simp only [eraseIdxEquiv, appendEquiv, Equiv.symm_trans_apply, congr_symm, Equiv.prodCongr_symm,
|
||||
Equiv.refl_symm, Equiv.prodCongr_apply, Prod.map_apply, Equiv.symm_symm, Equiv.coe_fn_mk,
|
||||
take_congr, congr_trans_apply, drop_congr, Equiv.prodAssoc_symm_apply, Equiv.coe_refl,
|
||||
Equiv.prodComm_symm, Equiv.prodComm_apply, Prod.swap_prod_mk, Equiv.coe_fn_symm_mk,
|
||||
congr_fst_append, id_eq, congr_snd_append]
|
||||
rw [append_assoc', singletonEquiv_append_eq_cons]
|
||||
simp only [List.singleton_append, congr_trans_apply]
|
||||
apply Subtype.ext
|
||||
simp only [congr_fst]
|
||||
have hn : (List.take n φs).length = n := by
|
||||
rw [@List.length_take]
|
||||
simp only [inf_eq_left]
|
||||
exact Nat.le_of_succ_le hn
|
||||
rw [hn]
|
||||
|
||||
@[simp]
|
||||
lemma eraseIdxEquiv_symm_getElem {n : ℕ} (φs : List 𝓕.States) (hn : n < φs.length)
|
||||
(a : 𝓕.statesToCrAnType φs[n]) (s : CrAnSection (φs.eraseIdx n)) :
|
||||
getElem ((eraseIdxEquiv n φs hn).symm ⟨a,s⟩).1 n
|
||||
(by rw [length_eq]; exact hn) = ⟨φs[n], a⟩ := by
|
||||
rw [eraseIdxEquiv_symm_eq_take_cons_drop]
|
||||
simp only [append, take, cons, drop, congr_fst]
|
||||
rw [List.getElem_append]
|
||||
simp only [List.length_take, length_eq, lt_inf_iff, lt_self_iff_false, false_and, ↓reduceDIte]
|
||||
have h0 : n ⊓ (φs.eraseIdx n).length = n := by
|
||||
simp only [inf_eq_left]
|
||||
rw [← HepLean.List.eraseIdx_length _ ⟨n, hn⟩] at hn
|
||||
exact Nat.le_of_lt_succ hn
|
||||
simp [h0]
|
||||
|
||||
@[simp]
|
||||
lemma eraseIdxEquiv_symm_eraseIdx {n : ℕ} (φs : List 𝓕.States) (hn : n < φs.length)
|
||||
(a : 𝓕.statesToCrAnType φs[n]) (s : CrAnSection (φs.eraseIdx n)) :
|
||||
((eraseIdxEquiv n φs hn).symm ⟨a, s⟩).1.eraseIdx n = s.1 := by
|
||||
change (((eraseIdxEquiv n φs hn).symm ⟨a, s⟩).eraseIdx n).1 = _
|
||||
rw [← eraseIdxEquiv_apply_snd _ hn]
|
||||
simp
|
||||
|
||||
lemma sum_eraseIdxEquiv (n : ℕ) (φs : List 𝓕.States) (hn : n < φs.length)
|
||||
(f : CrAnSection φs → M) [AddCommMonoid M] : ∑ (s : CrAnSection φs), f s =
|
||||
∑ (a : 𝓕.statesToCrAnType φs[n]), ∑ (s : CrAnSection (φs.eraseIdx n)),
|
||||
f ((eraseIdxEquiv n φs hn).symm ⟨a, s⟩) := by
|
||||
rw [← (eraseIdxEquiv n φs hn).symm.sum_comp]
|
||||
rw [Fintype.sum_prod_type]
|
||||
|
||||
end CrAnSection
|
||||
|
||||
end FieldStruct
|
|
@ -7,30 +7,30 @@ import HepLean.PerturbationTheory.FieldStruct.Basic
|
|||
import HepLean.PerturbationTheory.CreateAnnihilate
|
||||
/-!
|
||||
|
||||
# Creation and annihilation parts of fields
|
||||
# Creation and annihlation parts of fields
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable (𝓕 : FieldStruct)
|
||||
|
||||
/-- To each state the specification of the type of creation and annihilation parts.
|
||||
For asymptotic states there is only one allowed part, whilst for position states
|
||||
/-- To each state the specificaition of the type of creation and annihlation parts.
|
||||
For asymptotic staes there is only one allowed part, whilst for position states
|
||||
there is two. -/
|
||||
def statesToCreateAnnihilateType : 𝓕.States → Type
|
||||
def statesToCrAnType : 𝓕.States → Type
|
||||
| States.negAsymp _ => Unit
|
||||
| States.position _ => CreateAnnihilate
|
||||
| States.posAsymp _ => Unit
|
||||
|
||||
/-- The instance of a finite type on `𝓕.statesToCreateAnnihilateType i`. -/
|
||||
instance : ∀ i, Fintype (𝓕.statesToCreateAnnihilateType i) := fun i =>
|
||||
instance : ∀ i, Fintype (𝓕.statesToCrAnType i) := fun i =>
|
||||
match i with
|
||||
| States.negAsymp _ => inferInstanceAs (Fintype Unit)
|
||||
| States.position _ => inferInstanceAs (Fintype CreateAnnihilate)
|
||||
| States.posAsymp _ => inferInstanceAs (Fintype Unit)
|
||||
|
||||
/-- The instance of a decidable equality on `𝓕.statesToCreateAnnihilateType i`. -/
|
||||
instance : ∀ i, DecidableEq (𝓕.statesToCreateAnnihilateType i) := fun i =>
|
||||
instance : ∀ i, DecidableEq (𝓕.statesToCrAnType i) := fun i =>
|
||||
match i with
|
||||
| States.negAsymp _ => inferInstanceAs (DecidableEq Unit)
|
||||
| States.position _ => inferInstanceAs (DecidableEq CreateAnnihilate)
|
||||
|
@ -39,40 +39,44 @@ instance : ∀ i, DecidableEq (𝓕.statesToCreateAnnihilateType i) := fun i =>
|
|||
/-- The equivalence between `𝓕.statesToCreateAnnihilateType i` and
|
||||
`𝓕.statesToCreateAnnihilateType j` from an equality `i = j`. -/
|
||||
def statesToCreateAnnihilateTypeCongr : {i j : 𝓕.States} → i = j →
|
||||
𝓕.statesToCreateAnnihilateType i ≃ 𝓕.statesToCreateAnnihilateType j
|
||||
𝓕.statesToCrAnType i ≃ 𝓕.statesToCrAnType j
|
||||
| _, _, rfl => Equiv.refl _
|
||||
|
||||
/-- A creation and annihilation state is a state plus an valid specification of the
|
||||
/-- A creation and annihlation state is a state plus an valid specification of the
|
||||
creation or annihliation part of that state. (For asympotic states there is only one valid
|
||||
choice). -/
|
||||
def CreateAnnihilateStates : Type := Σ (s : 𝓕.States), 𝓕.statesToCreateAnnihilateType s
|
||||
def CrAnStates : Type := Σ (s : 𝓕.States), 𝓕.statesToCrAnType s
|
||||
|
||||
/-- The map from creation and annihilation states to their underlying states. -/
|
||||
def createAnnihilateStatesToStates : 𝓕.CreateAnnihilateStates → 𝓕.States := Sigma.fst
|
||||
/-- The map from creation and annihlation states to their underlying states. -/
|
||||
def crAnStatesToStates : 𝓕.CrAnStates → 𝓕.States := Sigma.fst
|
||||
|
||||
@[simp]
|
||||
lemma createAnnihilateStatesToStates_prod (s : 𝓕.States) (t : 𝓕.statesToCreateAnnihilateType s) :
|
||||
𝓕.createAnnihilateStatesToStates ⟨s, t⟩ = s := rfl
|
||||
lemma crAnStatesToStates_prod (s : 𝓕.States) (t : 𝓕.statesToCrAnType s) :
|
||||
𝓕.crAnStatesToStates ⟨s, t⟩ = s := rfl
|
||||
|
||||
/-- The map from creation and annihilation states to the type `CreateAnnihilate`
|
||||
/-- The map from creation and annihlation states to the type `CreateAnnihilate`
|
||||
specifying if a state is a creation or an annihilation state. -/
|
||||
def createAnnihlateStatesToCreateAnnihilate : 𝓕.CreateAnnihilateStates → CreateAnnihilate
|
||||
def crAnStatesToCreateAnnihilate : 𝓕.CrAnStates → CreateAnnihilate
|
||||
| ⟨States.negAsymp _, _⟩ => CreateAnnihilate.create
|
||||
| ⟨States.position _, CreateAnnihilate.create⟩ => CreateAnnihilate.create
|
||||
| ⟨States.position _, CreateAnnihilate.annihilate⟩ => CreateAnnihilate.annihilate
|
||||
| ⟨States.posAsymp _, _⟩ => CreateAnnihilate.annihilate
|
||||
|
||||
/-- The normal ordering on creation and annihilation states. -/
|
||||
def normalOrder : 𝓕.CreateAnnihilateStates → 𝓕.CreateAnnihilateStates → Prop :=
|
||||
fun a b => CreateAnnihilate.normalOrder (𝓕.createAnnihlateStatesToCreateAnnihilate a)
|
||||
(𝓕.createAnnihlateStatesToCreateAnnihilate b)
|
||||
/-- Takes a `CrAnStates` state to its corresponding fields statistic (bosonic or fermionic). -/
|
||||
def crAnStatistics : 𝓕.CrAnStates → FieldStatistic :=
|
||||
𝓕.statesStatistic ∘ 𝓕.crAnStatesToStates
|
||||
|
||||
/-- Normal ordering is total. -/
|
||||
instance : IsTotal 𝓕.CreateAnnihilateStates 𝓕.normalOrder where
|
||||
total _ _ := total_of CreateAnnihilate.normalOrder _ _
|
||||
/-- The field statistic of a `CrAnState`. -/
|
||||
scoped[FieldStruct] notation 𝓕 "|>ₛ" φ =>
|
||||
(crAnStatistics 𝓕) φ
|
||||
|
||||
/-- Normal ordering is transitive. -/
|
||||
instance : IsTrans 𝓕.CreateAnnihilateStates 𝓕.normalOrder where
|
||||
trans _ _ _ := fun h h' => IsTrans.trans (α := CreateAnnihilate) _ _ _ h h'
|
||||
/-- The field statistic of a list of `CrAnState`s. -/
|
||||
scoped[FieldStruct] notation 𝓕 "|>ₛ" φ => FieldStatistic.ofList
|
||||
(crAnStatistics 𝓕) φ
|
||||
|
||||
/-- The `CreateAnnihilate` value of a `CrAnState`s, i.e. whether it is a creation or
|
||||
annihilation operator. -/
|
||||
scoped[FieldStruct] infixl:80 "|>ᶜ" =>
|
||||
crAnStatesToCreateAnnihilate
|
||||
|
||||
end FieldStruct
|
||||
|
|
|
@ -1,216 +0,0 @@
|
|||
/-
|
||||
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.FieldStruct.CreateAnnihilate
|
||||
import HepLean.PerturbationTheory.CreateAnnihilate
|
||||
/-!
|
||||
|
||||
# Creation and annihilation sections
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
/-- The sections in `𝓕.CreateAnnihilateStates` over a list `φs : List 𝓕.States`.
|
||||
In terms of physics, given some fields `φ₁...φₙ`, the different ways one can associate
|
||||
each field as a `creation` or an `annihilation` operator. E.g. the number of terms
|
||||
`φ₁⁰φ₂¹...φₙ⁰` `φ₁¹φ₂¹...φₙ⁰` etc. If some fields are exclusively creation or annihilation
|
||||
operators at this point (e.g. asymptotic states) this is accounted for. -/
|
||||
def CreateAnnihilateSect (φs : List 𝓕.States) : Type :=
|
||||
{ψs : List 𝓕.CreateAnnihilateStates //
|
||||
List.map 𝓕.createAnnihilateStatesToStates ψs = φs}
|
||||
-- Π i, 𝓕.statesToCreateAnnihilateType (φs.get i)
|
||||
|
||||
namespace CreateAnnihilateSect
|
||||
|
||||
variable {𝓕 : FieldStruct} {φs : List 𝓕.States}
|
||||
|
||||
@[simp]
|
||||
lemma length_eq (ψs : CreateAnnihilateSect φs) : ψs.1.length = φs.length := by
|
||||
simpa using congrArg List.length ψs.2
|
||||
|
||||
/-- The tail of a section for `φs`. -/
|
||||
def tail : {φs : List 𝓕.States} → (ψs : CreateAnnihilateSect φs) → CreateAnnihilateSect φs.tail
|
||||
| [], ⟨[], h⟩ => ⟨[], h⟩
|
||||
| φ :: φs, ⟨[], h⟩ => False.elim (by simp at h)
|
||||
| φ :: φs, ⟨ψ :: ψs, h⟩ => ⟨ψs, by rw [List.map_cons, List.cons.injEq] at h; exact h.2⟩
|
||||
|
||||
lemma head_state_eq {φ : 𝓕.States} : (ψs : CreateAnnihilateSect (φ :: φs)) →
|
||||
(ψs.1.head (by simp [← List.length_pos_iff_ne_nil])).1 = φ
|
||||
| ⟨[], h⟩ => False.elim (by simp at h)
|
||||
| ⟨ψ :: ψs, h⟩ => by
|
||||
simp at h
|
||||
exact h.1
|
||||
|
||||
/-- The head of a section for `φ :: φs` as an element in `𝓕.statesToCreateAnnihilateType φ`. -/
|
||||
def head : {φ : 𝓕.States} → (ψs : CreateAnnihilateSect (φ :: φs)) →
|
||||
𝓕.statesToCreateAnnihilateType φ
|
||||
| φ, ⟨[], h⟩ => False.elim (by simp at h)
|
||||
| φ, ⟨ψ :: ψs, h⟩ => 𝓕.statesToCreateAnnihilateTypeCongr (by
|
||||
simpa using head_state_eq ⟨ψ :: ψs, h⟩) ψ.2
|
||||
|
||||
lemma eq_head_cons_tail {φ : 𝓕.States} {ψs : CreateAnnihilateSect (φ :: φs)} :
|
||||
ψs.1 = ⟨φ, head ψs⟩ :: ψs.tail.1 := by
|
||||
match ψs with
|
||||
| ⟨[], h⟩ => exact False.elim (by simp at h)
|
||||
| ⟨ψ :: ψs, h⟩ =>
|
||||
have h2 := head_state_eq ⟨ψ :: ψs, h⟩
|
||||
simp at h2
|
||||
subst h2
|
||||
rfl
|
||||
|
||||
/-- The creation of a section from for `φ : φs` from a section for `φs` and a
|
||||
element of `𝓕.statesToCreateAnnihilateType φ`. -/
|
||||
def cons {φ : 𝓕.States} (ψ : 𝓕.statesToCreateAnnihilateType φ) (ψs : CreateAnnihilateSect φs) :
|
||||
CreateAnnihilateSect (φ :: φs) := ⟨⟨φ, ψ⟩ :: ψs.1, by
|
||||
simp [List.map_cons, ψs.2]⟩
|
||||
|
||||
/-- The creation and annihilation sections for a singleton list is given by
|
||||
a choice of `𝓕.statesToCreateAnnihilateType φ`. If `φ` is a asymptotic state
|
||||
there is no choice here, else there are two choices. -/
|
||||
def singletonEquiv {φ : 𝓕.States} : CreateAnnihilateSect [φ] ≃
|
||||
𝓕.statesToCreateAnnihilateType φ where
|
||||
toFun ψs := ψs.head
|
||||
invFun ψ := ⟨[⟨φ, ψ⟩], by simp⟩
|
||||
left_inv ψs := by
|
||||
apply Subtype.ext
|
||||
simp only
|
||||
have h1 := eq_head_cons_tail (ψs := ψs)
|
||||
rw [h1]
|
||||
have h2 := ψs.tail.2
|
||||
simp at h2
|
||||
simp [h2]
|
||||
right_inv ψ := by
|
||||
simp [head]
|
||||
rfl
|
||||
|
||||
/-- An equivalence separating the head of a creation and annihilation section
|
||||
from the tail. -/
|
||||
def consEquiv {φ : 𝓕.States} {φs : List 𝓕.States} : CreateAnnihilateSect (φ :: φs) ≃
|
||||
𝓕.statesToCreateAnnihilateType φ × CreateAnnihilateSect φs where
|
||||
toFun ψs := ⟨ψs.head, ψs.tail⟩
|
||||
invFun ψψs :=
|
||||
match ψψs with
|
||||
| (ψ, ψs) => cons ψ ψs
|
||||
left_inv ψs := by
|
||||
apply Subtype.ext
|
||||
exact Eq.symm eq_head_cons_tail
|
||||
right_inv ψψs := by
|
||||
match ψψs with
|
||||
| (ψ, ψs) => rfl
|
||||
|
||||
/-- The equivalence between `CreateAnnihilateSect φs` and
|
||||
`CreateAnnihilateSect φs'` induced by an equality `φs = φs'`. -/
|
||||
def congr : {φs φs' : List 𝓕.States} → (h : φs = φs') →
|
||||
CreateAnnihilateSect φs ≃ CreateAnnihilateSect φs'
|
||||
| _, _, rfl => Equiv.refl _
|
||||
|
||||
@[simp]
|
||||
lemma congr_fst {φs φs' : List 𝓕.States} (h : φs = φs') (ψs : CreateAnnihilateSect φs) :
|
||||
(congr h ψs).1 = ψs.1 := by
|
||||
cases h
|
||||
rfl
|
||||
|
||||
/-- Returns the first `n` elements of a section and its underlying list. -/
|
||||
def take (n : ℕ) (ψs : CreateAnnihilateSect φs) : CreateAnnihilateSect (φs.take n) :=
|
||||
⟨ψs.1.take n, by simp [ψs.2]⟩
|
||||
|
||||
/-- Removes the first `n` elements of a section and its underlying list. -/
|
||||
def drop (n : ℕ) (ψs : CreateAnnihilateSect φs) : CreateAnnihilateSect (φs.drop n) :=
|
||||
⟨ψs.1.drop n, by simp [ψs.2]⟩
|
||||
|
||||
/-- Appends two sections and their underlying lists. -/
|
||||
def append {φs φs' : List 𝓕.States} (ψs : CreateAnnihilateSect φs)
|
||||
(ψs' : CreateAnnihilateSect φs') : CreateAnnihilateSect (φs ++ φs') :=
|
||||
⟨ψs.1 ++ ψs'.1, by simp [ψs.2, ψs'.2]⟩
|
||||
|
||||
@[simp]
|
||||
lemma take_append_drop {n : ℕ} (ψs : CreateAnnihilateSect φs) :
|
||||
append (take n ψs) (drop n ψs) = congr (List.take_append_drop n φs).symm ψs := by
|
||||
apply Subtype.ext
|
||||
simp [take, drop, append]
|
||||
|
||||
@[simp]
|
||||
lemma congr_append {φs1 φs1' φs2 φs2' : List 𝓕.States}
|
||||
(h1 : φs1 = φs1') (h2 : φs2 = φs2')
|
||||
(ψs1 : CreateAnnihilateSect φs1) (ψs2 : CreateAnnihilateSect φs2) :
|
||||
(append (congr h1 ψs1) (congr h2 ψs2)) = congr (by rw [h1, h2]) (append ψs1 ψs2) := by
|
||||
subst h1 h2
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma take_left {φs φs' : List 𝓕.States} (ψs : CreateAnnihilateSect φs)
|
||||
(ψs' : CreateAnnihilateSect φs') :
|
||||
take φs.length (ψs.append ψs') = congr (by simp) ψs := by
|
||||
apply Subtype.ext
|
||||
simp [take, append]
|
||||
|
||||
@[simp]
|
||||
lemma drop_left {φs φs' : List 𝓕.States} (ψs : CreateAnnihilateSect φs)
|
||||
(ψs' : CreateAnnihilateSect φs') :
|
||||
drop φs.length (ψs.append ψs') = congr (by simp) ψs' := by
|
||||
apply Subtype.ext
|
||||
simp [drop, append]
|
||||
|
||||
/-- The equivalence between `CreateAnnihilateSect (φs ++ φs')` and
|
||||
`CreateAnnihilateSect φs × CreateAnnihilateSect φs` formed by `append`, `take`
|
||||
and `drop` and their interrelationship. -/
|
||||
def appendEquiv {φs φs' : List 𝓕.States} : CreateAnnihilateSect (φs ++ φs') ≃
|
||||
CreateAnnihilateSect φs × CreateAnnihilateSect φs' where
|
||||
toFun ψs := (congr (List.take_left φs φs') (take φs.length ψs),
|
||||
congr (List.drop_left φs φs') (drop φs.length ψs))
|
||||
invFun ψsψs' := append ψsψs'.1 ψsψs'.2
|
||||
left_inv ψs := by
|
||||
apply Subtype.ext
|
||||
simp
|
||||
right_inv ψsψs' := by
|
||||
match ψsψs' with
|
||||
| (ψs, ψs') =>
|
||||
simp only [take_left, drop_left, Prod.mk.injEq]
|
||||
refine And.intro (Subtype.ext ?_) (Subtype.ext ?_)
|
||||
· simp
|
||||
· simp
|
||||
|
||||
@[simp]
|
||||
lemma _root_.List.map_eraseIdx {α β : Type} (f : α → β) : (l : List α) → (n : ℕ) →
|
||||
List.map f (l.eraseIdx n) = (List.map f l).eraseIdx n
|
||||
| [], _ => rfl
|
||||
| a :: l, 0 => rfl
|
||||
| a :: l, n+1 => by
|
||||
simp only [List.eraseIdx, List.map_cons, List.cons.injEq, true_and]
|
||||
exact List.map_eraseIdx f l n
|
||||
|
||||
/-- Erasing an element from a section and it's underlying list. -/
|
||||
def eraseIdx (n : ℕ) (ψs : CreateAnnihilateSect φs) : CreateAnnihilateSect (φs.eraseIdx n) :=
|
||||
⟨ψs.1.eraseIdx n, by simp [ψs.2]⟩
|
||||
|
||||
/-- The equivalence formed by extracting an element from a section. -/
|
||||
def eraseIdxEquiv (n : ℕ) (φs : List 𝓕.States) (hn : n < φs.length) :
|
||||
CreateAnnihilateSect φs ≃ 𝓕.statesToCreateAnnihilateType φs[n] ×
|
||||
CreateAnnihilateSect (φs.eraseIdx n) :=
|
||||
(congr (by simp only [List.take_concat_get', List.take_append_drop])).trans <|
|
||||
appendEquiv.trans <|
|
||||
(Equiv.prodCongr (appendEquiv.trans (Equiv.prodComm _ _)) (Equiv.refl _)).trans <|
|
||||
(Equiv.prodAssoc _ _ _).trans <|
|
||||
Equiv.prodCongr singletonEquiv <|
|
||||
appendEquiv.symm.trans <|
|
||||
congr (List.eraseIdx_eq_take_drop_succ φs n).symm
|
||||
|
||||
@[simp]
|
||||
lemma eraseIdxEquiv_apply_snd {n : ℕ} (ψs : CreateAnnihilateSect φs) (hn : n < φs.length) :
|
||||
(eraseIdxEquiv n φs hn ψs).snd = eraseIdx n ψs := by
|
||||
apply Subtype.ext
|
||||
simp only [eraseIdxEquiv, appendEquiv, take, List.take_concat_get', List.length_take, drop,
|
||||
append, Equiv.trans_apply, Equiv.coe_fn_mk, congr_fst, Equiv.prodCongr_apply, Equiv.coe_trans,
|
||||
Equiv.coe_prodComm, Equiv.coe_refl, Prod.map_apply, Function.comp_apply, Prod.swap_prod_mk,
|
||||
id_eq, Equiv.prodAssoc_apply, Equiv.coe_fn_symm_mk, eraseIdx]
|
||||
rw [Nat.min_eq_left (Nat.le_of_succ_le hn), Nat.min_eq_left hn, List.take_take]
|
||||
simp only [Nat.succ_eq_add_one, le_add_iff_nonneg_right, zero_le, inf_of_le_left]
|
||||
exact Eq.symm (List.eraseIdx_eq_take_drop_succ ψs.1 n)
|
||||
|
||||
end CreateAnnihilateSect
|
||||
|
||||
end FieldStruct
|
89
HepLean/PerturbationTheory/FieldStruct/Filters.lean
Normal file
89
HepLean/PerturbationTheory/FieldStruct/Filters.lean
Normal file
|
@ -0,0 +1,89 @@
|
|||
/-
|
||||
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.Algebras.OperatorAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
/-!
|
||||
|
||||
# Filters of lists of CrAnStates
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
/-- Given a list of creation and annihilation states, the filtered list only containing
|
||||
the creation states. As a schematic example, for the list:
|
||||
- `[φ1c, φ1a, φ2c, φ2a]` this will return `[φ1c, φ2c]`.
|
||||
-/
|
||||
def createFilter (φs : List 𝓕.CrAnStates) : List 𝓕.CrAnStates :=
|
||||
List.filter (fun φ => 𝓕 |>ᶜ φ = CreateAnnihilate.create) φs
|
||||
|
||||
lemma createFilter_cons_create {φ : 𝓕.CrAnStates}
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create) (φs : List 𝓕.CrAnStates) :
|
||||
createFilter (φ :: φs) = φ :: createFilter φs := by
|
||||
simp only [createFilter]
|
||||
rw [List.filter_cons_of_pos]
|
||||
simp [hφ]
|
||||
|
||||
lemma createFilter_cons_annihilate {φ : 𝓕.CrAnStates}
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) (φs : List 𝓕.CrAnStates) :
|
||||
createFilter (φ :: φs) = createFilter φs := by
|
||||
simp only [createFilter]
|
||||
rw [List.filter_cons_of_neg]
|
||||
simp [hφ]
|
||||
|
||||
lemma createFilter_append (φs φs' : List 𝓕.CrAnStates) :
|
||||
createFilter (φs ++ φs') = createFilter φs ++ createFilter φs' := by
|
||||
rw [createFilter, List.filter_append]
|
||||
rfl
|
||||
|
||||
lemma createFilter_singleton_create (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create) :
|
||||
createFilter [φ] = [φ] := by
|
||||
simp [createFilter, hφ]
|
||||
|
||||
lemma createFilter_singleton_annihilate (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) : createFilter [φ] = [] := by
|
||||
simp [createFilter, hφ]
|
||||
|
||||
/-- Given a list of creation and annihilation states, the filtered list only containing
|
||||
the annihilation states.
|
||||
As a schematic example, for the list:
|
||||
- `[φ1c, φ1a, φ2c, φ2a]` this will return `[φ1a, φ2a]`.
|
||||
-/
|
||||
def annihilateFilter (φs : List 𝓕.CrAnStates) : List 𝓕.CrAnStates :=
|
||||
List.filter (fun φ => 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) φs
|
||||
|
||||
lemma annihilateFilter_cons_create {φ : 𝓕.CrAnStates}
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create) (φs : List 𝓕.CrAnStates) :
|
||||
annihilateFilter (φ :: φs) = annihilateFilter φs := by
|
||||
simp only [annihilateFilter]
|
||||
rw [List.filter_cons_of_neg]
|
||||
simp [hφ]
|
||||
|
||||
lemma annihilateFilter_cons_annihilate {φ : 𝓕.CrAnStates}
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) (φs : List 𝓕.CrAnStates) :
|
||||
annihilateFilter (φ :: φs) = φ :: annihilateFilter φs := by
|
||||
simp only [annihilateFilter]
|
||||
rw [List.filter_cons_of_pos]
|
||||
simp [hφ]
|
||||
|
||||
lemma annihilateFilter_append (φs φs' : List 𝓕.CrAnStates) :
|
||||
annihilateFilter (φs ++ φs') = annihilateFilter φs ++ annihilateFilter φs' := by
|
||||
rw [annihilateFilter, List.filter_append]
|
||||
rfl
|
||||
|
||||
lemma annihilateFilter_singleton_create (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create) :
|
||||
annihilateFilter [φ] = [] := by
|
||||
simp [annihilateFilter, hφ]
|
||||
|
||||
lemma annihilateFilter_singleton_annihilate (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) :
|
||||
annihilateFilter [φ] = [φ] := by
|
||||
simp [annihilateFilter, hφ]
|
||||
|
||||
end FieldStruct
|
428
HepLean/PerturbationTheory/FieldStruct/NormalOrder.lean
Normal file
428
HepLean/PerturbationTheory/FieldStruct/NormalOrder.lean
Normal file
|
@ -0,0 +1,428 @@
|
|||
/-
|
||||
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.Algebras.OperatorAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
import HepLean.PerturbationTheory.FieldStruct.Filters
|
||||
/-!
|
||||
|
||||
# Normal Ordering
|
||||
|
||||
The normal ordering puts all creation operators to the left and all annihilation operators to the
|
||||
right. It acts on `CrAnStates` and defines a linear map from the `CrAnAlgebra` to itself.
|
||||
|
||||
The normal ordering satisfies a number of nice properties with relation to the operator
|
||||
algebra 𝓞.A.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
/-- The normal ordering relation on creation and annihlation operators.
|
||||
For a list of creation and annihlation states, this relation is designed
|
||||
to move all creation states to the left, and all annihlation operators to the right.
|
||||
We have that `normalOrderRel φ1 φ2` is true if
|
||||
- `φ1` is a creation operator
|
||||
or
|
||||
- `φ2` is an annihlate operator. -/
|
||||
def normalOrderRel : 𝓕.CrAnStates → 𝓕.CrAnStates → Prop :=
|
||||
fun a b => CreateAnnihilate.normalOrder (𝓕 |>ᶜ a) (𝓕 |>ᶜ b)
|
||||
|
||||
/-- Normal ordering is total. -/
|
||||
instance : IsTotal 𝓕.CrAnStates 𝓕.normalOrderRel where
|
||||
total _ _ := total_of CreateAnnihilate.normalOrder _ _
|
||||
|
||||
/-- Normal ordering is transitive. -/
|
||||
instance : IsTrans 𝓕.CrAnStates 𝓕.normalOrderRel where
|
||||
trans _ _ _ := fun h h' => IsTrans.trans (α := CreateAnnihilate) _ _ _ h h'
|
||||
|
||||
/-- A decidable instance on the normal ordering relation. -/
|
||||
instance (φ φ' : 𝓕.CrAnStates) : Decidable (normalOrderRel φ φ') :=
|
||||
CreateAnnihilate.instDecidableNormalOrder (𝓕 |>ᶜ φ) (𝓕 |>ᶜ φ')
|
||||
|
||||
/-!
|
||||
|
||||
## Normal order sign.
|
||||
|
||||
-/
|
||||
|
||||
/-- The sign associated with putting a list of creation and annihlation states into normal order
|
||||
(with the creation operators on the left).
|
||||
We pick up a minus sign for every fermion paired crossed. -/
|
||||
def normalOrderSign (φs : List 𝓕.CrAnStates) : ℂ :=
|
||||
Wick.koszulSign 𝓕.crAnStatistics 𝓕.normalOrderRel φs
|
||||
|
||||
@[simp]
|
||||
lemma normalOrderSign_mul_self (φs : List 𝓕.CrAnStates) :
|
||||
normalOrderSign φs * normalOrderSign φs = 1 := by
|
||||
simp [normalOrderSign, Wick.koszulSign, Wick.koszulSign_mul_self]
|
||||
|
||||
lemma koszulSignInsert_create (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create) : (φs : List 𝓕.CrAnStates) →
|
||||
Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φ φs = 1
|
||||
| [] => rfl
|
||||
| φ' :: φs => by
|
||||
dsimp only [Wick.koszulSignInsert]
|
||||
rw [if_pos]
|
||||
· exact koszulSignInsert_create φ hφ φs
|
||||
· dsimp only [normalOrderRel]
|
||||
rw [hφ]
|
||||
dsimp [CreateAnnihilate.normalOrder]
|
||||
|
||||
lemma normalOrderSign_cons_create (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create) (φs : List 𝓕.CrAnStates) :
|
||||
normalOrderSign (φ :: φs) = normalOrderSign φs := by
|
||||
dsimp only [normalOrderSign, Wick.koszulSign]
|
||||
rw [koszulSignInsert_create φ hφ φs]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma normalOrderSign_singleton (φ : 𝓕.CrAnStates) :
|
||||
normalOrderSign [φ] = 1 := by
|
||||
simp [normalOrderSign]
|
||||
|
||||
@[simp]
|
||||
lemma normalOrderSign_nil :
|
||||
normalOrderSign (𝓕 := 𝓕) [] = 1 := by
|
||||
simp [normalOrderSign, Wick.koszulSign]
|
||||
|
||||
lemma koszulSignInsert_append_annihilate (φ' φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) :
|
||||
(φs : List 𝓕.CrAnStates) →
|
||||
Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φ' (φs ++ [φ]) =
|
||||
Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φ' φs
|
||||
| [] => by
|
||||
simp only [Wick.koszulSignInsert, normalOrderRel, hφ, ite_eq_left_iff,
|
||||
CreateAnnihilate.not_normalOrder_annihilate_iff_false, ite_eq_right_iff, and_imp,
|
||||
IsEmpty.forall_iff]
|
||||
| φ'' :: φs => by
|
||||
dsimp only [List.cons_append, Wick.koszulSignInsert]
|
||||
rw [koszulSignInsert_append_annihilate φ' φ hφ φs]
|
||||
|
||||
lemma normalOrderSign_append_annihlate (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) :
|
||||
(φs : List 𝓕.CrAnStates) →
|
||||
normalOrderSign (φs ++ [φ]) = normalOrderSign φs
|
||||
| [] => by simp
|
||||
| φ' :: φs => by
|
||||
dsimp only [List.cons_append, normalOrderSign, Wick.koszulSign]
|
||||
have hi := normalOrderSign_append_annihlate φ hφ φs
|
||||
dsimp only [normalOrderSign] at hi
|
||||
rw [hi, koszulSignInsert_append_annihilate φ' φ hφ φs]
|
||||
|
||||
lemma koszulSignInsert_annihilate_cons_create (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(φs : List 𝓕.CrAnStates) :
|
||||
Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φa (φc :: φs)
|
||||
= FieldStatistic.exchangeSign (𝓕.crAnStatistics φc) (𝓕.crAnStatistics φa) *
|
||||
Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φa φs := by
|
||||
rw [Wick.koszulSignInsert_cons]
|
||||
simp only [FieldStatistic.instCommGroup.eq_1, mul_eq_mul_right_iff]
|
||||
apply Or.inl
|
||||
rw [Wick.koszulSignCons, if_neg, FieldStatistic.exchangeSign_symm,
|
||||
FieldStatistic.exchangeSign_eq_if]
|
||||
rw [normalOrderRel, hφa, hφc]
|
||||
simp [CreateAnnihilate.normalOrder]
|
||||
|
||||
lemma normalOrderSign_swap_create_annihlate_fst (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(φs : List 𝓕.CrAnStates) :
|
||||
normalOrderSign (φc :: φa :: φs) =
|
||||
FieldStatistic.exchangeSign (𝓕.crAnStatistics φc) (𝓕.crAnStatistics φa) *
|
||||
normalOrderSign (φa :: φc :: φs) := by
|
||||
rw [normalOrderSign_cons_create φc hφc (φa :: φs)]
|
||||
conv_rhs =>
|
||||
rw [normalOrderSign, Wick.koszulSign, ← normalOrderSign]
|
||||
rw [koszulSignInsert_annihilate_cons_create φc φa hφc hφa φs]
|
||||
rw [← mul_assoc, ← mul_assoc, FieldStatistic.exchangeSign_mul_self]
|
||||
rw [one_mul, normalOrderSign_cons_create φc hφc φs]
|
||||
rfl
|
||||
|
||||
lemma koszulSignInsert_swap (φ φc φa : 𝓕.CrAnStates) (φs φs' : List 𝓕.CrAnStates) :
|
||||
Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φ (φs ++ φa :: φc :: φs')
|
||||
= Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φ (φs ++ φc :: φa :: φs') := by
|
||||
apply Wick.koszulSignInsert_eq_perm
|
||||
refine List.Perm.append_left φs ?h.a
|
||||
exact List.Perm.swap φc φa φs'
|
||||
|
||||
lemma normalOrderSign_swap_create_annihlate (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate) :
|
||||
(φs φs' : List 𝓕.CrAnStates) → normalOrderSign (φs ++ φc :: φa :: φs') =
|
||||
FieldStatistic.exchangeSign (𝓕.crAnStatistics φc) (𝓕.crAnStatistics φa) *
|
||||
normalOrderSign (φs ++ φa :: φc :: φs')
|
||||
| [], φs' => normalOrderSign_swap_create_annihlate_fst φc φa hφc hφa φs'
|
||||
| φ :: φs, φs' => by
|
||||
rw [normalOrderSign]
|
||||
dsimp only [List.cons_append, Wick.koszulSign, FieldStatistic.instCommGroup.eq_1]
|
||||
rw [← normalOrderSign, normalOrderSign_swap_create_annihlate φc φa hφc hφa φs φs']
|
||||
rw [← mul_assoc, mul_comm _ (FieldStatistic.exchangeSign _ _), mul_assoc]
|
||||
simp only [FieldStatistic.instCommGroup.eq_1, mul_eq_mul_left_iff]
|
||||
apply Or.inl
|
||||
conv_rhs =>
|
||||
rw [normalOrderSign]
|
||||
dsimp [Wick.koszulSign]
|
||||
rw [← normalOrderSign]
|
||||
simp only [mul_eq_mul_right_iff]
|
||||
apply Or.inl
|
||||
rw [koszulSignInsert_swap]
|
||||
|
||||
lemma normalOrderSign_swap_create_create_fst (φc φc' : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφc' : 𝓕 |>ᶜ φc' = CreateAnnihilate.create)
|
||||
(φs : List 𝓕.CrAnStates) :
|
||||
normalOrderSign (φc :: φc' :: φs) = normalOrderSign (φc' :: φc :: φs) := by
|
||||
rw [normalOrderSign_cons_create φc hφc, normalOrderSign_cons_create φc' hφc']
|
||||
rw [normalOrderSign_cons_create φc' hφc', normalOrderSign_cons_create φc hφc]
|
||||
|
||||
lemma normalOrderSign_swap_create_create (φc φc' : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφc' : 𝓕 |>ᶜ φc' = CreateAnnihilate.create) :
|
||||
(φs φs' : List 𝓕.CrAnStates) →
|
||||
normalOrderSign (φs ++ φc :: φc' :: φs') = normalOrderSign (φs ++ φc' :: φc :: φs')
|
||||
| [], φs' => by
|
||||
exact normalOrderSign_swap_create_create_fst φc φc' hφc hφc' φs'
|
||||
| φ :: φs, φs' => by
|
||||
rw [normalOrderSign]
|
||||
dsimp only [List.cons_append, Wick.koszulSign]
|
||||
rw [← normalOrderSign, normalOrderSign_swap_create_create φc φc' hφc hφc']
|
||||
dsimp only [normalOrderSign, Wick.koszulSign]
|
||||
rw [← normalOrderSign]
|
||||
simp only [mul_eq_mul_right_iff]
|
||||
apply Or.inl (Wick.koszulSignInsert_eq_perm _ _ _ _ _ _)
|
||||
exact List.Perm.append_left φs (List.Perm.swap φc' φc φs')
|
||||
|
||||
lemma normalOrderSign_swap_annihilate_annihilate_fst (φa φa' : 𝓕.CrAnStates)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(hφa' : 𝓕 |>ᶜ φa' = CreateAnnihilate.annihilate)
|
||||
(φs : List 𝓕.CrAnStates) :
|
||||
normalOrderSign (φa :: φa' :: φs) =
|
||||
normalOrderSign (φa' :: φa :: φs) := by
|
||||
rw [normalOrderSign, normalOrderSign]
|
||||
dsimp only [Wick.koszulSign]
|
||||
rw [← mul_assoc, ← mul_assoc]
|
||||
congr 1
|
||||
rw [Wick.koszulSignInsert_cons, Wick.koszulSignInsert_cons, mul_assoc, mul_assoc]
|
||||
congr 1
|
||||
· dsimp only [Wick.koszulSignCons]
|
||||
rw [if_pos, if_pos]
|
||||
· simp [normalOrderRel, hφa, hφa', CreateAnnihilate.normalOrder]
|
||||
· simp [normalOrderRel, hφa, hφa', CreateAnnihilate.normalOrder]
|
||||
· rw [NonUnitalNormedCommRing.mul_comm]
|
||||
|
||||
lemma normalOrderSign_swap_annihilate_annihilate (φa φa' : 𝓕.CrAnStates)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(hφa' : 𝓕 |>ᶜ φa' = CreateAnnihilate.annihilate) : (φs φs' : List 𝓕.CrAnStates) →
|
||||
normalOrderSign (φs ++ φa :: φa' :: φs') = normalOrderSign (φs ++ φa' :: φa :: φs')
|
||||
| [], φs' => by
|
||||
exact normalOrderSign_swap_annihilate_annihilate_fst φa φa' hφa hφa' φs'
|
||||
| φ :: φs, φs' => by
|
||||
rw [normalOrderSign]
|
||||
dsimp only [List.cons_append, Wick.koszulSign]
|
||||
rw [← normalOrderSign]
|
||||
rw [normalOrderSign_swap_annihilate_annihilate φa φa' hφa hφa']
|
||||
dsimp only [normalOrderSign, Wick.koszulSign]
|
||||
rw [← normalOrderSign]
|
||||
simp only [mul_eq_mul_right_iff]
|
||||
apply Or.inl
|
||||
apply Wick.koszulSignInsert_eq_perm
|
||||
refine List.Perm.append_left φs ?h.h.a
|
||||
exact List.Perm.swap φa' φa φs'
|
||||
open FieldStatistic
|
||||
|
||||
/-!
|
||||
|
||||
## Normal order of lists
|
||||
|
||||
-/
|
||||
|
||||
/-- The normal ordering of a list of creation and annihilation states.
|
||||
To give some schematic. For example:
|
||||
- `normalOrderList [φ1c, φ1a, φ2c, φ2a] = [φ1c, φ2c, φ1a, φ2a]`
|
||||
-/
|
||||
def normalOrderList (φs : List 𝓕.CrAnStates) : List 𝓕.CrAnStates :=
|
||||
List.insertionSort 𝓕.normalOrderRel φs
|
||||
|
||||
@[simp]
|
||||
lemma normalOrderList_nil : normalOrderList (𝓕 := 𝓕) [] = [] := by
|
||||
simp [normalOrderList]
|
||||
|
||||
@[simp]
|
||||
lemma normalOrderList_statistics (φs : List 𝓕.CrAnStates) :
|
||||
(𝓕 |>ₛ (normalOrderList φs)) = 𝓕 |>ₛ φs := by
|
||||
simp [normalOrderList, List.insertionSort]
|
||||
|
||||
lemma orderedInsert_create (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create) :
|
||||
(φs : List 𝓕.CrAnStates) → List.orderedInsert normalOrderRel φ φs = φ :: φs
|
||||
| [] => rfl
|
||||
| φ' :: φs => by
|
||||
dsimp only [List.orderedInsert.eq_2]
|
||||
rw [if_pos]
|
||||
dsimp only [normalOrderRel]
|
||||
rw [hφ]
|
||||
dsimp [CreateAnnihilate.normalOrder]
|
||||
|
||||
lemma normalOrderList_cons_create (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create) (φs : List 𝓕.CrAnStates) :
|
||||
normalOrderList (φ :: φs) = φ :: normalOrderList φs := by
|
||||
simp only [normalOrderList, List.insertionSort]
|
||||
rw [orderedInsert_create φ hφ]
|
||||
|
||||
lemma orderedInsert_append_annihilate (φ' φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) :
|
||||
(φs : List 𝓕.CrAnStates) → List.orderedInsert normalOrderRel φ' (φs ++ [φ]) =
|
||||
List.orderedInsert normalOrderRel φ' φs ++ [φ]
|
||||
| [] => by
|
||||
simp [Wick.koszulSignInsert, normalOrderRel, hφ]
|
||||
| φ'' :: φs => by
|
||||
dsimp only [List.cons_append, List.orderedInsert.eq_2]
|
||||
have hi := orderedInsert_append_annihilate φ' φ hφ φs
|
||||
rw [hi]
|
||||
split
|
||||
next h => simp_all only [List.cons_append]
|
||||
next h => simp_all only [List.cons_append]
|
||||
|
||||
lemma normalOrderList_append_annihilate (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) :
|
||||
(φs : List 𝓕.CrAnStates) →
|
||||
normalOrderList (φs ++ [φ]) = normalOrderList φs ++ [φ]
|
||||
| [] => by simp [normalOrderList]
|
||||
| φ' :: φs => by
|
||||
simp only [normalOrderList, List.insertionSort, List.append_eq]
|
||||
have hi := normalOrderList_append_annihilate φ hφ φs
|
||||
dsimp only [normalOrderList] at hi
|
||||
rw [hi, orderedInsert_append_annihilate φ' φ hφ]
|
||||
|
||||
lemma normalOrder_swap_create_annihlate_fst (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
|
||||
(φs : List 𝓕.CrAnStates) :
|
||||
normalOrderList (φc :: φa :: φs) = normalOrderList (φa :: φc :: φs) := by
|
||||
rw [normalOrderList_cons_create φc hφc (φa :: φs)]
|
||||
conv_rhs =>
|
||||
rw [normalOrderList, List.insertionSort]
|
||||
have hi := normalOrderList_cons_create φc hφc φs
|
||||
rw [normalOrderList] at hi
|
||||
rw [hi]
|
||||
dsimp only [List.orderedInsert.eq_2]
|
||||
split
|
||||
· rename_i h
|
||||
rw [normalOrderRel, hφa, hφc] at h
|
||||
dsimp [CreateAnnihilate.normalOrder] at h
|
||||
· rfl
|
||||
|
||||
lemma normalOrderList_swap_create_annihlate (φc φa : 𝓕.CrAnStates)
|
||||
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
|
||||
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate) :
|
||||
(φs φs' : List 𝓕.CrAnStates) →
|
||||
normalOrderList (φs ++ φc :: φa :: φs') = normalOrderList (φs ++ φa :: φc :: φs')
|
||||
| [], φs' => normalOrder_swap_create_annihlate_fst φc φa hφc hφa φs'
|
||||
| φ :: φs, φs' => by
|
||||
dsimp only [List.cons_append, normalOrderList, List.insertionSort]
|
||||
have hi := normalOrderList_swap_create_annihlate φc φa hφc hφa φs φs'
|
||||
dsimp only [normalOrderList] at hi
|
||||
rw [hi]
|
||||
|
||||
/-- For a list of creation and annihlation states, the equivalence between
|
||||
`Fin φs.length` and `Fin (normalOrderList φs).length` taking each position in `φs`
|
||||
to it's corresponding position in the normal ordered list. This assumes that
|
||||
we are using the insertion sort method.
|
||||
For example:
|
||||
- For `[φ1c, φ1a, φ2c, φ2a]` this equivalence sends `0 ↦ 0`, `1 ↦ 2`, `2 ↦ 1`, `3 ↦ 3`.
|
||||
-/
|
||||
def normalOrderEquiv {φs : List 𝓕.CrAnStates} : Fin φs.length ≃ Fin (normalOrderList φs).length :=
|
||||
HepLean.List.insertionSortEquiv 𝓕.normalOrderRel φs
|
||||
|
||||
lemma sum_normalOrderList_length {M : Type} [AddCommMonoid M]
|
||||
(φs : List 𝓕.CrAnStates) (f : Fin (normalOrderList φs).length → M) :
|
||||
∑ (n : Fin (normalOrderList φs).length), f n = ∑ (n : Fin φs.length), f (normalOrderEquiv n) :=
|
||||
Eq.symm (Equiv.sum_comp normalOrderEquiv f)
|
||||
|
||||
@[simp]
|
||||
lemma normalOrderList_get_normalOrderEquiv {φs : List 𝓕.CrAnStates} (n : Fin φs.length) :
|
||||
(normalOrderList φs)[(normalOrderEquiv n).val] = φs[n.val] := by
|
||||
change (normalOrderList φs).get (normalOrderEquiv n) = _
|
||||
simp only [normalOrderList, normalOrderEquiv]
|
||||
erw [← HepLean.List.insertionSortEquiv_get]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma normalOrderList_eraseIdx_normalOrderEquiv {φs : List 𝓕.CrAnStates} (n : Fin φs.length) :
|
||||
(normalOrderList φs).eraseIdx (normalOrderEquiv n).val =
|
||||
normalOrderList (φs.eraseIdx n.val) := by
|
||||
simp only [normalOrderList, normalOrderEquiv]
|
||||
rw [HepLean.List.eraseIdx_insertionSort_fin]
|
||||
|
||||
lemma normalOrderSign_eraseIdx (φs : List 𝓕.CrAnStates) (n : Fin φs.length) :
|
||||
normalOrderSign (φs.eraseIdx n) = normalOrderSign φs *
|
||||
𝓢(𝓕 |>ₛ (φs.get n), 𝓕 |>ₛ (φs.take n)) *
|
||||
𝓢(𝓕 |>ₛ (φs.get n), 𝓕 |>ₛ ((normalOrderList φs).take (normalOrderEquiv n))) := by
|
||||
rw [normalOrderSign, Wick.koszulSign_eraseIdx, ← normalOrderSign]
|
||||
rfl
|
||||
|
||||
lemma orderedInsert_createFilter_append_annihilate (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) : (φs φs' : List 𝓕.CrAnStates) →
|
||||
List.orderedInsert normalOrderRel φ (createFilter φs ++ φs') =
|
||||
createFilter φs ++ List.orderedInsert normalOrderRel φ φs'
|
||||
| [], φs' => by simp [createFilter]
|
||||
| φ' :: φs, φs' => by
|
||||
rcases CreateAnnihilate.eq_create_or_annihilate (𝓕 |>ᶜ φ') with hφ' | hφ'
|
||||
· rw [createFilter_cons_create hφ']
|
||||
dsimp only [List.cons_append, List.orderedInsert.eq_2]
|
||||
rw [if_neg, orderedInsert_createFilter_append_annihilate φ hφ φs φs']
|
||||
simp [normalOrderRel, hφ, hφ', CreateAnnihilate.normalOrder]
|
||||
· rw [createFilter_cons_annihilate hφ', orderedInsert_createFilter_append_annihilate φ hφ φs]
|
||||
|
||||
lemma orderedInsert_annihilateFilter (φ : 𝓕.CrAnStates) : (φs : List 𝓕.CrAnStates) →
|
||||
List.orderedInsert normalOrderRel φ (annihilateFilter φs) =
|
||||
φ :: annihilateFilter φs
|
||||
| [] => by simp [annihilateFilter]
|
||||
| φ' :: φs => by
|
||||
rcases CreateAnnihilate.eq_create_or_annihilate (𝓕 |>ᶜ φ') with hφ' | hφ'
|
||||
· rw [annihilateFilter_cons_create hφ', orderedInsert_annihilateFilter φ φs]
|
||||
· rw [annihilateFilter_cons_annihilate hφ']
|
||||
dsimp only [List.orderedInsert.eq_2]
|
||||
rw [if_pos]
|
||||
dsimp only [normalOrderRel]
|
||||
rw [hφ']
|
||||
rcases CreateAnnihilate.eq_create_or_annihilate (𝓕 |>ᶜ φ) with hφ | hφ
|
||||
· rw [hφ]
|
||||
simp only [CreateAnnihilate.normalOrder]
|
||||
· rw [hφ]
|
||||
simp [CreateAnnihilate.normalOrder]
|
||||
|
||||
lemma orderedInsert_createFilter_append_annihilateFilter_annihilate (φ : 𝓕.CrAnStates)
|
||||
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) (φs : List 𝓕.CrAnStates) :
|
||||
List.orderedInsert normalOrderRel φ (createFilter φs ++ annihilateFilter φs) =
|
||||
createFilter φs ++ φ :: annihilateFilter φs := by
|
||||
rw [orderedInsert_createFilter_append_annihilate φ hφ, orderedInsert_annihilateFilter]
|
||||
|
||||
lemma normalOrderList_eq_createFilter_append_annihilateFilter : (φs : List 𝓕.CrAnStates) →
|
||||
normalOrderList φs = createFilter φs ++ annihilateFilter φs
|
||||
| [] => by simp [normalOrderList, createFilter, annihilateFilter]
|
||||
| φ :: φs => by
|
||||
by_cases hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.create
|
||||
· rw [normalOrderList_cons_create φ hφ φs]
|
||||
dsimp only [createFilter]
|
||||
rw [List.filter_cons_of_pos]
|
||||
swap
|
||||
simp only [hφ, decide_true]
|
||||
dsimp only [annihilateFilter, List.cons_append]
|
||||
rw [List.filter_cons_of_neg]
|
||||
swap
|
||||
simp only [hφ, reduceCtorEq, decide_false, Bool.false_eq_true, not_false_eq_true]
|
||||
rw [normalOrderList_eq_createFilter_append_annihilateFilter φs]
|
||||
rfl
|
||||
· dsimp only [normalOrderList, List.insertionSort]
|
||||
rw [← normalOrderList]
|
||||
have hφ' : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate := by
|
||||
have hx := CreateAnnihilate.eq_create_or_annihilate (𝓕 |>ᶜ φ)
|
||||
simp_all
|
||||
rw [normalOrderList_eq_createFilter_append_annihilateFilter φs]
|
||||
rw [orderedInsert_createFilter_append_annihilateFilter_annihilate φ hφ']
|
||||
rw [createFilter_cons_annihilate hφ', annihilateFilter_cons_annihilate hφ']
|
||||
|
||||
end FieldStruct
|
176
HepLean/PerturbationTheory/FieldStruct/TimeOrder.lean
Normal file
176
HepLean/PerturbationTheory/FieldStruct/TimeOrder.lean
Normal file
|
@ -0,0 +1,176 @@
|
|||
/-
|
||||
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.Mathematics.List.InsertionSort
|
||||
import HepLean.PerturbationTheory.Algebras.StateAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
/-!
|
||||
|
||||
# State algebra
|
||||
|
||||
We define the state algebra of a field structure to be the free algebra
|
||||
generated by the states.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
/-- The time ordering relation on states. We have that `timeOrderRel φ0 φ1` is true
|
||||
if and only if `φ1` has a time less-then or equal to `φ0`, or `φ1` is a negative
|
||||
asymptotic state, or `φ0` is a positive asymptotic state. -/
|
||||
def timeOrderRel : 𝓕.States → 𝓕.States → Prop
|
||||
| States.posAsymp _, _ => True
|
||||
| States.position φ0, States.position φ1 => φ1.2 0 ≤ φ0.2 0
|
||||
| States.position _, States.negAsymp _ => True
|
||||
| States.position _, States.posAsymp _ => False
|
||||
| States.negAsymp _, States.posAsymp _ => False
|
||||
| States.negAsymp _, States.position _ => False
|
||||
| States.negAsymp _, States.negAsymp _ => True
|
||||
|
||||
/-- The relation `timeOrderRel` is decidable, but not computablly so due to
|
||||
`Real.decidableLE`. -/
|
||||
noncomputable instance : (φ φ' : 𝓕.States) → Decidable (timeOrderRel φ φ')
|
||||
| States.posAsymp _, _ => isTrue True.intro
|
||||
| States.position φ0, States.position φ1 => inferInstanceAs (Decidable (φ1.2 0 ≤ φ0.2 0))
|
||||
| States.position _, States.negAsymp _ => isTrue True.intro
|
||||
| States.position _, States.posAsymp _ => isFalse (fun a => a)
|
||||
| States.negAsymp _, States.posAsymp _ => isFalse (fun a => a)
|
||||
| States.negAsymp _, States.position _ => isFalse (fun a => a)
|
||||
| States.negAsymp _, States.negAsymp _ => isTrue True.intro
|
||||
|
||||
/-- Time ordering is total. -/
|
||||
instance : IsTotal 𝓕.States 𝓕.timeOrderRel where
|
||||
total a b := by
|
||||
cases a <;> cases b <;>
|
||||
simp only [or_self, or_false, or_true, timeOrderRel, Fin.isValue, implies_true, imp_self,
|
||||
IsEmpty.forall_iff]
|
||||
exact LinearOrder.le_total _ _
|
||||
|
||||
/-- Time ordering is transitive. -/
|
||||
instance : IsTrans 𝓕.States 𝓕.timeOrderRel where
|
||||
trans a b c := by
|
||||
cases a <;> cases b <;> cases c <;>
|
||||
simp only [timeOrderRel, Fin.isValue, implies_true, imp_self, IsEmpty.forall_iff]
|
||||
exact fun h1 h2 => Preorder.le_trans _ _ _ h2 h1
|
||||
|
||||
noncomputable section
|
||||
|
||||
open FieldStatistic
|
||||
open HepLean.List
|
||||
|
||||
/-- Given a list `φ :: φs` of states, the (zero-based) position of the state which is
|
||||
of maximum time. For example
|
||||
- for the list `[φ1(t = 4), φ2(t = 5), φ3(t = 3), φ4(t = 5)]` this would return `1`.
|
||||
This is defined for a list `φ :: φs` instead of `φs` to ensure that such a position exists.
|
||||
-/
|
||||
def maxTimeFieldPos (φ : 𝓕.States) (φs : List 𝓕.States) : ℕ :=
|
||||
insertionSortMinPos timeOrderRel φ φs
|
||||
|
||||
lemma maxTimeFieldPos_lt_length (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
maxTimeFieldPos φ φs < (φ :: φs).length := by
|
||||
simp [maxTimeFieldPos]
|
||||
|
||||
/-- Given a list `φ :: φs` of states, the left-most state of maximum time, if there are more.
|
||||
As an example:
|
||||
- for the list `[φ1(t = 4), φ2(t = 5), φ3(t = 3), φ4(t = 5)]` this would return `φ2(t = 5)`.
|
||||
It is the state at the position `maxTimeFieldPos φ φs`.
|
||||
-/
|
||||
def maxTimeField (φ : 𝓕.States) (φs : List 𝓕.States) : 𝓕.States :=
|
||||
insertionSortMin timeOrderRel φ φs
|
||||
|
||||
/-- Given a list `φ :: φs` of states, the list with the left-most state of maximum
|
||||
time removed.
|
||||
As an example:
|
||||
- for the list `[φ1(t = 4), φ2(t = 5), φ3(t = 3), φ4(t = 5)]` this would return
|
||||
`[φ1(t = 4), φ3(t = 3), φ4(t = 5)]`.
|
||||
-/
|
||||
def eraseMaxTimeField (φ : 𝓕.States) (φs : List 𝓕.States) : List 𝓕.States :=
|
||||
insertionSortDropMinPos timeOrderRel φ φs
|
||||
|
||||
@[simp]
|
||||
lemma eraseMaxTimeField_length (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
(eraseMaxTimeField φ φs).length = φs.length := by
|
||||
simp [eraseMaxTimeField, insertionSortDropMinPos, eraseIdx_length']
|
||||
|
||||
lemma maxTimeFieldPos_lt_eraseMaxTimeField_length_succ (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
maxTimeFieldPos φ φs < (eraseMaxTimeField φ φs).length.succ := by
|
||||
simp only [eraseMaxTimeField_length, Nat.succ_eq_add_one]
|
||||
exact maxTimeFieldPos_lt_length φ φs
|
||||
|
||||
/-- Given a list `φ :: φs` of states, the position of the left-most state of maximum
|
||||
time as an eement of `Fin (eraseMaxTimeField φ φs).length.succ`.
|
||||
As an example:
|
||||
- for the list `[φ1(t = 4), φ2(t = 5), φ3(t = 3), φ4(t = 5)]` this would return `⟨1,...⟩`.
|
||||
-/
|
||||
def maxTimeFieldPosFin (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
Fin (eraseMaxTimeField φ φs).length.succ :=
|
||||
insertionSortMinPosFin timeOrderRel φ φs
|
||||
|
||||
lemma lt_maxTimeFieldPosFin_not_timeOrder (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(i : Fin (eraseMaxTimeField φ φs).length)
|
||||
(hi : (maxTimeFieldPosFin φ φs).succAbove i < maxTimeFieldPosFin φ φs) :
|
||||
¬ timeOrderRel ((eraseMaxTimeField φ φs)[i.val]) (maxTimeField φ φs) := by
|
||||
exact insertionSortMin_lt_mem_insertionSortDropMinPos_of_lt timeOrderRel φ φs i hi
|
||||
|
||||
lemma timeOrder_maxTimeField (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(i : Fin (eraseMaxTimeField φ φs).length) :
|
||||
timeOrderRel (maxTimeField φ φs) ((eraseMaxTimeField φ φs)[i.val]) := by
|
||||
exact insertionSortMin_lt_mem_insertionSortDropMinPos timeOrderRel φ φs _
|
||||
|
||||
/-- The sign associated with putting a list of states into time order (with
|
||||
the state of greatest time to the left).
|
||||
We pick up a minus sign for every fermion paired crossed. -/
|
||||
def timeOrderSign (φs : List 𝓕.States) : ℂ :=
|
||||
Wick.koszulSign 𝓕.statesStatistic 𝓕.timeOrderRel φs
|
||||
|
||||
lemma timeOrderSign_pair_ordered {φ ψ : 𝓕.States} (h : timeOrderRel φ ψ) :
|
||||
timeOrderSign [φ, ψ] = 1 := by
|
||||
simp only [timeOrderSign, Wick.koszulSign, Wick.koszulSignInsert, mul_one, ite_eq_left_iff,
|
||||
ite_eq_right_iff, and_imp]
|
||||
exact fun h' => False.elim (h' h)
|
||||
|
||||
lemma timeOrderSign_pair_not_ordered {φ ψ : 𝓕.States} (h : ¬ timeOrderRel φ ψ) :
|
||||
timeOrderSign [φ, ψ] = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ψ) := by
|
||||
simp only [timeOrderSign, Wick.koszulSign, Wick.koszulSignInsert, mul_one, instCommGroup.eq_1]
|
||||
rw [if_neg h]
|
||||
simp [FieldStatistic.exchangeSign_eq_if]
|
||||
|
||||
lemma timerOrderSign_of_eraseMaxTimeField (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
timeOrderSign (eraseMaxTimeField φ φs) = timeOrderSign (φ :: φs) *
|
||||
𝓢(𝓕 |>ₛ maxTimeField φ φs, 𝓕 |>ₛ (φ :: φs).take (maxTimeFieldPos φ φs)) := by
|
||||
rw [eraseMaxTimeField, insertionSortDropMinPos, timeOrderSign,
|
||||
Wick.koszulSign_eraseIdx_insertionSortMinPos]
|
||||
rw [← timeOrderSign, ← maxTimeField]
|
||||
rfl
|
||||
|
||||
/-- The time ordering of a list of states. A schematic example is:
|
||||
- `normalOrderList [φ1(t = 4), φ2(t = 5), φ3(t = 3), φ4(t = 5)]` is equal to
|
||||
`[φ2(t = 5), φ4(t = 5), φ1(t = 4), φ3(t = 3)]` -/
|
||||
def timeOrderList (φs : List 𝓕.States) : List 𝓕.States :=
|
||||
List.insertionSort 𝓕.timeOrderRel φs
|
||||
|
||||
lemma timeOrderList_pair_ordered {φ ψ : 𝓕.States} (h : timeOrderRel φ ψ) :
|
||||
timeOrderList [φ, ψ] = [φ, ψ] := by
|
||||
simp only [timeOrderList, List.insertionSort, List.orderedInsert, ite_eq_left_iff,
|
||||
List.cons.injEq, and_true]
|
||||
exact fun h' => False.elim (h' h)
|
||||
|
||||
lemma timeOrderList_pair_not_ordered {φ ψ : 𝓕.States} (h : ¬ timeOrderRel φ ψ) :
|
||||
timeOrderList [φ, ψ] = [ψ, φ] := by
|
||||
simp only [timeOrderList, List.insertionSort, List.orderedInsert, ite_eq_right_iff,
|
||||
List.cons.injEq, and_true]
|
||||
exact fun h' => False.elim (h h')
|
||||
|
||||
@[simp]
|
||||
lemma timeOrderList_nil : timeOrderList (𝓕 := 𝓕) [] = [] := by
|
||||
simp [timeOrderList]
|
||||
|
||||
lemma timeOrderList_eq_maxTimeField_timeOrderList (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
timeOrderList (φ :: φs) = maxTimeField φ φs :: timeOrderList (eraseMaxTimeField φ φs) := by
|
||||
exact insertionSort_eq_insertionSortMin_cons timeOrderRel φ φs
|
||||
|
||||
end
|
||||
end FieldStruct
|
|
@ -3,7 +3,7 @@ Copyright (c) 2024 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.Wick.Signs.KoszulSignInsert
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSignInsert
|
||||
/-!
|
||||
|
||||
# Koszul sign
|
||||
|
@ -24,6 +24,11 @@ def koszulSign (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Prop) [Deci
|
|||
| [] => 1
|
||||
| a :: l => koszulSignInsert q le a l * koszulSign q le l
|
||||
|
||||
@[simp]
|
||||
lemma koszulSign_singleton (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Prop) [DecidableRel le] (φ : 𝓕) :
|
||||
koszulSign q le [φ] = 1 := by
|
||||
simp [koszulSign, koszulSignInsert]
|
||||
|
||||
lemma koszulSign_mul_self (l : List 𝓕) : koszulSign q le l * koszulSign q le l = 1 := by
|
||||
induction l with
|
||||
| nil => simp [koszulSign]
|
||||
|
@ -80,12 +85,14 @@ lemma koszulSign_erase_boson {𝓕 : Type} (q : 𝓕 → FieldStatistic) (le :
|
|||
|
||||
lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) :
|
||||
(φs : List 𝓕) → (n : ℕ) → (hn : n ≤ φs.length) →
|
||||
koszulSign q le (List.insertIdx n φ φs) = insertSign q n φ φs * koszulSign q le φs *
|
||||
insertSign q (insertionSortEquiv le (List.insertIdx n φ φs) ⟨n, by
|
||||
rw [List.length_insertIdx, if_pos hn]
|
||||
exact Nat.succ_le_succ hn⟩) φ (List.insertionSort le (List.insertIdx n φ φs))
|
||||
koszulSign q le (List.insertIdx n φ φs) = 𝓢(q φ, ofList q (φs.take n)) * koszulSign q le φs *
|
||||
𝓢(q φ, ofList q ((List.insertionSort le (List.insertIdx n φ φs)).take
|
||||
(insertionSortEquiv le (List.insertIdx n φ φs) ⟨n, by
|
||||
rw [List.length_insertIdx _ _]
|
||||
simp only [hn, ↓reduceIte]
|
||||
omega⟩)))
|
||||
| [], 0, h => by
|
||||
simp [koszulSign, insertSign, superCommuteCoef, koszulSignInsert]
|
||||
simp [koszulSign, koszulSignInsert]
|
||||
| [], n + 1, h => by
|
||||
simp at h
|
||||
| φ1 :: φs, 0, h => by
|
||||
|
@ -99,15 +106,15 @@ lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) :
|
|||
Fin.isValue, HepLean.Fin.finExtractOne_symm_inl_apply, RelIso.coe_fn_toEquiv,
|
||||
Fin.castOrderIso_apply, Fin.cast_mk, Fin.eta]
|
||||
conv_rhs =>
|
||||
enter [2, 4]
|
||||
enter [2,2, 2, 2]
|
||||
rw [orderedInsert_eq_insertIdx_orderedInsertPos]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [← insertSign_insert]
|
||||
change insertSign q (↑(orderedInsertPos le ((List.insertionSort le (φ1 :: φs))) φ)) φ
|
||||
(List.insertionSort le (φ1 :: φs))
|
||||
rw [← koszulSignInsert_eq_insertSign q le]
|
||||
rw [insertSign_zero]
|
||||
rw [← ofList_take_insert]
|
||||
change 𝓢(q φ, ofList q ((List.insertionSort le (φ1 :: φs)).take
|
||||
(↑(orderedInsertPos le ((List.insertionSort le (φ1 :: φs))) φ))))
|
||||
rw [← koszulSignInsert_eq_exchangeSign_take q le]
|
||||
rw [ofList_take_zero]
|
||||
simp
|
||||
| φ1 :: φs, n + 1, h => by
|
||||
conv_lhs =>
|
||||
|
@ -122,10 +129,10 @@ lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) :
|
|||
erw [orderedInsertEquiv_fin_succ]
|
||||
simp only [Fin.eta, Fin.coe_cast]
|
||||
rhs
|
||||
rw [orderedInsert_eq_insertIdx_orderedInsertPos]
|
||||
simp [orderedInsert_eq_insertIdx_orderedInsertPos]
|
||||
conv_rhs =>
|
||||
lhs
|
||||
rw [insertSign_succ_cons, koszulSign]
|
||||
rw [ofList_take_succ_cons, map_mul, koszulSign]
|
||||
ring_nf
|
||||
conv_lhs =>
|
||||
lhs
|
||||
|
@ -138,18 +145,19 @@ lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) :
|
|||
have hnsL : n < (List.insertIdx n φ φs).length := by
|
||||
rw [List.length_insertIdx _ _]
|
||||
simp only [List.length_cons, add_le_add_iff_right] at h
|
||||
rw [if_pos h]
|
||||
exact Nat.succ_le_succ h
|
||||
simp only [h, ↓reduceIte]
|
||||
omega
|
||||
let ni : Fin rs.length := (insertionSortEquiv le (List.insertIdx n φ φs))
|
||||
⟨n, hnsL⟩
|
||||
let nro : Fin (rs.length + 1) :=
|
||||
⟨↑(orderedInsertPos le rs φ1), orderedInsertPos_lt_length le rs φ1⟩
|
||||
rw [koszulSignInsert_insertIdx, koszulSignInsert_cons]
|
||||
trans koszulSignInsert q le φ1 φs * (koszulSignCons q le φ1 φ *insertSign q ni φ rs)
|
||||
trans koszulSignInsert q le φ1 φs * (koszulSignCons q le φ1 φ *
|
||||
𝓢(q φ, ofList q (rs.take ni)))
|
||||
· simp only [rs, ni]
|
||||
ring
|
||||
trans koszulSignInsert q le φ1 φs * (superCommuteCoef q [φ] [φ1] *
|
||||
insertSign q (nro.succAbove ni) φ (List.insertIdx nro φ1 rs))
|
||||
trans koszulSignInsert q le φ1 φs * (𝓢(q φ, q φ1) *
|
||||
𝓢(q φ, ofList q ((List.insertIdx nro φ1 rs).take (nro.succAbove ni))))
|
||||
swap
|
||||
· simp only [rs, nro, ni]
|
||||
ring
|
||||
|
@ -169,21 +177,86 @@ lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) :
|
|||
exact List.sorted_insertionSort le (List.insertIdx n φ φs)
|
||||
by_cases hn : ni.castSucc < nro
|
||||
· simp only [hn, ↓reduceIte, Fin.coe_castSucc]
|
||||
rw [insertSign_insert_gt]
|
||||
rw [ofList_take_insertIdx_gt]
|
||||
swap
|
||||
· exact hn
|
||||
congr 1
|
||||
rw [koszulSignCons_eq_superComuteCoef]
|
||||
rw [koszulSignCons_eq_exchangeSign]
|
||||
simp only [hc1 hn, ↓reduceIte]
|
||||
rw [superCommuteCoef_comm]
|
||||
rw [exchangeSign_symm]
|
||||
· simp only [hn, ↓reduceIte, Fin.val_succ]
|
||||
rw [insertSign_insert_lt, ← mul_assoc]
|
||||
rw [ofList_take_insertIdx_le, map_mul, ← mul_assoc]
|
||||
congr 1
|
||||
rw [superCommuteCoef_mul_self, koszulSignCons]
|
||||
rw [exchangeSign_mul_self, koszulSignCons]
|
||||
simp only [hc2 hn, ↓reduceIte]
|
||||
exact Nat.le_of_not_lt hn
|
||||
exact Nat.le_of_lt_succ (orderedInsertPos_lt_length le rs φ1)
|
||||
· exact Nat.le_of_lt_succ h
|
||||
· exact Nat.le_of_lt_succ h
|
||||
|
||||
lemma insertIdx_eraseIdx {I : Type} : (n : ℕ) → (r : List I) → (hn : n < r.length) →
|
||||
List.insertIdx n (r.get ⟨n, hn⟩) (r.eraseIdx n) = r
|
||||
| n, [], hn => by
|
||||
simp at hn
|
||||
| 0, r0 :: r, hn => by
|
||||
simp
|
||||
| n + 1, r0 :: r, hn => by
|
||||
simp only [List.length_cons, List.get_eq_getElem, List.getElem_cons_succ,
|
||||
List.eraseIdx_cons_succ, List.insertIdx_succ_cons, List.cons.injEq, true_and]
|
||||
exact insertIdx_eraseIdx n r _
|
||||
|
||||
lemma koszulSign_eraseIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φs : List 𝓕) (n : Fin φs.length) :
|
||||
koszulSign q le (φs.eraseIdx n) = koszulSign q le φs * 𝓢(q (φs.get n), ofList q (φs.take n)) *
|
||||
𝓢(q (φs.get n), ofList q (List.take (↑(insertionSortEquiv le φs n))
|
||||
(List.insertionSort le φs))) := by
|
||||
let φs' := φs.eraseIdx ↑n
|
||||
have hφs : List.insertIdx n (φs.get n) φs' = φs := by
|
||||
exact insertIdx_eraseIdx n.1 φs n.prop
|
||||
conv_rhs =>
|
||||
lhs
|
||||
lhs
|
||||
rw [← hφs]
|
||||
rw [koszulSign_insertIdx q le (φs.get n) ((φs.eraseIdx ↑n)) n (by
|
||||
rw [List.length_eraseIdx]
|
||||
simp only [Fin.is_lt, ↓reduceIte]
|
||||
omega)]
|
||||
rhs
|
||||
enter [2, 2, 2]
|
||||
rw [hφs]
|
||||
conv_rhs =>
|
||||
enter [1, 1, 2, 2, 2, 1, 1]
|
||||
rw [insertionSortEquiv_congr _ _ hφs]
|
||||
simp only [instCommGroup.eq_1, List.get_eq_getElem, Equiv.trans_apply, RelIso.coe_fn_toEquiv,
|
||||
Fin.castOrderIso_apply, Fin.cast_mk, Fin.eta, Fin.coe_cast]
|
||||
trans koszulSign q le (φs.eraseIdx ↑n) *
|
||||
(𝓢(q φs[↑n], ofList q ((φs.eraseIdx ↑n).take n)) * 𝓢(q φs[↑n], ofList q (List.take (↑n) φs))) *
|
||||
(𝓢(q φs[↑n], ofList q ((List.insertionSort le φs).take (↑((insertionSortEquiv le φs) n)))) *
|
||||
𝓢(q φs[↑n], ofList q (List.take (↑((insertionSortEquiv le φs) n)) (List.insertionSort le φs))))
|
||||
swap
|
||||
· simp only [Fin.getElem_fin]
|
||||
ring
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [exchangeSign_mul_self]
|
||||
simp only [instCommGroup.eq_1, Fin.getElem_fin, mul_one]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [ofList_take_eraseIdx, exchangeSign_mul_self]
|
||||
simp
|
||||
|
||||
lemma koszulSign_eraseIdx_insertionSortMinPos [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) (φs : List 𝓕) :
|
||||
koszulSign q le ((φ :: φs).eraseIdx (insertionSortMinPos le φ φs)) = koszulSign q le (φ :: φs)
|
||||
* 𝓢(q (insertionSortMin le φ φs), ofList q ((φ :: φs).take (insertionSortMinPos le φ φs))) := by
|
||||
rw [koszulSign_eraseIdx]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
lhs
|
||||
simp [insertionSortMinPos]
|
||||
erw [Equiv.apply_symm_apply]
|
||||
simp only [instCommGroup.eq_1, List.get_eq_getElem, List.length_cons, List.insertionSort,
|
||||
List.take_zero, ofList_empty, exchangeSign_bosonic, mul_one, mul_eq_mul_left_iff]
|
||||
apply Or.inl
|
||||
rfl
|
||||
|
||||
end Wick
|
|
@ -3,7 +3,7 @@ Copyright (c) 2024 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.Wick.Signs.InsertSign
|
||||
import HepLean.PerturbationTheory.FieldStatistics.ExchangeSign
|
||||
/-!
|
||||
|
||||
# Koszul sign insert
|
||||
|
@ -153,11 +153,19 @@ lemma koszulSignInsert_eq_sort (φs : List 𝓕) (φ : 𝓕) :
|
|||
apply koszulSignInsert_eq_perm
|
||||
exact List.Perm.symm (List.perm_insertionSort le φs)
|
||||
|
||||
lemma koszulSignInsert_eq_insertSign [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) (φs : List 𝓕) :
|
||||
koszulSignInsert q le φ φs = insertSign q (orderedInsertPos le (List.insertionSort le φs) φ)
|
||||
φ (List.insertionSort le φs) := by
|
||||
lemma koszulSignInsert_eq_exchangeSign_take [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) (φs : List 𝓕) :
|
||||
koszulSignInsert q le φ φs = 𝓢(q φ, ofList q
|
||||
((List.insertionSort le φs).take (orderedInsertPos le (List.insertionSort le φs) φ))) := by
|
||||
rw [koszulSignInsert_eq_cons, koszulSignInsert_eq_sort, koszulSignInsert_eq_filter,
|
||||
koszulSignInsert_eq_grade, insertSign, superCommuteCoef]
|
||||
koszulSignInsert_eq_grade]
|
||||
have hx : (exchangeSign (q φ))
|
||||
(ofList q (List.take (↑(orderedInsertPos le (List.insertionSort le φs) φ))
|
||||
(List.insertionSort le φs))) = if FieldStatistic.ofList q [φ] = fermionic ∧
|
||||
FieldStatistic.ofList q (List.take (↑(orderedInsertPos le (List.insertionSort le φs) φ))
|
||||
(List.insertionSort le φs)) = fermionic then - 1 else 1 := by
|
||||
rw [exchangeSign_eq_if]
|
||||
simp
|
||||
rw [hx]
|
||||
congr
|
||||
simp only [List.filter_filter, Bool.and_self]
|
||||
rw [List.insertionSort]
|
||||
|
@ -206,14 +214,14 @@ def koszulSignCons (φ0 φ1 : 𝓕) : ℂ :=
|
|||
if le φ0 φ1 then 1 else
|
||||
if q φ0 = fermionic ∧ q φ1 = fermionic then -1 else 1
|
||||
|
||||
lemma koszulSignCons_eq_superComuteCoef (φ0 φ1 : 𝓕) : koszulSignCons q le φ0 φ1 =
|
||||
if le φ0 φ1 then 1 else superCommuteCoef q [φ0] [φ1] := by
|
||||
simp only [koszulSignCons, Fin.isValue, superCommuteCoef, ofList, ite_eq_right_iff, zero_ne_one,
|
||||
lemma koszulSignCons_eq_exchangeSign (φ0 φ1 : 𝓕) : koszulSignCons q le φ0 φ1 =
|
||||
if le φ0 φ1 then 1 else 𝓢(q φ0, q φ1) := by
|
||||
simp only [koszulSignCons, Fin.isValue, ofList, ite_eq_right_iff, zero_ne_one,
|
||||
imp_false]
|
||||
congr 1
|
||||
by_cases h0 : q φ0 = fermionic
|
||||
· by_cases h1 : q φ1 = fermionic
|
||||
· simp [h0, h1]
|
||||
· simp [h0, h1, exchangeSign]
|
||||
· have h1 : q φ1 = bosonic := (neq_fermionic_iff_eq_bosonic (q φ1)).mp h1
|
||||
simp [h0, h1]
|
||||
· have h0 : q φ0 = bosonic := (neq_fermionic_iff_eq_bosonic (q φ0)).mp h0
|
|
@ -1,435 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.Wick.Signs.StaticWickCoef
|
||||
/-!
|
||||
|
||||
# Create and annihilate sections (of bundles)
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
open HepLean.List
|
||||
open FieldStatistic
|
||||
|
||||
/-- The sections of `Σ i, f i` over a list `φs : List 𝓕`.
|
||||
In terms of physics, given some fields `φ₁...φₙ`, the different ways one can associate
|
||||
each field as a `creation` or an `annihilation` operator. E.g. the number of terms
|
||||
`φ₁⁰φ₂¹...φₙ⁰` `φ₁¹φ₂¹...φₙ⁰` etc. If some fields are exclusively creation or annihilation
|
||||
operators at this point (e.g. asymptotic states) this is accounted for. -/
|
||||
def CreateAnnihilateSect {𝓕 : Type} (f : 𝓕 → Type) (φs : List 𝓕) : Type :=
|
||||
Π i, f (φs.get i)
|
||||
|
||||
namespace CreateAnnihilateSect
|
||||
|
||||
section basic_defs
|
||||
|
||||
variable {𝓕 : Type} {f : 𝓕 → Type} [∀ i, Fintype (f i)] {φs : List 𝓕}
|
||||
(a : CreateAnnihilateSect f φs)
|
||||
|
||||
/-- The type `CreateAnnihilateSect f φs` is finite. -/
|
||||
instance fintype : Fintype (CreateAnnihilateSect f φs) := Pi.instFintype
|
||||
|
||||
/-- The section got by dropping the first element of `φs` if it exists. -/
|
||||
def tail : {φs : List 𝓕} → (a : CreateAnnihilateSect f φs) → CreateAnnihilateSect f φs.tail
|
||||
| [], a => a
|
||||
| _ :: _, a => fun i => a (Fin.succ i)
|
||||
|
||||
/-- For a list of fields `i :: l` the value of the section at the head `i`. -/
|
||||
def head {φ : 𝓕} (a : CreateAnnihilateSect f (φ :: φs)) : f φ := a ⟨0, Nat.zero_lt_succ φs.length⟩
|
||||
|
||||
end basic_defs
|
||||
|
||||
section toList_basic
|
||||
|
||||
variable {𝓕 : Type} {f : 𝓕 → Type} (q : 𝓕 → FieldStatistic)
|
||||
{φs : List 𝓕} (a : CreateAnnihilateSect f φs)
|
||||
|
||||
/-- The list `List (Σ i, f i)` defined by `a`. -/
|
||||
def toList : {l : List 𝓕} → (a : CreateAnnihilateSect f l) → List (Σ i, f i)
|
||||
| [], _ => []
|
||||
| i :: _, a => ⟨i, a.head⟩ :: toList a.tail
|
||||
|
||||
@[simp]
|
||||
lemma toList_length : (toList a).length = φs.length := by
|
||||
induction φs with
|
||||
| nil => rfl
|
||||
| cons i l ih =>
|
||||
simp only [toList, List.length_cons, Fin.zero_eta]
|
||||
rw [ih]
|
||||
|
||||
lemma toList_tail : {l : List 𝓕} → (a : CreateAnnihilateSect f l) → toList a.tail = (toList a).tail
|
||||
| [], _ => rfl
|
||||
| i :: l, a => by
|
||||
simp [toList]
|
||||
|
||||
lemma toList_cons {i : 𝓕} (a : CreateAnnihilateSect f (i :: φs)) :
|
||||
(toList a) = ⟨i, a.head⟩ :: toList a.tail := by
|
||||
rfl
|
||||
|
||||
lemma toList_get (a : CreateAnnihilateSect f φs) :
|
||||
(toList a).get = (fun i => ⟨φs.get i, a i⟩) ∘ Fin.cast (by simp) := by
|
||||
induction φs with
|
||||
| nil =>
|
||||
funext i
|
||||
exact Fin.elim0 i
|
||||
| cons i l ih =>
|
||||
simp only [toList_cons, List.get_eq_getElem, Fin.zero_eta, List.getElem_cons_succ,
|
||||
Function.comp_apply, Fin.cast_mk]
|
||||
funext x
|
||||
match x with
|
||||
| ⟨0, h⟩ => rfl
|
||||
| ⟨x + 1, h⟩ =>
|
||||
simp only [List.get_eq_getElem, Prod.mk.eta, List.getElem_cons_succ, Function.comp_apply]
|
||||
change (toList a.tail).get _ = _
|
||||
rw [ih]
|
||||
simp [tail]
|
||||
|
||||
@[simp]
|
||||
lemma toList_grade :
|
||||
FieldStatistic.ofList (fun i => q i.fst) a.toList = fermionic ↔
|
||||
FieldStatistic.ofList q φs = fermionic := by
|
||||
induction φs with
|
||||
| nil =>
|
||||
simp [toList]
|
||||
| cons i r ih =>
|
||||
simp only [ofList, Fin.isValue, ite_eq_right_iff, zero_ne_one, imp_false]
|
||||
have ih' := ih (fun i => a i.succ)
|
||||
have h1 : ofList (fun i => q i.fst) a.tail.toList = ofList q r := by
|
||||
by_cases h : ofList q r = fermionic
|
||||
· simp_all
|
||||
· have h0 : ofList q r = bosonic := (neq_fermionic_iff_eq_bosonic (ofList q r)).mp h
|
||||
rw [h0] at ih'
|
||||
simp only [reduceCtorEq, iff_false, neq_fermionic_iff_eq_bosonic] at ih'
|
||||
have h0' : ofList (fun i => q i.fst) a.tail.toList = bosonic := ih'
|
||||
rw [h0, h0']
|
||||
rw [h1]
|
||||
|
||||
@[simp]
|
||||
lemma toList_grade_take (q : 𝓕 → FieldStatistic) :
|
||||
(r : List 𝓕) → (a : CreateAnnihilateSect f r) → (n : ℕ) →
|
||||
ofList (fun i => q i.fst) (List.take n a.toList) = ofList q (List.take n r)
|
||||
| [], _, _ => by
|
||||
simp [toList]
|
||||
| i :: r, a, 0 => by
|
||||
simp
|
||||
| i :: r, a, Nat.succ n => by
|
||||
simp only [ofList, Fin.isValue]
|
||||
rw [toList_grade_take q r a.tail n]
|
||||
|
||||
end toList_basic
|
||||
|
||||
section toList_erase
|
||||
|
||||
variable {𝓕 : Type} {f : 𝓕 → Type} {l : List 𝓕}
|
||||
|
||||
/-- The equivalence between `CreateAnnihilateSect f l` and
|
||||
`f (l.get n) × CreateAnnihilateSect f (l.eraseIdx n)` obtained by extracting the `n`th field
|
||||
from `l`. -/
|
||||
def extractEquiv (n : Fin l.length) : CreateAnnihilateSect f l ≃
|
||||
f (l.get n) × CreateAnnihilateSect f (l.eraseIdx n) := by
|
||||
match l with
|
||||
| [] => exact Fin.elim0 n
|
||||
| l0 :: l =>
|
||||
let e1 : CreateAnnihilateSect f ((l0 :: l).eraseIdx n) ≃
|
||||
Π i, f ((l0 :: l).get (n.succAbove i)) :=
|
||||
Equiv.piCongr (Fin.castOrderIso (by rw [eraseIdx_cons_length])).toEquiv
|
||||
fun x => Equiv.cast (congrArg f (by
|
||||
rw [HepLean.List.eraseIdx_get]
|
||||
simp only [List.length_cons, Function.comp_apply, List.get_eq_getElem, Fin.coe_cast,
|
||||
RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply]
|
||||
congr 1
|
||||
simp only [Fin.succAbove]
|
||||
split
|
||||
next h =>
|
||||
simp_all only [Fin.coe_castSucc]
|
||||
split
|
||||
next h_1 => simp_all only [Fin.coe_castSucc, Fin.coe_cast]
|
||||
next h_1 =>
|
||||
simp_all only [not_lt, Fin.val_succ, Fin.coe_cast, self_eq_add_right, one_ne_zero]
|
||||
simp only [Fin.le_def, List.length_cons, Fin.coe_castSucc, Fin.coe_cast] at h_1
|
||||
simp only [Fin.lt_def, Fin.coe_castSucc, Fin.coe_cast] at h
|
||||
omega
|
||||
next h =>
|
||||
simp_all only [not_lt, Fin.val_succ]
|
||||
split
|
||||
next h_1 =>
|
||||
simp_all only [Fin.coe_castSucc, Fin.coe_cast, add_right_eq_self, one_ne_zero]
|
||||
simp only [Fin.lt_def, Fin.coe_castSucc, Fin.coe_cast] at h_1
|
||||
simp only [Fin.le_def, Fin.coe_cast, Fin.coe_castSucc] at h
|
||||
omega
|
||||
next h_1 => simp_all only [not_lt, Fin.val_succ, Fin.coe_cast]))
|
||||
exact (Fin.insertNthEquiv _ _).symm.trans (Equiv.prodCongr (Equiv.refl _) e1.symm)
|
||||
|
||||
lemma extractEquiv_symm_toList_get_same (n : Fin l.length) (a0 : f (l.get n))
|
||||
(a : CreateAnnihilateSect f (l.eraseIdx n)) :
|
||||
((extractEquiv n).symm (a0, a)).toList[n] = ⟨l[n], a0⟩ := by
|
||||
match l with
|
||||
| [] => exact Fin.elim0 n
|
||||
| l0 :: l =>
|
||||
trans (((CreateAnnihilateSect.extractEquiv n).symm (a0, a)).toList).get (Fin.cast (by simp) n)
|
||||
· simp only [List.length_cons, List.get_eq_getElem, Fin.coe_cast]
|
||||
rfl
|
||||
rw [CreateAnnihilateSect.toList_get]
|
||||
simp only [List.get_eq_getElem, List.length_cons, extractEquiv, RelIso.coe_fn_toEquiv,
|
||||
Fin.castOrderIso_apply, Equiv.symm_trans_apply, Equiv.symm_symm, Equiv.prodCongr_symm,
|
||||
Equiv.refl_symm, Equiv.prodCongr_apply, Equiv.coe_refl, Prod.map_apply, id_eq,
|
||||
Function.comp_apply, Fin.cast_trans, Fin.cast_eq_self, Sigma.mk.inj_iff, heq_eq_eq]
|
||||
apply And.intro
|
||||
· rfl
|
||||
erw [Fin.insertNthEquiv_apply]
|
||||
simp only [Fin.insertNth_apply_same]
|
||||
|
||||
/-- The section obtained by dropping the `n`th field. -/
|
||||
def eraseIdx (a : CreateAnnihilateSect f l) (n : Fin l.length) :
|
||||
CreateAnnihilateSect f (l.eraseIdx n) :=
|
||||
(extractEquiv n a).2
|
||||
|
||||
@[simp]
|
||||
lemma eraseIdx_zero_tail {i : 𝓕} (a : CreateAnnihilateSect f (i :: l)) :
|
||||
(eraseIdx a (@OfNat.ofNat (Fin (l.length + 1)) 0 Fin.instOfNat : Fin (l.length + 1))) =
|
||||
a.tail := by
|
||||
simp only [List.length_cons, Fin.val_zero, List.eraseIdx_cons_zero, eraseIdx, List.get_eq_getElem,
|
||||
List.getElem_cons_zero, extractEquiv, Fin.zero_succAbove, Fin.val_succ, List.getElem_cons_succ,
|
||||
Fin.insertNthEquiv_zero, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.cast_eq_self,
|
||||
Equiv.cast_refl, Equiv.trans_apply, Equiv.prodCongr_apply, Equiv.coe_refl, Prod.map_snd]
|
||||
rfl
|
||||
|
||||
lemma eraseIdx_succ_head {i : 𝓕} (n : ℕ) (hn : n + 1 < (i :: l).length)
|
||||
(a : CreateAnnihilateSect f (i :: l)) : (eraseIdx a ⟨n + 1, hn⟩).head = a.head := by
|
||||
rw [eraseIdx, extractEquiv]
|
||||
simp only [List.length_cons, List.get_eq_getElem, List.getElem_cons_succ, List.eraseIdx_cons_succ,
|
||||
RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Equiv.trans_apply, Equiv.prodCongr_apply,
|
||||
Equiv.coe_refl, Prod.map_snd]
|
||||
conv_lhs =>
|
||||
enter [1, 2, 1]
|
||||
erw [Fin.insertNthEquiv_symm_apply]
|
||||
simp only [head, Equiv.piCongr, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Equiv.piCongrRight,
|
||||
Equiv.cast_symm, Equiv.piCongrLeft, OrderIso.toEquiv_symm, OrderIso.symm_symm,
|
||||
Equiv.piCongrLeft', List.length_cons, Fin.zero_eta, Equiv.symm_trans_apply, Equiv.symm_symm,
|
||||
Equiv.coe_fn_mk, Equiv.coe_fn_symm_mk, Pi.map_apply, Fin.cast_zero, Fin.val_zero,
|
||||
List.getElem_cons_zero, Equiv.cast_apply]
|
||||
simp only [Fin.succAbove, Fin.castSucc_zero', Fin.removeNth]
|
||||
refine cast_eq_iff_heq.mpr ?_
|
||||
congr
|
||||
simp [Fin.ext_iff]
|
||||
|
||||
lemma eraseIdx_succ_tail {i : 𝓕} (n : ℕ) (hn : n + 1 < (i :: l).length)
|
||||
(a : CreateAnnihilateSect f (i :: l)) :
|
||||
(eraseIdx a ⟨n + 1, hn⟩).tail = eraseIdx a.tail ⟨n, Nat.succ_lt_succ_iff.mp hn⟩ := by
|
||||
match l with
|
||||
| [] =>
|
||||
simp at hn
|
||||
| r0 :: r =>
|
||||
rw [eraseIdx, extractEquiv]
|
||||
simp only [List.length_cons, List.eraseIdx_cons_succ, List.tail_cons, List.get_eq_getElem,
|
||||
List.getElem_cons_succ, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Equiv.trans_apply,
|
||||
Equiv.prodCongr_apply, Equiv.coe_refl, Prod.map_snd, Nat.succ_eq_add_one]
|
||||
conv_lhs =>
|
||||
enter [1, 2, 1]
|
||||
erw [Fin.insertNthEquiv_symm_apply]
|
||||
rw [eraseIdx]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [extractEquiv]
|
||||
simp only [List.get_eq_getElem, List.length_cons, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
|
||||
Equiv.trans_apply, Equiv.prodCongr_apply, Equiv.coe_refl, Prod.map_snd]
|
||||
erw [Fin.insertNthEquiv_symm_apply]
|
||||
simp only [tail, List.tail_cons, Equiv.piCongr, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
|
||||
Equiv.piCongrRight, Equiv.cast_symm, Equiv.piCongrLeft, OrderIso.toEquiv_symm,
|
||||
OrderIso.symm_symm, Equiv.piCongrLeft', Equiv.symm_trans_apply, Equiv.symm_symm,
|
||||
Equiv.coe_fn_mk, Equiv.coe_fn_symm_mk, Pi.map_apply, Fin.cast_succ_eq, Fin.val_succ,
|
||||
List.getElem_cons_succ, Equiv.cast_apply, List.get_eq_getElem, List.length_cons, Fin.succ_mk,
|
||||
Prod.map_apply, id_eq]
|
||||
funext i
|
||||
simp only [Pi.map_apply, Equiv.cast_apply]
|
||||
have hcast {α β : Type} (h : α = β) (a : α) (b : β) : cast h a = b ↔ a = cast (Eq.symm h) b := by
|
||||
cases h
|
||||
simp
|
||||
rw [hcast]
|
||||
simp only [cast_cast]
|
||||
refine eq_cast_iff_heq.mpr ?_
|
||||
simp only [Fin.succAbove, Fin.removeNth]
|
||||
congr
|
||||
simp only [List.length_cons, Fin.ext_iff, Fin.val_succ]
|
||||
split
|
||||
next h =>
|
||||
simp_all only [Fin.coe_castSucc, Fin.val_succ, Fin.coe_cast, add_left_inj]
|
||||
split
|
||||
next h_1 => simp_all only [Fin.coe_castSucc, Fin.coe_cast]
|
||||
next h_1 =>
|
||||
simp_all only [not_lt, Fin.val_succ, Fin.coe_cast, self_eq_add_right, one_ne_zero]
|
||||
simp only [Fin.lt_def, Fin.coe_castSucc, Fin.val_succ, Fin.coe_cast, add_lt_add_iff_right]
|
||||
at h
|
||||
simp only [Fin.le_def, Fin.coe_castSucc, Fin.coe_cast] at h_1
|
||||
omega
|
||||
next h =>
|
||||
simp_all only [not_lt, Fin.val_succ, Fin.coe_cast, add_left_inj]
|
||||
split
|
||||
next h_1 =>
|
||||
simp_all only [Fin.coe_castSucc, Fin.coe_cast, add_right_eq_self, one_ne_zero]
|
||||
simp only [Fin.le_def, Fin.coe_castSucc, Fin.val_succ, Fin.coe_cast, add_le_add_iff_right]
|
||||
at h
|
||||
simp only [Fin.lt_def, Fin.coe_castSucc, Fin.coe_cast] at h_1
|
||||
omega
|
||||
next h_1 => simp_all only [not_lt, Fin.val_succ, Fin.coe_cast]
|
||||
|
||||
lemma eraseIdx_toList : {l : List 𝓕} → {n : Fin l.length} → (a : CreateAnnihilateSect f l) →
|
||||
(eraseIdx a n).toList = a.toList.eraseIdx n
|
||||
| [], n, _ => Fin.elim0 n
|
||||
| r0 :: r, ⟨0, h⟩, a => by
|
||||
simp [toList_tail]
|
||||
| r0 :: r, ⟨n + 1, h⟩, a => by
|
||||
simp only [toList, List.length_cons, List.tail_cons, List.eraseIdx_cons_succ, List.cons.injEq,
|
||||
Sigma.mk.inj_iff, heq_eq_eq, true_and]
|
||||
apply And.intro
|
||||
· rw [eraseIdx_succ_head]
|
||||
· conv_rhs => rw [← eraseIdx_toList (l := r) (n := ⟨n, Nat.succ_lt_succ_iff.mp h⟩) a.tail]
|
||||
rw [eraseIdx_succ_tail]
|
||||
|
||||
lemma extractEquiv_symm_eraseIdx {I : Type} {f : I → Type}
|
||||
{l : List I} (n : Fin l.length) (a0 : f l[↑n]) (a : CreateAnnihilateSect f (l.eraseIdx n)) :
|
||||
((extractEquiv n).symm (a0, a)).eraseIdx n = a := by
|
||||
match l with
|
||||
| [] => exact Fin.elim0 n
|
||||
| l0 :: l =>
|
||||
rw [eraseIdx, extractEquiv]
|
||||
simp
|
||||
|
||||
end toList_erase
|
||||
|
||||
section toList_sign_conditions
|
||||
|
||||
variable {𝓕 : Type} {f : 𝓕 → Type} (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Prop) [DecidableRel le]
|
||||
{l : List 𝓕} (a : CreateAnnihilateSect f l)
|
||||
|
||||
lemma toList_koszulSignInsert (x : (i : 𝓕) × f i) :
|
||||
koszulSignInsert (fun i => q i.fst) (fun i j => le i.fst j.fst) x a.toList =
|
||||
koszulSignInsert q le x.1 l := by
|
||||
induction l with
|
||||
| nil => simp [koszulSignInsert]
|
||||
| cons b l ih =>
|
||||
simp only [koszulSignInsert, List.tail_cons, Fin.isValue]
|
||||
rw [ih]
|
||||
|
||||
lemma toList_koszulSign :
|
||||
koszulSign (fun i => q i.fst) (fun i j => le i.fst j.fst) a.toList =
|
||||
koszulSign q le l := by
|
||||
induction l with
|
||||
| nil => simp [koszulSign]
|
||||
| cons i l ih =>
|
||||
simp only [koszulSign, List.tail_cons]
|
||||
rw [ih]
|
||||
congr 1
|
||||
rw [toList_koszulSignInsert]
|
||||
|
||||
lemma insertionSortEquiv_toList :
|
||||
insertionSortEquiv (fun i j => le i.fst j.fst) a.toList =
|
||||
(Fin.castOrderIso (by simp)).toEquiv.trans ((insertionSortEquiv le l).trans
|
||||
(Fin.castOrderIso (by simp)).toEquiv) := by
|
||||
induction l with
|
||||
| nil =>
|
||||
simp [liftM, HepLean.List.insertionSortEquiv]
|
||||
| cons i l ih =>
|
||||
simp only [liftM, List.length_cons, Fin.zero_eta, List.insertionSort]
|
||||
conv_lhs => simp [HepLean.List.insertionSortEquiv]
|
||||
erw [orderedInsertEquiv_sigma]
|
||||
rw [ih]
|
||||
simp only [HepLean.Fin.equivCons_trans, Nat.succ_eq_add_one,
|
||||
HepLean.Fin.equivCons_castOrderIso, List.length_cons, Nat.add_zero, Nat.zero_eq,
|
||||
Fin.zero_eta]
|
||||
ext x
|
||||
conv_rhs => simp [HepLean.List.insertionSortEquiv]
|
||||
simp only [Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.cast_trans,
|
||||
Fin.coe_cast]
|
||||
have h2' (i : Σ i, f i) (l' : List (Σ i, f i)) :
|
||||
List.map (fun i => i.1) (List.orderedInsert (fun i j => le i.fst j.fst) i l') =
|
||||
List.orderedInsert le i.1 (List.map (fun i => i.1) l') := by
|
||||
induction l' with
|
||||
| nil =>
|
||||
simp [HepLean.List.orderedInsertEquiv]
|
||||
| cons j l' ih' =>
|
||||
by_cases hij : (fun i j => le i.fst j.fst) i j
|
||||
· rw [List.orderedInsert_of_le]
|
||||
· erw [List.orderedInsert_of_le]
|
||||
· simp
|
||||
· exact hij
|
||||
· exact hij
|
||||
· simp only [List.orderedInsert, hij, ↓reduceIte, List.unzip_snd, List.map_cons]
|
||||
simp only [↓reduceIte, List.cons.injEq, true_and]
|
||||
simpa using ih'
|
||||
have h2 (l' : List (Σ i, f i)) :
|
||||
List.map (fun i => i.1) (List.insertionSort (fun i j => le i.fst j.fst) l') =
|
||||
List.insertionSort le (List.map (fun i => i.1) l') := by
|
||||
induction l' with
|
||||
| nil =>
|
||||
simp [HepLean.List.orderedInsertEquiv]
|
||||
| cons i l' ih' =>
|
||||
simp only [List.insertionSort, List.unzip_snd]
|
||||
simp only [List.unzip_snd] at h2'
|
||||
rw [h2']
|
||||
congr
|
||||
rw [HepLean.List.orderedInsertEquiv_congr _ _ _ (h2 _)]
|
||||
simp only [List.length_cons, Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
|
||||
Fin.cast_trans, Fin.coe_cast]
|
||||
have h3 : (List.insertionSort le (List.map (fun i => i.1) a.tail.toList)) =
|
||||
List.insertionSort le l := by
|
||||
congr
|
||||
have h3' (l : List 𝓕) (a : CreateAnnihilateSect f l) :
|
||||
List.map (fun i => i.1) a.toList = l := by
|
||||
induction l with
|
||||
| nil => rfl
|
||||
| cons i l ih' =>
|
||||
simp only [toList, List.length_cons, Fin.zero_eta, Prod.mk.eta,
|
||||
List.unzip_snd, List.map_cons, List.cons.injEq, true_and]
|
||||
simpa using ih' _
|
||||
rw [h3']
|
||||
rfl
|
||||
rw [HepLean.List.orderedInsertEquiv_congr _ _ _ h3]
|
||||
simp only [List.length_cons, Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
|
||||
Fin.cast_trans, Fin.cast_eq_self, Fin.coe_cast]
|
||||
rfl
|
||||
|
||||
/-- Given a section for `l` the corresponding section
|
||||
for `List.insertionSort le1 l`. -/
|
||||
def sort :
|
||||
CreateAnnihilateSect f (List.insertionSort le l) :=
|
||||
Equiv.piCongr (HepLean.List.insertionSortEquiv le l) (fun i => (Equiv.cast (by
|
||||
congr 1
|
||||
rw [← HepLean.List.insertionSortEquiv_get]
|
||||
simp))) a
|
||||
|
||||
lemma sort_toList :
|
||||
(a.sort le).toList = List.insertionSort (fun i j => le i.fst j.fst) a.toList := by
|
||||
let l1 := List.insertionSort (fun i j => le i.fst j.fst) a.toList
|
||||
let l2 := (a.sort le).toList
|
||||
symm
|
||||
change l1 = l2
|
||||
have hlen : l1.length = l2.length := by
|
||||
simp [l1, l2]
|
||||
have hget : l1.get = l2.get ∘ Fin.cast hlen := by
|
||||
rw [← HepLean.List.insertionSortEquiv_get]
|
||||
rw [toList_get, toList_get]
|
||||
funext i
|
||||
rw [insertionSortEquiv_toList]
|
||||
simp only [Function.comp_apply, Equiv.symm_trans_apply,
|
||||
OrderIso.toEquiv_symm, Fin.symm_castOrderIso, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
|
||||
Fin.cast_trans, Fin.cast_eq_self, id_eq, eq_mpr_eq_cast, Fin.coe_cast, Sigma.mk.inj_iff]
|
||||
apply And.intro
|
||||
· have h1 := congrFun (HepLean.List.insertionSortEquiv_get (r := le) l) (Fin.cast (by simp) i)
|
||||
rw [← h1]
|
||||
simp
|
||||
· simp only [List.get_eq_getElem, sort, Equiv.piCongr, Equiv.trans_apply, Fin.coe_cast,
|
||||
Equiv.piCongrLeft_apply, Equiv.piCongrRight_apply, Pi.map_apply, Equiv.cast_apply,
|
||||
heq_eqRec_iff_heq]
|
||||
exact (cast_heq _ _).symm
|
||||
apply List.ext_get hlen
|
||||
rw [hget]
|
||||
simp
|
||||
|
||||
end toList_sign_conditions
|
||||
end CreateAnnihilateSect
|
||||
|
||||
end Wick
|
|
@ -1,245 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.Wick.Signs.StaticWickCoef
|
||||
/-!
|
||||
|
||||
# Koszul signs and ordering for lists and algebras
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
|
||||
noncomputable section
|
||||
open FieldStatistic
|
||||
|
||||
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Prop) [DecidableRel le]
|
||||
|
||||
/-- Given a relation `r` on `I` sorts elements of `MonoidAlgebra ℂ (FreeMonoid I)` by `r` giving it
|
||||
a signed dependent on `q`. -/
|
||||
def koszulOrderMonoidAlgebra :
|
||||
MonoidAlgebra ℂ (FreeMonoid 𝓕) →ₗ[ℂ] MonoidAlgebra ℂ (FreeMonoid 𝓕) :=
|
||||
Finsupp.lift (MonoidAlgebra ℂ (FreeMonoid 𝓕)) ℂ (List 𝓕)
|
||||
(fun i => Finsupp.lsingle (R := ℂ) (List.insertionSort le i) (koszulSign q le i))
|
||||
|
||||
lemma koszulOrderMonoidAlgebra_ofList (i : List 𝓕) :
|
||||
koszulOrderMonoidAlgebra q le (MonoidAlgebra.of ℂ (FreeMonoid 𝓕) i) =
|
||||
(koszulSign q le i) • (MonoidAlgebra.of ℂ (FreeMonoid 𝓕) (List.insertionSort le i)) := by
|
||||
simp only [koszulOrderMonoidAlgebra, Finsupp.lsingle_apply, MonoidAlgebra.of_apply,
|
||||
MonoidAlgebra.smul_single', mul_one]
|
||||
rw [MonoidAlgebra.ext_iff]
|
||||
intro x
|
||||
erw [Finsupp.lift_apply]
|
||||
simp only [MonoidAlgebra.smul_single', zero_mul, Finsupp.single_zero, Finsupp.sum_single_index,
|
||||
one_mul]
|
||||
|
||||
@[simp]
|
||||
lemma koszulOrderMonoidAlgebra_single (l : FreeMonoid 𝓕) (x : ℂ) :
|
||||
koszulOrderMonoidAlgebra q le (MonoidAlgebra.single l x)
|
||||
= (koszulSign q le l) • (MonoidAlgebra.single (List.insertionSort le l) x) := by
|
||||
simp only [koszulOrderMonoidAlgebra, Finsupp.lsingle_apply, MonoidAlgebra.smul_single']
|
||||
rw [MonoidAlgebra.ext_iff]
|
||||
intro x'
|
||||
erw [Finsupp.lift_apply]
|
||||
simp only [MonoidAlgebra.smul_single', zero_mul, Finsupp.single_zero, Finsupp.sum_single_index,
|
||||
one_mul, MonoidAlgebra.single]
|
||||
congr 2
|
||||
rw [NonUnitalNormedCommRing.mul_comm]
|
||||
|
||||
/-- Given a relation `r` on `I` sorts elements of `FreeAlgebra ℂ I` by `r` giving it
|
||||
a signed dependent on `q`. -/
|
||||
def koszulOrder : FreeAlgebra ℂ 𝓕 →ₗ[ℂ] FreeAlgebra ℂ 𝓕 :=
|
||||
FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm.toAlgHom.toLinearMap
|
||||
∘ₗ koszulOrderMonoidAlgebra q le
|
||||
∘ₗ FreeAlgebra.equivMonoidAlgebraFreeMonoid.toAlgHom.toLinearMap
|
||||
|
||||
@[simp]
|
||||
lemma koszulOrder_ι (i : 𝓕) : koszulOrder q le (FreeAlgebra.ι ℂ i) = FreeAlgebra.ι ℂ i := by
|
||||
simp only [koszulOrder, AlgEquiv.toAlgHom_eq_coe, AlgEquiv.toAlgHom_toLinearMap,
|
||||
koszulOrderMonoidAlgebra, Finsupp.lsingle_apply, LinearMap.coe_comp, Function.comp_apply,
|
||||
AlgEquiv.toLinearMap_apply]
|
||||
rw [AlgEquiv.symm_apply_eq]
|
||||
simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
AlgEquiv.ofAlgHom_apply, FreeAlgebra.lift_ι_apply]
|
||||
rw [@MonoidAlgebra.ext_iff]
|
||||
intro x
|
||||
erw [Finsupp.lift_apply]
|
||||
simp only [MonoidAlgebra.smul_single', List.insertionSort, List.orderedInsert,
|
||||
koszulSign_freeMonoid_of, mul_one, Finsupp.single_zero, Finsupp.sum_single_index]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma koszulOrder_single (l : FreeMonoid 𝓕) :
|
||||
koszulOrder q le (FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single l x))
|
||||
= FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm
|
||||
(MonoidAlgebra.single (List.insertionSort le l) (koszulSign q le l * x)) := by
|
||||
simp [koszulOrder]
|
||||
|
||||
@[simp]
|
||||
lemma koszulOrder_ι_pair (i j : 𝓕) : koszulOrder q le (FreeAlgebra.ι ℂ i * FreeAlgebra.ι ℂ j) =
|
||||
if le i j then FreeAlgebra.ι ℂ i * FreeAlgebra.ι ℂ j else
|
||||
(koszulSign q le [i, j]) • (FreeAlgebra.ι ℂ j * FreeAlgebra.ι ℂ i) := by
|
||||
have h1 : FreeAlgebra.ι ℂ i * FreeAlgebra.ι ℂ j =
|
||||
FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single [i, j] 1) := by
|
||||
simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
AlgEquiv.ofAlgHom_symm_apply, MonoidAlgebra.lift_single, one_smul]
|
||||
rfl
|
||||
conv_lhs => rw [h1]
|
||||
simp only [koszulOrder, AlgEquiv.toAlgHom_eq_coe, AlgEquiv.toAlgHom_toLinearMap,
|
||||
LinearMap.coe_comp, Function.comp_apply, AlgEquiv.toLinearMap_apply, AlgEquiv.apply_symm_apply,
|
||||
koszulOrderMonoidAlgebra_single, List.insertionSort, List.orderedInsert,
|
||||
MonoidAlgebra.smul_single', mul_one]
|
||||
by_cases hr : le i j
|
||||
· rw [if_pos hr, if_pos hr]
|
||||
simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
AlgEquiv.ofAlgHom_symm_apply, MonoidAlgebra.lift_single]
|
||||
have hKS : koszulSign q le [i, j] = 1 := by
|
||||
simp only [koszulSign, koszulSignInsert, Fin.isValue, mul_one, ite_eq_left_iff,
|
||||
ite_eq_right_iff, and_imp]
|
||||
exact fun a a_1 a_2 => False.elim (a hr)
|
||||
rw [hKS]
|
||||
simp only [one_smul]
|
||||
rfl
|
||||
· rw [if_neg hr, if_neg hr]
|
||||
simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
AlgEquiv.ofAlgHom_symm_apply, MonoidAlgebra.lift_single]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma koszulOrder_one : koszulOrder q le 1 = 1 := by
|
||||
trans koszulOrder q le (FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single [] 1))
|
||||
congr
|
||||
· simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
AlgEquiv.ofAlgHom_symm_apply, MonoidAlgebra.lift_single, one_smul]
|
||||
rfl
|
||||
· simp only [koszulOrder_single, List.insertionSort, mul_one, EmbeddingLike.map_eq_one_iff]
|
||||
rfl
|
||||
|
||||
lemma mul_koszulOrder_le (i : 𝓕) (A : FreeAlgebra ℂ 𝓕) (hi : ∀ j, le i j) :
|
||||
FreeAlgebra.ι ℂ i * koszulOrder q le A = koszulOrder q le (FreeAlgebra.ι ℂ i * A) := by
|
||||
let f : FreeAlgebra ℂ 𝓕 →ₗ[ℂ] FreeAlgebra ℂ 𝓕 := {
|
||||
toFun := fun x => FreeAlgebra.ι ℂ i * x,
|
||||
map_add' := fun x y => by
|
||||
simp [mul_add],
|
||||
map_smul' := by simp}
|
||||
change (f ∘ₗ koszulOrder q le) A = (koszulOrder q le ∘ₗ f) _
|
||||
have f_single (l : FreeMonoid 𝓕) (x : ℂ) :
|
||||
f ((FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single l x)))
|
||||
= (FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single (i :: l) x)) := by
|
||||
simp only [LinearMap.coe_mk, AddHom.coe_mk, f]
|
||||
have hf : FreeAlgebra.ι ℂ i = FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm
|
||||
(MonoidAlgebra.single [i] 1) := by
|
||||
simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
AlgEquiv.ofAlgHom_symm_apply, MonoidAlgebra.lift_single, one_smul]
|
||||
rfl
|
||||
rw [hf]
|
||||
rw [@AlgEquiv.eq_symm_apply]
|
||||
simp only [map_mul, AlgEquiv.apply_symm_apply, MonoidAlgebra.single_mul_single, one_mul]
|
||||
rfl
|
||||
have h1 : f ∘ₗ koszulOrder q le = koszulOrder q le ∘ₗ f := by
|
||||
let e : FreeAlgebra ℂ 𝓕 ≃ₗ[ℂ] MonoidAlgebra ℂ (FreeMonoid 𝓕) :=
|
||||
FreeAlgebra.equivMonoidAlgebraFreeMonoid.toLinearEquiv
|
||||
apply (LinearEquiv.eq_comp_toLinearMap_iff (e₁₂ := e.symm) _ _).mp
|
||||
apply MonoidAlgebra.lhom_ext'
|
||||
intro l
|
||||
apply LinearMap.ext
|
||||
intro x
|
||||
simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply,
|
||||
MonoidAlgebra.lsingle_apply]
|
||||
erw [koszulOrder_single]
|
||||
rw [f_single]
|
||||
erw [f_single]
|
||||
rw [koszulOrder_single]
|
||||
congr 2
|
||||
· simp only [List.insertionSort]
|
||||
have hi (l : List 𝓕) : i :: l = List.orderedInsert le i l := by
|
||||
induction l with
|
||||
| nil => rfl
|
||||
| cons j l ih =>
|
||||
refine (List.orderedInsert_of_le le l (hi j)).symm
|
||||
exact hi _
|
||||
· congr 1
|
||||
rw [koszulSign]
|
||||
have h1 (l : List 𝓕) : koszulSignInsert q le i l = 1 := by
|
||||
exact koszulSignInsert_le_forall q le i l hi
|
||||
rw [h1]
|
||||
simp
|
||||
rw [h1]
|
||||
|
||||
lemma koszulOrder_mul_ge (i : 𝓕) (A : FreeAlgebra ℂ 𝓕) (hi : ∀ j, le j i) :
|
||||
koszulOrder q le A * FreeAlgebra.ι ℂ i = koszulOrder q le (A * FreeAlgebra.ι ℂ i) := by
|
||||
let f : FreeAlgebra ℂ 𝓕 →ₗ[ℂ] FreeAlgebra ℂ 𝓕 := {
|
||||
toFun := fun x => x * FreeAlgebra.ι ℂ i,
|
||||
map_add' := fun x y => by
|
||||
simp [add_mul],
|
||||
map_smul' := by simp}
|
||||
change (f ∘ₗ koszulOrder q le) A = (koszulOrder q le ∘ₗ f) A
|
||||
have f_single (l : FreeMonoid 𝓕) (x : ℂ) :
|
||||
f ((FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single l x)))
|
||||
= (FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm
|
||||
(MonoidAlgebra.single (l.toList ++ [i]) x)) := by
|
||||
simp only [LinearMap.coe_mk, AddHom.coe_mk, f]
|
||||
have hf : FreeAlgebra.ι ℂ i = FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm
|
||||
(MonoidAlgebra.single [i] 1) := by
|
||||
simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
AlgEquiv.ofAlgHom_symm_apply, MonoidAlgebra.lift_single, one_smul]
|
||||
rfl
|
||||
rw [hf]
|
||||
rw [@AlgEquiv.eq_symm_apply]
|
||||
simp only [map_mul, AlgEquiv.apply_symm_apply, MonoidAlgebra.single_mul_single, mul_one]
|
||||
rfl
|
||||
have h1 : f ∘ₗ koszulOrder q le = koszulOrder q le ∘ₗ f := by
|
||||
let e : FreeAlgebra ℂ 𝓕 ≃ₗ[ℂ] MonoidAlgebra ℂ (FreeMonoid 𝓕) :=
|
||||
FreeAlgebra.equivMonoidAlgebraFreeMonoid.toLinearEquiv
|
||||
apply (LinearEquiv.eq_comp_toLinearMap_iff (e₁₂ := e.symm) _ _).mp
|
||||
apply MonoidAlgebra.lhom_ext'
|
||||
intro l
|
||||
apply LinearMap.ext
|
||||
intro x
|
||||
simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply,
|
||||
MonoidAlgebra.lsingle_apply]
|
||||
erw [koszulOrder_single]
|
||||
rw [f_single]
|
||||
erw [f_single]
|
||||
rw [koszulOrder_single]
|
||||
congr 3
|
||||
· change (List.insertionSort le l) ++ [i] = List.insertionSort le (l.toList ++ [i])
|
||||
have hoi (l : List 𝓕) (j : 𝓕) : List.orderedInsert le j (l ++ [i]) =
|
||||
List.orderedInsert le j l ++ [i] := by
|
||||
induction l with
|
||||
| nil => simp [hi]
|
||||
| cons b l ih =>
|
||||
simp only [List.orderedInsert, List.append_eq]
|
||||
by_cases hr : le j b
|
||||
· rw [if_pos hr, if_pos hr]
|
||||
rfl
|
||||
· rw [if_neg hr, if_neg hr]
|
||||
rw [ih]
|
||||
rfl
|
||||
have hI (l : List 𝓕) : (List.insertionSort le l) ++ [i] =
|
||||
List.insertionSort le (l ++ [i]) := by
|
||||
induction l with
|
||||
| nil => rfl
|
||||
| cons j l ih =>
|
||||
simp only [List.insertionSort, List.append_eq]
|
||||
rw [← ih]
|
||||
rw [hoi]
|
||||
rw [hI]
|
||||
rfl
|
||||
· have hI (l : List 𝓕) : koszulSign q le l = koszulSign q le (l ++ [i]) := by
|
||||
induction l with
|
||||
| nil => simp [koszulSign, koszulSignInsert]
|
||||
| cons j l ih =>
|
||||
simp only [koszulSign, List.append_eq]
|
||||
rw [ih]
|
||||
simp only [mul_eq_mul_right_iff]
|
||||
apply Or.inl
|
||||
rw [koszulSignInsert_ge_forall_append q le l j i hi]
|
||||
rw [hI]
|
||||
rfl
|
||||
rw [h1]
|
||||
|
||||
end
|
||||
end Wick
|
|
@ -1,205 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.Wick.CreateAnnihilateSection
|
||||
import HepLean.PerturbationTheory.Wick.KoszulOrder
|
||||
/-!
|
||||
|
||||
# Koszul signs and ordering for lists and algebras
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
open HepLean.List
|
||||
open FieldStatistic
|
||||
|
||||
noncomputable section
|
||||
|
||||
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Prop) [DecidableRel le]
|
||||
|
||||
/-- The element of the free algebra `FreeAlgebra ℂ I` associated with a `List I`. -/
|
||||
def ofList (l : List 𝓕) (x : ℂ) : FreeAlgebra ℂ 𝓕 :=
|
||||
FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single l x)
|
||||
|
||||
lemma ofList_pair (l r : List 𝓕) (x y : ℂ) :
|
||||
ofList (l ++ r) (x * y) = ofList l x * ofList r y := by
|
||||
simp only [ofList, ← map_mul, MonoidAlgebra.single_mul_single, EmbeddingLike.apply_eq_iff_eq]
|
||||
rfl
|
||||
|
||||
lemma ofList_triple (la lb lc : List 𝓕) (xa xb xc : ℂ) :
|
||||
ofList (la ++ lb ++ lc) (xa * xb * xc) = ofList la xa * ofList lb xb * ofList lc xc := by
|
||||
rw [ofList_pair, ofList_pair]
|
||||
|
||||
lemma ofList_triple_assoc (la lb lc : List 𝓕) (xa xb xc : ℂ) :
|
||||
ofList (la ++ (lb ++ lc)) (xa * (xb * xc)) = ofList la xa * ofList lb xb * ofList lc xc := by
|
||||
rw [ofList_pair, ofList_pair]
|
||||
exact Eq.symm (mul_assoc (ofList la xa) (ofList lb xb) (ofList lc xc))
|
||||
|
||||
lemma ofList_cons_eq_ofList (l : List 𝓕) (i : 𝓕) (x : ℂ) :
|
||||
ofList (i :: l) x = ofList [i] 1 * ofList l x := by
|
||||
simp only [ofList, ← map_mul, MonoidAlgebra.single_mul_single, one_mul,
|
||||
EmbeddingLike.apply_eq_iff_eq]
|
||||
rfl
|
||||
|
||||
lemma ofList_singleton (i : 𝓕) :
|
||||
ofList [i] 1 = FreeAlgebra.ι ℂ i := by
|
||||
simp only [ofList, FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
MonoidAlgebra.single, AlgEquiv.ofAlgHom_symm_apply, MonoidAlgebra.lift_single, one_smul]
|
||||
rfl
|
||||
|
||||
lemma ofList_eq_smul_one (l : List 𝓕) (x : ℂ) :
|
||||
ofList l x = x • ofList l 1 := by
|
||||
simp only [ofList]
|
||||
rw [← map_smul]
|
||||
simp
|
||||
|
||||
lemma ofList_empty : ofList [] 1 = (1 : FreeAlgebra ℂ 𝓕) := by
|
||||
simp only [ofList, EmbeddingLike.map_eq_one_iff]
|
||||
rfl
|
||||
|
||||
lemma ofList_empty' : ofList [] x = x • (1 : FreeAlgebra ℂ 𝓕) := by
|
||||
rw [ofList_eq_smul_one, ofList_empty]
|
||||
|
||||
lemma koszulOrder_ofList
|
||||
(l : List 𝓕) (x : ℂ) :
|
||||
koszulOrder q le (ofList l x) = (koszulSign q le l) • ofList (List.insertionSort le l) x := by
|
||||
rw [ofList]
|
||||
rw [koszulOrder_single]
|
||||
change ofList (List.insertionSort le l) _ = _
|
||||
rw [ofList_eq_smul_one]
|
||||
conv_rhs => rw [ofList_eq_smul_one]
|
||||
rw [smul_smul]
|
||||
|
||||
lemma ofList_insertionSort_eq_koszulOrder (l : List 𝓕) (x : ℂ) :
|
||||
ofList (List.insertionSort le l) x = (koszulSign q le l) • koszulOrder q le (ofList l x) := by
|
||||
rw [koszulOrder_ofList]
|
||||
rw [smul_smul]
|
||||
rw [koszulSign_mul_self]
|
||||
simp
|
||||
|
||||
/-- The map of algebras from `FreeAlgebra ℂ I` to `FreeAlgebra ℂ (Σ i, f i)` taking
|
||||
the monomial `i` to the sum of elements in `(Σ i, f i)` above `i`, i.e. the sum of the fiber
|
||||
above `i`. -/
|
||||
def sumFiber (f : 𝓕 → Type) [∀ i, Fintype (f i)] :
|
||||
FreeAlgebra ℂ 𝓕 →ₐ[ℂ] FreeAlgebra ℂ (Σ i, f i) :=
|
||||
FreeAlgebra.lift ℂ fun i => ∑ (j : f i), FreeAlgebra.ι ℂ ⟨i, j⟩
|
||||
|
||||
lemma sumFiber_ι (f : 𝓕 → Type) [∀ i, Fintype (f i)] (i : 𝓕) :
|
||||
sumFiber f (FreeAlgebra.ι ℂ i) = ∑ (b : f i), FreeAlgebra.ι ℂ ⟨i, b⟩ := by
|
||||
simp [sumFiber]
|
||||
|
||||
/-- Given a list `l : List I` the corresponding element of `FreeAlgebra ℂ (Σ i, f i)`
|
||||
by mapping each `i : I` to the sum of it's fiber in `Σ i, f i` and taking the product of the
|
||||
result.
|
||||
For example, in terms of creation and annihilation opperators,
|
||||
`[φ₁, φ₂, φ₃]` gets taken to `(φ₁⁰ + φ₁¹)(φ₂⁰ + φ₂¹)(φ₃⁰ + φ₃¹)`.
|
||||
-/
|
||||
def ofListLift (f : 𝓕 → Type) [∀ i, Fintype (f i)] (l : List 𝓕) (x : ℂ) :
|
||||
FreeAlgebra ℂ (Σ i, f i) :=
|
||||
sumFiber f (ofList l x)
|
||||
|
||||
lemma ofListLift_empty (f : 𝓕 → Type) [∀ i, Fintype (f i)] :
|
||||
ofListLift f [] 1 = 1 := by
|
||||
simp only [ofListLift, EmbeddingLike.map_eq_one_iff]
|
||||
rw [ofList_empty]
|
||||
exact map_one (sumFiber f)
|
||||
|
||||
lemma ofListLift_empty_smul (f : 𝓕 → Type) [∀ i, Fintype (f i)] (x : ℂ) :
|
||||
ofListLift f [] x = x • 1 := by
|
||||
simp only [ofListLift, EmbeddingLike.map_eq_one_iff]
|
||||
rw [ofList_eq_smul_one]
|
||||
rw [ofList_empty]
|
||||
simp
|
||||
|
||||
lemma ofListLift_cons (f : 𝓕 → Type) [∀ i, Fintype (f i)] (i : 𝓕) (r : List 𝓕) (x : ℂ) :
|
||||
ofListLift f (i :: r) x = (∑ j : f i, FreeAlgebra.ι ℂ ⟨i, j⟩) * (ofListLift f r x) := by
|
||||
rw [ofListLift, ofList_cons_eq_ofList, ofList_singleton, map_mul]
|
||||
conv_lhs => lhs; rw [sumFiber]
|
||||
rw [ofListLift]
|
||||
simp
|
||||
|
||||
lemma ofListLift_singleton (f : 𝓕 → Type) [∀ i, Fintype (f i)] (i : 𝓕) (x : ℂ) :
|
||||
ofListLift f [i] x = ∑ j : f i, x • FreeAlgebra.ι ℂ ⟨i, j⟩ := by
|
||||
simp only [ofListLift]
|
||||
rw [ofList_eq_smul_one, ofList_singleton, map_smul]
|
||||
rw [sumFiber_ι]
|
||||
rw [Finset.smul_sum]
|
||||
|
||||
lemma ofListLift_singleton_one (f : 𝓕 → Type) [∀ i, Fintype (f i)] (i : 𝓕) :
|
||||
ofListLift f [i] 1 = ∑ j : f i, FreeAlgebra.ι ℂ ⟨i, j⟩ := by
|
||||
simp only [ofListLift]
|
||||
rw [ofList_eq_smul_one, ofList_singleton, map_smul]
|
||||
rw [sumFiber_ι]
|
||||
rw [Finset.smul_sum]
|
||||
simp
|
||||
|
||||
lemma ofListLift_cons_eq_ofListLift (f : 𝓕 → Type) [∀ i, Fintype (f i)] (i : 𝓕)
|
||||
(r : List 𝓕) (x : ℂ) :
|
||||
ofListLift f (i :: r) x = ofListLift f [i] 1 * ofListLift f r x := by
|
||||
rw [ofListLift_cons, ofListLift_singleton]
|
||||
simp only [one_smul]
|
||||
|
||||
lemma ofListLift_expand (f : 𝓕 → Type) [∀ i, Fintype (f i)] (x : ℂ) :
|
||||
(l : List 𝓕) → ofListLift f l x = ∑ (a : CreateAnnihilateSect f l), ofList a.toList x
|
||||
| [] => by
|
||||
simp only [ofListLift, CreateAnnihilateSect, List.length_nil, List.get_eq_getElem,
|
||||
Finset.univ_unique, CreateAnnihilateSect.toList, Finset.sum_const, Finset.card_singleton,
|
||||
one_smul]
|
||||
rw [ofList_eq_smul_one, map_smul, ofList_empty, ofList_eq_smul_one, ofList_empty, map_one]
|
||||
| i :: l => by
|
||||
rw [ofListLift_cons, ofListLift_expand f x l]
|
||||
conv_rhs => rw [← (CreateAnnihilateSect.extractEquiv
|
||||
⟨0, by exact Nat.zero_lt_succ l.length⟩).symm.sum_comp (α := FreeAlgebra ℂ _)]
|
||||
erw [Finset.sum_product]
|
||||
rw [Finset.sum_mul]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
intro n
|
||||
rw [Finset.mul_sum]
|
||||
congr
|
||||
funext j
|
||||
congr
|
||||
funext n
|
||||
rw [← ofList_singleton, ← ofList_pair, one_mul]
|
||||
rfl
|
||||
|
||||
lemma koszulOrder_ofListLift {f : 𝓕 → Type} [∀ i, Fintype (f i)]
|
||||
(l : List 𝓕) (x : ℂ) :
|
||||
koszulOrder (fun i => q i.fst) (fun i j => le i.1 j.1) (ofListLift f l x) =
|
||||
sumFiber f (koszulOrder q le (ofList l x)) := by
|
||||
rw [koszulOrder_ofList]
|
||||
rw [map_smul]
|
||||
change _ = _ • ofListLift _ _ _
|
||||
rw [ofListLift_expand]
|
||||
rw [map_sum]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
intro a
|
||||
rw [koszulOrder_ofList]
|
||||
rw [CreateAnnihilateSect.toList_koszulSign]
|
||||
rw [← Finset.smul_sum]
|
||||
apply congrArg
|
||||
conv_lhs =>
|
||||
rhs
|
||||
intro n
|
||||
rw [← CreateAnnihilateSect.sort_toList]
|
||||
rw [ofListLift_expand]
|
||||
refine Fintype.sum_equiv
|
||||
((HepLean.List.insertionSortEquiv le l).piCongr fun i => Equiv.cast ?_) _ _ ?_
|
||||
congr 1
|
||||
· rw [← HepLean.List.insertionSortEquiv_get]
|
||||
simp
|
||||
· intro x
|
||||
rfl
|
||||
|
||||
lemma koszulOrder_ofListLift_eq_ofListLift {f : 𝓕 → Type} [∀ i, Fintype (f i)]
|
||||
(l : List 𝓕) (x : ℂ) : koszulOrder (fun i => q i.fst) (fun i j => le i.1 j.1)
|
||||
(ofListLift f l x) =
|
||||
koszulSign q le l • ofListLift f (List.insertionSort le l) x := by
|
||||
rw [koszulOrder_ofListLift, koszulOrder_ofList, map_smul]
|
||||
rfl
|
||||
|
||||
end
|
||||
end Wick
|
|
@ -1,349 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.Wick.SuperCommute
|
||||
/-!
|
||||
|
||||
# Operator map
|
||||
|
||||
-/
|
||||
namespace Wick
|
||||
|
||||
noncomputable section
|
||||
|
||||
open FieldStatistic
|
||||
|
||||
variable {𝓕 : Type}
|
||||
|
||||
/-- A map from the free algebra of fields `FreeAlgebra ℂ 𝓕` to an algebra `A`, to be
|
||||
thought of as the operator algebra is said to be an operator map if
|
||||
all super commutors of fields land in the center of `A`,
|
||||
if two fields are of a different grade then their super commutor lands on zero,
|
||||
and the `koszulOrder` (normal order) of any term with a super commutor of two fields present
|
||||
is zero.
|
||||
This can be thought as as a condition on the operator algebra `A` as much as it can
|
||||
on `F`. -/
|
||||
class OperatorMap {A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Prop)
|
||||
[DecidableRel le] (F : FreeAlgebra ℂ 𝓕 →ₐ[ℂ] A) : Prop where
|
||||
superCommute_mem_center : ∀ i j, F (superCommute q (FreeAlgebra.ι ℂ i) (FreeAlgebra.ι ℂ j)) ∈
|
||||
Subalgebra.center ℂ A
|
||||
superCommute_diff_grade_zero : ∀ i j, q i ≠ q j →
|
||||
F (superCommute q (FreeAlgebra.ι ℂ i) (FreeAlgebra.ι ℂ j)) = 0
|
||||
superCommute_ordered_zero : ∀ i j, ∀ a b,
|
||||
F (koszulOrder q le (a * superCommute q (FreeAlgebra.ι ℂ i) (FreeAlgebra.ι ℂ j) * b)) = 0
|
||||
|
||||
namespace OperatorMap
|
||||
|
||||
variable {A : Type} [Semiring A] [Algebra ℂ A]
|
||||
{q : 𝓕 → FieldStatistic} {le : 𝓕 → 𝓕 → Prop}
|
||||
[DecidableRel le] (F : FreeAlgebra ℂ 𝓕 →ₐ[ℂ] A)
|
||||
|
||||
lemma superCommute_ofList_singleton_ι_center [OperatorMap q le F] (i j : 𝓕) :
|
||||
F (superCommute q (ofList [i] xa) (FreeAlgebra.ι ℂ j)) ∈ Subalgebra.center ℂ A := by
|
||||
have h1 : F (superCommute q (ofList [i] xa) (FreeAlgebra.ι ℂ j)) =
|
||||
xa • F (superCommute q (FreeAlgebra.ι ℂ i) (FreeAlgebra.ι ℂ j)) := by
|
||||
rw [← map_smul]
|
||||
congr
|
||||
rw [ofList_eq_smul_one, ofList_singleton]
|
||||
rw [map_smul]
|
||||
rfl
|
||||
rw [h1]
|
||||
refine Subalgebra.smul_mem (Subalgebra.center ℂ A) ?_ xa
|
||||
exact superCommute_mem_center (le := le) i j
|
||||
|
||||
end OperatorMap
|
||||
|
||||
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Prop) [DecidableRel le]
|
||||
|
||||
lemma superCommuteSplit_operatorMap (lb : List 𝓕) (xa xb : ℂ) (n : ℕ)
|
||||
(hn : n < lb.length) {A : Type} [Semiring A] [Algebra ℂ A] (f : FreeAlgebra ℂ 𝓕 →ₐ[ℂ] A)
|
||||
[OperatorMap q le f] (i : 𝓕) :
|
||||
f (superCommuteSplit q [i] lb xa xb n hn) =
|
||||
f (superCommute q (ofList [i] xa) (FreeAlgebra.ι ℂ (lb.get ⟨n, hn⟩)))
|
||||
* (superCommuteCoef q [i] (List.take n lb) •
|
||||
f (ofList (List.eraseIdx lb n) xb)) := by
|
||||
have hn : f ((superCommute q) (ofList [i] xa) (FreeAlgebra.ι ℂ (lb.get ⟨n, hn⟩))) ∈
|
||||
Subalgebra.center ℂ A :=
|
||||
OperatorMap.superCommute_ofList_singleton_ι_center (le := le) f i (lb.get ⟨n, hn⟩)
|
||||
rw [Subalgebra.mem_center_iff] at hn
|
||||
rw [superCommuteSplit, map_mul, map_mul, map_smul, hn, mul_assoc, smul_mul_assoc,
|
||||
← map_mul, ← ofList_pair]
|
||||
congr
|
||||
· exact Eq.symm (List.eraseIdx_eq_take_drop_succ lb n)
|
||||
· exact one_mul xb
|
||||
|
||||
lemma superCommuteLiftSplit_operatorMap {f : 𝓕 → Type} [∀ i, Fintype (f i)]
|
||||
(c : (Σ i, f i)) (r : List 𝓕) (x y : ℂ) (n : ℕ)
|
||||
(hn : n < r.length)
|
||||
(le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le]
|
||||
{A : Type} [Semiring A] [Algebra ℂ A] (F : FreeAlgebra ℂ (Σ i, f i) →ₐ[ℂ] A)
|
||||
[OperatorMap (fun i => q i.1) le F] :
|
||||
F (superCommuteLiftSplit q [c] r x y n hn) = superCommuteLiftCoef q [c] (List.take n r) •
|
||||
(F (superCommute (fun i => q i.1) (ofList [c] x)
|
||||
(sumFiber f (FreeAlgebra.ι ℂ (r.get ⟨n, hn⟩))))
|
||||
* F (ofListLift f (List.eraseIdx r n) y)) := by
|
||||
rw [superCommuteLiftSplit]
|
||||
rw [map_smul]
|
||||
congr
|
||||
rw [map_mul, map_mul]
|
||||
have h1 : F ((superCommute fun i => q i.fst) (ofList [c] x) ((sumFiber f)
|
||||
(FreeAlgebra.ι ℂ (r.get ⟨n, hn⟩)))) ∈ Subalgebra.center ℂ A := by
|
||||
rw [sumFiber_ι]
|
||||
rw [map_sum, map_sum]
|
||||
refine Subalgebra.sum_mem _ ?_
|
||||
intro n
|
||||
exact fun a => OperatorMap.superCommute_ofList_singleton_ι_center (le := le) F c _
|
||||
rw [Subalgebra.mem_center_iff] at h1
|
||||
rw [h1, mul_assoc, ← map_mul]
|
||||
congr
|
||||
rw [ofListLift, ofListLift, ofListLift, ← map_mul]
|
||||
congr
|
||||
rw [← ofList_pair, one_mul]
|
||||
congr
|
||||
exact Eq.symm (List.eraseIdx_eq_take_drop_succ r n)
|
||||
|
||||
lemma superCommute_koszulOrder_le_ofList [IsTotal 𝓕 le] [IsTrans 𝓕 le] (r : List 𝓕) (x : ℂ)
|
||||
(i : 𝓕) {A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(F : FreeAlgebra ℂ 𝓕 →ₐ A) [OperatorMap q le F] :
|
||||
F ((superCommute q (FreeAlgebra.ι ℂ i) (koszulOrder q le (ofList r x)))) =
|
||||
∑ n : Fin r.length, (superCommuteCoef q [r.get n] (r.take n)) •
|
||||
(F (((superCommute q) (ofList [i] 1)) (FreeAlgebra.ι ℂ (r.get n))) *
|
||||
F ((koszulOrder q le) (ofList (r.eraseIdx ↑n) x))) := by
|
||||
rw [koszulOrder_ofList, map_smul, map_smul, ← ofList_singleton, superCommute_ofList_sum]
|
||||
rw [map_sum, ← (HepLean.List.insertionSortEquiv le r).sum_comp]
|
||||
conv_lhs =>
|
||||
enter [2, 2]
|
||||
intro n
|
||||
rw [superCommuteSplit_operatorMap (le := le)]
|
||||
enter [1, 2, 2, 2]
|
||||
change ((List.insertionSort le r).get ∘ (HepLean.List.insertionSortEquiv le r)) n
|
||||
rw [HepLean.List.insertionSort_get_comp_insertionSortEquiv]
|
||||
conv_lhs =>
|
||||
enter [2, 2]
|
||||
intro n
|
||||
rw [HepLean.List.eraseIdx_insertionSort_fin le r n]
|
||||
rw [ofList_insertionSort_eq_koszulOrder q le]
|
||||
rw [Finset.smul_sum]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
intro n
|
||||
rw [map_smul, smul_smul, Algebra.mul_smul_comm, smul_smul]
|
||||
congr
|
||||
funext n
|
||||
by_cases hq : q i ≠ q (r.get n)
|
||||
· have hn := OperatorMap.superCommute_diff_grade_zero (q := q) (F := F) le i (r.get n) hq
|
||||
conv_lhs =>
|
||||
enter [2, 1]
|
||||
rw [ofList_singleton, hn]
|
||||
conv_rhs =>
|
||||
enter [2, 1]
|
||||
rw [ofList_singleton, hn]
|
||||
simp
|
||||
· congr 1
|
||||
trans staticWickCoef q le r i n
|
||||
· rw [staticWickCoef, mul_assoc]
|
||||
refine staticWickCoef_eq_get q le r i n ?_
|
||||
simpa using hq
|
||||
|
||||
lemma koszulOrder_of_le_all_ofList (r : List 𝓕) (x : ℂ) (i : 𝓕)
|
||||
{A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(F : FreeAlgebra ℂ 𝓕 →ₐ A) [OperatorMap q le F] :
|
||||
F (koszulOrder q le (ofList r x * FreeAlgebra.ι ℂ i))
|
||||
= superCommuteCoef q [i] r • F (koszulOrder q le (FreeAlgebra.ι ℂ i * ofList r x)) := by
|
||||
conv_lhs =>
|
||||
enter [2, 2]
|
||||
rw [← ofList_singleton]
|
||||
rw [ofListLift_ofList_superCommute' q]
|
||||
rw [map_sub]
|
||||
rw [sub_eq_add_neg]
|
||||
rw [map_add]
|
||||
conv_lhs =>
|
||||
enter [2, 2]
|
||||
rw [map_smul]
|
||||
rw [← neg_smul]
|
||||
rw [map_smul, map_smul, map_smul]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
rw [superCommute_ofList_sum]
|
||||
rw [map_sum, map_sum]
|
||||
dsimp [superCommuteSplit]
|
||||
rw [ofList_singleton]
|
||||
rhs
|
||||
intro n
|
||||
rw [Algebra.smul_mul_assoc, Algebra.smul_mul_assoc]
|
||||
rw [map_smul, map_smul]
|
||||
rw [OperatorMap.superCommute_ordered_zero]
|
||||
simp only [smul_zero, Finset.sum_const_zero, add_zero]
|
||||
rw [ofList_singleton]
|
||||
|
||||
lemma le_all_mul_koszulOrder_ofList (r : List 𝓕) (x : ℂ)
|
||||
(i : 𝓕) (hi : ∀ (j : 𝓕), le j i)
|
||||
{A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(F : FreeAlgebra ℂ 𝓕 →ₐ A) [OperatorMap q le F] :
|
||||
F (FreeAlgebra.ι ℂ i * koszulOrder q le (ofList r x)) =
|
||||
F ((koszulOrder q le) (FreeAlgebra.ι ℂ i * ofList r x)) +
|
||||
F (((superCommute q) (ofList [i] 1)) ((koszulOrder q le) (ofList r x))) := by
|
||||
rw [koszulOrder_ofList, Algebra.mul_smul_comm, map_smul, ← ofList_singleton,
|
||||
ofList_ofList_superCommute q, map_add, smul_add, ← map_smul]
|
||||
conv_lhs =>
|
||||
enter [1, 2]
|
||||
rw [← Algebra.smul_mul_assoc, smul_smul, mul_comm, ← smul_smul, ← koszulOrder_ofList,
|
||||
Algebra.smul_mul_assoc, ofList_singleton]
|
||||
rw [koszulOrder_mul_ge, map_smul]
|
||||
congr
|
||||
· rw [koszulOrder_of_le_all_ofList]
|
||||
rw [superCommuteCoef_perm_snd q [i] (List.insertionSort le r) r
|
||||
(List.perm_insertionSort le r)]
|
||||
rw [smul_smul]
|
||||
rw [superCommuteCoef_mul_self]
|
||||
simp [ofList_singleton]
|
||||
· rw [map_smul, map_smul]
|
||||
· exact fun j => hi j
|
||||
|
||||
/-- In the expansions of `F (FreeAlgebra.ι ℂ i * koszulOrder q le (ofList r x))`
|
||||
the ter multiplying `F ((koszulOrder q le) (ofList (optionEraseZ r i n) x))`. -/
|
||||
def superCommuteCenterOrder (r : List 𝓕) (i : 𝓕)
|
||||
{A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(F : FreeAlgebra ℂ 𝓕 →ₐ[ℂ] A)
|
||||
(n : Option (Fin r.length)) : A :=
|
||||
match n with
|
||||
| none => 1
|
||||
| some n => superCommuteCoef q [r.get n] (r.take n) • F (((superCommute q) (ofList [i] 1))
|
||||
(FreeAlgebra.ι ℂ (r.get n)))
|
||||
|
||||
@[simp]
|
||||
lemma superCommuteCenterOrder_none (r : List 𝓕) (i : 𝓕)
|
||||
{A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(F : FreeAlgebra ℂ 𝓕 →ₐ[ℂ] A) :
|
||||
superCommuteCenterOrder q r i F none = 1 := by
|
||||
simp [superCommuteCenterOrder]
|
||||
|
||||
open HepLean.List
|
||||
|
||||
lemma le_all_mul_koszulOrder_ofList_expand [IsTotal 𝓕 le] [IsTrans 𝓕 le] (r : List 𝓕) (x : ℂ)
|
||||
(i : 𝓕) (hi : ∀ (j : 𝓕), le j i)
|
||||
{A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(F : FreeAlgebra ℂ 𝓕 →ₐ[ℂ] A) [OperatorMap q le F] :
|
||||
F (FreeAlgebra.ι ℂ i * koszulOrder q le (ofList r x)) =
|
||||
∑ n, superCommuteCenterOrder q r i F n *
|
||||
F ((koszulOrder q le) (ofList (optionEraseZ r i n) x)) := by
|
||||
rw [le_all_mul_koszulOrder_ofList]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rw [ofList_singleton]
|
||||
rw [superCommute_koszulOrder_le_ofList]
|
||||
simp only [List.get_eq_getElem, Fintype.sum_option, superCommuteCenterOrder_none, one_mul]
|
||||
congr
|
||||
· rw [← ofList_singleton, ← ofList_pair]
|
||||
simp only [List.singleton_append, one_mul]
|
||||
rfl
|
||||
· funext n
|
||||
simp only [superCommuteCenterOrder, List.get_eq_getElem, Algebra.smul_mul_assoc]
|
||||
rfl
|
||||
exact fun j => hi j
|
||||
|
||||
lemma le_all_mul_koszulOrder_ofListLift_expand {f : 𝓕 → Type} [∀ i, Fintype (f i)]
|
||||
(r : List 𝓕) (x : ℂ)
|
||||
(le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le]
|
||||
[IsTotal (Σ i, f i) le] [IsTrans (Σ i, f i) le]
|
||||
(i : (Σ i, f i)) (hi : ∀ (j : (Σ i, f i)), le j i)
|
||||
{A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(F : FreeAlgebra ℂ (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le F] :
|
||||
F (ofList [i] 1 * koszulOrder (fun i => q i.1) le (ofListLift f r x)) =
|
||||
F ((koszulOrder (fun i => q i.fst) le) (ofList [i] 1 * ofListLift f r x)) +
|
||||
∑ n : (Fin r.length), superCommuteCoef q [r.get n] (List.take (↑n) r) •
|
||||
F (((superCommute fun i => q i.fst) (ofList [i] 1)) (ofListLift f [r.get n] 1)) *
|
||||
F ((koszulOrder (fun i => q i.fst) le) (ofListLift f (r.eraseIdx ↑n) x)) := by
|
||||
match r with
|
||||
| [] =>
|
||||
simp only [map_mul, List.length_nil, Finset.univ_eq_empty, List.get_eq_getElem, List.take_nil,
|
||||
List.eraseIdx_nil, Algebra.smul_mul_assoc, Finset.sum_empty, add_zero]
|
||||
rw [ofListLift_empty_smul]
|
||||
simp only [map_smul, koszulOrder_one, map_one, Algebra.mul_smul_comm, mul_one]
|
||||
rw [ofList_singleton, koszulOrder_ι]
|
||||
| r0 :: r =>
|
||||
rw [ofListLift_expand, map_sum, Finset.mul_sum, map_sum]
|
||||
let e1 (a : CreateAnnihilateSect f (r0 :: r)) :
|
||||
Option (Fin a.toList.length) ≃ Option (Fin (r0 :: r).length) :=
|
||||
Equiv.optionCongr (Fin.castOrderIso (CreateAnnihilateSect.toList_length a)).toEquiv
|
||||
conv_lhs =>
|
||||
rhs
|
||||
intro a
|
||||
rw [ofList_singleton, le_all_mul_koszulOrder_ofList_expand _ _ _ _ _ hi]
|
||||
rw [← (e1 a).symm.sum_comp]
|
||||
rhs
|
||||
intro n
|
||||
rw [Finset.sum_comm]
|
||||
simp only [Fintype.sum_option]
|
||||
congr 1
|
||||
· simp only [List.length_cons, List.get_eq_getElem, superCommuteCenterOrder,
|
||||
Equiv.optionCongr_symm, OrderIso.toEquiv_symm, Fin.symm_castOrderIso, Equiv.optionCongr_apply,
|
||||
RelIso.coe_fn_toEquiv, Option.map_none', optionEraseZ, one_mul, e1]
|
||||
rw [← map_sum, Finset.mul_sum, ← map_sum]
|
||||
apply congrArg
|
||||
apply congrArg
|
||||
congr
|
||||
funext x
|
||||
rw [ofList_cons_eq_ofList]
|
||||
· congr
|
||||
funext n
|
||||
rw [← (CreateAnnihilateSect.extractEquiv n).symm.sum_comp]
|
||||
simp only [List.get_eq_getElem, List.length_cons, Equiv.optionCongr_symm, OrderIso.toEquiv_symm,
|
||||
Fin.symm_castOrderIso, Equiv.optionCongr_apply, RelIso.coe_fn_toEquiv, Option.map_some',
|
||||
Fin.castOrderIso_apply, Algebra.smul_mul_assoc, e1]
|
||||
erw [Finset.sum_product]
|
||||
have h1 (a0 : f (r0 :: r)[↑n]) (a : CreateAnnihilateSect f ((r0 :: r).eraseIdx ↑n)) :
|
||||
superCommuteCenterOrder (fun i => q i.fst)
|
||||
((CreateAnnihilateSect.extractEquiv n).symm (a0, a)).toList i F
|
||||
(some (Fin.cast (by simp) n)) =
|
||||
superCommuteCoef q [(r0 :: r).get n] (List.take (↑n) (r0 :: r)) •
|
||||
F (((superCommute fun i => q i.fst) (ofList [i] 1))
|
||||
(FreeAlgebra.ι ℂ ⟨(r0 :: r).get n, a0⟩)) := by
|
||||
simp only [superCommuteCenterOrder, List.get_eq_getElem, List.length_cons, Fin.coe_cast]
|
||||
erw [CreateAnnihilateSect.extractEquiv_symm_toList_get_same]
|
||||
have hsc : superCommuteCoef (fun i => q i.fst) [⟨(r0 :: r).get n, a0⟩]
|
||||
(List.take (↑n) ((CreateAnnihilateSect.extractEquiv n).symm (a0, a)).toList) =
|
||||
superCommuteCoef q [(r0 :: r).get n] (List.take (↑n) ((r0 :: r))) := by
|
||||
simp only [superCommuteCoef, List.get_eq_getElem, List.length_cons, Fin.isValue,
|
||||
CreateAnnihilateSect.toList_grade_take]
|
||||
rfl
|
||||
erw [hsc]
|
||||
rfl
|
||||
conv_lhs =>
|
||||
rhs
|
||||
intro a0
|
||||
rhs
|
||||
intro a
|
||||
erw [h1]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
intro a0
|
||||
rw [← Finset.mul_sum]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
intro a0
|
||||
enter [2, 2]
|
||||
intro a
|
||||
simp [optionEraseZ]
|
||||
enter [2, 2, 1]
|
||||
rw [← CreateAnnihilateSect.eraseIdx_toList]
|
||||
erw [CreateAnnihilateSect.extractEquiv_symm_eraseIdx]
|
||||
rw [← Finset.sum_mul]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
rw [← Finset.smul_sum]
|
||||
erw [← map_sum, ← map_sum, ← ofListLift_singleton_one]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rw [← map_sum, ← map_sum]
|
||||
simp only [List.get_eq_getElem, List.length_cons, Equiv.symm_apply_apply,
|
||||
Algebra.smul_mul_assoc]
|
||||
erw [← ofListLift_expand]
|
||||
simp only [List.get_eq_getElem, List.length_cons, Algebra.smul_mul_assoc]
|
||||
|
||||
end
|
||||
end Wick
|
|
@ -1,135 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.Mathematics.List
|
||||
import HepLean.PerturbationTheory.Wick.Signs.SuperCommuteCoef
|
||||
/-!
|
||||
|
||||
# Insert sign
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
open HepLean.List
|
||||
open FieldStatistic
|
||||
|
||||
section
|
||||
/-!
|
||||
|
||||
## Basic properties of lists
|
||||
|
||||
To be replaced with Mathlib or Lean definitions when/where appropriate.
|
||||
-/
|
||||
|
||||
lemma take_insert_same {I : Type} (i : I) :
|
||||
(n : ℕ) → (r : List I) →
|
||||
List.take n (List.insertIdx n i r) = List.take n r
|
||||
| 0, _ => by simp
|
||||
| _+1, [] => by simp
|
||||
| n+1, a::as => by
|
||||
simp only [List.insertIdx_succ_cons, List.take_succ_cons, List.cons.injEq, true_and]
|
||||
exact take_insert_same i n as
|
||||
|
||||
lemma take_eraseIdx_same {I : Type} :
|
||||
(n : ℕ) → (r : List I) →
|
||||
List.take n (List.eraseIdx r n) = List.take n r
|
||||
| 0, _ => by simp
|
||||
| _+1, [] => by simp
|
||||
| n+1, a::as => by
|
||||
simp only [List.eraseIdx_cons_succ, List.take_succ_cons, List.cons.injEq, true_and]
|
||||
exact take_eraseIdx_same n as
|
||||
|
||||
lemma take_insert_gt {I : Type} (i : I) :
|
||||
(n m : ℕ) → (h : n < m) → (r : List I) →
|
||||
List.take n (List.insertIdx m i r) = List.take n r
|
||||
| 0, 0, _, _ => by simp
|
||||
| 0, m + 1, _, _ => by simp
|
||||
| n+1, m + 1, _, [] => by simp
|
||||
| n+1, m + 1, h, a::as => by
|
||||
simp only [List.insertIdx_succ_cons, List.take_succ_cons, List.cons.injEq, true_and]
|
||||
refine take_insert_gt i n m (Nat.succ_lt_succ_iff.mp h) as
|
||||
|
||||
lemma take_insert_let {I : Type} (i : I) :
|
||||
(n m : ℕ) → (h : m ≤ n) → (r : List I) → (hm : m ≤ r.length) →
|
||||
(List.take (n + 1) (List.insertIdx m i r)).Perm (i :: List.take n r)
|
||||
| 0, 0, h, _, _ => by simp
|
||||
| m + 1, 0, h, r, _ => by simp
|
||||
| n + 1, m + 1, h, [], hm => by
|
||||
simp at hm
|
||||
| n + 1, m + 1, h, a::as, hm => by
|
||||
simp only [List.insertIdx_succ_cons, List.take_succ_cons]
|
||||
have hp : (i :: a :: List.take n as).Perm (a :: i :: List.take n as) := by
|
||||
exact List.Perm.swap a i (List.take n as)
|
||||
refine List.Perm.trans ?_ hp.symm
|
||||
refine List.Perm.cons a ?_
|
||||
exact take_insert_let i n m (Nat.le_of_succ_le_succ h) as (Nat.le_of_succ_le_succ hm)
|
||||
|
||||
end
|
||||
|
||||
/-!
|
||||
|
||||
## Insert sign
|
||||
|
||||
-/
|
||||
|
||||
section InsertSign
|
||||
|
||||
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic)
|
||||
|
||||
/-- The sign associated with inserting a field `φ` into a list of fields `φs` at
|
||||
the `n`th position. -/
|
||||
def insertSign (n : ℕ) (φ : 𝓕) (φs : List 𝓕) : ℂ :=
|
||||
superCommuteCoef q [φ] (List.take n φs)
|
||||
|
||||
/-- If `φ` is bosonic, there is no sign associated with inserting it into a list of fields. -/
|
||||
lemma insertSign_bosonic (n : ℕ) (φ : 𝓕) (φs : List 𝓕) (hφ : q φ = bosonic) :
|
||||
insertSign q n φ φs = 1 := by
|
||||
simp only [insertSign, superCommuteCoef, ofList_singleton, hφ, reduceCtorEq, false_and,
|
||||
↓reduceIte]
|
||||
|
||||
lemma insertSign_insert (n : ℕ) (φ : 𝓕) (φs : List 𝓕) :
|
||||
insertSign q n φ φs = insertSign q n φ (List.insertIdx n φ φs) := by
|
||||
simp only [insertSign]
|
||||
congr 1
|
||||
rw [take_insert_same]
|
||||
|
||||
lemma insertSign_eraseIdx (n : ℕ) (φ : 𝓕) (φs : List 𝓕) :
|
||||
insertSign q n φ (φs.eraseIdx n) = insertSign q n φ φs := by
|
||||
simp only [insertSign]
|
||||
congr 1
|
||||
rw [take_eraseIdx_same]
|
||||
|
||||
lemma insertSign_zero (φ : 𝓕) (φs : List 𝓕) : insertSign q 0 φ φs = 1 := by
|
||||
simp [insertSign, superCommuteCoef]
|
||||
|
||||
lemma insertSign_succ_cons (n : ℕ) (φ0 φ1 : 𝓕) (φs : List 𝓕) : insertSign q (n + 1) φ0 (φ1 :: φs) =
|
||||
superCommuteCoef q [φ0] [φ1] * insertSign q n φ0 φs := by
|
||||
simp only [insertSign, List.take_succ_cons]
|
||||
rw [superCommuteCoef_cons]
|
||||
|
||||
lemma insertSign_insert_gt (n m : ℕ) (φ0 φ1 : 𝓕) (φs : List 𝓕) (hn : n < m) :
|
||||
insertSign q n φ0 (List.insertIdx m φ1 φs) = insertSign q n φ0 φs := by
|
||||
rw [insertSign, insertSign]
|
||||
congr 1
|
||||
exact take_insert_gt φ1 n m hn φs
|
||||
|
||||
lemma insertSign_insert_lt_eq_insertSort (n m : ℕ) (φ0 φ1 : 𝓕) (φs : List 𝓕) (hn : m ≤ n)
|
||||
(hm : m ≤ φs.length) :
|
||||
insertSign q (n + 1) φ0 (List.insertIdx m φ1 φs) = insertSign q (n + 1) φ0 (φ1 :: φs) := by
|
||||
rw [insertSign, insertSign]
|
||||
apply superCommuteCoef_perm_snd
|
||||
simp only [List.take_succ_cons]
|
||||
refine take_insert_let φ1 n m hn φs hm
|
||||
|
||||
lemma insertSign_insert_lt (n m : ℕ) (φ0 φ1 : 𝓕) (φs : List 𝓕) (hn : m ≤ n) (hm : m ≤ φs.length) :
|
||||
insertSign q (n + 1) φ0 (List.insertIdx m φ1 φs) = superCommuteCoef q [φ0] [φ1] *
|
||||
insertSign q n φ0 φs := by
|
||||
rw [insertSign_insert_lt_eq_insertSort, insertSign_succ_cons]
|
||||
· exact hn
|
||||
· exact hm
|
||||
|
||||
end InsertSign
|
||||
|
||||
end Wick
|
|
@ -1,90 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.Wick.Signs.KoszulSign
|
||||
/-!
|
||||
|
||||
# Static wick coefficient
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
|
||||
open HepLean.List
|
||||
|
||||
open FieldStatistic
|
||||
|
||||
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic) (le :𝓕 → 𝓕 → Prop) [DecidableRel le]
|
||||
|
||||
/-- The sign that appears in the static version of Wicks theorem.
|
||||
This is actually equal to `superCommuteCoef q [r.get n] (r.take n)`, something
|
||||
which will be proved in a lemma. -/
|
||||
def staticWickCoef (r : List 𝓕) (i : 𝓕) (n : Fin r.length) : ℂ :=
|
||||
koszulSign q le r *
|
||||
superCommuteCoef q [i] (List.take (↑((HepLean.List.insertionSortEquiv le r) n))
|
||||
(List.insertionSort le r)) *
|
||||
koszulSign q le (r.eraseIdx ↑n)
|
||||
|
||||
lemma staticWickCoef_eq_q (r : List 𝓕) (i : 𝓕) (n : Fin r.length)
|
||||
(hq : q i = q (r.get n)) :
|
||||
staticWickCoef q le r i n =
|
||||
koszulSign q le r *
|
||||
superCommuteCoef q [r.get n] (List.take (↑(insertionSortEquiv le r n))
|
||||
(List.insertionSort le r)) *
|
||||
koszulSign q le (r.eraseIdx ↑n) := by
|
||||
simp [staticWickCoef, superCommuteCoef, ofList, hq]
|
||||
|
||||
lemma insertIdx_eraseIdx {I : Type} : (n : ℕ) → (r : List I) → (hn : n < r.length) →
|
||||
List.insertIdx n (r.get ⟨n, hn⟩) (r.eraseIdx n) = r
|
||||
| n, [], hn => by
|
||||
simp at hn
|
||||
| 0, r0 :: r, hn => by
|
||||
simp
|
||||
| n + 1, r0 :: r, hn => by
|
||||
simp only [List.length_cons, List.get_eq_getElem, List.getElem_cons_succ,
|
||||
List.eraseIdx_cons_succ, List.insertIdx_succ_cons, List.cons.injEq, true_and]
|
||||
exact insertIdx_eraseIdx n r _
|
||||
|
||||
lemma staticWickCoef_eq_get [IsTotal 𝓕 le] [IsTrans 𝓕 le] (r : List 𝓕) (i : 𝓕) (n : Fin r.length)
|
||||
(heq : q i = q (r.get n)) :
|
||||
staticWickCoef q le r i n = superCommuteCoef q [r.get n] (r.take n) := by
|
||||
rw [staticWickCoef_eq_q]
|
||||
let r' := r.eraseIdx ↑n
|
||||
have hr : List.insertIdx n (r.get n) (r.eraseIdx n) = r := by
|
||||
exact insertIdx_eraseIdx n.1 r n.prop
|
||||
conv_lhs =>
|
||||
lhs
|
||||
lhs
|
||||
rw [← hr]
|
||||
rw [koszulSign_insertIdx q le (r.get n) ((r.eraseIdx ↑n)) n (by
|
||||
rw [List.length_eraseIdx]
|
||||
simp only [Fin.is_lt, ↓reduceIte]
|
||||
omega)]
|
||||
rhs
|
||||
rhs
|
||||
rw [hr]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
lhs
|
||||
rhs
|
||||
enter [2, 1, 1]
|
||||
rw [insertionSortEquiv_congr _ _ hr]
|
||||
simp only [List.get_eq_getElem, Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
|
||||
Fin.cast_mk, Fin.eta, Fin.coe_cast]
|
||||
conv_lhs =>
|
||||
lhs
|
||||
rw [mul_assoc]
|
||||
rhs
|
||||
rw [insertSign]
|
||||
rw [superCommuteCoef_mul_self]
|
||||
simp only [mul_one]
|
||||
rw [mul_assoc]
|
||||
rw [koszulSign_mul_self]
|
||||
simp only [mul_one]
|
||||
rw [insertSign_eraseIdx]
|
||||
rfl
|
||||
exact heq
|
||||
|
||||
end Wick
|
|
@ -1,96 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.FieldStatistics
|
||||
/-!
|
||||
|
||||
# Super commutation coefficient.
|
||||
|
||||
This is a complex number which is `-1` when commuting two fermionic operators and `1` otherwise.
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
open FieldStatistic
|
||||
|
||||
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic)
|
||||
|
||||
/-- Given two lists `la` and `lb` returns `-1` if they are both of grade `1` and
|
||||
`1` otherwise. This corresponds to the sign associated with the super commutator
|
||||
when commuting `la` and `lb` in the free algebra.
|
||||
In terms of physics it is `-1` if commuting two fermionic operators and `1` otherwise. -/
|
||||
def superCommuteCoef (la lb : List 𝓕) : ℂ :=
|
||||
if FieldStatistic.ofList q la = fermionic ∧
|
||||
FieldStatistic.ofList q lb = fermionic then - 1 else 1
|
||||
|
||||
lemma superCommuteCoef_comm (la lb : List 𝓕) :
|
||||
superCommuteCoef q la lb = superCommuteCoef q lb la := by
|
||||
simp only [superCommuteCoef, Fin.isValue]
|
||||
congr 1
|
||||
exact Eq.propIntro (fun a => id (And.symm a)) fun a => id (And.symm a)
|
||||
|
||||
/-- Given a list `l : List (Σ i, f i)` and a list `r : List I` returns `-1` if the
|
||||
grade of `l` is `1` and the grade of `r` is `1` and `1` otherwise. This corresponds
|
||||
to the sign associated with the super commutator when commuting
|
||||
the lift of `l` and `r` (by summing over fibers) in the
|
||||
free algebra over `Σ i, f i`.
|
||||
In terms of physics it is `-1` if commuting two fermionic operators and `1` otherwise. -/
|
||||
def superCommuteLiftCoef {f : 𝓕 → Type} (φs : List (Σ i, f i)) (φs' : List 𝓕) : ℂ :=
|
||||
(if FieldStatistic.ofList (fun i => q i.fst) φs = fermionic ∧
|
||||
FieldStatistic.ofList q φs' = fermionic then -1 else 1)
|
||||
|
||||
lemma superCommuteLiftCoef_empty {f : 𝓕 → Type} (φs : List (Σ i, f i)) :
|
||||
superCommuteLiftCoef q φs [] = 1 := by
|
||||
simp [superCommuteLiftCoef]
|
||||
|
||||
lemma superCommuteCoef_perm_snd (la lb lb' : List 𝓕)
|
||||
(h : lb.Perm lb') :
|
||||
superCommuteCoef q la lb = superCommuteCoef q la lb' := by
|
||||
rw [superCommuteCoef, superCommuteCoef, FieldStatistic.ofList_perm q h]
|
||||
|
||||
lemma superCommuteCoef_mul_self (l lb : List 𝓕) :
|
||||
superCommuteCoef q l lb * superCommuteCoef q l lb = 1 := by
|
||||
simp only [superCommuteCoef, Fin.isValue, mul_ite, mul_neg, mul_one]
|
||||
have ha (a b : FieldStatistic) : (if a = fermionic ∧ b = fermionic then
|
||||
-if a = fermionic ∧ b = fermionic then -1 else 1
|
||||
else if a = fermionic ∧ b = fermionic then -1 else 1) = (1 : ℂ) := by
|
||||
fin_cases a <;> fin_cases b
|
||||
any_goals rfl
|
||||
simp
|
||||
exact ha (FieldStatistic.ofList q l) (FieldStatistic.ofList q lb)
|
||||
|
||||
lemma superCommuteCoef_empty (la : List 𝓕) :
|
||||
superCommuteCoef q la [] = 1 := by
|
||||
simp only [superCommuteCoef, ofList_empty, reduceCtorEq, and_false, ↓reduceIte]
|
||||
|
||||
lemma superCommuteCoef_append (la lb lc : List 𝓕) :
|
||||
superCommuteCoef q la (lb ++ lc) = superCommuteCoef q la lb * superCommuteCoef q la lc := by
|
||||
simp only [superCommuteCoef, Fin.isValue, ofList_append, ite_eq_right_iff, zero_ne_one, imp_false,
|
||||
mul_ite, mul_neg, mul_one]
|
||||
by_cases hla : ofList q la = fermionic
|
||||
· by_cases hlb : ofList q lb = fermionic
|
||||
· by_cases hlc : ofList q lc = fermionic
|
||||
· simp [hlc, hlb, hla]
|
||||
· have hc : ofList q lc = bosonic := by
|
||||
exact (neq_fermionic_iff_eq_bosonic (ofList q lc)).mp hlc
|
||||
simp [hc, hlb, hla]
|
||||
· have hb : ofList q lb = bosonic := by
|
||||
exact (neq_fermionic_iff_eq_bosonic (ofList q lb)).mp hlb
|
||||
by_cases hlc : ofList q lc = fermionic
|
||||
· simp [hlc, hb]
|
||||
· have hc : ofList q lc = bosonic := by
|
||||
exact (neq_fermionic_iff_eq_bosonic (ofList q lc)).mp hlc
|
||||
simp [hc, hb]
|
||||
· have ha : ofList q la = bosonic := by
|
||||
exact (neq_fermionic_iff_eq_bosonic (ofList q la)).mp hla
|
||||
simp [ha]
|
||||
|
||||
lemma superCommuteCoef_cons (i : 𝓕) (la lb : List 𝓕) :
|
||||
superCommuteCoef q la (i :: lb) = superCommuteCoef q la [i] * superCommuteCoef q la lb := by
|
||||
trans superCommuteCoef q la ([i] ++ lb)
|
||||
simp only [List.singleton_append]
|
||||
rw [superCommuteCoef_append]
|
||||
|
||||
end Wick
|
|
@ -1,97 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.Contractions.Basic
|
||||
/-!
|
||||
|
||||
# Static Wick's theorem
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
|
||||
noncomputable section
|
||||
|
||||
open HepLean.List
|
||||
open FieldStatistic
|
||||
|
||||
variable {𝓕 : Type} {f : 𝓕 → Type} [∀ i, Fintype (f i)] (q : 𝓕 → FieldStatistic)
|
||||
(le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le]
|
||||
|
||||
lemma static_wick_nil {A : Type} [Semiring A] [Algebra ℂ A]
|
||||
(F : FreeAlgebra ℂ (Σ i, f i) →ₐ A)
|
||||
(S : Contractions.Splitting f le) :
|
||||
F (ofListLift f [] 1) = ∑ c : Contractions [],
|
||||
c.toCenterTerm f q le F S *
|
||||
F (koszulOrder (fun i => q i.fst) le (ofListLift f c.uncontracted 1)) := by
|
||||
rw [← Contractions.nilEquiv.symm.sum_comp]
|
||||
simp only [Finset.univ_unique, PUnit.default_eq_unit, Contractions.nilEquiv, Equiv.coe_fn_symm_mk,
|
||||
Finset.sum_const, Finset.card_singleton, one_smul]
|
||||
dsimp only [Contractions.toCenterTerm, Contractions.uncontracted]
|
||||
simp [ofListLift_empty]
|
||||
|
||||
lemma static_wick_cons [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) × f i) le]
|
||||
{A : Type} [Semiring A] [Algebra ℂ A] (φs : List 𝓕) (φ : 𝓕)
|
||||
(F : FreeAlgebra ℂ (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le F]
|
||||
(S : Contractions.Splitting f le)
|
||||
(ih : F (ofListLift f φs 1) =
|
||||
∑ c : Contractions φs, c.toCenterTerm f q le F S * F (koszulOrder (fun i => q i.fst) le
|
||||
(ofListLift f c.uncontracted 1))) :
|
||||
F (ofListLift f (φ :: φs) 1) = ∑ c : Contractions (φ :: φs),
|
||||
c.toCenterTerm f q le F S *
|
||||
F (koszulOrder (fun i => q i.fst) le (ofListLift f c.uncontracted 1)) := by
|
||||
rw [ofListLift_cons_eq_ofListLift, map_mul, ih, Finset.mul_sum,
|
||||
← Contractions.consEquiv.symm.sum_comp]
|
||||
erw [Finset.sum_sigma]
|
||||
congr
|
||||
funext c
|
||||
have hb := S.h𝓑 φ
|
||||
rw [← mul_assoc]
|
||||
have hi := c.toCenterTerm_center f q le F S
|
||||
rw [Subalgebra.mem_center_iff] at hi
|
||||
rw [hi, mul_assoc, ← map_mul, hb, add_mul, map_add]
|
||||
conv_lhs =>
|
||||
enter [2, 1]
|
||||
rw [ofList_eq_smul_one, Algebra.smul_mul_assoc, ofList_singleton]
|
||||
rw [mul_koszulOrder_le]
|
||||
conv_lhs =>
|
||||
enter [2, 1]
|
||||
rw [← map_smul, ← Algebra.smul_mul_assoc]
|
||||
rw [← ofList_singleton, ← ofList_eq_smul_one]
|
||||
conv_lhs =>
|
||||
enter [2, 2]
|
||||
rw [ofList_eq_smul_one, Algebra.smul_mul_assoc, map_smul]
|
||||
rw [le_all_mul_koszulOrder_ofListLift_expand]
|
||||
conv_lhs =>
|
||||
enter [2, 2]
|
||||
rw [smul_add, Finset.smul_sum]
|
||||
rw [← map_smul, ← map_smul, ← Algebra.smul_mul_assoc, ← ofList_eq_smul_one]
|
||||
enter [2, 2]
|
||||
intro n
|
||||
rw [← Algebra.smul_mul_assoc, smul_comm, ← map_smul, ← LinearMap.map_smul₂,
|
||||
← ofList_eq_smul_one]
|
||||
rw [← add_assoc, ← map_add, ← map_add, ← add_mul, ← hb, ← ofListLift_cons_eq_ofListLift, mul_add]
|
||||
rw [Fintype.sum_option]
|
||||
congr 1
|
||||
rw [Finset.mul_sum]
|
||||
congr
|
||||
funext n
|
||||
rw [← mul_assoc]
|
||||
rfl
|
||||
exact S.h𝓑p φ
|
||||
exact S.h𝓑n φ
|
||||
|
||||
theorem static_wick_theorem [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) × f i) le]
|
||||
{A : Type} [Semiring A] [Algebra ℂ A] (φs : List 𝓕)
|
||||
(F : FreeAlgebra ℂ (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le F]
|
||||
(S : Contractions.Splitting f le) :
|
||||
F (ofListLift f φs 1) = ∑ c : Contractions φs, c.toCenterTerm f q le F S *
|
||||
F (koszulOrder (fun i => q i.fst) le (ofListLift f c.uncontracted 1)) := by
|
||||
induction φs with
|
||||
| nil => exact static_wick_nil q le F S
|
||||
| cons a r ih => exact static_wick_cons q le r a F S ih
|
||||
|
||||
end
|
||||
end Wick
|
|
@ -1,478 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 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.Wick.OfList
|
||||
/-!
|
||||
|
||||
# Koszul signs and ordering for lists and algebras
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
|
||||
noncomputable section
|
||||
open FieldStatistic
|
||||
|
||||
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic)
|
||||
|
||||
/-- Given a grading `q : I → Fin 2` and a list `l : List I` the super-commutor on the free algebra
|
||||
`FreeAlgebra ℂ I` corresponding to commuting with `l`
|
||||
as a linear map from `MonoidAlgebra ℂ (FreeMonoid I)` (the module of lists in `I`)
|
||||
to itself. -/
|
||||
def superCommuteMonoidAlgebra (l : List 𝓕) :
|
||||
MonoidAlgebra ℂ (FreeMonoid 𝓕) →ₗ[ℂ] MonoidAlgebra ℂ (FreeMonoid 𝓕) :=
|
||||
Finsupp.lift (MonoidAlgebra ℂ (FreeMonoid 𝓕)) ℂ (List 𝓕)
|
||||
(fun r =>
|
||||
Finsupp.lsingle (R := ℂ) (l ++ r) 1 +
|
||||
if FieldStatistic.ofList q l = fermionic ∧ FieldStatistic.ofList q r = fermionic then
|
||||
Finsupp.lsingle (R := ℂ) (r ++ l) 1
|
||||
else
|
||||
- Finsupp.lsingle (R := ℂ) (r ++ l) 1)
|
||||
|
||||
/-- Given a grading `q : I → Fin 2` the super-commutor on the free algebra `FreeAlgebra ℂ I`
|
||||
as a linear map from `MonoidAlgebra ℂ (FreeMonoid I)` (the module of lists in `I`)
|
||||
to `FreeAlgebra ℂ I →ₗ[ℂ] FreeAlgebra ℂ I`. -/
|
||||
def superCommuteAlgebra :
|
||||
MonoidAlgebra ℂ (FreeMonoid 𝓕) →ₗ[ℂ] FreeAlgebra ℂ 𝓕 →ₗ[ℂ] FreeAlgebra ℂ 𝓕 :=
|
||||
Finsupp.lift (FreeAlgebra ℂ 𝓕 →ₗ[ℂ] FreeAlgebra ℂ 𝓕) ℂ (List 𝓕) fun l =>
|
||||
(FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm.toAlgHom.toLinearMap
|
||||
∘ₗ superCommuteMonoidAlgebra q l
|
||||
∘ₗ FreeAlgebra.equivMonoidAlgebraFreeMonoid.toAlgHom.toLinearMap)
|
||||
|
||||
/-- Given a grading `q : I → Fin 2` the super-commutor on the free algebra `FreeAlgebra ℂ I`
|
||||
as a bi-linear map. -/
|
||||
def superCommute :
|
||||
FreeAlgebra ℂ 𝓕 →ₗ[ℂ] FreeAlgebra ℂ 𝓕 →ₗ[ℂ] FreeAlgebra ℂ 𝓕 :=
|
||||
superCommuteAlgebra q
|
||||
∘ₗ FreeAlgebra.equivMonoidAlgebraFreeMonoid.toAlgHom.toLinearMap
|
||||
|
||||
lemma equivMonoidAlgebraFreeMonoid_freeAlgebra {I : Type} (i : I) :
|
||||
(FreeAlgebra.equivMonoidAlgebraFreeMonoid (FreeAlgebra.ι ℂ i)) =
|
||||
Finsupp.single (FreeMonoid.of i) 1 := by
|
||||
simp [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.single]
|
||||
|
||||
@[simp]
|
||||
lemma superCommute_ι (i j : 𝓕) :
|
||||
superCommute q (FreeAlgebra.ι ℂ i) (FreeAlgebra.ι ℂ j) =
|
||||
FreeAlgebra.ι ℂ i * FreeAlgebra.ι ℂ j +
|
||||
if q i = fermionic ∧ q j = fermionic then
|
||||
FreeAlgebra.ι ℂ j * FreeAlgebra.ι ℂ i
|
||||
else
|
||||
- FreeAlgebra.ι ℂ j * FreeAlgebra.ι ℂ i := by
|
||||
simp only [superCommute, superCommuteAlgebra, AlgEquiv.toAlgHom_eq_coe,
|
||||
AlgEquiv.toAlgHom_toLinearMap, LinearMap.coe_comp, Function.comp_apply,
|
||||
AlgEquiv.toLinearMap_apply, equivMonoidAlgebraFreeMonoid_freeAlgebra, Fin.isValue, neg_mul]
|
||||
erw [Finsupp.lift_apply]
|
||||
simp only [superCommuteMonoidAlgebra, Finsupp.lsingle_apply, Fin.isValue, ofList_freeMonoid,
|
||||
zero_smul, Finsupp.sum_single_index, one_smul, LinearMap.coe_comp, Function.comp_apply,
|
||||
AlgEquiv.toLinearMap_apply, equivMonoidAlgebraFreeMonoid_freeAlgebra]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
erw [Finsupp.lift_apply]
|
||||
simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply, Fin.isValue,
|
||||
smul_add, MonoidAlgebra.smul_single', mul_one, smul_ite, smul_neg, Finsupp.sum_add,
|
||||
Finsupp.single_zero, Finsupp.sum_single_index, ofList_freeMonoid, neg_zero, ite_self,
|
||||
AlgEquiv.ofAlgHom_symm_apply, map_add, MonoidAlgebra.lift_single, one_smul]
|
||||
congr
|
||||
by_cases hq : q i = fermionic ∧ q j = fermionic
|
||||
· rw [if_pos hq, if_pos hq]
|
||||
simp only [MonoidAlgebra.lift_single, one_smul]
|
||||
obtain ⟨left, right⟩ := hq
|
||||
rfl
|
||||
· rw [if_neg hq, if_neg hq]
|
||||
simp only [map_neg, MonoidAlgebra.lift_single, one_smul, neg_inj]
|
||||
rfl
|
||||
|
||||
lemma superCommute_ofList_ofList (l r : List 𝓕) (x y : ℂ) :
|
||||
superCommute q (ofList l x) (ofList r y) =
|
||||
ofList (l ++ r) (x * y) + (if FieldStatistic.ofList q l = fermionic ∧
|
||||
FieldStatistic.ofList q r = fermionic then
|
||||
ofList (r ++ l) (y * x) else - ofList (r ++ l) (y * x)) := by
|
||||
simp only [superCommute, superCommuteAlgebra, AlgEquiv.toAlgHom_eq_coe,
|
||||
AlgEquiv.toAlgHom_toLinearMap, ofList, LinearMap.coe_comp, Function.comp_apply,
|
||||
AlgEquiv.toLinearMap_apply, AlgEquiv.apply_symm_apply, Fin.isValue]
|
||||
erw [Finsupp.lift_apply]
|
||||
simp only [superCommuteMonoidAlgebra, Finsupp.lsingle_apply, Fin.isValue, zero_smul,
|
||||
Finsupp.sum_single_index, LinearMap.smul_apply, LinearMap.coe_comp, Function.comp_apply,
|
||||
AlgEquiv.toLinearMap_apply, AlgEquiv.apply_symm_apply]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
erw [Finsupp.lift_apply]
|
||||
simp only [Fin.isValue, smul_add, MonoidAlgebra.smul_single', mul_one, smul_ite, smul_neg,
|
||||
Finsupp.sum_add, Finsupp.single_zero, Finsupp.sum_single_index, neg_zero, ite_self, map_add]
|
||||
by_cases hg : FieldStatistic.ofList q l = fermionic ∧ FieldStatistic.ofList q r = fermionic
|
||||
· simp only [hg, and_self, ↓reduceIte]
|
||||
congr
|
||||
· rw [← map_smul]
|
||||
congr
|
||||
exact MonoidAlgebra.smul_single' x (l ++ r) y
|
||||
· rw [← map_smul]
|
||||
congr
|
||||
rw [mul_comm]
|
||||
exact MonoidAlgebra.smul_single' x (r ++ l) y
|
||||
· simp only [Fin.isValue, hg, ↓reduceIte, map_neg, smul_neg]
|
||||
congr
|
||||
· rw [← map_smul]
|
||||
congr
|
||||
exact MonoidAlgebra.smul_single' x (l ++ r) y
|
||||
· rw [← map_smul]
|
||||
congr
|
||||
rw [mul_comm]
|
||||
exact MonoidAlgebra.smul_single' x (r ++ l) y
|
||||
|
||||
@[simp]
|
||||
lemma superCommute_zero (a : FreeAlgebra ℂ 𝓕) :
|
||||
superCommute q a 0 = 0 := by
|
||||
simp [superCommute]
|
||||
|
||||
@[simp]
|
||||
lemma superCommute_one (a : FreeAlgebra ℂ 𝓕) :
|
||||
superCommute q a 1 = 0 := by
|
||||
let f : FreeAlgebra ℂ 𝓕 →ₗ[ℂ] FreeAlgebra ℂ 𝓕 := (LinearMap.flip (superCommute q)) 1
|
||||
have h1 : FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single [] 1) =
|
||||
(1 : FreeAlgebra ℂ 𝓕) := by
|
||||
simp_all only [EmbeddingLike.map_eq_one_iff]
|
||||
rfl
|
||||
have f_single (l : FreeMonoid 𝓕) (x : ℂ) :
|
||||
f ((FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single l x)))
|
||||
= 0 := by
|
||||
simp only [superCommute, superCommuteAlgebra, AlgEquiv.toAlgHom_eq_coe,
|
||||
AlgEquiv.toAlgHom_toLinearMap, LinearMap.flip_apply, LinearMap.coe_comp, Function.comp_apply,
|
||||
AlgEquiv.toLinearMap_apply, AlgEquiv.apply_symm_apply, f]
|
||||
rw [← h1]
|
||||
erw [Finsupp.lift_apply]
|
||||
simp only [superCommuteMonoidAlgebra, Finsupp.lsingle_apply, Fin.isValue, zero_smul,
|
||||
Finsupp.sum_single_index, LinearMap.smul_apply, LinearMap.coe_comp, Function.comp_apply,
|
||||
AlgEquiv.toLinearMap_apply, AlgEquiv.apply_symm_apply, smul_eq_zero,
|
||||
EmbeddingLike.map_eq_zero_iff]
|
||||
apply Or.inr
|
||||
conv_lhs =>
|
||||
erw [Finsupp.lift_apply]
|
||||
simp
|
||||
have hf : f = 0 := by
|
||||
let e : FreeAlgebra ℂ 𝓕 ≃ₗ[ℂ] MonoidAlgebra ℂ (FreeMonoid 𝓕) :=
|
||||
FreeAlgebra.equivMonoidAlgebraFreeMonoid.toLinearEquiv
|
||||
apply (LinearEquiv.eq_comp_toLinearMap_iff (e₁₂ := e.symm) _ _).mp
|
||||
apply MonoidAlgebra.lhom_ext'
|
||||
intro l
|
||||
apply LinearMap.ext
|
||||
intro x
|
||||
simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply,
|
||||
MonoidAlgebra.lsingle_apply, LinearMap.zero_comp, LinearMap.zero_apply]
|
||||
erw [f_single]
|
||||
change f a = _
|
||||
rw [hf]
|
||||
simp
|
||||
|
||||
lemma superCommute_ofList_mul (la lb lc : List 𝓕) (xa xb xc : ℂ) :
|
||||
superCommute q (ofList la xa) (ofList lb xb * ofList lc xc) =
|
||||
(superCommute q (ofList la xa) (ofList lb xb) * ofList lc xc +
|
||||
superCommuteCoef q la lb • ofList lb xb * superCommute q (ofList la xa) (ofList lc xc)) := by
|
||||
simp only [Algebra.smul_mul_assoc]
|
||||
conv_lhs => rw [← ofList_pair]
|
||||
simp only [superCommute_ofList_ofList, Fin.isValue, ofList_append, ite_eq_right_iff, zero_ne_one,
|
||||
imp_false]
|
||||
simp only [superCommute_ofList_ofList, Fin.isValue, ofList_append, ite_eq_right_iff, zero_ne_one,
|
||||
imp_false, ofList_triple_assoc, ofList_triple, ofList_pair, superCommuteCoef]
|
||||
by_cases hla : FieldStatistic.ofList q la = fermionic
|
||||
· simp only [hla, Fin.isValue, true_and, ite_not, ite_smul, neg_smul, one_smul]
|
||||
by_cases hlb : FieldStatistic.ofList q lb = fermionic
|
||||
· simp only [hlb, Fin.isValue, ↓reduceIte]
|
||||
by_cases hlc : FieldStatistic.ofList q lc = fermionic
|
||||
· simp only [hlc, reduceCtorEq, imp_false, not_true_eq_false, ↓reduceIte]
|
||||
simp only [mul_assoc, add_mul, mul_add]
|
||||
abel
|
||||
· have hc : FieldStatistic.ofList q lc = bosonic := by
|
||||
exact (neq_fermionic_iff_eq_bosonic (FieldStatistic.ofList q lc)).mp hlc
|
||||
simp only [hc, fermionic_not_eq_bonsic, reduceCtorEq, imp_self, ↓reduceIte]
|
||||
simp only [mul_assoc, add_mul, mul_add, mul_neg, neg_add_rev, neg_neg]
|
||||
abel
|
||||
· have hb : FieldStatistic.ofList q lb = bosonic := by
|
||||
exact (neq_fermionic_iff_eq_bosonic (FieldStatistic.ofList q lb)).mp hlb
|
||||
simp only [hb, Fin.isValue, zero_ne_one, ↓reduceIte]
|
||||
by_cases hlc : FieldStatistic.ofList q lc = fermionic
|
||||
· simp only [hlc, reduceCtorEq, imp_self, ↓reduceIte]
|
||||
simp only [mul_assoc, add_mul, neg_mul, mul_add]
|
||||
abel
|
||||
· have hc : FieldStatistic.ofList q lc = bosonic := by
|
||||
exact (neq_fermionic_iff_eq_bosonic (FieldStatistic.ofList q lc)).mp hlc
|
||||
simp only [hc, reduceCtorEq, imp_false, not_true_eq_false, ↓reduceIte]
|
||||
simp only [mul_assoc, add_mul, neg_mul, mul_add, mul_neg]
|
||||
abel
|
||||
· simp only [Fin.isValue, hla, false_and, ↓reduceIte, mul_assoc, add_mul, neg_mul, mul_add,
|
||||
mul_neg, smul_add, one_smul, smul_neg]
|
||||
abel
|
||||
|
||||
/-- Given two lists `la lb : List I`, in the expansion of the supercommutor of `la` and `lb`
|
||||
via elements of `lb`the term associated with the `n`th element.
|
||||
E.g. in the commutator
|
||||
`[a, bc] = [a, b] c + b [a, c] ` the `superCommuteSplit` for `n=0` is `[a, b] c`
|
||||
and for `n=1` is `b [a, c]`. -/
|
||||
def superCommuteSplit (la lb : List 𝓕) (xa xb : ℂ) (n : ℕ)
|
||||
(hn : n < lb.length) : FreeAlgebra ℂ 𝓕 :=
|
||||
superCommuteCoef q la (List.take n lb) •
|
||||
ofList (List.take n lb) 1 *
|
||||
superCommute q (ofList la xa) (FreeAlgebra.ι ℂ (lb.get ⟨n, hn⟩))
|
||||
* ofList (List.drop (n + 1) lb) xb
|
||||
|
||||
lemma superCommute_ofList_cons (la lb : List 𝓕) (xa xb : ℂ) (b1 : 𝓕) :
|
||||
superCommute q (ofList la xa) (ofList (b1 :: lb) xb) =
|
||||
superCommute q (ofList la xa) (FreeAlgebra.ι ℂ b1) * ofList lb xb +
|
||||
superCommuteCoef q la [b1] •
|
||||
(ofList [b1] 1) * superCommute q (ofList la xa) (ofList lb xb) := by
|
||||
rw [ofList_cons_eq_ofList]
|
||||
rw [superCommute_ofList_mul]
|
||||
congr
|
||||
· exact ofList_singleton b1
|
||||
|
||||
lemma superCommute_ofList_sum (la lb : List 𝓕) (xa xb : ℂ) :
|
||||
superCommute q (ofList la xa) (ofList lb xb) =
|
||||
∑ (n : Fin lb.length), superCommuteSplit q la lb xa xb n n.prop := by
|
||||
induction lb with
|
||||
| nil =>
|
||||
simp only [superCommute_ofList_ofList, List.append_nil, FieldStatistic.ofList_empty,
|
||||
reduceCtorEq, and_false, ↓reduceIte, List.nil_append, List.length_nil, Finset.univ_eq_empty,
|
||||
Finset.sum_empty]
|
||||
ring_nf
|
||||
abel
|
||||
| cons b lb ih =>
|
||||
rw [superCommute_ofList_cons, ih]
|
||||
have h0 : ((superCommute q) (ofList la xa)) (FreeAlgebra.ι ℂ b) * ofList lb xb =
|
||||
superCommuteSplit q la (b :: lb) xa xb 0 (Nat.zero_lt_succ lb.length) := by
|
||||
simp [superCommuteSplit, superCommuteCoef_empty, ofList_empty]
|
||||
rw [h0]
|
||||
have hf (f : Fin (b :: lb).length → FreeAlgebra ℂ 𝓕) : ∑ n, f n = f ⟨0,
|
||||
Nat.zero_lt_succ lb.length⟩ + ∑ n, f (Fin.succ n) := by
|
||||
exact Fin.sum_univ_succAbove f ⟨0, Nat.zero_lt_succ lb.length⟩
|
||||
rw [hf]
|
||||
congr
|
||||
rw [Finset.mul_sum]
|
||||
congr
|
||||
funext n
|
||||
simp only [superCommuteSplit, Fin.eta, List.get_eq_getElem, Algebra.smul_mul_assoc,
|
||||
Algebra.mul_smul_comm, smul_smul, List.length_cons, Fin.val_succ, List.take_succ_cons,
|
||||
List.getElem_cons_succ, List.drop_succ_cons]
|
||||
congr 1
|
||||
· rw [mul_comm, ← superCommuteCoef_append]
|
||||
rfl
|
||||
· simp only [← mul_assoc, mul_eq_mul_right_iff]
|
||||
exact Or.inl (Or.inl (ofList_cons_eq_ofList (List.take (↑n) lb) b 1).symm)
|
||||
|
||||
lemma superCommute_ofList_ofList_superCommuteCoef (la lb : List 𝓕)
|
||||
(xa xb : ℂ) : superCommute q (ofList la xa) (ofList lb xb) =
|
||||
ofList la xa * ofList lb xb - superCommuteCoef q la lb • ofList lb xb * ofList la xa := by
|
||||
rw [superCommute_ofList_ofList, superCommuteCoef]
|
||||
by_cases hq : FieldStatistic.ofList q la = fermionic ∧ FieldStatistic.ofList q lb = fermionic
|
||||
· simp [hq, ofList_pair]
|
||||
· simp only [ofList_pair, Fin.isValue, hq, ↓reduceIte, one_smul]
|
||||
abel
|
||||
|
||||
lemma ofList_ofList_superCommute (la lb : List 𝓕) (xa xb : ℂ) :
|
||||
ofList la xa * ofList lb xb = superCommuteCoef q la lb • ofList lb xb * ofList la xa
|
||||
+ superCommute q (ofList la xa) (ofList lb xb) := by
|
||||
rw [superCommute_ofList_ofList_superCommuteCoef]
|
||||
abel
|
||||
|
||||
lemma ofListLift_ofList_superCommute' (l : List 𝓕) (r : List 𝓕) (x y : ℂ) :
|
||||
ofList r y * ofList l x = superCommuteCoef q l r • (ofList l x * ofList r y)
|
||||
- superCommuteCoef q l r • superCommute q (ofList l x) (ofList r y) := by
|
||||
nth_rewrite 2 [ofList_ofList_superCommute q]
|
||||
rw [superCommuteCoef]
|
||||
by_cases hq : FieldStatistic.ofList q l = fermionic ∧ FieldStatistic.ofList q r = fermionic
|
||||
· simp [hq, superCommuteCoef]
|
||||
· simp [hq]
|
||||
|
||||
section lift
|
||||
|
||||
variable {𝓕 : Type} {f : 𝓕 → Type} [∀ i, Fintype (f i)] (q : 𝓕 → FieldStatistic)
|
||||
|
||||
lemma superCommute_ofList_ofListLift (l : List (Σ i, f i)) (r : List 𝓕) (x y : ℂ) :
|
||||
superCommute (fun i => q i.1) (ofList l x) (ofListLift f r y) =
|
||||
ofList l x * ofListLift f r y +
|
||||
(if FieldStatistic.ofList (fun i => q i.1) l = fermionic ∧
|
||||
FieldStatistic.ofList q r = fermionic then
|
||||
ofListLift f r y * ofList l x else - ofListLift f r y * ofList l x) := by
|
||||
conv_lhs => rw [ofListLift_expand]
|
||||
rw [map_sum]
|
||||
conv_rhs =>
|
||||
lhs
|
||||
rw [ofListLift_expand, Finset.mul_sum]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rhs
|
||||
rw [ofListLift_expand, ← Finset.sum_neg_distrib, Finset.sum_mul]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
lhs
|
||||
rw [ofListLift_expand, Finset.sum_mul]
|
||||
rw [← Finset.sum_ite_irrel]
|
||||
rw [← Finset.sum_add_distrib]
|
||||
congr
|
||||
funext a
|
||||
rw [superCommute_ofList_ofList]
|
||||
congr 1
|
||||
· exact ofList_pair l a.toList x y
|
||||
congr 1
|
||||
· simp
|
||||
· exact ofList_pair a.toList l y x
|
||||
· rw [ofList_pair]
|
||||
simp only [neg_mul]
|
||||
|
||||
lemma superCommute_ofList_ofListLift_superCommuteLiftCoef
|
||||
(l : List (Σ i, f i)) (r : List 𝓕) (x y : ℂ) :
|
||||
superCommute (fun i => q i.1) (ofList l x) (ofListLift f r y) =
|
||||
ofList l x * ofListLift f r y - superCommuteLiftCoef q l r • ofListLift f r y * ofList l x := by
|
||||
rw [superCommute_ofList_ofListLift, superCommuteLiftCoef]
|
||||
by_cases hq : FieldStatistic.ofList (fun i => q i.fst) l = fermionic ∧
|
||||
FieldStatistic.ofList q r = fermionic
|
||||
· simp [hq]
|
||||
· simp only [Fin.isValue, hq, ↓reduceIte, neg_mul, one_smul]
|
||||
abel
|
||||
|
||||
lemma ofList_ofListLift_superCommute (l : List (Σ i, f i)) (r : List 𝓕) (x y : ℂ) :
|
||||
ofList l x * ofListLift f r y = superCommuteLiftCoef q l r • ofListLift f r y * ofList l x
|
||||
+ superCommute (fun i => q i.1) (ofList l x) (ofListLift f r y) := by
|
||||
rw [superCommute_ofList_ofListLift_superCommuteLiftCoef]
|
||||
abel
|
||||
|
||||
lemma ofListLift_ofList_superCommute (l : List (Σ i, f i)) (r : List 𝓕) (x y : ℂ) :
|
||||
ofListLift f r y * ofList l x = superCommuteLiftCoef q l r • (ofList l x * ofListLift f r y)
|
||||
- superCommuteLiftCoef q l r •
|
||||
superCommute (fun i => q i.1) (ofList l x) (ofListLift f r y) := by
|
||||
rw [ofList_ofListLift_superCommute, superCommuteLiftCoef]
|
||||
by_cases hq : FieldStatistic.ofList (fun i => q i.fst) l = fermionic ∧
|
||||
FieldStatistic.ofList q r = fermionic
|
||||
· simp [hq]
|
||||
· simp [hq]
|
||||
|
||||
omit [(i : 𝓕) → Fintype (f i)] in
|
||||
lemma superCommuteLiftCoef_append (l : List (Σ i, f i)) (r1 r2 : List 𝓕) :
|
||||
superCommuteLiftCoef q l (r1 ++ r2) =
|
||||
superCommuteLiftCoef q l r1 * superCommuteLiftCoef q l r2 := by
|
||||
simp only [superCommuteLiftCoef, Fin.isValue, ofList_append, ite_eq_right_iff, zero_ne_one,
|
||||
imp_false, mul_ite, mul_neg, mul_one]
|
||||
by_cases hla : FieldStatistic.ofList (fun i => q i.1) l = fermionic
|
||||
· by_cases hlb : FieldStatistic.ofList q r1 = fermionic
|
||||
· by_cases hlc : FieldStatistic.ofList q r2 = fermionic
|
||||
· simp [hlc, hlb, hla]
|
||||
· have hc : FieldStatistic.ofList q r2 = bosonic := by
|
||||
exact (neq_fermionic_iff_eq_bosonic (FieldStatistic.ofList q r2)).mp hlc
|
||||
simp [hc, hlb, hla]
|
||||
· have hb : FieldStatistic.ofList q r1 = bosonic := by
|
||||
exact (neq_fermionic_iff_eq_bosonic (FieldStatistic.ofList q r1)).mp hlb
|
||||
by_cases hlc : FieldStatistic.ofList q r2 = fermionic
|
||||
· simp [hlc, hb]
|
||||
· have hc : FieldStatistic.ofList q r2 = bosonic := by
|
||||
exact (neq_fermionic_iff_eq_bosonic (FieldStatistic.ofList q r2)).mp hlc
|
||||
simp [hc, hb]
|
||||
· have ha : FieldStatistic.ofList (fun i => q i.1) l = bosonic := by
|
||||
exact (neq_fermionic_iff_eq_bosonic (FieldStatistic.ofList (fun i => q i.fst) l)).mp hla
|
||||
simp [ha]
|
||||
|
||||
/-- Given two lists `l : List (Σ i, f i)` and `r : List I`, on
|
||||
in the expansion of the supercommutor of `l` and the lift of `r`
|
||||
via elements of `r`the term associated with the `n`th element.
|
||||
E.g. in the commutator
|
||||
`[a, bc] = [a, b] c + b [a, c] ` the `superCommuteSplit` for `n=0` is `[a, b] c`
|
||||
and for `n=1` is `b [a, c]`. -/
|
||||
def superCommuteLiftSplit (l : List (Σ i, f i)) (r : List 𝓕) (x y : ℂ) (n : ℕ)
|
||||
(hn : n < r.length) : FreeAlgebra ℂ (Σ i, f i) :=
|
||||
superCommuteLiftCoef q l (List.take n r) •
|
||||
(ofListLift f (List.take n r) 1 *
|
||||
superCommute (fun i => q i.1) (ofList l x) (sumFiber f (FreeAlgebra.ι ℂ (r.get ⟨n, hn⟩)))
|
||||
* ofListLift f (List.drop (n + 1) r) y)
|
||||
|
||||
lemma superCommute_ofList_ofListLift_cons (l : List (Σ i, f i)) (r : List 𝓕) (x y : ℂ) (b1 : 𝓕) :
|
||||
superCommute (fun i => q i.1) (ofList l x) (ofListLift f (b1 :: r) y) =
|
||||
superCommute (fun i => q i.1) (ofList l x) (sumFiber f (FreeAlgebra.ι ℂ b1))
|
||||
* ofListLift f r y + superCommuteLiftCoef q l [b1] •
|
||||
(ofListLift f [b1] 1) * superCommute (fun i => q i.1) (ofList l x) (ofListLift f r y) := by
|
||||
rw [ofListLift_cons]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rw [ofListLift_expand]
|
||||
rw [Finset.mul_sum]
|
||||
rw [map_sum]
|
||||
trans ∑ (n : CreateAnnihilateSect f r), ∑ j : f b1, ((superCommute fun i => q i.fst) (ofList l x))
|
||||
((FreeAlgebra.ι ℂ ⟨b1, j⟩) * ofList n.toList y)
|
||||
· apply congrArg
|
||||
funext n
|
||||
rw [← map_sum]
|
||||
congr
|
||||
rw [Finset.sum_mul]
|
||||
conv_rhs =>
|
||||
lhs
|
||||
rw [ofListLift_expand, Finset.mul_sum]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rhs
|
||||
rw [ofListLift_expand]
|
||||
rw [map_sum]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [Finset.mul_sum]
|
||||
rw [← Finset.sum_add_distrib]
|
||||
congr
|
||||
funext n
|
||||
rw [sumFiber_ι, map_sum, Finset.sum_mul]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rw [ofListLift_singleton]
|
||||
rw [Finset.smul_sum, Finset.sum_mul]
|
||||
rw [← Finset.sum_add_distrib]
|
||||
congr
|
||||
funext b
|
||||
trans ((superCommute fun i => q i.fst) (ofList l x)) (ofList (⟨b1, b⟩ :: n.toList) y)
|
||||
· congr
|
||||
rw [ofList_cons_eq_ofList]
|
||||
rw [ofList_singleton]
|
||||
rw [superCommute_ofList_cons]
|
||||
congr
|
||||
rw [ofList_singleton]
|
||||
simp
|
||||
|
||||
lemma superCommute_ofList_ofListLift_sum (l : List (Σ i, f i)) (r : List 𝓕) (x y : ℂ) :
|
||||
superCommute (fun i => q i.1) (ofList l x) (ofListLift f r y) =
|
||||
∑ (n : Fin r.length), superCommuteLiftSplit q l r x y n n.prop := by
|
||||
induction r with
|
||||
| nil =>
|
||||
simp only [superCommute_ofList_ofListLift, Fin.isValue, ofList_empty, zero_ne_one, and_false,
|
||||
↓reduceIte, neg_mul, List.length_nil, Finset.univ_eq_empty, Finset.sum_empty]
|
||||
rw [ofListLift, ofList_empty']
|
||||
simp
|
||||
| cons b r ih =>
|
||||
rw [superCommute_ofList_ofListLift_cons]
|
||||
have h0 : ((superCommute fun i => q i.fst) (ofList l x))
|
||||
((sumFiber f) (FreeAlgebra.ι ℂ b)) * ofListLift f r y =
|
||||
superCommuteLiftSplit q l (b :: r) x y 0 (Nat.zero_lt_succ r.length) := by
|
||||
simp [superCommuteLiftSplit, superCommuteLiftCoef_empty, ofListLift_empty]
|
||||
rw [h0]
|
||||
have hf (g : Fin (b :: r).length → FreeAlgebra ℂ ((i : 𝓕) × f i)) : ∑ n, g n = g ⟨0,
|
||||
Nat.zero_lt_succ r.length⟩ + ∑ n, g (Fin.succ n) := by
|
||||
exact Fin.sum_univ_succAbove g ⟨0, Nat.zero_lt_succ r.length⟩
|
||||
rw [hf]
|
||||
congr
|
||||
rw [ih]
|
||||
rw [Finset.mul_sum]
|
||||
congr
|
||||
funext n
|
||||
simp only [superCommuteLiftSplit, Fin.eta, List.get_eq_getElem, Algebra.mul_smul_comm,
|
||||
Algebra.smul_mul_assoc, smul_smul, List.length_cons, Fin.val_succ, List.take_succ_cons,
|
||||
List.getElem_cons_succ, List.drop_succ_cons]
|
||||
congr 1
|
||||
· rw [mul_comm, ← superCommuteLiftCoef_append]
|
||||
rfl
|
||||
· simp only [← mul_assoc, mul_eq_mul_right_iff]
|
||||
apply Or.inl
|
||||
apply Or.inl
|
||||
rw [ofListLift, ofListLift, ofListLift]
|
||||
rw [← map_mul]
|
||||
congr
|
||||
rw [← ofList_pair, one_mul]
|
||||
rfl
|
||||
end lift
|
||||
end
|
||||
end Wick
|
525
HepLean/PerturbationTheory/WickContraction/Basic.lean
Normal file
525
HepLean/PerturbationTheory/WickContraction/Basic.lean
Normal file
|
@ -0,0 +1,525 @@
|
|||
/-
|
||||
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.Algebras.OperatorAlgebra.NormalOrder
|
||||
import HepLean.Mathematics.List.InsertIdx
|
||||
/-!
|
||||
|
||||
# Contractions
|
||||
|
||||
-/
|
||||
open FieldStruct
|
||||
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
/--
|
||||
Given a natural number `n` corresponding to the number of fields, a Wick contraction
|
||||
is a finite set of pairs of `Fin n`, such that no element of `Fin n` occurs in more then one pair.
|
||||
For example for `n = 3` there are `4` Wick contractions:
|
||||
- `∅`, corresponding to the case where no fields are contracted.
|
||||
- `{{0, 1}}`, corresponding to the case where the field at position `0` and `1` are contracted.
|
||||
- `{{0, 2}}`, corresponding to the case where the field at position `0` and `2` are contracted.
|
||||
- `{{1, 2}}`, corresponding to the case where the field at position `1` and `2` are contracted.
|
||||
For `n=4` some possible Wick contractions are
|
||||
- `∅`, corresponding to the case where no fields are contracted.
|
||||
- `{{0, 1}, {2, 3}}`, corresponding to the case where the field at position `0` and `1` are
|
||||
contracted and the field at position `2` and `3` are contracted.
|
||||
- `{{0, 2}, {1, 3}}`, corresponding to the case where the field at position `0` and `2` are
|
||||
contracted and the field at position `1` and `3` are contracted.
|
||||
etc.
|
||||
-/
|
||||
def WickContraction (n : ℕ) : Type :=
|
||||
{f : Finset ((Finset (Fin n))) // (∀ a ∈ f, a.card = 2) ∧
|
||||
(∀ a ∈ f, ∀ b ∈ f, a = b ∨ Disjoint a b)}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
|
||||
/-- The contraction consisting of no contracted pairs. -/
|
||||
def empty : WickContraction n := ⟨∅, by simp, by simp⟩
|
||||
|
||||
/-- 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
|
||||
| n, .(n), rfl => Equiv.refl _
|
||||
|
||||
@[simp]
|
||||
lemma congr_refl : c.congr rfl = c := by
|
||||
cases c
|
||||
rfl
|
||||
|
||||
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
|
||||
simp only [congr_refl, Finset.le_eq_subset, finCongr_refl, Equiv.refl_toEmbedding]
|
||||
ext a
|
||||
apply Iff.intro
|
||||
· intro ha
|
||||
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
use a
|
||||
simp only [ha, true_and]
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
simp
|
||||
· intro ha
|
||||
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding] at ha
|
||||
obtain ⟨b, hb, hab⟩ := ha
|
||||
rw [Finset.mapEmbedding_apply] at hab
|
||||
simp only [Finset.map_refl] at hab
|
||||
subst hab
|
||||
exact hb
|
||||
|
||||
@[simp]
|
||||
lemma congr_trans {n m o : ℕ} (h1 : n = m) (h2 : m = o) :
|
||||
(congr h1).trans (congr h2) = congr (h1.trans h2) := by
|
||||
subst h1 h2
|
||||
simp [congr]
|
||||
|
||||
@[simp]
|
||||
lemma congr_trans_apply {n m o : ℕ} (h1 : n = m) (h2 : m = o) (c : WickContraction n) :
|
||||
(congr h2) ((congr h1) c) = congr (h1.trans h2) c := by
|
||||
subst h1 h2
|
||||
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 :=
|
||||
⟨a.1.map (finCongr h).toEmbedding, by
|
||||
subst h
|
||||
simp⟩
|
||||
|
||||
@[simp]
|
||||
lemma congrLift_rfl {n : ℕ} {c : WickContraction n} :
|
||||
c.congrLift rfl = id := by
|
||||
funext a
|
||||
simp [congrLift]
|
||||
|
||||
lemma congrLift_injective {n m : ℕ} {c : WickContraction n} (h : n = m) :
|
||||
Function.Injective (c.congrLift h) := by
|
||||
subst h
|
||||
simp only [congrLift_rfl]
|
||||
exact fun ⦃a₁ a₂⦄ a => a
|
||||
|
||||
lemma congrLift_surjective {n m : ℕ} {c : WickContraction n} (h : n = m) :
|
||||
Function.Surjective (c.congrLift h) := by
|
||||
subst h
|
||||
simp only [congrLift_rfl]
|
||||
exact Function.surjective_id
|
||||
|
||||
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⟩
|
||||
|
||||
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)
|
||||
|
||||
/-- For a contraction `c : WickContraction n` and `i : Fin n` the `j` such that
|
||||
`{i, j}` is a contracted pair in `c`. If such an `j` does not exist, this returns `none`. -/
|
||||
def getDual? (i : Fin n) : Option (Fin n) := Fin.find (fun j => {i, j} ∈ c.1)
|
||||
|
||||
lemma getDual?_congr {n m : ℕ} (h : n = m) (c : WickContraction n) (i : Fin m) :
|
||||
(congr h c).getDual? i = Option.map (finCongr h) (c.getDual? (finCongr h.symm i)) := by
|
||||
subst h
|
||||
simp
|
||||
|
||||
lemma getDual?_congr_get {n m : ℕ} (h : n = m) (c : WickContraction n) (i : Fin m)
|
||||
(hg : ((congr h c).getDual? i).isSome) :
|
||||
((congr h c).getDual? i).get hg =
|
||||
(finCongr h ((c.getDual? (finCongr h.symm i)).get (by simpa [getDual?_congr] using hg))) := by
|
||||
simp only [getDual?_congr, finCongr_apply]
|
||||
exact Option.get_map
|
||||
|
||||
lemma getDual?_eq_some_iff_mem (i j : Fin n) :
|
||||
c.getDual? i = some j ↔ {i, j} ∈ c.1 := by
|
||||
simp only [getDual?]
|
||||
rw [Fin.find_eq_some_iff]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
exact h.1
|
||||
· intro h
|
||||
simp only [h, true_and]
|
||||
intro k hk
|
||||
have hc := c.2.2 _ h _ hk
|
||||
simp only [Finset.disjoint_insert_right, Finset.mem_insert, Finset.mem_singleton, true_or,
|
||||
not_true_eq_false, Finset.disjoint_singleton_right, not_or, false_and, or_false] at hc
|
||||
have hj : k ∈ ({i, j} : Finset (Fin n)) := by
|
||||
simp [hc]
|
||||
simp only [Finset.mem_insert, Finset.mem_singleton] at hj
|
||||
rcases hj with hj | hj
|
||||
· subst hj
|
||||
simp only [Finset.mem_singleton, Finset.insert_eq_of_mem] at hk
|
||||
have hc := c.2.1 _ hk
|
||||
simp at hc
|
||||
· subst hj
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_one_eq_none (c : WickContraction 1) (i : Fin 1) : c.getDual? i = none := by
|
||||
by_contra h
|
||||
have hn : (c.getDual? i).isSome := by
|
||||
rw [← Option.not_isSome_iff_eq_none] at h
|
||||
simpa [- Option.not_isSome, -Option.isNone_iff_eq_none] using h
|
||||
rw [@Option.isSome_iff_exists] at hn
|
||||
obtain ⟨a, hn⟩ := hn
|
||||
rw [getDual?_eq_some_iff_mem] at hn
|
||||
have hc := c.2.1 {i, a} hn
|
||||
fin_cases i
|
||||
fin_cases a
|
||||
simp at hc
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_get_self_mem (i : Fin n) (h : (c.getDual? i).isSome) :
|
||||
{(c.getDual? i).get h, i} ∈ c.1 := by
|
||||
rw [@Finset.pair_comm]
|
||||
rw [← getDual?_eq_some_iff_mem]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma self_getDual?_get_mem (i : Fin n) (h : (c.getDual? i).isSome) :
|
||||
{i, (c.getDual? i).get h} ∈ c.1 := by
|
||||
rw [← getDual?_eq_some_iff_mem]
|
||||
simp
|
||||
|
||||
lemma getDual?_eq_some_neq (i j : Fin n) (h : c.getDual? i = some j) :
|
||||
¬ i = j := by
|
||||
rw [getDual?_eq_some_iff_mem] at h
|
||||
by_contra hn
|
||||
subst hn
|
||||
have hc := c.2.1 _ h
|
||||
simp at hc
|
||||
|
||||
@[simp]
|
||||
lemma self_neq_getDual?_get (i : Fin n) (h : (c.getDual? i).isSome) :
|
||||
¬ i = (c.getDual? i).get h := by
|
||||
by_contra hn
|
||||
have hx : {i, (c.getDual? i).get h} ∈ c.1 := by simp
|
||||
have hc := c.2.1 _ hx
|
||||
nth_rewrite 1 [hn] at hc
|
||||
simp at hc
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_get_self_neq (i : Fin n) (h : (c.getDual? i).isSome) :
|
||||
¬ (c.getDual? i).get h = i := by
|
||||
by_contra hn
|
||||
have hx : {i, (c.getDual? i).get h} ∈ c.1 := by simp
|
||||
have hc := c.2.1 _ hx
|
||||
nth_rewrite 1 [hn] at hc
|
||||
simp at hc
|
||||
|
||||
lemma getDual?_isSome_iff (i : Fin n) : (c.getDual? i).isSome ↔ ∃ (a : c.1), i ∈ a.1 := by
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
simp only [getDual?] at h
|
||||
rw [Fin.isSome_find_iff] at h
|
||||
obtain ⟨a, ha⟩ := h
|
||||
use ⟨{i, a}, ha⟩
|
||||
simp
|
||||
· intro h
|
||||
obtain ⟨a, ha⟩ := h
|
||||
have ha := c.2.1 a a.2
|
||||
rw [@Finset.card_eq_two] at ha
|
||||
obtain ⟨x, y, hx, hy⟩ := ha
|
||||
rw [hy] at ha
|
||||
simp only [Finset.mem_insert, Finset.mem_singleton] at ha
|
||||
match ha with
|
||||
| Or.inl ha =>
|
||||
subst ha
|
||||
simp only [getDual?]
|
||||
rw [Fin.isSome_find_iff]
|
||||
use y
|
||||
rw [← hy]
|
||||
exact a.2
|
||||
| Or.inr ha =>
|
||||
subst ha
|
||||
simp only [getDual?]
|
||||
rw [Fin.isSome_find_iff]
|
||||
use x
|
||||
rw [Finset.pair_comm]
|
||||
rw [← hy]
|
||||
exact a.2
|
||||
|
||||
lemma getDual?_isSome_of_mem (a : c.1) (i : a.1) : (c.getDual? i).isSome := by
|
||||
rw [getDual?_isSome_iff]
|
||||
use ⟨a.1, a.2⟩
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_getDual?_get_get (i : Fin n) (h : (c.getDual? i).isSome) :
|
||||
c.getDual? ((c.getDual? i).get h) = some i := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
simp
|
||||
|
||||
lemma getDual?_getDual?_get_isSome (i : Fin n) (h : (c.getDual? i).isSome) :
|
||||
(c.getDual? ((c.getDual? i).get h)).isSome := by
|
||||
simp
|
||||
|
||||
lemma getDual?_getDual?_get_not_none (i : Fin n) (h : (c.getDual? i).isSome) :
|
||||
¬ (c.getDual? ((c.getDual? i).get h)) = none := by
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## Extracting parts from a contraction.
|
||||
|
||||
-/
|
||||
|
||||
/-- The smallest of the two positions in a contracted pair given a Wick contraction. -/
|
||||
def fstFieldOfContract (c : WickContraction n) (a : c.1) : Fin n :=
|
||||
(a.1.sort (· ≤ ·)).head (by
|
||||
have hx : (Finset.sort (fun x1 x2 => x1 ≤ x2) a.1).length = a.1.card := by
|
||||
exact Finset.length_sort fun x1 x2 => x1 ≤ x2
|
||||
by_contra hn
|
||||
rw [hn, c.2.1 a.1 a.2] at hx
|
||||
simp at hx)
|
||||
|
||||
@[simp]
|
||||
lemma fstFieldOfContract_congr {n m : ℕ} (h : n = m) (c : WickContraction n) (a : c.1) :
|
||||
(congr h c).fstFieldOfContract (c.congrLift h a) = (finCongr h) (c.fstFieldOfContract a) := by
|
||||
subst h
|
||||
simp [congr]
|
||||
|
||||
/-- The largest of the two positions in a contracted pair given a Wick contraction. -/
|
||||
def sndFieldOfContract (c : WickContraction n) (a : c.1) : Fin n :=
|
||||
(a.1.sort (· ≤ ·)).tail.head (by
|
||||
have hx : (Finset.sort (fun x1 x2 => x1 ≤ x2) a.1).length = a.1.card := by
|
||||
exact Finset.length_sort fun x1 x2 => x1 ≤ x2
|
||||
by_contra hn
|
||||
have hn := congrArg List.length hn
|
||||
simp only [List.length_tail, Finset.length_sort, List.length_nil] at hn
|
||||
rw [c.2.1 a.1 a.2] at hn
|
||||
omega)
|
||||
|
||||
@[simp]
|
||||
lemma sndFieldOfContract_congr {n m : ℕ} (h : n = m) (c : WickContraction n) (a : c.1) :
|
||||
(congr h c).sndFieldOfContract (c.congrLift h a) = (finCongr h) (c.sndFieldOfContract a) := by
|
||||
subst h
|
||||
simp [congr]
|
||||
|
||||
lemma finset_eq_fstFieldOfContract_sndFieldOfContract (c : WickContraction n) (a : c.1) :
|
||||
a.1 = {c.fstFieldOfContract a, c.sndFieldOfContract a} := by
|
||||
have h1 := c.2.1 a.1 a.2
|
||||
rw [Finset.card_eq_two] at h1
|
||||
obtain ⟨x, y, hxy, ha⟩ := h1
|
||||
rw [ha]
|
||||
by_cases hxyle : x ≤ y
|
||||
· have ha : a.1.sort (· ≤ ·) = [x, y] := by
|
||||
rw [ha]
|
||||
trans Finset.sort (· ≤ ·) (Finset.cons x {y} (by simp [hxy]))
|
||||
· congr
|
||||
simp
|
||||
rw [Finset.sort_cons]
|
||||
simp only [Finset.sort_singleton]
|
||||
intro b hb
|
||||
simp only [Finset.mem_singleton] at hb
|
||||
subst hb
|
||||
omega
|
||||
simp [fstFieldOfContract, ha, sndFieldOfContract]
|
||||
· have ha : a.1.sort (· ≤ ·) = [y, x] := by
|
||||
rw [ha]
|
||||
trans Finset.sort (· ≤ ·) (Finset.cons y {x} (by simp only [Finset.mem_singleton]; omega))
|
||||
· congr
|
||||
simp only [Finset.cons_eq_insert]
|
||||
rw [@Finset.pair_comm]
|
||||
rw [Finset.sort_cons]
|
||||
simp only [Finset.sort_singleton]
|
||||
intro b hb
|
||||
simp only [Finset.mem_singleton] at hb
|
||||
subst hb
|
||||
omega
|
||||
simp only [fstFieldOfContract, ha, List.head_cons, sndFieldOfContract, List.tail_cons]
|
||||
rw [Finset.pair_comm]
|
||||
|
||||
lemma fstFieldOfContract_neq_sndFieldOfContract (c : WickContraction n) (a : c.1) :
|
||||
c.fstFieldOfContract a ≠ c.sndFieldOfContract a := by
|
||||
have h1 := c.2.1 a.1 a.2
|
||||
have h2 := c.finset_eq_fstFieldOfContract_sndFieldOfContract a
|
||||
by_contra hn
|
||||
rw [h2, hn] at h1
|
||||
simp at h1
|
||||
|
||||
lemma fstFieldOfContract_le_sndFieldOfContract (c : WickContraction n) (a : c.1) :
|
||||
c.fstFieldOfContract a ≤ c.sndFieldOfContract a := by
|
||||
simp only [fstFieldOfContract, sndFieldOfContract, List.head_tail]
|
||||
have h1 (n : ℕ) (l : List (Fin n)) (h : l ≠ []) (hl : l.Sorted (· ≤ ·)) :
|
||||
∀ a ∈ l, l.head h ≤ a := by
|
||||
induction l with
|
||||
| nil => simp at h
|
||||
| cons i l ih =>
|
||||
simp only [List.sorted_cons] at hl
|
||||
simpa using hl.1
|
||||
apply h1
|
||||
· exact Finset.sort_sorted (fun x1 x2 => x1 ≤ x2) _
|
||||
· exact List.getElem_mem _
|
||||
|
||||
lemma fstFieldOfContract_lt_sndFieldOfContract (c : WickContraction n) (a : c.1) :
|
||||
c.fstFieldOfContract a < c.sndFieldOfContract a :=
|
||||
lt_of_le_of_ne (c.fstFieldOfContract_le_sndFieldOfContract a)
|
||||
(c.fstFieldOfContract_neq_sndFieldOfContract a)
|
||||
|
||||
@[simp]
|
||||
lemma fstFieldOfContract_mem (c : WickContraction n) (a : c.1) :
|
||||
c.fstFieldOfContract a ∈ a.1 := by
|
||||
rw [finset_eq_fstFieldOfContract_sndFieldOfContract]
|
||||
simp
|
||||
|
||||
lemma fstFieldOfContract_getDual?_isSome (c : WickContraction n) (a : c.1) :
|
||||
(c.getDual? (c.fstFieldOfContract a)).isSome := by
|
||||
rw [getDual?_isSome_iff]
|
||||
use a
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma fstFieldOfContract_getDual? (c : WickContraction n) (a : c.1) :
|
||||
c.getDual? (c.fstFieldOfContract a) = some (c.sndFieldOfContract a) := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
erw [← finset_eq_fstFieldOfContract_sndFieldOfContract]
|
||||
exact a.2
|
||||
|
||||
@[simp]
|
||||
lemma sndFieldOfContract_mem (c : WickContraction n) (a : c.1) :
|
||||
c.sndFieldOfContract a ∈ a.1 := by
|
||||
rw [finset_eq_fstFieldOfContract_sndFieldOfContract]
|
||||
simp
|
||||
|
||||
lemma sndFieldOfContract_getDual?_isSome (c : WickContraction n) (a : c.1) :
|
||||
(c.getDual? (c.sndFieldOfContract a)).isSome := by
|
||||
rw [getDual?_isSome_iff]
|
||||
use a
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma sndFieldOfContract_getDual? (c : WickContraction n) (a : c.1) :
|
||||
c.getDual? (c.sndFieldOfContract a) = some (c.fstFieldOfContract a) := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
rw [Finset.pair_comm]
|
||||
erw [← finset_eq_fstFieldOfContract_sndFieldOfContract]
|
||||
exact a.2
|
||||
|
||||
lemma eq_fstFieldOfContract_of_mem (c : WickContraction n) (a : c.1) (i j : Fin n)
|
||||
(hi : i ∈ a.1) (hj : j ∈ a.1) (hij : i < j) :
|
||||
c.fstFieldOfContract a = i := by
|
||||
rw [finset_eq_fstFieldOfContract_sndFieldOfContract] at hi hj
|
||||
simp_all only [Finset.mem_insert, Finset.mem_singleton]
|
||||
match hi, hj with
|
||||
| Or.inl hi, Or.inl hj =>
|
||||
subst hi hj
|
||||
simp at hij
|
||||
| Or.inl hi, Or.inr hj =>
|
||||
subst hi
|
||||
rfl
|
||||
| Or.inr hi, Or.inl hj =>
|
||||
subst hi hj
|
||||
have hn := fstFieldOfContract_lt_sndFieldOfContract c a
|
||||
omega
|
||||
| Or.inr hi, Or.inr hj =>
|
||||
subst hi hj
|
||||
simp at hij
|
||||
|
||||
lemma eq_sndFieldOfContract_of_mem (c : WickContraction n) (a : c.1) (i j : Fin n)
|
||||
(hi : i ∈ a.1) (hj : j ∈ a.1) (hij : i < j) :
|
||||
c.sndFieldOfContract a = j := by
|
||||
rw [finset_eq_fstFieldOfContract_sndFieldOfContract] at hi hj
|
||||
simp_all only [Finset.mem_insert, Finset.mem_singleton]
|
||||
match hi, hj with
|
||||
| Or.inl hi, Or.inl hj =>
|
||||
subst hi hj
|
||||
simp at hij
|
||||
| Or.inl hi, Or.inr hj =>
|
||||
subst hi hj
|
||||
omega
|
||||
| Or.inr hi, Or.inl hj =>
|
||||
subst hi hj
|
||||
have hn := fstFieldOfContract_lt_sndFieldOfContract c a
|
||||
omega
|
||||
| Or.inr hi, Or.inr hj =>
|
||||
subst hi hj
|
||||
simp at hij
|
||||
|
||||
/-- As a type, any pair of contractions is equivalent to `Fin 2`
|
||||
with `0` being associated with `c.fstFieldOfContract a` and `1` being associated with
|
||||
`c.sndFieldOfContract`. -/
|
||||
def contractEquivFinTwo (c : WickContraction n) (a : c.1) :
|
||||
a ≃ Fin 2 where
|
||||
toFun i := if i = c.fstFieldOfContract a then 0 else 1
|
||||
invFun i :=
|
||||
match i with
|
||||
| 0 => ⟨c.fstFieldOfContract a, fstFieldOfContract_mem c a⟩
|
||||
| 1 => ⟨c.sndFieldOfContract a, sndFieldOfContract_mem c a⟩
|
||||
left_inv i := by
|
||||
simp only [Fin.isValue]
|
||||
have hi := i.2
|
||||
have ha := c.finset_eq_fstFieldOfContract_sndFieldOfContract a
|
||||
simp only [ha, Finset.mem_insert, Finset.mem_singleton] at hi
|
||||
rcases hi with hi | hi
|
||||
· rw [hi]
|
||||
simp only [↓reduceIte, Fin.isValue]
|
||||
exact Subtype.eq (id (Eq.symm hi))
|
||||
· rw [hi]
|
||||
rw [if_neg]
|
||||
simp only
|
||||
exact Subtype.eq (id (Eq.symm hi))
|
||||
have h := fstFieldOfContract_neq_sndFieldOfContract c a
|
||||
exact (Ne.symm h)
|
||||
right_inv i := by
|
||||
fin_cases i
|
||||
· simp
|
||||
· simp only [Fin.isValue, Fin.mk_one, ite_eq_right_iff, zero_ne_one, imp_false]
|
||||
have h := fstFieldOfContract_neq_sndFieldOfContract c a
|
||||
exact (Ne.symm h)
|
||||
|
||||
lemma prod_finset_eq_mul_fst_snd (c : WickContraction n) (a : c.1)
|
||||
(f : a.1 → M) [CommMonoid M] :
|
||||
∏ (x : a), f x = f (⟨c.fstFieldOfContract a, fstFieldOfContract_mem c a⟩)
|
||||
* f (⟨c.sndFieldOfContract a, sndFieldOfContract_mem c a⟩) := by
|
||||
rw [← (c.contractEquivFinTwo a).symm.prod_comp]
|
||||
simp [contractEquivFinTwo]
|
||||
|
||||
/-- A Wick contraction associated with a list of states is said to be `GradingCompliant` if in any
|
||||
contracted pair of states they are either both fermionic or both bosonic. -/
|
||||
def GradingCompliant (φs : List 𝓕.States) (c : WickContraction φs.length) :=
|
||||
∀ (a : c.1), (𝓕 |>ₛ φs[c.fstFieldOfContract a]) = (𝓕 |>ₛ φs[c.sndFieldOfContract a])
|
||||
|
||||
/-- 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
|
||||
toFun := fun x => ⟨x.2, getDual?_isSome_of_mem c x.fst x.snd⟩
|
||||
invFun := fun x => ⟨
|
||||
⟨{x.1, (c.getDual? x.1).get x.2}, self_getDual?_get_mem c (↑x) x.prop⟩,
|
||||
⟨x.1, by simp⟩⟩
|
||||
left_inv x := by
|
||||
have hxa (x1 x2 : (a : c.1) × a) (h1 : x1.1 = x2.1)
|
||||
(h2 : x1.2.val = x2.2.val) : x1 = x2 := by
|
||||
cases x1
|
||||
cases x2
|
||||
simp_all only [Sigma.mk.inj_iff, true_and]
|
||||
subst h1
|
||||
rename_i fst snd snd_1
|
||||
simp_all only [heq_eq_eq]
|
||||
obtain ⟨val, property⟩ := fst
|
||||
obtain ⟨val_2, property_2⟩ := snd
|
||||
subst h2
|
||||
simp_all only
|
||||
match x with
|
||||
| ⟨a, i⟩ =>
|
||||
apply hxa
|
||||
· have hc := c.2.2 a.1 a.2 {i.1, (c.getDual? ↑i).get (getDual?_isSome_of_mem c a i)}
|
||||
(self_getDual?_get_mem c (↑i) (getDual?_isSome_of_mem c a i))
|
||||
have hn : ¬ Disjoint a.1 {i.1, (c.getDual? ↑i).get (getDual?_isSome_of_mem c a i)} := by
|
||||
rw [Finset.disjoint_iff_inter_eq_empty]
|
||||
rw [@Finset.eq_empty_iff_forall_not_mem]
|
||||
simp only [Finset.coe_mem, Finset.inter_insert_of_mem, Finset.mem_insert, Finset.mem_inter,
|
||||
Finset.mem_singleton, not_or, not_and, not_forall, Classical.not_imp, Decidable.not_not]
|
||||
use i
|
||||
simp
|
||||
simp_all only [or_false, disjoint_self, Finset.bot_eq_empty, Finset.insert_ne_empty,
|
||||
not_false_eq_true]
|
||||
exact Subtype.eq (id (Eq.symm hc))
|
||||
· simp
|
||||
right_inv := by
|
||||
intro x
|
||||
cases x
|
||||
rfl
|
||||
|
||||
end WickContraction
|
160
HepLean/PerturbationTheory/WickContraction/Erase.lean
Normal file
160
HepLean/PerturbationTheory/WickContraction/Erase.lean
Normal file
|
@ -0,0 +1,160 @@
|
|||
/-
|
||||
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.Uncontracted
|
||||
/-!
|
||||
|
||||
# Erasing an element from a contraction
|
||||
|
||||
-/
|
||||
|
||||
open FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
open HepLean.Fin
|
||||
|
||||
/-- Given a Wick contraction `WickContraction n.succ` and a `i : Fin n.succ` the
|
||||
Wick contraction associated with `n` obtained by removing `i`.
|
||||
If `i` is contracted with `j` in the new Wick contraction `j` will be uncontracted. -/
|
||||
def erase (c : WickContraction n.succ) (i : Fin n.succ) : WickContraction n := by
|
||||
refine ⟨Finset.filter (fun x => Finset.map i.succAboveEmb x ∈ c.1) Finset.univ, ?_, ?_⟩
|
||||
· intro a ha
|
||||
simpa using c.2.1 (Finset.map i.succAboveEmb a) (by simpa using ha)
|
||||
· intro a ha b hb
|
||||
simp only [Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, true_and] at ha hb
|
||||
rw [← Finset.disjoint_map i.succAboveEmb, ← (Finset.map_injective i.succAboveEmb).eq_iff]
|
||||
exact c.2.2 _ ha _ hb
|
||||
|
||||
lemma mem_erase_uncontracted_iff (c : WickContraction n.succ) (i : Fin n.succ) (j : Fin n) :
|
||||
j ∈ (c.erase i).uncontracted ↔
|
||||
i.succAbove j ∈ c.uncontracted ∨ c.getDual? (i.succAbove j) = some i := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
simp only [uncontracted, getDual?, erase, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ,
|
||||
Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton, true_and]
|
||||
rw [Fin.find_eq_none_iff, Fin.find_eq_none_iff]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
by_cases hi : {i.succAbove j, i} ∈ c.1
|
||||
· simp [hi]
|
||||
· apply Or.inl
|
||||
intro k
|
||||
by_cases hi' : k = i
|
||||
· subst hi'
|
||||
exact hi
|
||||
· simp only [← Fin.exists_succAbove_eq_iff] at hi'
|
||||
obtain ⟨z, hz⟩ := hi'
|
||||
subst hz
|
||||
exact h z
|
||||
· intro h
|
||||
intro k
|
||||
rcases h with h | h
|
||||
· exact h (i.succAbove k)
|
||||
· by_contra hn
|
||||
have hc := c.2.2 _ h _ hn
|
||||
simp only [Nat.succ_eq_add_one, Finset.disjoint_insert_right, Finset.mem_insert,
|
||||
Finset.mem_singleton, true_or, not_true_eq_false, Finset.disjoint_singleton_right, not_or,
|
||||
false_and, or_false] at hc
|
||||
have hi : i ∈ ({i.succAbove j, i.succAbove k} : Finset (Fin n.succ)) := by
|
||||
simp [← hc]
|
||||
simp only [Nat.succ_eq_add_one, Finset.mem_insert, Finset.mem_singleton] at hi
|
||||
rcases hi with hi | hi
|
||||
· exact False.elim (Fin.succAbove_ne _ _ hi.symm)
|
||||
· exact False.elim (Fin.succAbove_ne _ _ hi.symm)
|
||||
|
||||
lemma mem_not_eq_erase_of_isSome (c : WickContraction n.succ) (i : Fin n.succ)
|
||||
(h : (c.getDual? i).isSome) (ha : a ∈ c.1) (ha2 : a ≠ {i, (c.getDual? i).get h}) :
|
||||
∃ a', a' ∈ (c.erase i).1 ∧ a = Finset.map i.succAboveEmb a' := by
|
||||
have h2a := c.2.1 a ha
|
||||
rw [@Finset.card_eq_two] at h2a
|
||||
obtain ⟨x, y, hx,hy⟩ := h2a
|
||||
subst hy
|
||||
have hxn : ¬ x = i := by
|
||||
by_contra hx
|
||||
subst hx
|
||||
rw [← @getDual?_eq_some_iff_mem] at ha
|
||||
rw [(Option.get_of_mem h ha)] at ha2
|
||||
simp at ha2
|
||||
have hyn : ¬ y = i := by
|
||||
by_contra hy
|
||||
subst hy
|
||||
rw [@Finset.pair_comm] at ha
|
||||
rw [← @getDual?_eq_some_iff_mem] at ha
|
||||
rw [(Option.get_of_mem h ha)] at ha2
|
||||
simp [Finset.pair_comm] at ha2
|
||||
simp only [Nat.succ_eq_add_one, ← Fin.exists_succAbove_eq_iff] at hxn hyn
|
||||
obtain ⟨x', hx'⟩ := hxn
|
||||
obtain ⟨y', hy'⟩ := hyn
|
||||
use {x', y'}
|
||||
subst hx' hy'
|
||||
simp only [erase, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, Finset.map_insert,
|
||||
Fin.succAboveEmb_apply, Finset.map_singleton, true_and, and_true]
|
||||
exact ha
|
||||
|
||||
lemma mem_not_eq_erase_of_isNone (c : WickContraction n.succ) (i : Fin n.succ)
|
||||
(h : (c.getDual? i).isNone) (ha : a ∈ c.1) :
|
||||
∃ a', a' ∈ (c.erase i).1 ∧ a = Finset.map i.succAboveEmb a' := by
|
||||
have h2a := c.2.1 a ha
|
||||
rw [@Finset.card_eq_two] at h2a
|
||||
obtain ⟨x, y, hx,hy⟩ := h2a
|
||||
subst hy
|
||||
have hi : i ∈ c.uncontracted := by
|
||||
simp only [Nat.succ_eq_add_one, uncontracted, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
simp_all only [Nat.succ_eq_add_one, Option.isNone_iff_eq_none, ne_eq]
|
||||
rw [@mem_uncontracted_iff_not_contracted] at hi
|
||||
have hxn : ¬ x = i := by
|
||||
by_contra hx
|
||||
subst hx
|
||||
exact hi {x, y} ha (by simp)
|
||||
have hyn : ¬ y = i := by
|
||||
by_contra hy
|
||||
subst hy
|
||||
exact hi {x, y} ha (by simp)
|
||||
simp only [Nat.succ_eq_add_one, ← Fin.exists_succAbove_eq_iff] at hxn hyn
|
||||
obtain ⟨x', hx'⟩ := hxn
|
||||
obtain ⟨y', hy'⟩ := hyn
|
||||
use {x', y'}
|
||||
subst hx' hy'
|
||||
simp only [erase, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, Finset.map_insert,
|
||||
Fin.succAboveEmb_apply, Finset.map_singleton, true_and, and_true]
|
||||
exact ha
|
||||
|
||||
/-- Given a Wick contraction `c : WickContraction n.succ` and a `i : Fin n.succ` the (optional)
|
||||
element of `(erase c i).uncontracted` which comes from the element in `c` contracted
|
||||
with `i`. -/
|
||||
def getDualErase {n : ℕ} (c : WickContraction n.succ) (i : Fin n.succ) :
|
||||
Option ((erase c i).uncontracted) := by
|
||||
match n with
|
||||
| 0 => exact none
|
||||
| Nat.succ n =>
|
||||
refine if hj : (c.getDual? i).isSome then some ⟨(predAboveI i ((c.getDual? i).get hj)), ?_⟩
|
||||
else none
|
||||
rw [mem_erase_uncontracted_iff]
|
||||
apply Or.inr
|
||||
rw [succsAbove_predAboveI, getDual?_eq_some_iff_mem]
|
||||
· simp
|
||||
· apply c.getDual?_eq_some_neq _ _ _
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma getDualErase_isSome_iff_getDual?_isSome (c : WickContraction n.succ) (i : Fin n.succ) :
|
||||
(c.getDualErase i).isSome ↔ (c.getDual? i).isSome := by
|
||||
match n with
|
||||
| 0 =>
|
||||
fin_cases i
|
||||
simp [getDualErase]
|
||||
|
||||
| Nat.succ n =>
|
||||
simp [getDualErase]
|
||||
|
||||
@[simp]
|
||||
lemma getDualErase_one (c : WickContraction 1) (i : Fin 1) :
|
||||
c.getDualErase i = none := by
|
||||
fin_cases i
|
||||
simp [getDualErase]
|
||||
|
||||
end WickContraction
|
110
HepLean/PerturbationTheory/WickContraction/ExtractEquiv.lean
Normal file
110
HepLean/PerturbationTheory/WickContraction/ExtractEquiv.lean
Normal file
|
@ -0,0 +1,110 @@
|
|||
/-
|
||||
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.Insert
|
||||
/-!
|
||||
|
||||
# Equivalence extracting element from contraction
|
||||
|
||||
-/
|
||||
|
||||
open FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
open HepLean.Fin
|
||||
|
||||
lemma extractEquiv_equiv {c1 c2 : (c : WickContraction n) × Option c.uncontracted}
|
||||
(h : c1.1 = c2.1) (ho : c1.2 = uncontractedCongr (by rw [h]) c2.2) : c1 = c2 := by
|
||||
cases c1
|
||||
cases c2
|
||||
simp_all only [Sigma.mk.inj_iff]
|
||||
simp only at h
|
||||
subst h
|
||||
simp only [uncontractedCongr, Equiv.optionCongr_apply, heq_eq_eq, true_and]
|
||||
rename_i a
|
||||
match a with
|
||||
| none => simp
|
||||
| some a =>
|
||||
simp only [Option.map_some', Option.some.injEq]
|
||||
ext
|
||||
simp
|
||||
|
||||
/-- The equivalence between `WickContraction n.succ` and the sigma type
|
||||
`(c : WickContraction n) × Option c.uncontracted` formed by inserting
|
||||
and erasing elements from a contraction. -/
|
||||
def extractEquiv (i : Fin n.succ) : WickContraction n.succ ≃
|
||||
(c : WickContraction n) × Option c.uncontracted where
|
||||
toFun := fun c => ⟨erase c i, getDualErase c i⟩
|
||||
invFun := fun ⟨c, j⟩ => insert c i j
|
||||
left_inv f := by
|
||||
simp
|
||||
right_inv f := by
|
||||
refine extractEquiv_equiv ?_ ?_
|
||||
simp only [insert_erase]
|
||||
simp only [Nat.succ_eq_add_one]
|
||||
have h1 := insert_getDualErase f.fst i f.snd
|
||||
exact insert_getDualErase _ i _
|
||||
|
||||
lemma extractEquiv_symm_none_uncontracted (i : Fin n.succ) (c : WickContraction n) :
|
||||
((extractEquiv i).symm ⟨c, none⟩).uncontracted =
|
||||
(Insert.insert i (c.uncontracted.map i.succAboveEmb)) := by
|
||||
exact insert_none_uncontracted c i
|
||||
|
||||
@[simp]
|
||||
lemma extractEquiv_apply_congr_symm_apply {n m : ℕ} (k : ℕ)
|
||||
(hnk : k < n.succ) (hkm : k < m.succ) (hnm : n = m) (c : WickContraction n)
|
||||
(i : c.uncontracted) : congr (by rw [hnm]) ((extractEquiv ⟨k, hkm⟩
|
||||
(congr (by rw [hnm]) ((extractEquiv ⟨k, hnk⟩).symm ⟨c, i⟩)))).1 = c := by
|
||||
subst hnm
|
||||
simp
|
||||
|
||||
/-- The fintype instance of `WickContraction 0` defined through its single
|
||||
element `empty`. -/
|
||||
instance fintype_zero : Fintype (WickContraction 0) where
|
||||
elems := {empty}
|
||||
complete := by
|
||||
intro c
|
||||
simp only [Finset.mem_singleton]
|
||||
apply Subtype.eq
|
||||
simp only [empty]
|
||||
ext a
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
have hc := c.2.1 a h
|
||||
rw [Finset.card_eq_two] at hc
|
||||
obtain ⟨x, y, hxy, ha⟩ := hc
|
||||
exact Fin.elim0 x
|
||||
· simp
|
||||
|
||||
lemma sum_WickContraction_nil (f : WickContraction 0 → M) [AddCommMonoid M] :
|
||||
∑ c, f c = f empty := by
|
||||
rw [Finset.sum_eq_single_of_mem]
|
||||
simp only [Finset.mem_univ]
|
||||
intro b hb
|
||||
fin_cases b
|
||||
simp
|
||||
|
||||
/-- The fintype instance of `WickContraction n`, for `n.succ` this is defined
|
||||
through the equivalence `extractEquiv`. -/
|
||||
instance fintype_succ : (n : ℕ) → Fintype (WickContraction n)
|
||||
| 0 => fintype_zero
|
||||
| Nat.succ n => by
|
||||
letI := fintype_succ n
|
||||
exact Fintype.ofEquiv _ (extractEquiv 0).symm
|
||||
|
||||
lemma sum_extractEquiv_congr [AddCommMonoid M] {n m : ℕ} (i : Fin n) (f : WickContraction n → M)
|
||||
(h : n = m.succ) :
|
||||
∑ c, f c = ∑ (c : WickContraction m), ∑ (k : Option c.uncontracted),
|
||||
f (congr h.symm ((extractEquiv (finCongr h i)).symm ⟨c, k⟩)) := by
|
||||
subst h
|
||||
simp only [finCongr_refl, Equiv.refl_apply, congr_refl]
|
||||
rw [← (extractEquiv i).symm.sum_comp]
|
||||
rw [Finset.sum_sigma']
|
||||
rfl
|
||||
|
||||
end WickContraction
|
718
HepLean/PerturbationTheory/WickContraction/Insert.lean
Normal file
718
HepLean/PerturbationTheory/WickContraction/Insert.lean
Normal file
|
@ -0,0 +1,718 @@
|
|||
/-
|
||||
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.Erase
|
||||
/-!
|
||||
|
||||
# Inserting an element into a contraction
|
||||
|
||||
-/
|
||||
|
||||
open FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
open HepLean.Fin
|
||||
|
||||
/-!
|
||||
|
||||
## Inserting an element into a contraction
|
||||
|
||||
-/
|
||||
|
||||
/-- Given a Wick contraction `c` for `n`, a position `i : Fin n.succ` and
|
||||
an optional uncontracted element `j : Option (c.uncontracted)` of `c`.
|
||||
The Wick contraction for `n.succ` formed by 'inserting' `i` into `Fin n`
|
||||
and contracting it optionally with `j`. -/
|
||||
def insert (c : WickContraction n) (i : Fin n.succ) (j : Option (c.uncontracted)) :
|
||||
WickContraction n.succ := by
|
||||
let f := Finset.map (Finset.mapEmbedding i.succAboveEmb).toEmbedding c.1
|
||||
let f' := match j with | none => f | some j => Insert.insert {i, i.succAbove j} f
|
||||
refine ⟨f', ?_, ?_⟩
|
||||
· simp only [Nat.succ_eq_add_one, f']
|
||||
match j with
|
||||
| none =>
|
||||
simp only [Finset.le_eq_subset, Finset.mem_map, RelEmbedding.coe_toEmbedding,
|
||||
forall_exists_index, and_imp, forall_apply_eq_imp_iff₂, f]
|
||||
intro a ha
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
simp only [Finset.card_map]
|
||||
exact c.2.1 a ha
|
||||
| some j =>
|
||||
simp only [Finset.mem_insert, forall_eq_or_imp]
|
||||
apply And.intro
|
||||
· rw [@Finset.card_eq_two]
|
||||
use i
|
||||
use (i.succAbove j)
|
||||
simp only [ne_eq, and_true]
|
||||
exact Fin.ne_succAbove i j
|
||||
· simp only [Finset.le_eq_subset, Finset.mem_map, RelEmbedding.coe_toEmbedding,
|
||||
forall_exists_index, and_imp, forall_apply_eq_imp_iff₂, f]
|
||||
intro a ha
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
simp only [Finset.card_map]
|
||||
exact c.2.1 a ha
|
||||
· intro a ha b hb
|
||||
simp only [Nat.succ_eq_add_one, f'] at ha hb
|
||||
match j with
|
||||
| none =>
|
||||
simp_all only [f, Finset.le_eq_subset, Finset.mem_map, RelEmbedding.coe_toEmbedding,
|
||||
Nat.succ_eq_add_one]
|
||||
obtain ⟨a', ha', ha''⟩ := ha
|
||||
obtain ⟨b', hb', hb''⟩ := hb
|
||||
subst ha''
|
||||
subst hb''
|
||||
simp only [EmbeddingLike.apply_eq_iff_eq]
|
||||
rw [Finset.mapEmbedding_apply, Finset.mapEmbedding_apply, Finset.disjoint_map]
|
||||
exact c.2.2 a' ha' b' hb'
|
||||
| some j =>
|
||||
simp_all only [Finset.mem_insert, Nat.succ_eq_add_one]
|
||||
match ha, hb with
|
||||
| Or.inl ha, Or.inl hb =>
|
||||
rw [ha, hb]
|
||||
simp
|
||||
| Or.inl ha, Or.inr hb =>
|
||||
apply Or.inr
|
||||
subst ha
|
||||
simp only [Finset.disjoint_insert_left, Finset.disjoint_singleton_left]
|
||||
simp only [Finset.le_eq_subset, Finset.mem_map, RelEmbedding.coe_toEmbedding, f] at hb
|
||||
obtain ⟨a', hb', hb''⟩ := hb
|
||||
subst hb''
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
apply And.intro
|
||||
· simp only [Finset.mem_map, Fin.succAboveEmb_apply, not_exists, not_and]
|
||||
exact fun x _ => Fin.succAbove_ne i x
|
||||
· simp only [Finset.mem_map, Fin.succAboveEmb_apply, not_exists, not_and]
|
||||
have hj := j.2
|
||||
rw [mem_uncontracted_iff_not_contracted] at hj
|
||||
intro a ha hja
|
||||
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at hja
|
||||
subst hja
|
||||
exact False.elim (hj a' hb' ha)
|
||||
| Or.inr ha, Or.inl hb =>
|
||||
apply Or.inr
|
||||
subst hb
|
||||
simp only [Finset.disjoint_insert_right, Nat.succ_eq_add_one,
|
||||
Finset.disjoint_singleton_right]
|
||||
simp only [Finset.le_eq_subset, Finset.mem_map, RelEmbedding.coe_toEmbedding, f] at ha
|
||||
obtain ⟨a', ha', ha''⟩ := ha
|
||||
subst ha''
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
apply And.intro
|
||||
· simp only [Finset.mem_map, Fin.succAboveEmb_apply, not_exists, not_and]
|
||||
exact fun x _ => Fin.succAbove_ne i x
|
||||
· simp only [Finset.mem_map, Fin.succAboveEmb_apply, not_exists, not_and]
|
||||
have hj := j.2
|
||||
rw [mem_uncontracted_iff_not_contracted] at hj
|
||||
intro a ha hja
|
||||
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at hja
|
||||
subst hja
|
||||
exact False.elim (hj a' ha' ha)
|
||||
| Or.inr ha, Or.inr hb =>
|
||||
simp_all only [f, Finset.le_eq_subset,
|
||||
or_true, Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
obtain ⟨a', ha', ha''⟩ := ha
|
||||
obtain ⟨b', hb', hb''⟩ := hb
|
||||
subst ha''
|
||||
subst hb''
|
||||
simp only [EmbeddingLike.apply_eq_iff_eq]
|
||||
rw [Finset.mapEmbedding_apply, Finset.mapEmbedding_apply, Finset.disjoint_map]
|
||||
exact c.2.2 a' ha' b' hb'
|
||||
|
||||
lemma insert_of_isSome (c : WickContraction n) (i : Fin n.succ) (j : Option c.uncontracted)
|
||||
(hj : j.isSome) : (insert c i j).1 = Insert.insert {i, i.succAbove (j.get hj)}
|
||||
(Finset.map (Finset.mapEmbedding i.succAboveEmb).toEmbedding c.1) := by
|
||||
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset]
|
||||
rw [@Option.isSome_iff_exists] at hj
|
||||
obtain ⟨j, hj⟩ := hj
|
||||
subst hj
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma self_mem_uncontracted_of_insert_none (c : WickContraction n) (i : Fin n.succ) :
|
||||
i ∈ (insert c i none).uncontracted := by
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
intro p hp
|
||||
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding] at hp
|
||||
obtain ⟨a, ha, ha'⟩ := hp
|
||||
have hc := c.2.1 a ha
|
||||
rw [@Finset.card_eq_two] at hc
|
||||
obtain ⟨x, y, hxy, ha⟩ := hc
|
||||
subst ha
|
||||
subst ha'
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
simp only [Nat.succ_eq_add_one, Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton,
|
||||
Finset.mem_insert, Finset.mem_singleton, not_or]
|
||||
apply And.intro
|
||||
· exact Fin.ne_succAbove i x
|
||||
· exact Fin.ne_succAbove i y
|
||||
|
||||
@[simp]
|
||||
lemma self_not_mem_uncontracted_of_insert_some (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : c.uncontracted) : i ∉ (insert c i (some j)).uncontracted := by
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
simp [insert]
|
||||
|
||||
lemma insert_succAbove_mem_uncontracted_iff (c : WickContraction n) (i : Fin n.succ) (j : Fin n) :
|
||||
(i.succAbove j) ∈ (insert c i none).uncontracted ↔ j ∈ c.uncontracted := by
|
||||
rw [mem_uncontracted_iff_not_contracted, mem_uncontracted_iff_not_contracted]
|
||||
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂]
|
||||
apply Iff.intro
|
||||
· intro h p hp
|
||||
have hp' := h p hp
|
||||
have hc := c.2.1 p hp
|
||||
rw [Finset.card_eq_two] at hc
|
||||
obtain ⟨x, y, hxy, hp⟩ := hc
|
||||
subst hp
|
||||
rw [Finset.mapEmbedding_apply] at hp'
|
||||
simp only [Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton, Finset.mem_insert,
|
||||
Finset.mem_singleton, not_or] at hp'
|
||||
simp only [Finset.mem_insert, Finset.mem_singleton, not_or]
|
||||
exact And.intro (fun a => hp'.1 (congrArg i.succAbove a))
|
||||
(fun a => hp'.2 (congrArg i.succAbove a))
|
||||
· intro h p hp
|
||||
have hc := c.2.1 p hp
|
||||
rw [Finset.card_eq_two] at hc
|
||||
obtain ⟨x, y, hxy, hp⟩ := hc
|
||||
subst hp
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
simp only [Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton, Finset.mem_insert,
|
||||
Finset.mem_singleton, not_or]
|
||||
have hp' := h {x, y} hp
|
||||
simp only [Finset.mem_insert, Finset.mem_singleton, not_or] at hp'
|
||||
apply And.intro
|
||||
(fun a => hp'.1 (i.succAbove_right_injective a))
|
||||
(fun a => hp'.2 (i.succAbove_right_injective a))
|
||||
|
||||
@[simp]
|
||||
lemma mem_uncontracted_insert_none_iff (c : WickContraction n) (i : Fin n.succ) (k : Fin n.succ) :
|
||||
k ∈ (insert c i none).uncontracted ↔
|
||||
k = i ∨ ∃ j, k = i.succAbove j ∧ j ∈ c.uncontracted := by
|
||||
by_cases hi : k = i
|
||||
· subst hi
|
||||
simp
|
||||
· simp only [Nat.succ_eq_add_one, ← Fin.exists_succAbove_eq_iff] at hi
|
||||
obtain ⟨z, hk⟩ := hi
|
||||
subst hk
|
||||
have hn : ¬ i.succAbove z = i := by exact Fin.succAbove_ne i z
|
||||
simp only [Nat.succ_eq_add_one, insert_succAbove_mem_uncontracted_iff, hn, false_or]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
exact ⟨z, rfl, h⟩
|
||||
· intro h
|
||||
obtain ⟨j, hk⟩ := h
|
||||
have hjk : z = j := Fin.succAbove_right_inj.mp hk.1
|
||||
subst hjk
|
||||
exact hk.2
|
||||
|
||||
lemma insert_none_uncontracted (c : WickContraction n) (i : Fin n.succ) :
|
||||
(insert c i none).uncontracted = Insert.insert i (c.uncontracted.map i.succAboveEmb) := by
|
||||
ext a
|
||||
simp only [Nat.succ_eq_add_one, mem_uncontracted_insert_none_iff, Finset.mem_insert,
|
||||
Finset.mem_map, Fin.succAboveEmb_apply]
|
||||
apply Iff.intro
|
||||
· intro a_1
|
||||
cases a_1 with
|
||||
| inl h =>
|
||||
subst h
|
||||
simp_all only [true_or]
|
||||
| inr h_1 =>
|
||||
obtain ⟨w, h⟩ := h_1
|
||||
obtain ⟨left, right⟩ := h
|
||||
subst left
|
||||
apply Or.inr
|
||||
apply Exists.intro
|
||||
· apply And.intro
|
||||
on_goal 2 => {rfl
|
||||
}
|
||||
· simp_all only
|
||||
· intro a_1
|
||||
cases a_1 with
|
||||
| inl h =>
|
||||
subst h
|
||||
simp_all only [true_or]
|
||||
| inr h_1 =>
|
||||
obtain ⟨w, h⟩ := h_1
|
||||
obtain ⟨left, right⟩ := h
|
||||
subst right
|
||||
apply Or.inr
|
||||
apply Exists.intro
|
||||
· apply And.intro
|
||||
on_goal 2 => {exact left
|
||||
}
|
||||
· simp_all only
|
||||
|
||||
@[simp]
|
||||
lemma mem_uncontracted_insert_some_iff (c : WickContraction n) (i : Fin n.succ)
|
||||
(k : Fin n.succ) (j : c.uncontracted) :
|
||||
k ∈ (insert c i (some j)).uncontracted ↔
|
||||
∃ z, k = i.succAbove z ∧ z ∈ c.uncontracted ∧ z ≠ j := by
|
||||
by_cases hki : k = i
|
||||
· subst hki
|
||||
simp only [Nat.succ_eq_add_one, self_not_mem_uncontracted_of_insert_some, ne_eq, false_iff,
|
||||
not_exists, not_and, Decidable.not_not]
|
||||
exact fun x hx => False.elim (Fin.ne_succAbove k x hx)
|
||||
· simp only [Nat.succ_eq_add_one, ← Fin.exists_succAbove_eq_iff] at hki
|
||||
obtain ⟨z, hk⟩ := hki
|
||||
subst hk
|
||||
by_cases hjz : j = z
|
||||
· subst hjz
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_insert,
|
||||
Finset.mem_map, RelEmbedding.coe_toEmbedding, forall_eq_or_imp, Finset.mem_singleton,
|
||||
or_true, not_true_eq_false, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂,
|
||||
false_and, ne_eq, false_iff, not_exists, not_and, Decidable.not_not]
|
||||
intro x
|
||||
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)]
|
||||
exact fun a _a => a.symm
|
||||
· apply Iff.intro
|
||||
· intro h
|
||||
use z
|
||||
simp only [Nat.succ_eq_add_one, ne_eq, true_and]
|
||||
refine And.intro ?_ (fun a => hjz a.symm)
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
intro p hp
|
||||
rw [mem_uncontracted_iff_not_contracted] at h
|
||||
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_insert,
|
||||
Finset.mem_map, RelEmbedding.coe_toEmbedding, forall_eq_or_imp, Finset.mem_singleton,
|
||||
not_or, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] at h
|
||||
have hc := h.2 p hp
|
||||
rw [Finset.mapEmbedding_apply] at hc
|
||||
exact (Finset.mem_map' (i.succAboveEmb)).mpr.mt hc
|
||||
· intro h
|
||||
obtain ⟨z', hz'1, hz'⟩ := h
|
||||
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at hz'1
|
||||
subst hz'1
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_insert,
|
||||
Finset.mem_map, RelEmbedding.coe_toEmbedding, forall_eq_or_imp, Finset.mem_singleton,
|
||||
not_or, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂]
|
||||
apply And.intro
|
||||
· rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)]
|
||||
exact And.intro (Fin.succAbove_ne i z) (fun a => hjz a.symm)
|
||||
· rw [mem_uncontracted_iff_not_contracted] at hz'
|
||||
exact fun a ha hc => hz'.1 a ha ((Finset.mem_map' (i.succAboveEmb)).mp hc)
|
||||
|
||||
lemma insert_some_uncontracted (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted) :
|
||||
(insert c i (some j)).uncontracted = (c.uncontracted.erase j).map i.succAboveEmb := by
|
||||
ext a
|
||||
simp only [Nat.succ_eq_add_one, mem_uncontracted_insert_some_iff, ne_eq, Finset.map_erase,
|
||||
Fin.succAboveEmb_apply, Finset.mem_erase, Finset.mem_map]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
obtain ⟨z, h1, h2, h3⟩ := h
|
||||
subst h1
|
||||
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)]
|
||||
simp only [h3, not_false_eq_true, true_and]
|
||||
use z
|
||||
· intro h
|
||||
obtain ⟨z, h1, h2⟩ := h.2
|
||||
use z
|
||||
subst h2
|
||||
simp only [true_and]
|
||||
obtain ⟨a, ha1, ha2⟩ := h.2
|
||||
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at ha2
|
||||
subst ha2
|
||||
simp_all only [true_and]
|
||||
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at h
|
||||
exact h.1
|
||||
|
||||
/-!
|
||||
|
||||
## Insert and getDual?
|
||||
|
||||
-/
|
||||
|
||||
lemma insert_none_getDual?_isNone (c : WickContraction n) (i : Fin n.succ) :
|
||||
((insert c i none).getDual? i).isNone := by
|
||||
have hi : i ∈ (insert c i none).uncontracted := by
|
||||
simp
|
||||
simp only [Nat.succ_eq_add_one, uncontracted, Finset.mem_filter, Finset.mem_univ, true_and] at hi
|
||||
rw [hi]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma insert_succAbove_getDual?_eq_none_iff (c : WickContraction n) (i : Fin n.succ) (j : Fin n) :
|
||||
(insert c i none).getDual? (i.succAbove j) = none ↔ c.getDual? j = none := by
|
||||
have h1 := insert_succAbove_mem_uncontracted_iff c i j
|
||||
simpa [uncontracted] using h1
|
||||
|
||||
@[simp]
|
||||
lemma insert_succAbove_getDual?_isSome_iff (c : WickContraction n) (i : Fin n.succ) (j : Fin n) :
|
||||
((insert c i none).getDual? (i.succAbove j)).isSome ↔ (c.getDual? j).isSome := by
|
||||
rw [← not_iff_not]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma insert_succAbove_getDual?_get (c : WickContraction n) (i : Fin n.succ) (j : Fin n)
|
||||
(h : ((insert c i none).getDual? (i.succAbove j)).isSome) :
|
||||
((insert c i none).getDual? (i.succAbove j)).get h =
|
||||
i.succAbove ((c.getDual? j).get (by simpa using h)) := by
|
||||
have h1 : (insert c i none).getDual? (i.succAbove j) = some
|
||||
(i.succAbove ((c.getDual? j).get (by simpa using h))) := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding]
|
||||
use {j, ((c.getDual? j).get (by simpa using h))}
|
||||
simp only [self_getDual?_get_mem, true_and]
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
simp
|
||||
exact Option.get_of_mem h h1
|
||||
|
||||
@[simp]
|
||||
lemma insert_some_getDual?_eq (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted) :
|
||||
(insert c i (some j)).getDual? i = some (i.succAbove j) := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
simp [insert]
|
||||
|
||||
@[simp]
|
||||
lemma insert_some_getDual?_neq_none (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted)
|
||||
(k : Fin n) (hkj : k ≠ j.1) :
|
||||
(insert c i (some j)).getDual? (i.succAbove k) = none ↔ c.getDual? k = none := by
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
have hk : (i.succAbove k) ∈ (insert c i (some j)).uncontracted := by
|
||||
simp only [Nat.succ_eq_add_one, uncontracted, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
exact h
|
||||
simp only [Nat.succ_eq_add_one, mem_uncontracted_insert_some_iff, ne_eq] at hk
|
||||
obtain ⟨z, hz1, hz2, hz3⟩ := hk
|
||||
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at hz1
|
||||
subst hz1
|
||||
simpa [uncontracted] using hz2
|
||||
· intro h
|
||||
have hk : (i.succAbove k) ∈ (insert c i (some j)).uncontracted := by
|
||||
simp only [Nat.succ_eq_add_one, mem_uncontracted_insert_some_iff, ne_eq]
|
||||
use k
|
||||
simp only [hkj, not_false_eq_true, and_true, true_and]
|
||||
simpa [uncontracted] using h
|
||||
simpa [uncontracted, -mem_uncontracted_insert_some_iff, ne_eq] using hk
|
||||
|
||||
@[simp]
|
||||
lemma insert_some_getDual?_neq_isSome (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted)
|
||||
(k : Fin n) (hkj : k ≠ j.1) :
|
||||
((insert c i (some j)).getDual? (i.succAbove k)).isSome ↔ (c.getDual? k).isSome := by
|
||||
rw [← not_iff_not]
|
||||
simp [hkj]
|
||||
|
||||
@[simp]
|
||||
lemma insert_some_getDual?_neq_isSome_get (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : c.uncontracted) (k : Fin n) (hkj : k ≠ j.1)
|
||||
(h : ((insert c i (some j)).getDual? (i.succAbove k)).isSome) :
|
||||
((insert c i (some j)).getDual? (i.succAbove k)).get h =
|
||||
i.succAbove ((c.getDual? k).get (by simpa [hkj] using h)) := by
|
||||
have h1 : ((insert c i (some j)).getDual? (i.succAbove k))
|
||||
= some (i.succAbove ((c.getDual? k).get (by simpa [hkj] using h))) := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_insert, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding]
|
||||
apply Or.inr
|
||||
use { k, ((c.getDual? k).get (by simpa [hkj] using h))}
|
||||
simp only [self_getDual?_get_mem, true_and]
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
simp
|
||||
exact Option.get_of_mem h h1
|
||||
|
||||
@[simp]
|
||||
lemma insert_some_getDual?_of_neq (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted)
|
||||
(k : Fin n) (hkj : k ≠ j.1) :
|
||||
(insert c i (some j)).getDual? (i.succAbove k) =
|
||||
Option.map i.succAbove (c.getDual? k) := by
|
||||
by_cases h : (c.getDual? k).isSome
|
||||
· have h1 : (c.insert i (some j)).getDual? (i.succAbove k) =
|
||||
some (i.succAbove ((c.getDual? k).get h)) := by
|
||||
rw [← insert_some_getDual?_neq_isSome_get c i j k hkj]
|
||||
refine Eq.symm (Option.some_get ?_)
|
||||
simpa [hkj] using h
|
||||
rw [h1]
|
||||
have h2 :(c.getDual? k) = some ((c.getDual? k).get h) := by simp
|
||||
conv_rhs => rw [h2]
|
||||
rw [@Option.map_coe']
|
||||
· simp only [Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none] at h
|
||||
simp only [Nat.succ_eq_add_one, h, Option.map_none']
|
||||
simp only [ne_eq, hkj, not_false_eq_true, insert_some_getDual?_neq_none]
|
||||
exact h
|
||||
|
||||
/-!
|
||||
|
||||
## Interaction with erase.
|
||||
|
||||
-/
|
||||
@[simp]
|
||||
lemma insert_erase (c : WickContraction n) (i : Fin n.succ) (j : Option (c.uncontracted)) :
|
||||
erase (insert c i j) i = c := by
|
||||
refine Subtype.eq ?_
|
||||
simp only [erase, Nat.succ_eq_add_one, insert, Finset.le_eq_subset]
|
||||
conv_rhs => rw [c.eq_filter_mem_self]
|
||||
refine Finset.filter_inj'.mpr ?_
|
||||
intro a _
|
||||
match j with
|
||||
| none =>
|
||||
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
apply Iff.intro
|
||||
· intro ha
|
||||
obtain ⟨a', ha', ha''⟩ := ha
|
||||
rw [Finset.mapEmbedding_apply] at ha''
|
||||
simp only [Finset.map_inj] at ha''
|
||||
subst ha''
|
||||
exact ha'
|
||||
· intro ha
|
||||
use a
|
||||
simp only [ha, true_and]
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
| some j =>
|
||||
simp only [Finset.mem_insert, Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
apply Iff.intro
|
||||
· intro ha
|
||||
rcases ha with ha | ha
|
||||
· have hin : ¬ i ∈ Finset.map i.succAboveEmb a := by
|
||||
simp only [Nat.succ_eq_add_one, Finset.mem_map, Fin.succAboveEmb_apply, not_exists,
|
||||
not_and]
|
||||
intro x
|
||||
exact fun a => Fin.succAbove_ne i x
|
||||
refine False.elim (hin ?_)
|
||||
rw [ha]
|
||||
simp
|
||||
· obtain ⟨a', ha', ha''⟩ := ha
|
||||
rw [Finset.mapEmbedding_apply] at ha''
|
||||
simp only [Finset.map_inj] at ha''
|
||||
subst ha''
|
||||
exact ha'
|
||||
· intro ha
|
||||
apply Or.inr
|
||||
use a
|
||||
simp only [ha, true_and]
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
|
||||
lemma insert_getDualErase (c : WickContraction n) (i : Fin n.succ) (j : Option c.uncontracted) :
|
||||
(insert c i j).getDualErase i =
|
||||
uncontractedCongr (c := c) (c' := (c.insert i j).erase i) (by simp) j := by
|
||||
match n with
|
||||
| 0 =>
|
||||
simp only [insert, Nat.succ_eq_add_one, Nat.reduceAdd, Finset.le_eq_subset, getDualErase]
|
||||
fin_cases j
|
||||
simp
|
||||
| Nat.succ n =>
|
||||
match j with
|
||||
| none =>
|
||||
have hi := insert_none_getDual?_isNone c i
|
||||
simp [getDualErase, hi]
|
||||
| some j =>
|
||||
simp only [Nat.succ_eq_add_one, getDualErase, insert_some_getDual?_eq, Option.isSome_some,
|
||||
↓reduceDIte, Option.get_some, predAboveI_succAbove, uncontractedCongr_some, Option.some.injEq]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma erase_insert (c : WickContraction n.succ) (i : Fin n.succ) :
|
||||
insert (erase c i) i (getDualErase c i) = c := by
|
||||
match n with
|
||||
| 0 =>
|
||||
apply Subtype.eq
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, insert, getDualErase, Finset.le_eq_subset]
|
||||
ext a
|
||||
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
simp only [erase, Nat.reduceAdd, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ,
|
||||
true_and] at h
|
||||
obtain ⟨a', ha', ha''⟩ := h
|
||||
subst ha''
|
||||
exact ha'
|
||||
· intro ha
|
||||
obtain ⟨a, ha⟩ := c.mem_not_eq_erase_of_isNone (a := a) i (by simp) ha
|
||||
simp_all only [Nat.succ_eq_add_one]
|
||||
obtain ⟨left, right⟩ := ha
|
||||
subst right
|
||||
apply Exists.intro
|
||||
· apply And.intro
|
||||
on_goal 2 => {rfl
|
||||
}
|
||||
· simp_all only
|
||||
| Nat.succ n =>
|
||||
apply Subtype.eq
|
||||
by_cases hi : (c.getDual? i).isSome
|
||||
· rw [insert_of_isSome]
|
||||
simp only [Nat.succ_eq_add_one, getDualErase, hi, ↓reduceDIte, Option.get_some,
|
||||
Finset.le_eq_subset]
|
||||
rw [succsAbove_predAboveI]
|
||||
ext a
|
||||
apply Iff.intro
|
||||
· simp only [Finset.mem_insert, Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
intro ha
|
||||
rcases ha with ha | ha
|
||||
· subst ha
|
||||
simp
|
||||
· obtain ⟨a', ha', ha''⟩ := ha
|
||||
subst ha''
|
||||
simpa [erase] using ha'
|
||||
· intro ha
|
||||
simp only [Finset.mem_insert, Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
by_cases hia : a = {i, (c.getDual? i).get hi}
|
||||
· subst hia
|
||||
simp
|
||||
· apply Or.inr
|
||||
simp only [erase, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
obtain ⟨a', ha'⟩ := c.mem_not_eq_erase_of_isSome (a := a) i hi ha hia
|
||||
use a'
|
||||
simp_all only [Nat.succ_eq_add_one, true_and]
|
||||
obtain ⟨left, right⟩ := ha'
|
||||
subst right
|
||||
rfl
|
||||
simp only [Nat.succ_eq_add_one, ne_eq, self_neq_getDual?_get, not_false_eq_true]
|
||||
exact (getDualErase_isSome_iff_getDual?_isSome c i).mpr hi
|
||||
· simp only [Nat.succ_eq_add_one, insert, getDualErase, hi, Bool.false_eq_true, ↓reduceDIte,
|
||||
Finset.le_eq_subset]
|
||||
ext a
|
||||
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
simp only [erase, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, true_and] at h
|
||||
obtain ⟨a', ha', ha''⟩ := h
|
||||
subst ha''
|
||||
exact ha'
|
||||
· intro ha
|
||||
obtain ⟨a, ha⟩ := c.mem_not_eq_erase_of_isNone (a := a) i (by simpa using hi) ha
|
||||
simp_all only [Nat.succ_eq_add_one, Bool.not_eq_true, Option.not_isSome,
|
||||
Option.isNone_iff_eq_none]
|
||||
obtain ⟨left, right⟩ := ha
|
||||
subst right
|
||||
apply Exists.intro
|
||||
· apply And.intro
|
||||
on_goal 2 => {rfl
|
||||
}
|
||||
· simp_all only
|
||||
|
||||
/-- Lifts a contraction in `c` to a contraction in `(c.insert i j)`. -/
|
||||
def insertLift {c : WickContraction n} (i : Fin n.succ) (j : Option (c.uncontracted))
|
||||
(a : c.1) : (c.insert i j).1 := ⟨a.1.map (Fin.succAboveEmb i), by
|
||||
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset]
|
||||
match j with
|
||||
| none =>
|
||||
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
use a
|
||||
simp only [a.2, true_and]
|
||||
rfl
|
||||
| some j =>
|
||||
simp only [Finset.mem_insert, Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
apply Or.inr
|
||||
use a
|
||||
simp only [a.2, true_and]
|
||||
rfl⟩
|
||||
|
||||
lemma insertLift_injective {c : WickContraction n} (i : Fin n.succ) (j : Option (c.uncontracted)) :
|
||||
Function.Injective (insertLift i j) := by
|
||||
intro a b hab
|
||||
simp only [Nat.succ_eq_add_one, insertLift, Subtype.mk.injEq, Finset.map_inj] at hab
|
||||
exact Subtype.eq hab
|
||||
|
||||
lemma insertLift_none_surjective {c : WickContraction n} (i : Fin n.succ) :
|
||||
Function.Surjective (c.insertLift i none) := by
|
||||
intro a
|
||||
have ha := a.2
|
||||
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding] at ha
|
||||
obtain ⟨a', ha', ha''⟩ := ha
|
||||
use ⟨a', ha'⟩
|
||||
exact Subtype.eq ha''
|
||||
|
||||
lemma insertLift_none_bijective {c : WickContraction n} (i : Fin n.succ) :
|
||||
Function.Bijective (c.insertLift i none) := by
|
||||
exact ⟨insertLift_injective i none, insertLift_none_surjective i⟩
|
||||
|
||||
@[simp]
|
||||
lemma insert_fstFieldOfContract (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : Option (c.uncontracted)) (a : c.1) : (c.insert i j).fstFieldOfContract (insertLift i j a) =
|
||||
i.succAbove (c.fstFieldOfContract a) := by
|
||||
refine (c.insert i j).eq_fstFieldOfContract_of_mem (a := (insertLift i j a))
|
||||
(i := i.succAbove (c.fstFieldOfContract a)) (j := i.succAbove (c.sndFieldOfContract a)) ?_ ?_ ?_
|
||||
· simp only [Nat.succ_eq_add_one, insertLift, Finset.mem_map, Fin.succAboveEmb_apply]
|
||||
use (c.fstFieldOfContract a)
|
||||
simp
|
||||
· simp only [Nat.succ_eq_add_one, insertLift, Finset.mem_map, Fin.succAboveEmb_apply]
|
||||
use (c.sndFieldOfContract a)
|
||||
simp
|
||||
· refine Fin.succAbove_lt_succAbove_iff.mpr ?_
|
||||
exact fstFieldOfContract_lt_sndFieldOfContract c a
|
||||
|
||||
@[simp]
|
||||
lemma insert_sndFieldOfContract (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : Option (c.uncontracted)) (a : c.1) : (c.insert i j).sndFieldOfContract (insertLift i j a) =
|
||||
i.succAbove (c.sndFieldOfContract a) := by
|
||||
refine (c.insert i j).eq_sndFieldOfContract_of_mem (a := (insertLift i j a))
|
||||
(i := i.succAbove (c.fstFieldOfContract a)) (j := i.succAbove (c.sndFieldOfContract a)) ?_ ?_ ?_
|
||||
· simp only [Nat.succ_eq_add_one, insertLift, Finset.mem_map, Fin.succAboveEmb_apply]
|
||||
use (c.fstFieldOfContract a)
|
||||
simp
|
||||
· simp only [Nat.succ_eq_add_one, insertLift, Finset.mem_map, Fin.succAboveEmb_apply]
|
||||
use (c.sndFieldOfContract a)
|
||||
simp
|
||||
· refine Fin.succAbove_lt_succAbove_iff.mpr ?_
|
||||
exact fstFieldOfContract_lt_sndFieldOfContract c a
|
||||
|
||||
/-- Given a contracted pair for a Wick contraction `WickContraction n`, the
|
||||
corresponding contracted pair of a wick contraction `(c.insert i (some j))` formed
|
||||
by inserting an element `i` into the contraction. -/
|
||||
def insertLiftSome {c : WickContraction n} (i : Fin n.succ) (j : c.uncontracted)
|
||||
(a : Unit ⊕ c.1) : (c.insert i (some j)).1 :=
|
||||
match a with
|
||||
| Sum.inl () => ⟨{i, i.succAbove j}, by
|
||||
simp [insert]⟩
|
||||
| Sum.inr a => c.insertLift i j a
|
||||
|
||||
lemma insertLiftSome_injective {c : WickContraction n} (i : Fin n.succ) (j : c.uncontracted) :
|
||||
Function.Injective (insertLiftSome i j) := by
|
||||
intro a b hab
|
||||
match a, b with
|
||||
| Sum.inl (), Sum.inl () => rfl
|
||||
| Sum.inl (), Sum.inr a =>
|
||||
simp only [Nat.succ_eq_add_one, insertLiftSome, insertLift, Subtype.mk.injEq] at hab
|
||||
rw [finset_eq_fstFieldOfContract_sndFieldOfContract] at hab
|
||||
simp only [Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton] at hab
|
||||
have hi : i ∈ ({i.succAbove (c.fstFieldOfContract a),
|
||||
i.succAbove (c.sndFieldOfContract a)} : Finset (Fin (n + 1))) := by
|
||||
rw [← hab]
|
||||
simp
|
||||
simp only [Nat.succ_eq_add_one, Finset.mem_insert, Finset.mem_singleton] at hi
|
||||
rcases hi with hi | hi
|
||||
· exact False.elim (Fin.ne_succAbove _ _ hi)
|
||||
· exact False.elim (Fin.ne_succAbove _ _ hi)
|
||||
| Sum.inr a, Sum.inl () =>
|
||||
simp only [Nat.succ_eq_add_one, insertLiftSome, insertLift, Subtype.mk.injEq] at hab
|
||||
rw [finset_eq_fstFieldOfContract_sndFieldOfContract] at hab
|
||||
simp only [Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton] at hab
|
||||
have hi : i ∈ ({i.succAbove (c.fstFieldOfContract a),
|
||||
i.succAbove (c.sndFieldOfContract a)} : Finset (Fin (n + 1))) := by
|
||||
rw [hab]
|
||||
simp
|
||||
simp only [Nat.succ_eq_add_one, Finset.mem_insert, Finset.mem_singleton] at hi
|
||||
rcases hi with hi | hi
|
||||
· exact False.elim (Fin.ne_succAbove _ _ hi)
|
||||
· exact False.elim (Fin.ne_succAbove _ _ hi)
|
||||
| Sum.inr a, Sum.inr b =>
|
||||
simp only [Nat.succ_eq_add_one, insertLiftSome] at hab
|
||||
simpa using insertLift_injective i (some j) hab
|
||||
|
||||
lemma insertLiftSome_surjective {c : WickContraction n} (i : Fin n.succ) (j : c.uncontracted) :
|
||||
Function.Surjective (insertLiftSome i j) := by
|
||||
intro a
|
||||
have ha := a.2
|
||||
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_insert, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding] at ha
|
||||
rcases ha with ha | ha
|
||||
· use Sum.inl ()
|
||||
exact Subtype.eq ha.symm
|
||||
· obtain ⟨a', ha', ha''⟩ := ha
|
||||
use Sum.inr ⟨a', ha'⟩
|
||||
simp only [Nat.succ_eq_add_one, insertLiftSome, insertLift]
|
||||
exact Subtype.eq ha''
|
||||
|
||||
lemma insertLiftSome_bijective {c : WickContraction n} (i : Fin n.succ) (j : c.uncontracted) :
|
||||
Function.Bijective (insertLiftSome i j) :=
|
||||
⟨insertLiftSome_injective i j, insertLiftSome_surjective i j⟩
|
||||
|
||||
end WickContraction
|
293
HepLean/PerturbationTheory/WickContraction/InsertList.lean
Normal file
293
HepLean/PerturbationTheory/WickContraction/InsertList.lean
Normal file
|
@ -0,0 +1,293 @@
|
|||
/-
|
||||
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.UncontractedList
|
||||
/-!
|
||||
|
||||
# Inserting an element into a contraction
|
||||
|
||||
-/
|
||||
|
||||
open FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
open HepLean.Fin
|
||||
|
||||
/-!
|
||||
|
||||
## Inserting an element into a list
|
||||
|
||||
-/
|
||||
|
||||
/-- Given a Wick contraction `c` associated to a list `φs`,
|
||||
a position `i : Fin n.succ`, an element `φ`, and an optional uncontracted element
|
||||
`j : Option (c.uncontracted)` of `c`.
|
||||
The Wick contraction associated with `(φs.insertIdx i φ).length` formed by 'inserting' `φ`
|
||||
into `φs` after the first `i` elements and contracting it optionally with j. -/
|
||||
def insertList (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Option (c.uncontracted)) :
|
||||
WickContraction (φs.insertIdx i φ).length :=
|
||||
congr (by simp) (c.insert i j)
|
||||
|
||||
@[simp]
|
||||
lemma insertList_fstFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Option (c.uncontracted))
|
||||
(a : c.1) : (insertList φ φs c i j).fstFieldOfContract
|
||||
(congrLift (insertIdx_length_fin φ φs i).symm (insertLift i j a)) =
|
||||
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove (c.fstFieldOfContract a)) := by
|
||||
simp [insertList]
|
||||
|
||||
@[simp]
|
||||
lemma insertList_sndFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Option (c.uncontracted))
|
||||
(a : c.1) : (insertList φ φs c i j).sndFieldOfContract
|
||||
(congrLift (insertIdx_length_fin φ φs i).symm (insertLift i j a)) =
|
||||
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove (c.sndFieldOfContract a)) := by
|
||||
simp [insertList]
|
||||
|
||||
@[simp]
|
||||
lemma insertList_fstFieldOfContract_some_incl (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
|
||||
(insertList φ φs c i (some j)).fstFieldOfContract
|
||||
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩) =
|
||||
if i < i.succAbove j.1 then
|
||||
finCongr (insertIdx_length_fin φ φs i).symm i else
|
||||
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j.1) := by
|
||||
split
|
||||
· rename_i h
|
||||
refine (insertList φ φs c i (some j)).eq_fstFieldOfContract_of_mem
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩)
|
||||
(i := finCongr (insertIdx_length_fin φ φs i).symm i) (j :=
|
||||
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j)) ?_ ?_ ?_
|
||||
· simp [congrLift]
|
||||
· simp [congrLift]
|
||||
· rw [Fin.lt_def] at h ⊢
|
||||
simp_all
|
||||
· rename_i h
|
||||
refine (insertList φ φs c i (some j)).eq_fstFieldOfContract_of_mem
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩)
|
||||
(i := finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j))
|
||||
(j := finCongr (insertIdx_length_fin φ φs i).symm i) ?_ ?_ ?_
|
||||
· simp [congrLift]
|
||||
· simp [congrLift]
|
||||
· rw [Fin.lt_def] at h ⊢
|
||||
simp_all only [Nat.succ_eq_add_one, Fin.val_fin_lt, not_lt, finCongr_apply, Fin.coe_cast]
|
||||
have hi : i.succAbove j ≠ i := by exact Fin.succAbove_ne i j
|
||||
omega
|
||||
|
||||
/-!
|
||||
|
||||
## insertList and getDual?
|
||||
|
||||
-/
|
||||
@[simp]
|
||||
lemma insertList_none_getDual?_self (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) :
|
||||
(insertList φ φs c i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm i) = none := by
|
||||
simp only [Nat.succ_eq_add_one, insertList, getDual?_congr, finCongr_apply, Fin.cast_trans,
|
||||
Fin.cast_eq_self, Option.map_eq_none']
|
||||
have h1 := c.insert_none_getDual?_isNone i
|
||||
simpa using h1
|
||||
|
||||
lemma insertList_isSome_getDual?_self (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
|
||||
((insertList φ φs c i (some j)).getDual?
|
||||
(Fin.cast (insertIdx_length_fin φ φs i).symm i)).isSome := by
|
||||
simp [insertList, getDual?_congr]
|
||||
|
||||
lemma insertList_some_getDual?_self_not_none (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
|
||||
¬ ((insertList φ φs c i (some j)).getDual?
|
||||
(Fin.cast (insertIdx_length_fin φ φs i).symm i)) = none := by
|
||||
simp [insertList, getDual?_congr]
|
||||
|
||||
@[simp]
|
||||
lemma insertList_some_getDual?_self_eq (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
|
||||
((insertList φ φs c i (some j)).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm i))
|
||||
= some (Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove j)) := by
|
||||
simp [insertList, getDual?_congr]
|
||||
|
||||
@[simp]
|
||||
lemma insertList_some_getDual?_some_eq (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
|
||||
((insertList φ φs c i (some j)).getDual?
|
||||
(Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove j)))
|
||||
= some (Fin.cast (insertIdx_length_fin φ φs i).symm i) := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
rw [@Finset.pair_comm]
|
||||
rw [← getDual?_eq_some_iff_mem]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma insertList_none_succAbove_getDual?_eq_none_iff (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length) :
|
||||
(insertList φ φs c i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
|
||||
(i.succAbove j)) = none ↔ c.getDual? j = none := by
|
||||
simp [insertList, getDual?_congr]
|
||||
|
||||
@[simp]
|
||||
lemma insertList_some_succAbove_getDual?_eq_option (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length)
|
||||
(k : c.uncontracted) (hkj : j ≠ k.1) :
|
||||
(insertList φ φs c i (some k)).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
|
||||
(i.succAbove j)) = Option.map (Fin.cast (insertIdx_length_fin φ φs i).symm ∘ i.succAbove)
|
||||
(c.getDual? j) := by
|
||||
simp only [Nat.succ_eq_add_one, insertList, getDual?_congr, finCongr_apply, Fin.cast_trans,
|
||||
Fin.cast_eq_self, ne_eq, hkj, not_false_eq_true, insert_some_getDual?_of_neq, Option.map_map]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma insertList_none_succAbove_getDual?_isSome_iff (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length) :
|
||||
((insertList φ φs c i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
|
||||
(i.succAbove j))).isSome ↔ (c.getDual? j).isSome := by
|
||||
rw [← not_iff_not]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma insertList_none_getDual?_get_eq (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length)
|
||||
(h : ((insertList φ φs c i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
|
||||
(i.succAbove j))).isSome) :
|
||||
((insertList φ φs c i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
|
||||
(i.succAbove j))).get h = Fin.cast (insertIdx_length_fin φ φs i).symm
|
||||
(i.succAbove ((c.getDual? j).get (by simpa using h))) := by
|
||||
simp [insertList, getDual?_congr_get]
|
||||
|
||||
/-........................................... -/
|
||||
@[simp]
|
||||
lemma insertList_sndFieldOfContract_some_incl (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
|
||||
(insertList φ φs c i (some j)).sndFieldOfContract
|
||||
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩) =
|
||||
if i < i.succAbove j.1 then
|
||||
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j.1) else
|
||||
finCongr (insertIdx_length_fin φ φs i).symm i := by
|
||||
split
|
||||
· rename_i h
|
||||
refine (insertList φ φs c i (some j)).eq_sndFieldOfContract_of_mem
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩)
|
||||
(i := finCongr (insertIdx_length_fin φ φs i).symm i) (j :=
|
||||
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j)) ?_ ?_ ?_
|
||||
· simp [congrLift]
|
||||
· simp [congrLift]
|
||||
· rw [Fin.lt_def] at h ⊢
|
||||
simp_all
|
||||
· rename_i h
|
||||
refine (insertList φ φs c i (some j)).eq_sndFieldOfContract_of_mem
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩)
|
||||
(i := finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j))
|
||||
(j := finCongr (insertIdx_length_fin φ φs i).symm i) ?_ ?_ ?_
|
||||
· simp [congrLift]
|
||||
· simp [congrLift]
|
||||
· rw [Fin.lt_def] at h ⊢
|
||||
simp_all only [Nat.succ_eq_add_one, Fin.val_fin_lt, not_lt, finCongr_apply, Fin.coe_cast]
|
||||
have hi : i.succAbove j ≠ i := by exact Fin.succAbove_ne i j
|
||||
omega
|
||||
|
||||
lemma insertList_none_prod_contractions (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ)
|
||||
(f : (c.insertList φ φs i none).1 → M) [CommMonoid M] :
|
||||
∏ a, f a = ∏ (a : c.1), f (congrLift (insertIdx_length_fin φ φs i).symm
|
||||
(insertLift i none a)) := by
|
||||
let e1 := Equiv.ofBijective (c.insertLift i none) (insertLift_none_bijective i)
|
||||
let e2 := Equiv.ofBijective (congrLift (insertIdx_length_fin φ φs i).symm)
|
||||
((c.insert i none).congrLift_bijective ((insertIdx_length_fin φ φs i).symm))
|
||||
erw [← e2.prod_comp]
|
||||
erw [← e1.prod_comp]
|
||||
rfl
|
||||
|
||||
lemma insertList_some_prod_contractions (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted)
|
||||
(f : (c.insertList φ φs i (some j)).1 → M) [CommMonoid M] :
|
||||
∏ a, f a = f (congrLift (insertIdx_length_fin φ φs i).symm
|
||||
⟨{i, i.succAbove j}, by simp [insert]⟩) *
|
||||
∏ (a : c.1), f (congrLift (insertIdx_length_fin φ φs i).symm (insertLift i (some j) a)) := by
|
||||
let e2 := Equiv.ofBijective (congrLift (insertIdx_length_fin φ φs i).symm)
|
||||
((c.insert i (some j)).congrLift_bijective ((insertIdx_length_fin φ φs i).symm))
|
||||
erw [← e2.prod_comp]
|
||||
let e1 := Equiv.ofBijective (c.insertLiftSome i j) (insertLiftSome_bijective i j)
|
||||
erw [← e1.prod_comp]
|
||||
rw [Fintype.prod_sum_type]
|
||||
simp only [Finset.univ_unique, PUnit.default_eq_unit, Nat.succ_eq_add_one, Finset.prod_singleton,
|
||||
Finset.univ_eq_attach]
|
||||
rfl
|
||||
|
||||
/-- Given a finite set of `Fin φs.length` the finite set of `(φs.insertIdx i φ).length`
|
||||
formed by mapping elements using `i.succAboveEmb` and `finCongr`. -/
|
||||
def insertListLiftFinset (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
(i : Fin φs.length.succ) (a : Finset (Fin φs.length)) :
|
||||
Finset (Fin (φs.insertIdx i φ).length) :=
|
||||
(a.map i.succAboveEmb).map (finCongr (insertIdx_length_fin φ φs i).symm).toEmbedding
|
||||
|
||||
@[simp]
|
||||
lemma self_not_mem_insertListLiftFinset (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
(i : Fin φs.length.succ) (a : Finset (Fin φs.length)) :
|
||||
Fin.cast (insertIdx_length_fin φ φs i).symm i ∉ insertListLiftFinset φ i a := by
|
||||
simp only [Nat.succ_eq_add_one, insertListLiftFinset, Finset.mem_map_equiv, finCongr_symm,
|
||||
finCongr_apply, Fin.cast_trans, Fin.cast_eq_self]
|
||||
simp only [Finset.mem_map, Fin.succAboveEmb_apply, not_exists, not_and]
|
||||
intro x
|
||||
exact fun a => Fin.succAbove_ne i x
|
||||
|
||||
lemma succAbove_mem_insertListLiftFinset (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
(i : Fin φs.length.succ) (a : Finset (Fin φs.length)) (j : Fin φs.length) :
|
||||
Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove j) ∈ insertListLiftFinset φ i a ↔
|
||||
j ∈ a := by
|
||||
simp only [insertListLiftFinset, Finset.mem_map_equiv, finCongr_symm, finCongr_apply,
|
||||
Fin.cast_trans, Fin.cast_eq_self]
|
||||
simp only [Finset.mem_map, Fin.succAboveEmb_apply]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
obtain ⟨x, hx1, hx2⟩ := h
|
||||
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at hx2
|
||||
simp_all
|
||||
· intro h
|
||||
use j
|
||||
|
||||
lemma insert_fin_eq_self (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
(i : Fin φs.length.succ) (j : Fin (List.insertIdx i φ φs).length) :
|
||||
j = Fin.cast (insertIdx_length_fin φ φs i).symm i
|
||||
∨ ∃ k, j = Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove k) := by
|
||||
obtain ⟨k, hk⟩ := (finCongr (insertIdx_length_fin φ φs i).symm).surjective j
|
||||
subst hk
|
||||
by_cases hi : k = i
|
||||
· simp [hi]
|
||||
· simp only [Nat.succ_eq_add_one, ← Fin.exists_succAbove_eq_iff] at hi
|
||||
obtain ⟨z, hk⟩ := hi
|
||||
subst hk
|
||||
right
|
||||
use z
|
||||
rfl
|
||||
|
||||
lemma insertList_uncontractedList_none_map (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) :
|
||||
List.map (List.insertIdx (↑i) φ φs).get (insertList φ φs c i none).uncontractedList =
|
||||
List.insertIdx (c.uncontractedListOrderPos i) φ (List.map φs.get c.uncontractedList) := by
|
||||
simp only [Nat.succ_eq_add_one, insertList]
|
||||
rw [congr_uncontractedList]
|
||||
erw [uncontractedList_extractEquiv_symm_none]
|
||||
rw [orderedInsert_succAboveEmb_uncontractedList_eq_insertIdx]
|
||||
rw [insertIdx_map, insertIdx_map]
|
||||
congr 1
|
||||
· simp
|
||||
rw [List.map_map, List.map_map]
|
||||
congr
|
||||
conv_rhs => rw [get_eq_insertIdx_succAbove φ φs i]
|
||||
rfl
|
||||
|
||||
lemma insertLift_sum (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
(i : Fin φs.length.succ) [AddCommMonoid M] (f : WickContraction (φs.insertIdx i φ).length → M) :
|
||||
∑ c, f c = ∑ (c : WickContraction φs.length), ∑ (k : Option (c.uncontracted)),
|
||||
f (insertList φ φs c i k) := by
|
||||
rw [sum_extractEquiv_congr (finCongr (insertIdx_length_fin φ φs i).symm i) f
|
||||
(insertIdx_length_fin φ φs i)]
|
||||
rfl
|
||||
|
||||
end WickContraction
|
162
HepLean/PerturbationTheory/WickContraction/Involutions.lean
Normal file
162
HepLean/PerturbationTheory/WickContraction/Involutions.lean
Normal file
|
@ -0,0 +1,162 @@
|
|||
/-
|
||||
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.Uncontracted
|
||||
import HepLean.PerturbationTheory.Algebras.StateAlgebra.TimeOrder
|
||||
import HepLean.PerturbationTheory.Algebras.OperatorAlgebra.TimeContraction
|
||||
import HepLean.PerturbationTheory.WickContraction.InsertList
|
||||
/-!
|
||||
|
||||
# Involution associated with a contraction
|
||||
|
||||
-/
|
||||
|
||||
open FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
open FieldStatistic
|
||||
|
||||
/-- The involution of `Fin n` associated with a Wick contraction `c : WickContraction n` as follows.
|
||||
If `i : Fin n` is contracted in `c` then it is taken to its dual, otherwise
|
||||
it is taken to itself. -/
|
||||
def toInvolution : {f : Fin n → Fin n // Function.Involutive f} :=
|
||||
⟨fun i => if h : (c.getDual? i).isSome then (c.getDual? i).get h else i, by
|
||||
intro i
|
||||
by_cases h : (c.getDual? i).isSome
|
||||
· simp [h]
|
||||
· simp [h]⟩
|
||||
|
||||
/-- The Wick contraction formed by an involution `f` of `Fin n` by taking as the contracted
|
||||
sets of the contraction the orbits of `f` of cardinality `2`. -/
|
||||
def fromInvolution (f : {f : Fin n → Fin n // Function.Involutive f}) : WickContraction n :=
|
||||
⟨Finset.univ.filter (fun a => a.card = 2 ∧ ∃ i, {i, f.1 i} = a), by
|
||||
intro a
|
||||
simp only [Finset.mem_filter, Finset.mem_univ, true_and, and_imp, forall_exists_index]
|
||||
intro h1 _ _
|
||||
exact h1, by
|
||||
intro a ha b hb
|
||||
simp only [Finset.mem_filter, Finset.mem_univ, true_and] at ha hb
|
||||
obtain ⟨i, hai⟩ := ha.2
|
||||
subst hai
|
||||
obtain ⟨j, hbj⟩ := hb.2
|
||||
subst hbj
|
||||
by_cases h : i = j
|
||||
· subst h
|
||||
simp
|
||||
· by_cases hi : i = f.1 j
|
||||
· subst hi
|
||||
simp only [Finset.disjoint_insert_right, Finset.mem_insert, Finset.mem_singleton, not_or,
|
||||
Finset.disjoint_singleton_right, true_or, not_true_eq_false, and_false, or_false]
|
||||
rw [f.2]
|
||||
rw [@Finset.pair_comm]
|
||||
· apply Or.inr
|
||||
simp only [Finset.disjoint_insert_right, Finset.mem_insert, Finset.mem_singleton, not_or,
|
||||
Finset.disjoint_singleton_right]
|
||||
apply And.intro
|
||||
· apply And.intro
|
||||
· exact fun a => h a.symm
|
||||
· by_contra hn
|
||||
subst hn
|
||||
rw [f.2 i] at hi
|
||||
simp at hi
|
||||
· apply And.intro
|
||||
· exact fun a => hi a.symm
|
||||
· rw [Function.Injective.eq_iff]
|
||||
exact fun a => h (id (Eq.symm a))
|
||||
exact Function.Involutive.injective f.2⟩
|
||||
|
||||
@[simp]
|
||||
lemma fromInvolution_getDual?_isSome (f : {f : Fin n → Fin n // Function.Involutive f})
|
||||
(i : Fin n) : ((fromInvolution f).getDual? i).isSome ↔ f.1 i ≠ i := by
|
||||
rw [getDual?_isSome_iff]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
obtain ⟨a, ha⟩ := h
|
||||
have ha2 := a.2
|
||||
simp only [fromInvolution, Finset.mem_filter, Finset.mem_univ, true_and] at ha2
|
||||
obtain ⟨j, h⟩ := ha2.2
|
||||
rw [← h] at ha
|
||||
have hj : f.1 j ≠ j := by
|
||||
by_contra hn
|
||||
rw [hn] at h
|
||||
rw [← h] at ha2
|
||||
simp at ha2
|
||||
simp only [Finset.mem_insert, Finset.mem_singleton] at ha
|
||||
rcases ha with ha | ha
|
||||
· subst ha
|
||||
exact hj
|
||||
· subst ha
|
||||
rw [f.2]
|
||||
exact id (Ne.symm hj)
|
||||
· intro hi
|
||||
use ⟨{i, f.1 i}, by
|
||||
simp only [fromInvolution, Finset.mem_filter, Finset.mem_univ, exists_apply_eq_apply,
|
||||
and_true, true_and]
|
||||
rw [Finset.card_pair (id (Ne.symm hi))]⟩
|
||||
simp
|
||||
|
||||
lemma fromInvolution_getDual?_eq_some (f : {f : Fin n → Fin n // Function.Involutive f}) (i : Fin n)
|
||||
(h : ((fromInvolution f).getDual? i).isSome) :
|
||||
((fromInvolution f).getDual? i) = some (f.1 i) := by
|
||||
rw [@getDual?_eq_some_iff_mem]
|
||||
simp only [fromInvolution, Finset.mem_filter, Finset.mem_univ, exists_apply_eq_apply, and_true,
|
||||
true_and]
|
||||
apply Finset.card_pair
|
||||
simp only [fromInvolution_getDual?_isSome, ne_eq] at h
|
||||
exact fun a => h (id (Eq.symm a))
|
||||
|
||||
@[simp]
|
||||
lemma fromInvolution_getDual?_get (f : {f : Fin n → Fin n // Function.Involutive f}) (i : Fin n)
|
||||
(h : ((fromInvolution f).getDual? i).isSome) :
|
||||
((fromInvolution f).getDual? i).get h = (f.1 i) := by
|
||||
have h1 := fromInvolution_getDual?_eq_some f i h
|
||||
exact Option.get_of_mem h h1
|
||||
|
||||
lemma toInvolution_fromInvolution : fromInvolution c.toInvolution = c := by
|
||||
apply Subtype.eq
|
||||
simp only [fromInvolution, toInvolution]
|
||||
ext a
|
||||
simp only [Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
obtain ⟨i, hi⟩ := h.2
|
||||
split at hi
|
||||
· subst hi
|
||||
simp
|
||||
· simp only [Finset.mem_singleton, Finset.insert_eq_of_mem] at hi
|
||||
subst hi
|
||||
simp at h
|
||||
· intro ha
|
||||
apply And.intro (c.2.1 a ha)
|
||||
use c.fstFieldOfContract ⟨a, ha⟩
|
||||
simp only [fstFieldOfContract_getDual?, Option.isSome_some, ↓reduceDIte, Option.get_some]
|
||||
change _ = (⟨a, ha⟩ : c.1).1
|
||||
conv_rhs => rw [finset_eq_fstFieldOfContract_sndFieldOfContract]
|
||||
|
||||
lemma fromInvolution_toInvolution (f : {f : Fin n → Fin n // Function.Involutive f}) :
|
||||
(fromInvolution f).toInvolution = f := by
|
||||
apply Subtype.eq
|
||||
funext i
|
||||
simp only [toInvolution, ne_eq, dite_not]
|
||||
split
|
||||
· rename_i h
|
||||
simp
|
||||
· rename_i h
|
||||
simp only [fromInvolution_getDual?_isSome, ne_eq, Decidable.not_not] at h
|
||||
exact id (Eq.symm h)
|
||||
|
||||
/-- The equivalence btween Wick contractions for `n` and involutions of `Fin n`.
|
||||
The involution of `Fin n` associated with a Wick contraction `c : WickContraction n` as follows.
|
||||
If `i : Fin n` is contracted in `c` then it is taken to its dual, otherwise
|
||||
it is taken to itself. -/
|
||||
def equivInvolution : WickContraction n ≃ {f : Fin n → Fin n // Function.Involutive f} where
|
||||
toFun := toInvolution
|
||||
invFun := fromInvolution
|
||||
left_inv := toInvolution_fromInvolution
|
||||
right_inv := fromInvolution_toInvolution
|
||||
|
||||
end WickContraction
|
64
HepLean/PerturbationTheory/WickContraction/IsFull.lean
Normal file
64
HepLean/PerturbationTheory/WickContraction/IsFull.lean
Normal file
|
@ -0,0 +1,64 @@
|
|||
/-
|
||||
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.Mathematics.Fin.Involutions
|
||||
import HepLean.PerturbationTheory.WickContraction.Involutions
|
||||
/-!
|
||||
|
||||
# Full contraction
|
||||
|
||||
We say that a contraction is full if it has no uncontracted fields.
|
||||
|
||||
-/
|
||||
|
||||
open FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
open FieldStatistic
|
||||
|
||||
/-- A contraction is full if there are no uncontracted fields, i.e. the finite set
|
||||
of uncontracted fields is empty. -/
|
||||
def IsFull : Prop := c.uncontracted = ∅
|
||||
|
||||
/-- The condition on whether or not a contraction is full is decidable. -/
|
||||
instance : Decidable (IsFull c) := decEq c.uncontracted ∅
|
||||
|
||||
lemma isFull_iff_equivInvolution_no_fixed_point :
|
||||
IsFull c ↔ ∀ (i : Fin n), (equivInvolution c).1 i ≠ i := by
|
||||
simp only [IsFull, ne_eq]
|
||||
rw [Finset.eq_empty_iff_forall_not_mem]
|
||||
simp [equivInvolution, toInvolution, uncontracted]
|
||||
|
||||
/-- The equivalence between full contractions and fixed-point free involutions. -/
|
||||
def isFullInvolutionEquiv : {c : WickContraction n // IsFull c} ≃
|
||||
{f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} where
|
||||
toFun c := ⟨equivInvolution c.1, by
|
||||
apply And.intro (equivInvolution c.1).2
|
||||
rw [← isFull_iff_equivInvolution_no_fixed_point]
|
||||
exact c.2⟩
|
||||
invFun f := ⟨equivInvolution.symm ⟨f.1, f.2.1⟩, by
|
||||
rw [isFull_iff_equivInvolution_no_fixed_point]
|
||||
simpa using f.2.2⟩
|
||||
left_inv c := by simp
|
||||
right_inv f := by simp
|
||||
|
||||
open Nat
|
||||
|
||||
/-- If `n` is even then the number of full contractions is `(n-1)!!`. -/
|
||||
theorem card_of_isfull_even (he : Even n) :
|
||||
Fintype.card {c : WickContraction n // IsFull c} = (n - 1)‼ := by
|
||||
rw [Fintype.card_congr (isFullInvolutionEquiv)]
|
||||
exact HepLean.Fin.involutionNoFixed_card_even n he
|
||||
|
||||
/-- If `n` is odd then there are no full contractions. This is because
|
||||
there will always be at least one element unpaired. -/
|
||||
theorem card_of_isfull_odd (ho : Odd n) :
|
||||
Fintype.card {c : WickContraction n // IsFull c} = 0 := by
|
||||
rw [Fintype.card_congr (isFullInvolutionEquiv)]
|
||||
exact HepLean.Fin.involutionNoFixed_card_odd n ho
|
||||
|
||||
end WickContraction
|
1143
HepLean/PerturbationTheory/WickContraction/Sign.lean
Normal file
1143
HepLean/PerturbationTheory/WickContraction/Sign.lean
Normal file
File diff suppressed because it is too large
Load diff
165
HepLean/PerturbationTheory/WickContraction/TimeContract.lean
Normal file
165
HepLean/PerturbationTheory/WickContraction/TimeContract.lean
Normal file
|
@ -0,0 +1,165 @@
|
|||
/-
|
||||
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.Sign
|
||||
import HepLean.PerturbationTheory.Algebras.OperatorAlgebra.TimeContraction
|
||||
/-!
|
||||
|
||||
# Time contractions
|
||||
|
||||
-/
|
||||
|
||||
open FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
|
||||
/-!
|
||||
|
||||
## Time contract.
|
||||
|
||||
-/
|
||||
|
||||
/-- Given a Wick contraction `c` associated with a list `φs`, the
|
||||
product of all time-contractions of pairs of contracted elements in `φs`,
|
||||
as a member of the center of `𝓞.A`. -/
|
||||
noncomputable def timeContract (𝓞 : 𝓕.OperatorAlgebra) {φs : List 𝓕.States}
|
||||
(c : WickContraction φs.length) :
|
||||
Subalgebra.center ℂ 𝓞.A :=
|
||||
∏ (a : c.1), ⟨𝓞.timeContract (φs.get (c.fstFieldOfContract a)) (φs.get (c.sndFieldOfContract a)),
|
||||
𝓞.timeContract_mem_center _ _⟩
|
||||
|
||||
@[simp]
|
||||
lemma timeContract_insertList_none (𝓞 : 𝓕.OperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) :
|
||||
(c.insertList φ φs i none).timeContract 𝓞 = c.timeContract 𝓞 := by
|
||||
rw [timeContract, insertList_none_prod_contractions]
|
||||
congr
|
||||
ext a
|
||||
simp
|
||||
|
||||
lemma timeConract_insertList_some (𝓞 : 𝓕.OperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
|
||||
(c.insertList φ φs i (some j)).timeContract 𝓞 =
|
||||
(if i < i.succAbove j then
|
||||
⟨𝓞.timeContract φ φs[j.1], 𝓞.timeContract_mem_center _ _⟩
|
||||
else ⟨𝓞.timeContract φs[j.1] φ, 𝓞.timeContract_mem_center _ _⟩) * c.timeContract 𝓞 := by
|
||||
rw [timeContract, insertList_some_prod_contractions]
|
||||
congr 1
|
||||
· simp only [Nat.succ_eq_add_one, insertList_fstFieldOfContract_some_incl, finCongr_apply,
|
||||
List.get_eq_getElem, insertList_sndFieldOfContract_some_incl, Fin.getElem_fin]
|
||||
split
|
||||
· simp
|
||||
· simp
|
||||
· congr
|
||||
ext a
|
||||
simp
|
||||
|
||||
open FieldStatistic
|
||||
|
||||
lemma timeConract_insertList_some_eq_mul_contractStateAtIndex_lt
|
||||
(𝓞 : 𝓕.OperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (k : c.uncontracted)
|
||||
(ht : 𝓕.timeOrderRel φ φs[k.1]) (hik : i < i.succAbove k) :
|
||||
(c.insertList φ φs i (some k)).timeContract 𝓞 =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (c.uncontracted.filter (fun x => x < k))⟩)
|
||||
• (𝓞.contractStateAtIndex φ (List.map φs.get c.uncontractedList)
|
||||
((uncontractedStatesEquiv φs c) (some k)) * c.timeContract 𝓞) := by
|
||||
rw [timeConract_insertList_some]
|
||||
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
|
||||
OperatorAlgebra.contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
|
||||
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
|
||||
List.getElem_map, uncontractedList_getElem_uncontractedFinEquiv_symm, List.get_eq_getElem,
|
||||
Algebra.smul_mul_assoc]
|
||||
· simp only [hik, ↓reduceIte, MulMemClass.coe_mul]
|
||||
rw [𝓞.timeContract_of_timeOrderRel]
|
||||
trans (1 : ℂ) • (𝓞.crAnF ((CrAnAlgebra.superCommute
|
||||
(CrAnAlgebra.anPart (StateAlgebra.ofState φ))) (CrAnAlgebra.ofState φs[k.1])) *
|
||||
↑(timeContract 𝓞 c))
|
||||
· simp
|
||||
simp only [smul_smul]
|
||||
congr
|
||||
have h1 : ofList 𝓕.statesStatistic (List.take (↑(c.uncontractedFinEquiv.symm k))
|
||||
(List.map φs.get c.uncontractedList))
|
||||
= (𝓕 |>ₛ ⟨φs.get, (Finset.filter (fun x => x < k) c.uncontracted)⟩) := by
|
||||
simp only [ofFinset]
|
||||
congr
|
||||
rw [← List.map_take]
|
||||
congr
|
||||
rw [take_uncontractedFinEquiv_symm]
|
||||
rw [filter_uncontractedList]
|
||||
rw [h1]
|
||||
simp only [exchangeSign_mul_self]
|
||||
· exact ht
|
||||
|
||||
lemma timeConract_insertList_some_eq_mul_contractStateAtIndex_not_lt
|
||||
(𝓞 : 𝓕.OperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (i : Fin φs.length.succ) (k : c.uncontracted)
|
||||
(ht : ¬ 𝓕.timeOrderRel φs[k.1] φ) (hik : ¬ i < i.succAbove k) :
|
||||
(c.insertList φ φs i (some k)).timeContract 𝓞 =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (c.uncontracted.filter (fun x => x ≤ k))⟩)
|
||||
• (𝓞.contractStateAtIndex φ (List.map φs.get c.uncontractedList)
|
||||
((uncontractedStatesEquiv φs c) (some k)) * c.timeContract 𝓞) := by
|
||||
rw [timeConract_insertList_some]
|
||||
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
|
||||
OperatorAlgebra.contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
|
||||
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
|
||||
List.getElem_map, uncontractedList_getElem_uncontractedFinEquiv_symm, List.get_eq_getElem,
|
||||
Algebra.smul_mul_assoc]
|
||||
simp only [hik, ↓reduceIte, MulMemClass.coe_mul]
|
||||
rw [𝓞.timeContract_of_not_timeOrderRel, 𝓞.timeContract_of_timeOrderRel]
|
||||
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc, smul_smul]
|
||||
congr
|
||||
have h1 : ofList 𝓕.statesStatistic (List.take (↑(c.uncontractedFinEquiv.symm k))
|
||||
(List.map φs.get c.uncontractedList))
|
||||
= (𝓕 |>ₛ ⟨φs.get, (Finset.filter (fun x => x < k) c.uncontracted)⟩) := by
|
||||
simp only [ofFinset]
|
||||
congr
|
||||
rw [← List.map_take]
|
||||
congr
|
||||
rw [take_uncontractedFinEquiv_symm, filter_uncontractedList]
|
||||
rw [h1]
|
||||
trans 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, {k.1}⟩)
|
||||
· rw [exchangeSign_symm, ofFinset_singleton]
|
||||
simp
|
||||
rw [← map_mul]
|
||||
congr
|
||||
rw [ofFinset_union]
|
||||
congr
|
||||
ext a
|
||||
simp only [Finset.mem_singleton, Finset.mem_sdiff, Finset.mem_union, Finset.mem_filter,
|
||||
Finset.mem_inter, not_and, not_lt, and_imp]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
subst h
|
||||
simp
|
||||
· intro h
|
||||
have h1 := h.1
|
||||
rcases h1 with h1 | h1
|
||||
· have h2' := h.2 h1.1 h1.2 h1.1
|
||||
omega
|
||||
· have h2' := h.2 h1.1 (by omega) h1.1
|
||||
omega
|
||||
have ht := IsTotal.total (r := timeOrderRel) φs[k.1] φ
|
||||
simp_all only [Fin.getElem_fin, Nat.succ_eq_add_one, not_lt, false_or]
|
||||
exact ht
|
||||
|
||||
lemma timeContract_of_not_gradingCompliant (𝓞 : 𝓕.OperatorAlgebra) (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (h : ¬ GradingCompliant φs c) :
|
||||
c.timeContract 𝓞 = 0 := by
|
||||
rw [timeContract]
|
||||
simp only [GradingCompliant, Fin.getElem_fin, Subtype.forall, not_forall] at h
|
||||
obtain ⟨a, ha⟩ := h
|
||||
obtain ⟨ha, ha2⟩ := ha
|
||||
apply Finset.prod_eq_zero (i := ⟨a, ha⟩)
|
||||
simp only [Finset.univ_eq_attach, Finset.mem_attach]
|
||||
apply Subtype.eq
|
||||
simp only [List.get_eq_getElem, ZeroMemClass.coe_zero]
|
||||
rw [OperatorAlgebra.timeContract_zero_of_diff_grade]
|
||||
simp [ha2]
|
||||
|
||||
end WickContraction
|
67
HepLean/PerturbationTheory/WickContraction/Uncontracted.lean
Normal file
67
HepLean/PerturbationTheory/WickContraction/Uncontracted.lean
Normal file
|
@ -0,0 +1,67 @@
|
|||
/-
|
||||
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.Basic
|
||||
/-!
|
||||
|
||||
# Uncontracted elements
|
||||
|
||||
-/
|
||||
open FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
|
||||
/-- Given a Wick contraction, the finset of elements of `Fin n` which are not contracted. -/
|
||||
def uncontracted : Finset (Fin n) := Finset.filter (fun i => c.getDual? i = none) (Finset.univ)
|
||||
|
||||
lemma congr_uncontracted {n m : ℕ} (c : WickContraction n) (h : n = m) :
|
||||
(c.congr h).uncontracted = Finset.map (finCongr h).toEmbedding c.uncontracted := by
|
||||
subst h
|
||||
simp
|
||||
|
||||
/-- 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 :=
|
||||
Equiv.optionCongr (Equiv.subtypeEquivRight (by rw [h]; simp))
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedCongr_none {c c': WickContraction n} (h : c = c') :
|
||||
(uncontractedCongr h) none = none := by
|
||||
simp [uncontractedCongr]
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedCongr_some {c c': WickContraction n} (h : c = c') (i : c.uncontracted) :
|
||||
(uncontractedCongr h) (some i) = some (Equiv.subtypeEquivRight (by rw [h]; simp) i) := by
|
||||
simp [uncontractedCongr]
|
||||
|
||||
lemma mem_uncontracted_iff_not_contracted (i : Fin n) :
|
||||
i ∈ c.uncontracted ↔ ∀ p ∈ c.1, i ∉ p := by
|
||||
simp only [uncontracted, getDual?, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
apply Iff.intro
|
||||
· intro h p hp
|
||||
have hp := c.2.1 p hp
|
||||
rw [Finset.card_eq_two] at hp
|
||||
obtain ⟨a, b, ha, hb, hab⟩ := hp
|
||||
rw [Fin.find_eq_none_iff] at h
|
||||
by_contra hn
|
||||
simp only [Finset.mem_insert, Finset.mem_singleton] at hn
|
||||
rcases hn with hn | hn
|
||||
· subst hn
|
||||
exact h b hp
|
||||
· subst hn
|
||||
rw [Finset.pair_comm] at hp
|
||||
exact h a hp
|
||||
· intro h
|
||||
rw [Fin.find_eq_none_iff]
|
||||
by_contra hn
|
||||
simp only [not_forall, Decidable.not_not] at hn
|
||||
obtain ⟨j, hj⟩ := hn
|
||||
apply h {i, j} hj
|
||||
simp
|
||||
|
||||
end WickContraction
|
490
HepLean/PerturbationTheory/WickContraction/UncontractedList.lean
Normal file
490
HepLean/PerturbationTheory/WickContraction/UncontractedList.lean
Normal file
|
@ -0,0 +1,490 @@
|
|||
/-
|
||||
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.ExtractEquiv
|
||||
/-!
|
||||
|
||||
# List of uncontracted elements
|
||||
|
||||
-/
|
||||
|
||||
open FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} (c : WickContraction n)
|
||||
open HepLean.List
|
||||
open HepLean.Fin
|
||||
|
||||
/-!
|
||||
|
||||
## Some properties of lists of fin
|
||||
|
||||
-/
|
||||
|
||||
lemma fin_list_sorted_monotone_sorted {n m : ℕ} (l: List (Fin n)) (hl : l.Sorted (· ≤ ·))
|
||||
(f : Fin n → Fin m) (hf : StrictMono f) : ((List.map f l)).Sorted (· ≤ ·) := by
|
||||
induction l with
|
||||
| nil => simp
|
||||
| cons a l ih =>
|
||||
simp only [List.map_cons, List.sorted_cons, List.mem_map, forall_exists_index, and_imp,
|
||||
forall_apply_eq_imp_iff₂]
|
||||
apply And.intro
|
||||
· simp only [List.sorted_cons] at hl
|
||||
intro b hb
|
||||
have hl1 := hl.1 b hb
|
||||
exact (StrictMono.le_iff_le hf).mpr hl1
|
||||
· simp only [List.sorted_cons] at hl
|
||||
exact ih hl.2
|
||||
|
||||
lemma fin_list_sorted_succAboveEmb_sorted (l: List (Fin n)) (hl : l.Sorted (· ≤ ·))
|
||||
(i : Fin n.succ) : ((List.map i.succAboveEmb l)).Sorted (· ≤ ·) := by
|
||||
apply fin_list_sorted_monotone_sorted
|
||||
exact hl
|
||||
simp only [Fin.coe_succAboveEmb]
|
||||
exact Fin.strictMono_succAbove i
|
||||
|
||||
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)
|
||||
| [], _, _ => by simp
|
||||
| a :: l, hl, i => by
|
||||
simp only [List.sorted_cons] at hl
|
||||
by_cases ha : a < i
|
||||
· conv_lhs => rw [fin_list_sorted_split l hl.2 i]
|
||||
rw [← List.cons_append]
|
||||
rw [List.filter_cons_of_pos, List.filter_cons_of_neg]
|
||||
simp only [decide_eq_true_eq, not_le, ha]
|
||||
simp [ha]
|
||||
· have hx : List.filter (fun x => decide (x.1 < i)) (a :: l) = [] := by
|
||||
simp only [ha, decide_false, Bool.false_eq_true, not_false_eq_true, List.filter_cons_of_neg,
|
||||
List.filter_eq_nil_iff, decide_eq_true_eq, not_lt]
|
||||
intro b hb
|
||||
have hb' := hl.1 b hb
|
||||
omega
|
||||
simp only [hx, List.nil_append]
|
||||
rw [List.filter_cons_of_pos]
|
||||
simp only [List.cons.injEq, true_and]
|
||||
have hl' := fin_list_sorted_split l hl.2 i
|
||||
have hx : List.filter (fun x => decide (x.1 < i)) (l) = [] := by
|
||||
simp only [List.filter_eq_nil_iff, decide_eq_true_eq, not_lt]
|
||||
intro b hb
|
||||
have hb' := hl.1 b hb
|
||||
omega
|
||||
simp only [hx, List.nil_append] at hl'
|
||||
conv_lhs => rw [hl']
|
||||
simp only [decide_eq_true_eq]
|
||||
omega
|
||||
|
||||
lemma fin_list_sorted_indexOf_filter_le_mem :
|
||||
(l : List (Fin n)) → (hl : l.Sorted (· ≤ ·)) → (i : Fin n) →
|
||||
(hl : i ∈ l) →
|
||||
List.indexOf i (List.filter (fun x => decide (↑i ≤ ↑x)) l) = 0
|
||||
| [], _, _, _ => by simp
|
||||
| a :: l, hl, i, hi => by
|
||||
simp only [List.sorted_cons] at hl
|
||||
by_cases ha : i ≤ a
|
||||
· simp only [ha, decide_true, List.filter_cons_of_pos]
|
||||
have ha : a = i := by
|
||||
simp only [List.mem_cons] at hi
|
||||
rcases hi with hi | hi
|
||||
· subst hi
|
||||
rfl
|
||||
· have hl' := hl.1 i hi
|
||||
exact Fin.le_antisymm hl' ha
|
||||
subst ha
|
||||
simp
|
||||
· simp only [not_le] at ha
|
||||
rw [List.filter_cons_of_neg (by simpa using ha)]
|
||||
rw [fin_list_sorted_indexOf_filter_le_mem l hl.2]
|
||||
simp only [List.mem_cons] at hi
|
||||
rcases hi with hi | hi
|
||||
· omega
|
||||
· exact hi
|
||||
|
||||
lemma fin_list_sorted_indexOf_mem :
|
||||
(l : List (Fin n)) → (hl : l.Sorted (· ≤ ·)) → (i : Fin n) →
|
||||
(hi : i ∈ l) →
|
||||
l.indexOf i = (l.filter (fun x => x.1 < i.1)).length := by
|
||||
intro l hl i hi
|
||||
conv_lhs => rw [fin_list_sorted_split l hl i]
|
||||
rw [List.indexOf_append_of_not_mem]
|
||||
erw [fin_list_sorted_indexOf_filter_le_mem l hl i hi]
|
||||
· simp
|
||||
· simp
|
||||
|
||||
lemma orderedInsert_of_fin_list_sorted :
|
||||
(l : List (Fin n)) → (hl : l.Sorted (· ≤ ·)) → (i : Fin n) →
|
||||
List.orderedInsert (· ≤ ·) i l = l.filter (fun x => x.1 < i.1) ++
|
||||
i :: l.filter (fun x => i.1 ≤ x.1)
|
||||
| [], _, _ => by simp
|
||||
| a :: l, hl, i => by
|
||||
simp only [List.sorted_cons] at hl
|
||||
by_cases ha : i ≤ a
|
||||
· simp only [List.orderedInsert, ha, ↓reduceIte, Fin.val_fin_lt, decide_eq_true_eq, not_lt,
|
||||
List.filter_cons_of_neg, Fin.val_fin_le, decide_true, List.filter_cons_of_pos]
|
||||
have h1 : List.filter (fun x => decide (↑x < ↑i)) l = [] := by
|
||||
simp only [List.filter_eq_nil_iff, decide_eq_true_eq, not_lt]
|
||||
intro a ha
|
||||
have ha' := hl.1 a ha
|
||||
omega
|
||||
have hl : l = List.filter (fun x => decide (i ≤ x)) l := by
|
||||
conv_lhs => rw [fin_list_sorted_split l hl.2 i]
|
||||
simp [h1]
|
||||
simp [← hl, h1]
|
||||
· simp only [List.orderedInsert, ha, ↓reduceIte, Fin.val_fin_lt, Fin.val_fin_le, decide_false,
|
||||
Bool.false_eq_true, not_false_eq_true, List.filter_cons_of_neg]
|
||||
rw [List.filter_cons_of_pos]
|
||||
rw [orderedInsert_of_fin_list_sorted l hl.2 i]
|
||||
simp only [Fin.val_fin_lt, Fin.val_fin_le, List.cons_append]
|
||||
simp only [decide_eq_true_eq]
|
||||
omega
|
||||
|
||||
lemma orderedInsert_eq_insertIdx_of_fin_list_sorted (l : List (Fin n)) (hl : l.Sorted (· ≤ ·))
|
||||
(i : Fin n) :
|
||||
List.orderedInsert (· ≤ ·) i l = l.insertIdx (l.filter (fun x => x.1 < i.1)).length i := by
|
||||
let n : Fin l.length.succ := ⟨(List.filter (fun x => decide (x < i)) l).length, by
|
||||
have h1 := l.length_filter_le (fun x => x.1 < i.1)
|
||||
simp only [Fin.val_fin_lt] at h1
|
||||
omega⟩
|
||||
simp only [Fin.val_fin_lt]
|
||||
conv_rhs => rw [insertIdx_eq_take_drop _ _ n]
|
||||
rw [orderedInsert_of_fin_list_sorted]
|
||||
congr
|
||||
· conv_rhs =>
|
||||
rhs
|
||||
rw [fin_list_sorted_split l hl i]
|
||||
simp [n]
|
||||
· conv_rhs =>
|
||||
rhs
|
||||
rw [fin_list_sorted_split l hl i]
|
||||
simp [n]
|
||||
exact hl
|
||||
|
||||
/-!
|
||||
|
||||
## Uncontracted List
|
||||
|
||||
-/
|
||||
|
||||
/-- Given a Wick contraction `c`, the ordered list of elements of `Fin n` which are not contracted,
|
||||
i.e. do not appera anywhere in `c.1`. -/
|
||||
def uncontractedList : List (Fin n) := List.filter (fun x => x ∈ c.uncontracted) (List.finRange n)
|
||||
|
||||
lemma uncontractedList_mem_iff (i : Fin n) :
|
||||
i ∈ c.uncontractedList ↔ i ∈ c.uncontracted := by
|
||||
simp [uncontractedList]
|
||||
|
||||
@[simp]
|
||||
lemma nil_zero_uncontractedList : (empty (n := 0)).uncontractedList = [] := by
|
||||
simp [empty, uncontractedList]
|
||||
|
||||
lemma congr_uncontractedList {n m : ℕ} (h : n = m) (c : WickContraction n) :
|
||||
((congr h) c).uncontractedList = List.map (finCongr h) c.uncontractedList := by
|
||||
subst h
|
||||
simp [congr]
|
||||
|
||||
lemma uncontractedList_get_mem_uncontracted (i : Fin c.uncontractedList.length) :
|
||||
c.uncontractedList.get i ∈ c.uncontracted := by
|
||||
rw [← uncontractedList_mem_iff]
|
||||
simp
|
||||
|
||||
lemma uncontractedList_sorted : List.Sorted (· ≤ ·) c.uncontractedList := by
|
||||
rw [uncontractedList]
|
||||
apply List.Sorted.filter
|
||||
rw [← List.ofFn_id]
|
||||
exact Monotone.ofFn_sorted fun ⦃a b⦄ a => a
|
||||
|
||||
lemma uncontractedList_nodup : c.uncontractedList.Nodup := by
|
||||
rw [uncontractedList]
|
||||
refine List.Nodup.filter (fun x => decide (x ∈ c.uncontracted)) ?_
|
||||
exact List.nodup_finRange n
|
||||
|
||||
lemma uncontractedList_toFinset (c : WickContraction n) :
|
||||
c.uncontractedList.toFinset = c.uncontracted := by
|
||||
simp [uncontractedList]
|
||||
|
||||
lemma uncontractedList_eq_sort (c : WickContraction n) :
|
||||
c.uncontractedList = c.uncontracted.sort (· ≤ ·) := by
|
||||
symm
|
||||
rw [← uncontractedList_toFinset]
|
||||
refine (List.toFinset_sort (α := Fin n) (· ≤ ·) ?_).mpr ?_
|
||||
· exact uncontractedList_nodup c
|
||||
· exact uncontractedList_sorted c
|
||||
|
||||
lemma uncontractedList_length_eq_card (c : WickContraction n) :
|
||||
c.uncontractedList.length = c.uncontracted.card := by
|
||||
rw [uncontractedList_eq_sort]
|
||||
exact Finset.length_sort fun x1 x2 => x1 ≤ x2
|
||||
|
||||
lemma uncontractedList_succAboveEmb_sorted (c : WickContraction n) (i : Fin n.succ) :
|
||||
((List.map i.succAboveEmb c.uncontractedList)).Sorted (· ≤ ·) := by
|
||||
apply fin_list_sorted_succAboveEmb_sorted
|
||||
exact uncontractedList_sorted c
|
||||
|
||||
lemma uncontractedList_succAboveEmb_nodup (c : WickContraction n) (i : Fin n.succ) :
|
||||
((List.map i.succAboveEmb c.uncontractedList)).Nodup := by
|
||||
refine List.Nodup.map ?_ ?_
|
||||
· exact Function.Embedding.injective i.succAboveEmb
|
||||
· exact uncontractedList_nodup c
|
||||
|
||||
lemma uncontractedList_succAboveEmb_toFinset (c : WickContraction n) (i : Fin n.succ) :
|
||||
(List.map i.succAboveEmb c.uncontractedList).toFinset =
|
||||
(Finset.map i.succAboveEmb c.uncontracted) := by
|
||||
ext a
|
||||
simp only [Fin.coe_succAboveEmb, List.mem_toFinset, List.mem_map, Finset.mem_map,
|
||||
Fin.succAboveEmb_apply]
|
||||
rw [← c.uncontractedList_toFinset]
|
||||
simp
|
||||
|
||||
lemma uncontractedList_succAboveEmb_eq_sort(c : WickContraction n) (i : Fin n.succ) :
|
||||
(List.map i.succAboveEmb c.uncontractedList) =
|
||||
(c.uncontracted.map i.succAboveEmb).sort (· ≤ ·) := by
|
||||
rw [← uncontractedList_succAboveEmb_toFinset]
|
||||
symm
|
||||
refine (List.toFinset_sort (α := Fin n.succ) (· ≤ ·) ?_).mpr ?_
|
||||
· exact uncontractedList_succAboveEmb_nodup c i
|
||||
· exact uncontractedList_succAboveEmb_sorted c i
|
||||
|
||||
lemma uncontractedList_succAboveEmb_eraseIdx_sorted (c : WickContraction n) (i : Fin n.succ)
|
||||
(k: ℕ) : ((List.map i.succAboveEmb c.uncontractedList).eraseIdx k).Sorted (· ≤ ·) := by
|
||||
apply HepLean.List.eraseIdx_sorted
|
||||
exact uncontractedList_succAboveEmb_sorted c i
|
||||
|
||||
lemma uncontractedList_succAboveEmb_eraseIdx_nodup (c : WickContraction n) (i : Fin n.succ) (k: ℕ) :
|
||||
((List.map i.succAboveEmb c.uncontractedList).eraseIdx k).Nodup := by
|
||||
refine List.Nodup.eraseIdx k ?_
|
||||
exact uncontractedList_succAboveEmb_nodup c i
|
||||
|
||||
lemma uncontractedList_succAboveEmb_eraseIdx_toFinset (c : WickContraction n) (i : Fin n.succ)
|
||||
(k : ℕ) (hk : k < c.uncontractedList.length) :
|
||||
((List.map i.succAboveEmb c.uncontractedList).eraseIdx k).toFinset =
|
||||
(c.uncontracted.map i.succAboveEmb).erase (i.succAboveEmb c.uncontractedList[k]) := by
|
||||
ext a
|
||||
simp only [Fin.coe_succAboveEmb, List.mem_toFinset, Fin.succAboveEmb_apply, Finset.mem_erase,
|
||||
ne_eq, Finset.mem_map]
|
||||
rw [mem_eraseIdx_nodup _ _ _ (by simpa using hk)]
|
||||
simp_all only [List.mem_map, List.getElem_map, ne_eq]
|
||||
apply Iff.intro
|
||||
· intro a_1
|
||||
simp_all only [not_false_eq_true, true_and]
|
||||
obtain ⟨left, right⟩ := a_1
|
||||
obtain ⟨w, h⟩ := left
|
||||
obtain ⟨left, right_1⟩ := h
|
||||
subst right_1
|
||||
use w
|
||||
simp_all [uncontractedList]
|
||||
· intro a_1
|
||||
simp_all only [not_false_eq_true, and_true]
|
||||
obtain ⟨left, right⟩ := a_1
|
||||
obtain ⟨w, h⟩ := right
|
||||
obtain ⟨left_1, right⟩ := h
|
||||
subst right
|
||||
use w
|
||||
simp_all [uncontractedList]
|
||||
exact uncontractedList_succAboveEmb_nodup c i
|
||||
|
||||
lemma uncontractedList_succAboveEmb_eraseIdx_eq_sort (c : WickContraction n) (i : Fin n.succ)
|
||||
(k : ℕ) (hk : k < c.uncontractedList.length) :
|
||||
((List.map i.succAboveEmb c.uncontractedList).eraseIdx k) =
|
||||
((c.uncontracted.map i.succAboveEmb).erase
|
||||
(i.succAboveEmb c.uncontractedList[k])).sort (· ≤ ·) := by
|
||||
rw [← uncontractedList_succAboveEmb_eraseIdx_toFinset]
|
||||
symm
|
||||
refine (List.toFinset_sort (α := Fin n.succ) (· ≤ ·) ?_).mpr ?_
|
||||
· exact uncontractedList_succAboveEmb_eraseIdx_nodup c i k
|
||||
· exact uncontractedList_succAboveEmb_eraseIdx_sorted c i k
|
||||
|
||||
lemma uncontractedList_succAbove_orderedInsert_sorted (c : WickContraction n) (i : Fin n.succ) :
|
||||
(List.orderedInsert (· ≤ ·) i
|
||||
(List.map i.succAboveEmb c.uncontractedList)).Sorted (· ≤ ·) := by
|
||||
refine List.Sorted.orderedInsert i (List.map (⇑i.succAboveEmb) c.uncontractedList) ?_
|
||||
exact uncontractedList_succAboveEmb_sorted c i
|
||||
|
||||
lemma uncontractedList_succAbove_orderedInsert_nodup (c : WickContraction n) (i : Fin n.succ) :
|
||||
(List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList)).Nodup := by
|
||||
have h1 : (List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList)).Perm
|
||||
(i :: List.map i.succAboveEmb c.uncontractedList) := by
|
||||
exact List.perm_orderedInsert (fun x1 x2 => x1 ≤ x2) i _
|
||||
apply List.Perm.nodup h1.symm
|
||||
simp only [Nat.succ_eq_add_one, List.nodup_cons, List.mem_map, not_exists,
|
||||
not_and]
|
||||
apply And.intro
|
||||
· intro x _
|
||||
exact Fin.succAbove_ne i x
|
||||
· exact uncontractedList_succAboveEmb_nodup c i
|
||||
|
||||
lemma uncontractedList_succAbove_orderedInsert_toFinset (c : WickContraction n) (i : Fin n.succ) :
|
||||
(List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList)).toFinset =
|
||||
(Insert.insert i (Finset.map i.succAboveEmb c.uncontracted)) := by
|
||||
ext a
|
||||
simp only [Nat.succ_eq_add_one, Fin.coe_succAboveEmb, List.mem_toFinset, List.mem_orderedInsert,
|
||||
List.mem_map, Finset.mem_insert, Finset.mem_map, Fin.succAboveEmb_apply]
|
||||
rw [← uncontractedList_toFinset]
|
||||
simp
|
||||
|
||||
lemma uncontractedList_succAbove_orderedInsert_eq_sort (c : WickContraction n) (i : Fin n.succ) :
|
||||
(List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList)) =
|
||||
(Insert.insert i (Finset.map i.succAboveEmb c.uncontracted)).sort (· ≤ ·) := by
|
||||
rw [← uncontractedList_succAbove_orderedInsert_toFinset]
|
||||
symm
|
||||
refine (List.toFinset_sort (α := Fin n.succ) (· ≤ ·) ?_).mpr ?_
|
||||
· exact uncontractedList_succAbove_orderedInsert_nodup c i
|
||||
· exact uncontractedList_succAbove_orderedInsert_sorted c i
|
||||
|
||||
lemma uncontractedList_extractEquiv_symm_none (c : WickContraction n) (i : Fin n.succ) :
|
||||
((extractEquiv i).symm ⟨c, none⟩).uncontractedList =
|
||||
List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList) := by
|
||||
rw [uncontractedList_eq_sort, extractEquiv_symm_none_uncontracted]
|
||||
rw [uncontractedList_succAbove_orderedInsert_eq_sort]
|
||||
|
||||
/-- Given a Wick contraction `c : WickContraction n` and a `Fin n.succ`, the number of elements
|
||||
of `c.uncontractedList` which are less then `i`.
|
||||
Suppose we want to insert into `c` at position `i`, then this is the position we would
|
||||
need to insert into `c.uncontractedList`. -/
|
||||
def uncontractedListOrderPos (c : WickContraction n) (i : Fin n.succ) : ℕ :=
|
||||
(List.filter (fun x => x.1 < i.1) c.uncontractedList).length
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedListOrderPos_lt_length_add_one (c : WickContraction n) (i : Fin n.succ) :
|
||||
c.uncontractedListOrderPos i < c.uncontractedList.length + 1 := by
|
||||
simp only [uncontractedListOrderPos, Nat.succ_eq_add_one]
|
||||
have h1 := c.uncontractedList.length_filter_le (fun x => x.1 < i.1)
|
||||
omega
|
||||
|
||||
lemma take_uncontractedListOrderPos_eq_filter (c : WickContraction n) (i : Fin n.succ) :
|
||||
(c.uncontractedList.take (c.uncontractedListOrderPos i)) =
|
||||
c.uncontractedList.filter (fun x => x.1 < i.1) := by
|
||||
nth_rewrite 1 [fin_list_sorted_split c.uncontractedList _ i]
|
||||
simp only [uncontractedListOrderPos, Nat.succ_eq_add_one, List.take_left']
|
||||
exact uncontractedList_sorted c
|
||||
|
||||
lemma take_uncontractedListOrderPos_eq_filter_sort (c : WickContraction n) (i : Fin n.succ) :
|
||||
(c.uncontractedList.take (c.uncontractedListOrderPos i)) =
|
||||
(c.uncontracted.filter (fun x => x.1 < i.1)).sort (· ≤ ·) := by
|
||||
rw [take_uncontractedListOrderPos_eq_filter]
|
||||
have h1 : (c.uncontractedList.filter (fun x => x.1 < i.1)).Sorted (· ≤ ·) := by
|
||||
apply List.Sorted.filter
|
||||
exact uncontractedList_sorted c
|
||||
have h2 : (c.uncontractedList.filter (fun x => x.1 < i.1)).Nodup := by
|
||||
refine List.Nodup.filter _ ?_
|
||||
exact uncontractedList_nodup c
|
||||
have h3 : (c.uncontractedList.filter (fun x => x.1 < i.1)).toFinset =
|
||||
(c.uncontracted.filter (fun x => x.1 < i.1)) := by
|
||||
rw [uncontractedList_eq_sort]
|
||||
simp
|
||||
rw [← h3]
|
||||
exact ((List.toFinset_sort (α := Fin n) (· ≤ ·) h2).mpr h1).symm
|
||||
|
||||
lemma orderedInsert_succAboveEmb_uncontractedList_eq_insertIdx (c : WickContraction n)
|
||||
(i : Fin n.succ) :
|
||||
(List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList)) =
|
||||
(List.map i.succAboveEmb c.uncontractedList).insertIdx (uncontractedListOrderPos c i) i := by
|
||||
rw [orderedInsert_eq_insertIdx_of_fin_list_sorted]
|
||||
swap
|
||||
exact uncontractedList_succAboveEmb_sorted c i
|
||||
congr 1
|
||||
simp only [Nat.succ_eq_add_one, Fin.val_fin_lt, Fin.coe_succAboveEmb, uncontractedListOrderPos]
|
||||
rw [List.filter_map]
|
||||
simp only [List.length_map]
|
||||
congr
|
||||
funext x
|
||||
simp only [Function.comp_apply, Fin.succAbove, decide_eq_decide]
|
||||
split
|
||||
simp only [Fin.lt_def, Fin.coe_castSucc]
|
||||
rename_i h
|
||||
simp_all only [Fin.lt_def, Fin.coe_castSucc, not_lt, Fin.val_succ]
|
||||
omega
|
||||
|
||||
/-- The equivalence between the positions of `c.uncontractedList` i.e. elements of
|
||||
`Fin (c.uncontractedList).length` and the finite set `c.uncontracted` considered as a finite type.
|
||||
-/
|
||||
def uncontractedFinEquiv (c : WickContraction n) :
|
||||
Fin (c.uncontractedList).length ≃ c.uncontracted where
|
||||
toFun i := ⟨c.uncontractedList.get i, c.uncontractedList_get_mem_uncontracted i⟩
|
||||
invFun i := ⟨List.indexOf i.1 c.uncontractedList,
|
||||
List.indexOf_lt_length.mpr ((c.uncontractedList_mem_iff i.1).mpr i.2)⟩
|
||||
left_inv i := by
|
||||
ext
|
||||
exact List.get_indexOf (uncontractedList_nodup c) _
|
||||
right_inv i := by
|
||||
ext
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedList_getElem_uncontractedFinEquiv_symm (k : c.uncontracted) :
|
||||
c.uncontractedList[(c.uncontractedFinEquiv.symm k).val] = k := by
|
||||
simp [uncontractedFinEquiv]
|
||||
|
||||
lemma uncontractedFinEquiv_symm_eq_filter_length (k : c.uncontracted) :
|
||||
(c.uncontractedFinEquiv.symm k).val =
|
||||
(List.filter (fun i => i < k.val) c.uncontractedList).length := by
|
||||
simp only [uncontractedFinEquiv, List.get_eq_getElem, Equiv.coe_fn_symm_mk]
|
||||
rw [fin_list_sorted_indexOf_mem]
|
||||
· simp
|
||||
· exact uncontractedList_sorted c
|
||||
· rw [uncontractedList_mem_iff]
|
||||
exact k.2
|
||||
|
||||
lemma take_uncontractedFinEquiv_symm (k : c.uncontracted) :
|
||||
c.uncontractedList.take (c.uncontractedFinEquiv.symm k).val =
|
||||
c.uncontractedList.filter (fun i => i < k.val) := by
|
||||
have hl := fin_list_sorted_split c.uncontractedList (uncontractedList_sorted c) k.val
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rw [hl]
|
||||
rw [uncontractedFinEquiv_symm_eq_filter_length]
|
||||
simp
|
||||
|
||||
/-- The equivalence between the type `Option c.uncontracted` for `WickContraction φs.length` and
|
||||
`Option (Fin (c.uncontractedList.map φs.get).length)`, that is optional positions of
|
||||
`c.uncontractedList.map φs.get` induced by `uncontractedFinEquiv`. -/
|
||||
def uncontractedStatesEquiv (φs : List 𝓕.States) (c : WickContraction φs.length) :
|
||||
Option c.uncontracted ≃ Option (Fin (c.uncontractedList.map φs.get).length) :=
|
||||
Equiv.optionCongr (c.uncontractedFinEquiv.symm.trans (finCongr (by simp)))
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedStatesEquiv_none (φs : List 𝓕.States) (c : WickContraction φs.length) :
|
||||
(uncontractedStatesEquiv φs c).toFun none = none := by
|
||||
simp [uncontractedStatesEquiv]
|
||||
|
||||
lemma uncontractedStatesEquiv_list_sum [AddCommMonoid α] (φs : List 𝓕.States)
|
||||
(c : WickContraction φs.length) (f : Option (Fin (c.uncontractedList.map φs.get).length) → α) :
|
||||
∑ (i : Option (Fin (c.uncontractedList.map φs.get).length)), f i =
|
||||
∑ (i : Option c.uncontracted), f (c.uncontractedStatesEquiv φs i) := by
|
||||
rw [(c.uncontractedStatesEquiv φs).sum_comp]
|
||||
|
||||
lemma uncontractedList_extractEquiv_symm_some (c : WickContraction n) (i : Fin n.succ)
|
||||
(k : c.uncontracted) : ((extractEquiv i).symm ⟨c, some k⟩).uncontractedList =
|
||||
((c.uncontractedList).map i.succAboveEmb).eraseIdx (c.uncontractedFinEquiv.symm k) := by
|
||||
rw [uncontractedList_eq_sort]
|
||||
rw [uncontractedList_succAboveEmb_eraseIdx_eq_sort]
|
||||
swap
|
||||
simp only [Fin.is_lt]
|
||||
congr
|
||||
simp only [Nat.succ_eq_add_one, extractEquiv, Equiv.coe_fn_symm_mk,
|
||||
uncontractedList_getElem_uncontractedFinEquiv_symm, Fin.succAboveEmb_apply]
|
||||
rw [insert_some_uncontracted]
|
||||
ext a
|
||||
simp
|
||||
|
||||
lemma filter_uncontractedList (c : WickContraction n) (p : Fin n → Prop) [DecidablePred p] :
|
||||
(c.uncontractedList.filter p) = (c.uncontracted.filter p).sort (· ≤ ·) := by
|
||||
have h1 : (c.uncontractedList.filter p).Sorted (· ≤ ·) := by
|
||||
apply List.Sorted.filter
|
||||
exact uncontractedList_sorted c
|
||||
have h2 : (c.uncontractedList.filter p).Nodup := by
|
||||
refine List.Nodup.filter _ ?_
|
||||
exact uncontractedList_nodup c
|
||||
have h3 : (c.uncontractedList.filter p).toFinset = (c.uncontracted.filter p) := by
|
||||
ext a
|
||||
simp only [List.toFinset_filter, decide_eq_true_eq, Finset.mem_filter, List.mem_toFinset,
|
||||
and_congr_left_iff]
|
||||
rw [uncontractedList_mem_iff]
|
||||
simp
|
||||
have hx := (List.toFinset_sort (· ≤ ·) h2).mpr h1
|
||||
rw [← hx, h3]
|
||||
|
||||
end WickContraction
|
364
HepLean/PerturbationTheory/WicksTheorem.lean
Normal file
364
HepLean/PerturbationTheory/WicksTheorem.lean
Normal file
|
@ -0,0 +1,364 @@
|
|||
/-
|
||||
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
|
||||
/-!
|
||||
|
||||
# Wick's theorem
|
||||
|
||||
This file contrains the time-dependent version of Wick's theorem
|
||||
for lists of fields containing both fermions and bosons.
|
||||
|
||||
Wick's theorem is related to Isserlis' theorem in mathematics.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct} {𝓞 : 𝓕.OperatorAlgebra}
|
||||
open CrAnAlgebra
|
||||
open StateAlgebra
|
||||
open OperatorAlgebra
|
||||
open HepLean.List
|
||||
open WickContraction
|
||||
open FieldStatistic
|
||||
|
||||
lemma insertList_none_normalOrder (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(i : Fin φs.length.succ) (c : WickContraction φs.length) :
|
||||
𝓞.crAnF (normalOrder (ofStateList (List.map (φs.insertIdx i φ).get
|
||||
(c.insertList φ φs i none).uncontractedList)))
|
||||
= 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, c.uncontracted.filter (fun x => i.succAbove x < i)⟩) •
|
||||
𝓞.crAnF (normalOrder (ofStateList (optionEraseZ (c.uncontractedList.map φs.get) φ none))) := by
|
||||
simp only [Nat.succ_eq_add_one, instCommGroup.eq_1, optionEraseZ]
|
||||
rw [crAnF_ofState_normalOrder_insert φ (c.uncontractedList.map φs.get)
|
||||
⟨(c.uncontractedListOrderPos i), by simp⟩, smul_smul]
|
||||
trans (1 : ℂ) • 𝓞.crAnF (normalOrder (ofStateList
|
||||
(List.map (List.insertIdx (↑i) φ φs).get (insertList φ φs c i none).uncontractedList)))
|
||||
· simp
|
||||
congr 1
|
||||
simp only [instCommGroup.eq_1]
|
||||
rw [← List.map_take, take_uncontractedListOrderPos_eq_filter]
|
||||
have h1 : (𝓕 |>ₛ List.map φs.get (List.filter (fun x => decide (↑x < i.1)) c.uncontractedList))
|
||||
= 𝓕 |>ₛ ⟨φs.get, (c.uncontracted.filter (fun x => x.val < i.1))⟩ := by
|
||||
simp only [Nat.succ_eq_add_one, ofFinset]
|
||||
congr
|
||||
rw [uncontractedList_eq_sort]
|
||||
have hdup : (List.filter (fun x => decide (x.1 < i.1))
|
||||
(Finset.sort (fun x1 x2 => x1 ≤ x2) c.uncontracted)).Nodup := by
|
||||
exact List.Nodup.filter _ (Finset.sort_nodup (fun x1 x2 => x1 ≤ x2) c.uncontracted)
|
||||
have hsort : (List.filter (fun x => decide (x.1 < i.1))
|
||||
(Finset.sort (fun x1 x2 => x1 ≤ x2) c.uncontracted)).Sorted (· ≤ ·) := by
|
||||
exact List.Sorted.filter _ (Finset.sort_sorted (fun x1 x2 => x1 ≤ x2) c.uncontracted)
|
||||
rw [← (List.toFinset_sort (· ≤ ·) hdup).mpr hsort]
|
||||
congr
|
||||
ext a
|
||||
simp
|
||||
rw [h1]
|
||||
simp only [Nat.succ_eq_add_one]
|
||||
have h2 : (Finset.filter (fun x => x.1 < i.1) c.uncontracted) =
|
||||
(Finset.filter (fun x => i.succAbove x < i) c.uncontracted) := by
|
||||
ext a
|
||||
simp only [Nat.succ_eq_add_one, Finset.mem_filter, and_congr_right_iff]
|
||||
intro ha
|
||||
simp only [Fin.succAbove]
|
||||
split
|
||||
· apply Iff.intro
|
||||
· intro h
|
||||
omega
|
||||
· intro h
|
||||
rename_i h
|
||||
rw [Fin.lt_def] at h
|
||||
simp only [Fin.coe_castSucc] at h
|
||||
omega
|
||||
· apply Iff.intro
|
||||
· intro h
|
||||
rename_i h'
|
||||
rw [Fin.lt_def]
|
||||
simp only [Fin.val_succ]
|
||||
rw [Fin.lt_def] at h'
|
||||
simp only [Fin.coe_castSucc, not_lt] at h'
|
||||
omega
|
||||
· intro h
|
||||
rename_i h
|
||||
rw [Fin.lt_def] at h
|
||||
simp only [Fin.val_succ] at h
|
||||
omega
|
||||
rw [h2]
|
||||
simp only [exchangeSign_mul_self]
|
||||
congr
|
||||
simp only [Nat.succ_eq_add_one]
|
||||
rw [insertList_uncontractedList_none_map]
|
||||
|
||||
lemma insertList_some_normalOrder (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(i : Fin φs.length.succ) (c : WickContraction φs.length) (k : c.uncontracted) :
|
||||
𝓞.crAnF (normalOrder (ofStateList (List.map (φs.insertIdx i φ).get
|
||||
(c.insertList φ φs i (some k)).uncontractedList)))
|
||||
= 𝓞.crAnF (normalOrder (ofStateList (optionEraseZ (c.uncontractedList.map φs.get) φ
|
||||
((uncontractedStatesEquiv φs c) k)))) := by
|
||||
simp only [Nat.succ_eq_add_one, insertList, optionEraseZ, uncontractedStatesEquiv,
|
||||
Equiv.optionCongr_apply, Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply,
|
||||
Fin.coe_cast]
|
||||
congr
|
||||
rw [congr_uncontractedList]
|
||||
erw [uncontractedList_extractEquiv_symm_some]
|
||||
simp only [Fin.coe_succAboveEmb, List.map_eraseIdx, List.map_map]
|
||||
congr
|
||||
conv_rhs => rw [get_eq_insertIdx_succAbove φ φs i]
|
||||
|
||||
lemma sign_timeContract_normalOrder_insertList_none (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(i : Fin φs.length.succ) (c : WickContraction φs.length) :
|
||||
(c.insertList φ φs i none).sign • (c.insertList φ φs i none).timeContract 𝓞
|
||||
* 𝓞.crAnF (normalOrder (ofStateList (List.map (φs.insertIdx i φ).get
|
||||
(c.insertList φ φs i none).uncontractedList))) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (Finset.univ.filter (fun k => i.succAbove k < i))⟩)
|
||||
• (c.sign • c.timeContract 𝓞 * 𝓞.crAnF (normalOrder
|
||||
(ofStateList (optionEraseZ (c.uncontractedList.map φs.get) φ none)))) := by
|
||||
by_cases hg : GradingCompliant φs c
|
||||
· rw [insertList_none_normalOrder, sign_insert_none]
|
||||
simp only [Nat.succ_eq_add_one, timeContract_insertList_none, instCommGroup.eq_1,
|
||||
Algebra.mul_smul_comm, Algebra.smul_mul_assoc, smul_smul]
|
||||
congr 1
|
||||
rw [← mul_assoc]
|
||||
congr 1
|
||||
rw [signInsertNone_eq_filterset _ _ _ _ hg, ← map_mul]
|
||||
congr
|
||||
rw [ofFinset_union]
|
||||
congr
|
||||
ext a
|
||||
simp only [Finset.mem_sdiff, Finset.mem_union, Finset.mem_filter, Finset.mem_univ, true_and,
|
||||
Finset.mem_inter, not_and, not_lt, and_imp]
|
||||
apply Iff.intro
|
||||
· intro ha
|
||||
have ha1 := ha.1
|
||||
rcases ha1 with ha1 | ha1
|
||||
· exact ha1.2
|
||||
· exact ha1.2
|
||||
· intro ha
|
||||
simp only [uncontracted, Finset.mem_filter, Finset.mem_univ, true_and, ha, and_true,
|
||||
forall_const]
|
||||
have hx : c.getDual? a = none ↔ ¬ (c.getDual? a).isSome := by
|
||||
simp
|
||||
rw [hx]
|
||||
simp only [Bool.not_eq_true, Bool.eq_false_or_eq_true_self, true_and]
|
||||
intro h1 h2
|
||||
simp_all
|
||||
· simp only [Nat.succ_eq_add_one, timeContract_insertList_none, Algebra.smul_mul_assoc,
|
||||
instCommGroup.eq_1]
|
||||
rw [timeContract_of_not_gradingCompliant]
|
||||
simp only [ZeroMemClass.coe_zero, zero_mul, smul_zero]
|
||||
exact hg
|
||||
|
||||
lemma sign_timeContract_normalOrder_insertList_some (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(i : Fin φs.length.succ) (c : WickContraction φs.length) (k : c.uncontracted)
|
||||
(hlt : ∀ (k : Fin φs.length), timeOrderRel φ φs[k])
|
||||
(hn : ∀ (k : Fin φs.length), i.succAbove k < i → ¬ timeOrderRel φs[k] φ) :
|
||||
(c.insertList φ φs i (some k)).sign • (c.insertList φ φs i (some k)).timeContract 𝓞
|
||||
* 𝓞.crAnF (normalOrder (ofStateList (List.map (φs.insertIdx i φ).get
|
||||
(c.insertList φ φs i (some k)).uncontractedList))) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (Finset.univ.filter (fun x => i.succAbove x < i))⟩)
|
||||
• (c.sign • (𝓞.contractStateAtIndex φ (List.map φs.get c.uncontractedList)
|
||||
((uncontractedStatesEquiv φs c) (some k)) * c.timeContract 𝓞)
|
||||
* 𝓞.crAnF (normalOrder (ofStateList (optionEraseZ (c.uncontractedList.map φs.get) φ
|
||||
((uncontractedStatesEquiv φs c) k))))) := by
|
||||
by_cases hg : GradingCompliant φs c ∧ (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[k.1])
|
||||
· by_cases hk : i.succAbove k < i
|
||||
· rw [WickContraction.timeConract_insertList_some_eq_mul_contractStateAtIndex_not_lt]
|
||||
swap
|
||||
· exact hn _ hk
|
||||
rw [insertList_some_normalOrder, sign_insert_some]
|
||||
simp only [instCommGroup.eq_1, smul_smul, Algebra.smul_mul_assoc]
|
||||
congr 1
|
||||
rw [mul_assoc, mul_comm (sign φs c), ← mul_assoc]
|
||||
congr 1
|
||||
exact signInsertSome_mul_filter_contracted_of_lt φ φs c i k hk hg
|
||||
· omega
|
||||
· have hik : i.succAbove ↑k ≠ i := by exact Fin.succAbove_ne i ↑k
|
||||
rw [WickContraction.timeConract_insertList_some_eq_mul_contractStateAtIndex_lt]
|
||||
swap
|
||||
· exact hlt _
|
||||
rw [insertList_some_normalOrder]
|
||||
rw [sign_insert_some]
|
||||
simp only [instCommGroup.eq_1, smul_smul, Algebra.smul_mul_assoc]
|
||||
congr 1
|
||||
rw [mul_assoc, mul_comm (sign φs c), ← mul_assoc]
|
||||
congr 1
|
||||
exact signInsertSome_mul_filter_contracted_of_not_lt φ φs c i k hk hg
|
||||
· omega
|
||||
· rw [timeConract_insertList_some]
|
||||
simp only [Fin.getElem_fin, not_and] at hg
|
||||
by_cases hg' : GradingCompliant φs c
|
||||
· have hg := hg hg'
|
||||
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, Algebra.smul_mul_assoc,
|
||||
instCommGroup.eq_1, contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
|
||||
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
|
||||
List.getElem_map, uncontractedList_getElem_uncontractedFinEquiv_symm, List.get_eq_getElem]
|
||||
by_cases h1 : i < i.succAbove ↑k
|
||||
· simp only [h1, ↓reduceIte, MulMemClass.coe_mul]
|
||||
rw [timeContract_zero_of_diff_grade]
|
||||
simp only [zero_mul, smul_zero]
|
||||
rw [crAnF_superCommute_anPart_ofState_diff_grade_zero]
|
||||
simp only [zero_mul, smul_zero]
|
||||
exact hg
|
||||
exact hg
|
||||
· simp only [h1, ↓reduceIte, MulMemClass.coe_mul]
|
||||
rw [timeContract_zero_of_diff_grade]
|
||||
simp only [zero_mul, smul_zero]
|
||||
rw [crAnF_superCommute_anPart_ofState_diff_grade_zero]
|
||||
simp only [zero_mul, smul_zero]
|
||||
exact hg
|
||||
exact fun a => hg (id (Eq.symm a))
|
||||
· rw [timeContract_of_not_gradingCompliant]
|
||||
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, mul_zero, ZeroMemClass.coe_zero, smul_zero,
|
||||
zero_mul, instCommGroup.eq_1]
|
||||
exact hg'
|
||||
|
||||
lemma mul_sum_contractions (φ : 𝓕.States) (φs : List 𝓕.States) (i : Fin φs.length.succ)
|
||||
(c : WickContraction φs.length) (hlt : ∀ (k : Fin φs.length), timeOrderRel φ φs[k])
|
||||
(hn : ∀ (k : Fin φs.length), i.succAbove k < i → ¬timeOrderRel φs[k] φ) :
|
||||
(c.sign • c.timeContract 𝓞) * 𝓞.crAnF ((CrAnAlgebra.ofState φ) *
|
||||
normalOrder (ofStateList (c.uncontractedList.map φs.get))) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (Finset.univ.filter (fun x => i.succAbove x < i))⟩) •
|
||||
∑ (k : Option (c.uncontracted)),
|
||||
((c.insertList φ φs i k).sign • (c.insertList φ φs i k).timeContract 𝓞
|
||||
* 𝓞.crAnF (normalOrder
|
||||
(ofStateList ((c.insertList φ φs i k).uncontractedList.map (φs.insertIdx i φ).get)))) := by
|
||||
rw [crAnF_ofState_mul_normalOrder_ofStatesList_eq_sum, Finset.mul_sum,
|
||||
uncontractedStatesEquiv_list_sum, Finset.smul_sum]
|
||||
simp only [Finset.univ_eq_attach, instCommGroup.eq_1,
|
||||
Nat.succ_eq_add_one, timeContract_insertList_none]
|
||||
congr 1
|
||||
funext n
|
||||
match n with
|
||||
| none =>
|
||||
rw [sign_timeContract_normalOrder_insertList_none]
|
||||
simp only [contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
|
||||
Equiv.coe_trans, Option.map_none', one_mul, Algebra.smul_mul_assoc, instCommGroup.eq_1,
|
||||
smul_smul]
|
||||
congr 1
|
||||
rw [← mul_assoc, exchangeSign_mul_self]
|
||||
simp
|
||||
| some n =>
|
||||
rw [sign_timeContract_normalOrder_insertList_some _ _ _ _ _
|
||||
(fun k => hlt k) (fun k a => hn k a)]
|
||||
simp only [uncontractedStatesEquiv, Equiv.optionCongr_apply, Equiv.coe_trans, Option.map_some',
|
||||
Function.comp_apply, finCongr_apply, Algebra.smul_mul_assoc, instCommGroup.eq_1, smul_smul]
|
||||
congr 1
|
||||
· rw [← mul_assoc, exchangeSign_mul_self]
|
||||
rw [one_mul]
|
||||
· rw [← mul_assoc]
|
||||
congr 1
|
||||
have ht := (WickContraction.timeContract 𝓞 c).prop
|
||||
rw [@Subalgebra.mem_center_iff] at ht
|
||||
rw [ht]
|
||||
|
||||
lemma wicks_theorem_congr {φs φs' : List 𝓕.States} (h : φs = φs') :
|
||||
∑ (c : WickContraction φs.length), (c.sign • c.timeContract 𝓞) *
|
||||
𝓞.crAnF (normalOrder (ofStateList (c.uncontractedList.map φs.get)))
|
||||
= ∑ (c : WickContraction φs'.length), (c.sign • c.timeContract 𝓞) *
|
||||
𝓞.crAnF (normalOrder (ofStateList (c.uncontractedList.map φs'.get))) := by
|
||||
subst h
|
||||
simp
|
||||
|
||||
/-- Wick's theorem for the empty list. -/
|
||||
lemma wicks_theorem_nil :
|
||||
𝓞.crAnF (ofStateAlgebra (timeOrder (ofList []))) = ∑ (c : WickContraction [].length),
|
||||
(c.sign [] • c.timeContract 𝓞) *
|
||||
𝓞.crAnF (normalOrder (ofStateList (c.uncontractedList.map [].get))) := by
|
||||
rw [timeOrder_ofList_nil]
|
||||
simp only [map_one, List.length_nil, Algebra.smul_mul_assoc]
|
||||
rw [sum_WickContraction_nil, nil_zero_uncontractedList]
|
||||
simp only [List.map_nil]
|
||||
have h1 : ofStateList (𝓕 := 𝓕) [] = CrAnAlgebra.ofCrAnList [] := by simp
|
||||
rw [h1, normalOrder_ofCrAnList]
|
||||
simp [WickContraction.timeContract, empty, sign]
|
||||
|
||||
lemma timeOrder_eq_maxTimeField_mul_finset (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
timeOrder (ofList (φ :: φs)) = 𝓢(𝓕 |>ₛ maxTimeField φ φs, 𝓕 |>ₛ ⟨(eraseMaxTimeField φ φs).get,
|
||||
(Finset.filter (fun x =>
|
||||
(maxTimeFieldPosFin φ φs).succAbove x < maxTimeFieldPosFin φ φs) Finset.univ)⟩) •
|
||||
StateAlgebra.ofState (maxTimeField φ φs) * timeOrder (ofList (eraseMaxTimeField φ φs)) := by
|
||||
rw [timeOrder_eq_maxTimeField_mul]
|
||||
congr 3
|
||||
apply FieldStatistic.ofList_perm
|
||||
nth_rewrite 1 [← List.finRange_map_get (φ :: φs)]
|
||||
simp only [List.length_cons, eraseMaxTimeField, insertionSortDropMinPos]
|
||||
rw [eraseIdx_get, ← List.map_take, ← List.map_map]
|
||||
refine List.Perm.map (φ :: φs).get ?_
|
||||
apply (List.perm_ext_iff_of_nodup _ _).mpr
|
||||
· intro i
|
||||
simp only [List.length_cons, maxTimeFieldPos, mem_take_finrange, Fin.val_fin_lt, List.mem_map,
|
||||
Finset.mem_sort, Finset.mem_filter, Finset.mem_univ, true_and, Function.comp_apply]
|
||||
refine Iff.intro (fun hi => ?_) (fun h => ?_)
|
||||
· have h2 := (maxTimeFieldPosFin φ φs).2
|
||||
simp only [eraseMaxTimeField, insertionSortDropMinPos, List.length_cons, Nat.succ_eq_add_one,
|
||||
maxTimeFieldPosFin, insertionSortMinPosFin] at h2
|
||||
use ⟨i, by omega⟩
|
||||
apply And.intro
|
||||
· simp only [Fin.succAbove, List.length_cons, Fin.castSucc_mk, maxTimeFieldPosFin,
|
||||
insertionSortMinPosFin, Nat.succ_eq_add_one, Fin.mk_lt_mk, Fin.val_fin_lt, Fin.succ_mk]
|
||||
rw [Fin.lt_def]
|
||||
split
|
||||
· simp only [Fin.val_fin_lt]
|
||||
omega
|
||||
· omega
|
||||
· simp only [Fin.succAbove, List.length_cons, Fin.castSucc_mk, Fin.succ_mk, Fin.ext_iff,
|
||||
Fin.coe_cast]
|
||||
split
|
||||
· simp
|
||||
· simp_all [Fin.lt_def]
|
||||
· obtain ⟨j, h1, h2⟩ := h
|
||||
subst h2
|
||||
simp only [Fin.lt_def, Fin.coe_cast]
|
||||
exact h1
|
||||
· exact List.Sublist.nodup (List.take_sublist _ _) <|
|
||||
List.nodup_finRange (φs.length + 1)
|
||||
· refine List.Nodup.map ?_ ?_
|
||||
· refine Function.Injective.comp ?hf.hg Fin.succAbove_right_injective
|
||||
exact Fin.cast_injective (eraseIdx_length (φ :: φs) (insertionSortMinPos timeOrderRel φ φs))
|
||||
· exact Finset.sort_nodup (fun x1 x2 => x1 ≤ x2)
|
||||
(Finset.filter (fun x => (maxTimeFieldPosFin φ φs).succAbove x < maxTimeFieldPosFin φ φs)
|
||||
Finset.univ)
|
||||
|
||||
/-- Wick's theorem for time-ordered products of bosonic and fermionic fields. -/
|
||||
theorem wicks_theorem : (φs : List 𝓕.States) → 𝓞.crAnF (ofStateAlgebra (timeOrder (ofList φs))) =
|
||||
∑ (c : WickContraction φs.length), (c.sign φs • c.timeContract 𝓞) *
|
||||
𝓞.crAnF (normalOrder (ofStateList (c.uncontractedList.map φs.get)))
|
||||
| [] => wicks_theorem_nil
|
||||
| φ :: φs => by
|
||||
have ih := wicks_theorem (eraseMaxTimeField φ φs)
|
||||
rw [timeOrder_eq_maxTimeField_mul_finset, map_mul, map_mul, ih, Finset.mul_sum]
|
||||
have h1 : φ :: φs =
|
||||
(eraseMaxTimeField φ φs).insertIdx (maxTimeFieldPosFin φ φs) (maxTimeField φ φs) := by
|
||||
simp only [eraseMaxTimeField, insertionSortDropMinPos, List.length_cons, Nat.succ_eq_add_one,
|
||||
maxTimeField, insertionSortMin, List.get_eq_getElem]
|
||||
erw [insertIdx_eraseIdx_fin]
|
||||
rw [wicks_theorem_congr h1]
|
||||
conv_rhs => rw [insertLift_sum]
|
||||
congr
|
||||
funext c
|
||||
have ht := Subalgebra.mem_center_iff.mp (Subalgebra.smul_mem (Subalgebra.center ℂ 𝓞.A)
|
||||
(WickContraction.timeContract 𝓞 c).2 (sign (eraseMaxTimeField φ φs) c))
|
||||
rw [map_smul, map_smul, Algebra.smul_mul_assoc, ← mul_assoc, ht, mul_assoc, ← map_mul]
|
||||
rw [ofStateAlgebra_ofState, mul_sum_contractions (𝓞 := 𝓞)
|
||||
(maxTimeField φ φs) (eraseMaxTimeField φ φs) (maxTimeFieldPosFin φ φs) c]
|
||||
trans (1 : ℂ) • ∑ k : Option { x // x ∈ c.uncontracted }, sign
|
||||
(List.insertIdx (↑(maxTimeFieldPosFin φ φs)) (maxTimeField φ φs) (eraseMaxTimeField φ φs))
|
||||
(insertList (maxTimeField φ φs) (eraseMaxTimeField φ φs) c (maxTimeFieldPosFin φ φs) k) •
|
||||
↑(WickContraction.timeContract 𝓞 (insertList (maxTimeField φ φs) (eraseMaxTimeField φ φs) c
|
||||
(maxTimeFieldPosFin φ φs) k)) *
|
||||
𝓞.crAnF (normalOrder (ofStateList (List.map (List.insertIdx (↑(maxTimeFieldPosFin φ φs))
|
||||
(maxTimeField φ φs) (eraseMaxTimeField φ φs)).get
|
||||
(insertList (maxTimeField φ φs) (eraseMaxTimeField φ φs) c
|
||||
(maxTimeFieldPosFin φ φs) k).uncontractedList)))
|
||||
swap
|
||||
· simp
|
||||
rw [smul_smul]
|
||||
simp only [instCommGroup.eq_1, exchangeSign_mul_self, Nat.succ_eq_add_one,
|
||||
Algebra.smul_mul_assoc, Fintype.sum_option, timeContract_insertList_none,
|
||||
Finset.univ_eq_attach, smul_add, one_smul]
|
||||
· exact fun k => timeOrder_maxTimeField _ _ k
|
||||
· exact fun k => lt_maxTimeFieldPosFin_not_timeOrder _ _ k
|
||||
termination_by φs => φs.length
|
||||
|
||||
end FieldStruct
|
|
@ -42,8 +42,7 @@ __BSM physics [🗂️](https://heplean.github.io/HepLean/docs/HepLean/BeyondThe
|
|||
|
||||
__Flavor physics [🗂️](https://heplean.github.io/HepLean/docs/HepLean/FlavorPhysics/CKMMatrix/Basic.html):__ Properties of the CKM matrix.
|
||||
|
||||
__Perturbation Theory [🗂️](https://heplean.github.io/HepLean/docs/HepLean/PerturbationTheory/Wick/Species.html) [🚧](https://heplean.github.io/HepLean/InformalGraph.html):__ Informal statements relating to Feynman diagrams, Wick contractions, Operator
|
||||
algebras.
|
||||
__Perturbation Theory [🗂️](https://heplean.github.io/HepLean/docs/HepLean/PerturbationTheory/WicksTheorem.html):__ Time-dependent version of Wick's theorem for both fermions and bosons.
|
||||
|
||||
## Associated media and publications
|
||||
- [📄](https://arxiv.org/abs/2405.08863) Joseph Tooby-Smith,
|
||||
|
|
|
@ -18,10 +18,6 @@ def pertubationTheory : NoteFile where
|
|||
abstract := "Notes on perturbation theory in quantum field theory."
|
||||
authors := ["Joseph Tooby-Smith"]
|
||||
files := [
|
||||
`HepLean.PerturbationTheory.Wick.Species,
|
||||
`HepLean.PerturbationTheory.Wick.Algebra,
|
||||
`HepLean.PerturbationTheory.Wick.Contract,
|
||||
`HepLean.PerturbationTheory.Wick.Theorem
|
||||
]
|
||||
|
||||
unsafe def main (_ : List String) : IO UInt32 := do
|
||||
|
|
|
@ -53,12 +53,24 @@ def expectedHepLeanImports : IO (Array Name) := do
|
|||
needed.push root
|
||||
pure needed
|
||||
|
||||
def listDif : (a: List String) → (b : List String) → (List String × List String)
|
||||
| [], [] => ([], [])
|
||||
| a :: as, [] => (a :: as, [])
|
||||
| [], b :: bs => ([], b :: bs)
|
||||
| a :: as, b :: bs =>
|
||||
if a = b then listDif as bs
|
||||
else (a :: (listDif as bs).1, b :: (listDif as bs).2)
|
||||
|
||||
/-- Checks whether an array `imports` is sorted after `Init` is removed. -/
|
||||
def arrayImportSorted (imports : Array Import) : IO Bool := do
|
||||
let X := (imports.map (fun x => x.module.toString)).filter (fun x => x != "Init")
|
||||
let mut warned := false
|
||||
if ! X = X.qsort (· < ·) then
|
||||
IO.print s!"Import file is not sorted. \n"
|
||||
let ldif := listDif X.toList (X.qsort (· < ·)).toList
|
||||
let lzip := List.zip ldif.1 ldif.2
|
||||
let lstring := String.intercalate "\n" (lzip.map (fun x => s!"{x.1} > {x.2}"))
|
||||
println! lstring
|
||||
warned := true
|
||||
pure warned
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue