Merge pull request #303 from HEPLean/FieldOpAlgebra
feat: Sorting property of Koszul Signs
This commit is contained in:
commit
d52abdd9d4
15 changed files with 2430 additions and 60 deletions
|
@ -129,6 +129,8 @@ import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.SuperCommute
|
|||
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.TimeOrder
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.NormalOrder
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.SuperCommute
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.TimeOrder
|
||||
import HepLean.PerturbationTheory.Algebras.ProtoOperatorAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Algebras.ProtoOperatorAlgebra.NormalOrder
|
||||
import HepLean.PerturbationTheory.Algebras.ProtoOperatorAlgebra.TimeContraction
|
||||
|
|
|
@ -118,6 +118,14 @@ lemma insertIdx_eraseIdx_fin {I : Type} :
|
|||
List.insertIdx_succ_cons, List.cons.injEq, true_and]
|
||||
exact insertIdx_eraseIdx_fin as ⟨n, Nat.lt_of_succ_lt_succ h⟩
|
||||
|
||||
lemma insertIdx_length_fst_append {I : Type} (φ : I) : (φs φs' : List I) →
|
||||
List.insertIdx φs.length φ (φs ++ φs') = (φs ++ φ :: φs')
|
||||
| [], φs' => by simp
|
||||
| φ' :: φs, φs' => by
|
||||
simp only [List.length_cons, List.cons_append, List.insertIdx_succ_cons, List.cons.injEq,
|
||||
true_and]
|
||||
exact insertIdx_length_fst_append φ φs φs'
|
||||
|
||||
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
|
||||
|
|
|
@ -91,4 +91,600 @@ lemma insertionSortMin_lt_mem_insertionSortDropMinPos_of_lt {α : Type} (r : α
|
|||
simp only [hl, Nat.succ_eq_add_one, Fin.val_eq_val, ne_eq]
|
||||
exact Fin.succAbove_ne (insertionSortMinPosFin r a l) i
|
||||
|
||||
lemma insertionSort_insertionSort {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] (l1 : List α) :
|
||||
List.insertionSort r (List.insertionSort r l1) = List.insertionSort r l1 := by
|
||||
apply List.Sorted.insertionSort_eq
|
||||
exact List.sorted_insertionSort r l1
|
||||
|
||||
lemma orderedInsert_commute {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] (a b : α) (hr : ¬ r a b) : (l : List α) →
|
||||
List.orderedInsert r a (List.orderedInsert r b l) =
|
||||
List.orderedInsert r b (List.orderedInsert r a l)
|
||||
| [] => by
|
||||
have hrb : r b a := by
|
||||
have ht := IsTotal.total (r := r) a b
|
||||
simp_all
|
||||
simp [hr, hrb]
|
||||
| c :: l => by
|
||||
have hrb : r b a := by
|
||||
have ht := IsTotal.total (r := r) a b
|
||||
simp_all
|
||||
simp only [List.orderedInsert]
|
||||
by_cases h : r a c
|
||||
· simp only [h, ↓reduceIte, List.orderedInsert.eq_2, hrb]
|
||||
rw [if_pos]
|
||||
simp only [List.orderedInsert, hr, ↓reduceIte, h]
|
||||
exact IsTrans.trans (r :=r) _ _ _ hrb h
|
||||
· simp only [h, ↓reduceIte, List.orderedInsert.eq_2]
|
||||
by_cases hbc : r b c
|
||||
· simp [hbc, hr, h]
|
||||
· simp only [hbc, ↓reduceIte, List.orderedInsert.eq_2, h, List.cons.injEq, true_and]
|
||||
exact orderedInsert_commute r a b hr l
|
||||
|
||||
lemma insertionSort_orderedInsert_append {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] (a : α) : (l1 l2 : List α) →
|
||||
List.insertionSort r (List.orderedInsert r a l1 ++ l2) = List.insertionSort r (a :: l1 ++ l2)
|
||||
| [], l2 => by
|
||||
simp
|
||||
| b :: l1, l2 => by
|
||||
conv_lhs => simp
|
||||
by_cases h : r a b
|
||||
· simp [h]
|
||||
conv_lhs => simp [h]
|
||||
rw [insertionSort_orderedInsert_append r a l1 l2]
|
||||
simp only [List.insertionSort, List.append_eq]
|
||||
rw [orderedInsert_commute r a b h]
|
||||
|
||||
lemma insertionSort_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] : (l1 l2 : List α) →
|
||||
List.insertionSort r (List.insertionSort r l1 ++ l2) = List.insertionSort r (l1 ++ l2)
|
||||
| [], l2 => by
|
||||
simp
|
||||
| a :: l1, l2 => by
|
||||
conv_lhs => simp
|
||||
rw [insertionSort_orderedInsert_append]
|
||||
simp only [List.insertionSort, List.append_eq]
|
||||
rw [insertionSort_insertionSort_append r l1 l2]
|
||||
|
||||
lemma insertionSort_append_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] : (l1 l2 l3 : List α) →
|
||||
List.insertionSort r (l1 ++ List.insertionSort r l2 ++ l3) =
|
||||
List.insertionSort r (l1 ++ l2 ++ l3)
|
||||
| [], l2, l3 => by
|
||||
simp only [List.nil_append]
|
||||
exact insertionSort_insertionSort_append r l2 l3
|
||||
| a :: l1, l2, l3 => by
|
||||
simp only [List.insertionSort, List.append_eq]
|
||||
rw [insertionSort_append_insertionSort_append r l1 l2 l3]
|
||||
|
||||
@[simp]
|
||||
lemma orderedInsert_length {α : Type} (r : α → α → Prop) [DecidableRel r] (a : α) (l : List α) :
|
||||
(List.orderedInsert r a l).length = (a :: l).length := by
|
||||
apply List.Perm.length_eq
|
||||
exact List.perm_orderedInsert r a l
|
||||
|
||||
lemma takeWhile_orderedInsert {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r]
|
||||
(a b : α) (hr : ¬ r a b) : (l : List α) →
|
||||
(List.takeWhile (fun c => !decide (r a c)) (List.orderedInsert r b l)).length =
|
||||
(List.takeWhile (fun c => !decide (r a c)) l).length + 1
|
||||
| [] => by
|
||||
simp [hr]
|
||||
| c :: l => by
|
||||
simp only [List.orderedInsert]
|
||||
by_cases h : r b c
|
||||
· simp only [h, ↓reduceIte]
|
||||
rw [List.takeWhile_cons_of_pos]
|
||||
simp only [List.length_cons]
|
||||
simp [hr]
|
||||
· simp only [h, ↓reduceIte]
|
||||
have hrba : r b a:= by
|
||||
have ht := IsTotal.total (r := r) a b
|
||||
simp_all
|
||||
have hl : ¬ r a c := by
|
||||
by_contra hn
|
||||
apply h
|
||||
exact IsTrans.trans _ _ _ hrba hn
|
||||
simp only [hl, decide_false, Bool.not_false, List.takeWhile_cons_of_pos, List.length_cons,
|
||||
add_left_inj]
|
||||
exact takeWhile_orderedInsert r a b hr l
|
||||
|
||||
lemma takeWhile_orderedInsert' {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r]
|
||||
(a b : α) (hr : ¬ r a b) : (l : List α) →
|
||||
(List.takeWhile (fun c => !decide (r b c)) (List.orderedInsert r a l)).length =
|
||||
(List.takeWhile (fun c => !decide (r b c)) l).length
|
||||
| [] => by
|
||||
simp only [List.orderedInsert, List.takeWhile_nil, List.length_nil, List.length_eq_zero,
|
||||
List.takeWhile_eq_nil_iff, List.length_singleton, zero_lt_one, Fin.zero_eta, Fin.isValue,
|
||||
List.get_eq_getElem, Fin.val_eq_zero, List.getElem_cons_zero, Bool.not_eq_eq_eq_not,
|
||||
Bool.not_true, decide_eq_false_iff_not, Decidable.not_not, forall_const]
|
||||
have ht := IsTotal.total (r := r) a b
|
||||
simp_all
|
||||
| c :: l => by
|
||||
have hrba : r b a:= by
|
||||
have ht := IsTotal.total (r := r) a b
|
||||
simp_all
|
||||
simp only [List.orderedInsert]
|
||||
by_cases h : r b c
|
||||
· simp only [h, decide_true, Bool.not_true, Bool.false_eq_true, not_false_eq_true,
|
||||
List.takeWhile_cons_of_neg, List.length_nil, List.length_eq_zero, List.takeWhile_eq_nil_iff,
|
||||
List.get_eq_getElem, Bool.not_eq_eq_eq_not, decide_eq_false_iff_not, Decidable.not_not]
|
||||
by_cases hac : r a c
|
||||
· simp [hac, hrba]
|
||||
· simp [hac, h]
|
||||
· by_cases hac : r a c
|
||||
· refine False.elim (h ?_)
|
||||
exact IsTrans.trans _ _ _ hrba hac
|
||||
· simp only [hac, ↓reduceIte, h, decide_false, Bool.not_false, List.takeWhile_cons_of_pos,
|
||||
List.length_cons, add_left_inj]
|
||||
exact takeWhile_orderedInsert' r a b hr l
|
||||
|
||||
lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] (a b : α) (hr : ¬ r a b) (n : ℕ) : (l : List α) →
|
||||
(hn : n + 2 < (a :: b :: l).length) →
|
||||
insertionSortEquiv r (a :: b :: l) ⟨n + 2, hn⟩ = (finCongr (by simp))
|
||||
(insertionSortEquiv r (b :: a :: l) ⟨n + 2, hn⟩) := by
|
||||
have hrba : r b a:= by
|
||||
have ht := IsTotal.total (r := r) a b
|
||||
simp_all
|
||||
intro l hn
|
||||
simp only [List.insertionSort, List.length_cons, insertionSortEquiv, Nat.succ_eq_add_one,
|
||||
equivCons_trans, Equiv.trans_apply, equivCons_succ, finCongr_apply]
|
||||
conv_lhs => erw [equivCons_succ]
|
||||
conv_rhs => erw [equivCons_succ]
|
||||
simp only [Equiv.toFun_as_coe]
|
||||
conv_lhs =>
|
||||
rhs
|
||||
rhs
|
||||
erw [orderedInsertEquiv_succ]
|
||||
conv_lhs => erw [orderedInsertEquiv_fin_succ]
|
||||
simp only [Fin.eta, Fin.coe_cast]
|
||||
conv_rhs =>
|
||||
rhs
|
||||
rhs
|
||||
erw [orderedInsertEquiv_succ]
|
||||
conv_rhs => erw [orderedInsertEquiv_fin_succ]
|
||||
ext
|
||||
simp only [Fin.coe_cast, Fin.eta, Fin.cast_trans]
|
||||
let a1 : Fin ((List.orderedInsert r b (List.insertionSort r l)).length + 1) :=
|
||||
⟨↑(orderedInsertPos r (List.orderedInsert r b (List.insertionSort r l)) a),
|
||||
orderedInsertPos_lt_length r (List.orderedInsert r b (List.insertionSort r l)) a⟩
|
||||
let b1 : Fin ((List.insertionSort r l).length + 1) :=
|
||||
⟨↑(orderedInsertPos r (List.insertionSort r l) b),
|
||||
orderedInsertPos_lt_length r (List.insertionSort r l) b⟩
|
||||
let a2 : Fin ((List.insertionSort r l).length + 1) :=
|
||||
⟨↑(orderedInsertPos r (List.insertionSort r l) a),
|
||||
orderedInsertPos_lt_length r (List.insertionSort r l) a⟩
|
||||
let b2 : Fin ((List.orderedInsert r a (List.insertionSort r l)).length + 1) :=
|
||||
⟨↑(orderedInsertPos r (List.orderedInsert r a (List.insertionSort r l)) b),
|
||||
orderedInsertPos_lt_length r (List.orderedInsert r a (List.insertionSort r l)) b⟩
|
||||
have ht : (List.takeWhile (fun c => !decide (r b c)) (List.insertionSort r l))
|
||||
= (List.takeWhile (fun c => !decide (r b c))
|
||||
((List.takeWhile (fun c => !decide (r a c)) (List.insertionSort r l)))) := by
|
||||
rw [List.takeWhile_takeWhile]
|
||||
simp only [Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not, Bool.decide_and,
|
||||
decide_not]
|
||||
congr
|
||||
funext c
|
||||
simp only [Bool.iff_self_and, Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not]
|
||||
intro hbc hac
|
||||
refine hbc ?_
|
||||
exact IsTrans.trans _ _ _ hrba hac
|
||||
have ha1 : b1.1 ≤ a2.1 := by
|
||||
simp only [orderedInsertPos, decide_not]
|
||||
rw [ht]
|
||||
apply List.Sublist.length_le
|
||||
exact List.takeWhile_sublist fun c => !decide (r b c)
|
||||
have ha2 : a1.1 = a2.1 + 1 := by
|
||||
simp only [orderedInsertPos, decide_not]
|
||||
rw [takeWhile_orderedInsert]
|
||||
exact hr
|
||||
have hb : b1.1 = b2.1 := by
|
||||
simp only [orderedInsertPos, decide_not]
|
||||
rw [takeWhile_orderedInsert']
|
||||
exact hr
|
||||
let n := ((insertionSortEquiv r l) ⟨n, by simpa using hn⟩)
|
||||
change (a1.succAbove ⟨b1.succAbove n, _⟩).1 = (b2.succAbove ⟨a2.succAbove n, _⟩).1
|
||||
trans if (b1.succAbove n).1 < a1.1 then (b1.succAbove n).1 else (b1.succAbove n).1 + 1
|
||||
· rw [Fin.succAbove]
|
||||
simp only [Fin.castSucc_mk, Fin.lt_def, Fin.succ_mk]
|
||||
by_cases ha : (b1.succAbove n).1 < a1.1
|
||||
· simp [ha]
|
||||
· simp [ha]
|
||||
trans if (a2.succAbove n).1 < b2.1 then (a2.succAbove n).1 else (a2.succAbove n).1 + 1
|
||||
swap
|
||||
· conv_rhs => rw [Fin.succAbove]
|
||||
simp only [Fin.castSucc_mk, Fin.lt_def, Fin.succ_mk]
|
||||
by_cases ha : (a2.succAbove n).1 < b2.1
|
||||
· simp [ha]
|
||||
· simp [ha]
|
||||
have hbs1 : (b1.succAbove n).1 = if n.1 < b1.1 then n.1 else n.1 + 1 := by
|
||||
rw [Fin.succAbove]
|
||||
simp only [Fin.castSucc_mk, Fin.lt_def, Fin.succ_mk]
|
||||
by_cases ha : n.1 < b1.1
|
||||
· simp [ha]
|
||||
· simp [ha]
|
||||
have has2 : (a2.succAbove n).1 = if n.1 < a2.1 then n.1 else n.1 + 1 := by
|
||||
rw [Fin.succAbove]
|
||||
simp only [Fin.castSucc_mk, Fin.lt_def, Fin.succ_mk]
|
||||
by_cases ha : n.1 < a2.1
|
||||
· simp [ha]
|
||||
· simp [ha]
|
||||
rw [hbs1, has2, hb, ha2]
|
||||
have hnat (a2 b2 n : ℕ) (h : b2 ≤ a2) : (if (if ↑n < ↑b2 then ↑n else ↑n + 1) < ↑a2 + 1 then
|
||||
if ↑n < ↑b2 then ↑n else ↑n + 1
|
||||
else (if ↑n < ↑b2 then ↑n else ↑n + 1) + 1) =
|
||||
if (if ↑n < ↑a2 then ↑n else ↑n + 1) < ↑b2 then if ↑n < ↑a2 then ↑n else ↑n + 1
|
||||
else (if ↑n < ↑a2 then ↑n else ↑n + 1) + 1 := by
|
||||
by_cases hnb2 : n < b2
|
||||
· simp only [hnb2, ↓reduceIte]
|
||||
have h1 : n < a2 + 1 := by omega
|
||||
have h2 : n < a2 := by omega
|
||||
simp [h1, h2, hnb2]
|
||||
· simp only [hnb2, ↓reduceIte, add_lt_add_iff_right]
|
||||
by_cases ha2 : n < a2
|
||||
· simp [ha2, hnb2]
|
||||
· simp only [ha2, ↓reduceIte]
|
||||
rw [if_neg]
|
||||
omega
|
||||
apply hnat
|
||||
rw [← hb]
|
||||
exact ha1
|
||||
|
||||
lemma insertionSortEquiv_orderedInsert_append {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] (a a2 : α) : (l1 l2 : List α) →
|
||||
(insertionSortEquiv r (List.orderedInsert r a l1 ++ a2 :: l2) ⟨l1.length + 1, by
|
||||
simp⟩)
|
||||
= (finCongr (by
|
||||
simp only [List.insertionSort, List.append_eq, orderedInsert_length, List.length_cons,
|
||||
List.length_insertionSort, List.length_append]
|
||||
omega))
|
||||
((insertionSortEquiv r (a :: l1 ++ a2 :: l2)) ⟨l1.length + 1, by simp⟩)
|
||||
| [], l2 => by
|
||||
simp
|
||||
| b :: l1, l2 => by
|
||||
by_cases h : r a b
|
||||
· have h1 : (List.orderedInsert r a (b :: l1) ++ a2 :: l2) = (a :: b :: l1 ++ a2 :: l2) := by
|
||||
simp [h]
|
||||
rw [insertionSortEquiv_congr _ _ h1]
|
||||
simp
|
||||
· have h1 : (List.orderedInsert r a (b :: l1) ++ a2 :: l2) =
|
||||
(b :: List.orderedInsert r a (l1) ++ a2 :: l2) := by
|
||||
simp [h]
|
||||
rw [insertionSortEquiv_congr _ _ h1]
|
||||
simp only [List.orderedInsert.eq_2, List.cons_append, List.length_cons, List.insertionSort,
|
||||
Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.cast_mk,
|
||||
finCongr_apply]
|
||||
conv_lhs => simp [insertionSortEquiv]
|
||||
rw [insertionSortEquiv_orderedInsert_append r a]
|
||||
have hl : (List.insertionSort r (List.orderedInsert r a l1 ++ a2 :: l2)) =
|
||||
List.insertionSort r (a :: l1 ++ a2 :: l2) := by
|
||||
exact insertionSort_orderedInsert_append r a l1 (a2 :: l2)
|
||||
rw [orderedInsertEquiv_congr _ _ _ hl]
|
||||
simp only [List.length_cons, List.cons_append, List.insertionSort, finCongr_apply,
|
||||
Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.cast_succ_eq,
|
||||
Fin.cast_trans, Fin.cast_eq_self]
|
||||
change Fin.cast _
|
||||
((insertionSortEquiv r (b :: a :: (l1 ++ a2 :: l2))) ⟨l1.length + 2, by simp⟩) = _
|
||||
have hl : l1.length + 1 +1 = l1.length + 2 := by omega
|
||||
simp only [List.insertionSort, List.length_cons, hl]
|
||||
conv_rhs =>
|
||||
erw [insertionSortEquiv_commute _ _ _ h _ _]
|
||||
simp
|
||||
|
||||
lemma insertionSortEquiv_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] (a : α) : (l1 l2 : List α) →
|
||||
(insertionSortEquiv r (List.insertionSort r l1 ++ a :: l2) ⟨l1.length, by simp⟩)
|
||||
= finCongr (by simp) (insertionSortEquiv r (l1 ++ a :: l2) ⟨l1.length, by simp⟩)
|
||||
| [], l2 => by
|
||||
simp only [List.insertionSort, List.nil_append, List.length_cons, List.length_nil, Fin.zero_eta,
|
||||
finCongr_refl, Equiv.refl_apply]
|
||||
| b :: l1, l2 => by
|
||||
simp only [List.insertionSort, List.length_cons, List.cons_append, finCongr_apply]
|
||||
have hl := insertionSortEquiv_orderedInsert_append r b a (List.insertionSort r l1) l2
|
||||
simp only [List.length_insertionSort, List.cons_append, List.insertionSort, List.length_cons,
|
||||
finCongr_apply] at hl
|
||||
rw [hl]
|
||||
have ih := insertionSortEquiv_insertionSort_append r a l1 l2
|
||||
simp only [insertionSortEquiv, Nat.succ_eq_add_one, List.insertionSort, Equiv.trans_apply,
|
||||
equivCons_succ]
|
||||
rw [ih]
|
||||
have hl : (List.insertionSort r (List.insertionSort r l1 ++ a :: l2)) =
|
||||
(List.insertionSort r (l1 ++ a :: l2)) := by
|
||||
exact insertionSort_insertionSort_append r l1 (a :: l2)
|
||||
rw [orderedInsertEquiv_congr _ _ _ hl]
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## Insertion sort with equal fields
|
||||
|
||||
-/
|
||||
|
||||
lemma orderedInsert_filter_of_pos {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTrans α r] (a : α) (p : α → Prop) [DecidablePred p] (h : p a) : (l : List α) →
|
||||
(hl : l.Sorted r) →
|
||||
List.filter p (List.orderedInsert r a l) = List.orderedInsert r a (List.filter p l)
|
||||
| [], hl => by
|
||||
simp only [List.orderedInsert, List.filter_eq_self, List.mem_singleton, decide_eq_true_eq,
|
||||
forall_eq]
|
||||
exact h
|
||||
| b :: l, hl => by
|
||||
simp only [List.orderedInsert]
|
||||
by_cases hpb : p b <;> by_cases hab : r a b
|
||||
· simp only [hab, ↓reduceIte, hpb, decide_true, List.filter_cons_of_pos,
|
||||
List.orderedInsert.eq_2]
|
||||
rw [List.filter_cons_of_pos (by simp [h])]
|
||||
rw [List.filter_cons_of_pos (by simp [hpb])]
|
||||
· simp only [hab, ↓reduceIte]
|
||||
rw [List.filter_cons_of_pos (by simp [hpb])]
|
||||
rw [List.filter_cons_of_pos (by simp [hpb])]
|
||||
simp only [List.orderedInsert, hab, ↓reduceIte, List.cons.injEq, true_and]
|
||||
simp only [List.sorted_cons] at hl
|
||||
exact orderedInsert_filter_of_pos r a p h l hl.2
|
||||
· simp only [hab, ↓reduceIte]
|
||||
rw [List.filter_cons_of_pos (by simp [h]),
|
||||
List.filter_cons_of_neg (by simp [hpb])]
|
||||
rw [List.orderedInsert_eq_take_drop]
|
||||
have hl : List.takeWhile (fun b => decide ¬r a b)
|
||||
(List.filter (fun b => decide (p b)) l) = [] := by
|
||||
rw [List.takeWhile_eq_nil_iff]
|
||||
intro c hc
|
||||
simp only [List.get_eq_getElem, decide_not, Bool.not_eq_eq_eq_not, Bool.not_true,
|
||||
decide_eq_false_iff_not] at hc
|
||||
apply hc
|
||||
apply IsTrans.trans a b _ hab
|
||||
simp only [List.sorted_cons] at hl
|
||||
apply hl.1
|
||||
have hlf : (List.filter (fun b => decide (p b)) l)[0] ∈
|
||||
(List.filter (fun b => decide (p b)) l) := by
|
||||
exact List.getElem_mem c
|
||||
simp only [List.mem_filter, decide_eq_true_eq] at hlf
|
||||
exact hlf.1
|
||||
rw [hl]
|
||||
simp only [decide_not, List.nil_append, List.cons.injEq, true_and]
|
||||
conv_lhs => rw [← List.takeWhile_append_dropWhile (fun b => decide ¬r a b)
|
||||
(List.filter (fun b => decide (p b)) l)]
|
||||
rw [hl]
|
||||
simp
|
||||
· simp only [hab, ↓reduceIte]
|
||||
rw [List.filter_cons_of_neg (by simp [hpb])]
|
||||
rw [List.filter_cons_of_neg (by simp [hpb])]
|
||||
simp only [List.sorted_cons] at hl
|
||||
exact orderedInsert_filter_of_pos r a p h l hl.2
|
||||
|
||||
lemma orderedInsert_filter_of_neg {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
(a : α) (p : α → Prop) [DecidablePred p] (h : ¬ p a) (l : List α) :
|
||||
List.filter p (List.orderedInsert r a l) = (List.filter p l) := by
|
||||
rw [List.orderedInsert_eq_take_drop]
|
||||
simp only [decide_not, List.filter_append]
|
||||
rw [List.filter_cons_of_neg]
|
||||
rw [← List.filter_append]
|
||||
congr
|
||||
exact List.takeWhile_append_dropWhile (fun b => !decide (r a b)) l
|
||||
simp [h]
|
||||
|
||||
lemma insertionSort_filter {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r]
|
||||
[IsTrans α r] (p : α → Prop) [DecidablePred p] : (l : List α) →
|
||||
List.insertionSort r (List.filter p l) =
|
||||
List.filter p (List.insertionSort r l)
|
||||
| [] => by simp
|
||||
| a :: l => by
|
||||
simp only [List.insertionSort]
|
||||
by_cases h : p a
|
||||
· rw [orderedInsert_filter_of_pos]
|
||||
rw [List.filter_cons_of_pos]
|
||||
simp only [List.insertionSort]
|
||||
rw [insertionSort_filter]
|
||||
simp_all only [decide_true]
|
||||
simp_all only
|
||||
exact List.sorted_insertionSort r l
|
||||
· rw [orderedInsert_filter_of_neg]
|
||||
rw [List.filter_cons_of_neg]
|
||||
rw [insertionSort_filter]
|
||||
simp_all only [decide_false, Bool.false_eq_true, not_false_eq_true]
|
||||
exact h
|
||||
|
||||
lemma takeWhile_sorted_eq_filter {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTrans α r] (a : α) : (l : List α) → (hl : l.Sorted r) →
|
||||
List.takeWhile (fun c => ¬ r a c) l = List.filter (fun c => ¬ r a c) l
|
||||
| [], _ => by simp
|
||||
| b :: l, hl => by
|
||||
simp only [List.sorted_cons] at hl
|
||||
by_cases hb : ¬ r a b
|
||||
· simp only [decide_not, hb, decide_false, Bool.not_false, List.takeWhile_cons_of_pos,
|
||||
List.filter_cons_of_pos, List.cons.injEq, true_and]
|
||||
simpa using takeWhile_sorted_eq_filter r a l hl.2
|
||||
· simp_all only [Decidable.not_not, decide_not, decide_true, Bool.not_true, Bool.false_eq_true,
|
||||
not_false_eq_true, List.takeWhile_cons_of_neg, List.filter_cons_of_neg, List.nil_eq,
|
||||
List.filter_eq_nil_iff, Bool.not_eq_eq_eq_not, decide_eq_false_iff_not]
|
||||
intro c hc
|
||||
apply IsTrans.trans a b c hb
|
||||
exact hl.1 c hc
|
||||
|
||||
lemma dropWhile_sorted_eq_filter {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTrans α r] (a : α) : (l : List α) → (hl : l.Sorted r) →
|
||||
List.dropWhile (fun c => ¬ r a c) l = List.filter (fun c => r a c) l
|
||||
| [], _ => by simp
|
||||
| b :: l, hl => by
|
||||
simp only [List.sorted_cons] at hl
|
||||
by_cases hb : ¬ r a b
|
||||
· simp only [decide_not, hb, decide_false, Bool.not_false, List.dropWhile_cons_of_pos,
|
||||
Bool.false_eq_true, not_false_eq_true, List.filter_cons_of_neg]
|
||||
simpa using dropWhile_sorted_eq_filter r a l hl.2
|
||||
· simp_all only [Decidable.not_not, decide_not, decide_true, Bool.not_true, Bool.false_eq_true,
|
||||
not_false_eq_true, List.dropWhile_cons_of_neg, List.filter_cons_of_pos, List.cons.injEq,
|
||||
true_and]
|
||||
symm
|
||||
rw [List.filter_eq_self]
|
||||
intro c hc
|
||||
simp only [decide_eq_true_eq]
|
||||
apply IsTrans.trans a b c hb
|
||||
exact hl.1 c hc
|
||||
|
||||
lemma dropWhile_sorted_eq_filter_filter {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTrans α r] (a : α) :(l : List α) → (hl : l.Sorted r) →
|
||||
List.filter (fun c => r a c) l =
|
||||
List.filter (fun c => r a c ∧ r c a) l ++ List.filter (fun c => r a c ∧ ¬ r c a) l
|
||||
| [], _ => by
|
||||
simp
|
||||
| b :: l, hl => by
|
||||
simp only [List.sorted_cons] at hl
|
||||
by_cases hb : ¬ r a b
|
||||
· simp only [hb, decide_false, Bool.false_eq_true, not_false_eq_true, List.filter_cons_of_neg,
|
||||
Bool.decide_and, Bool.false_and, decide_not]
|
||||
simpa using dropWhile_sorted_eq_filter_filter r a l hl.2
|
||||
· simp_all only [Decidable.not_not, decide_true, List.filter_cons_of_pos, Bool.decide_and,
|
||||
decide_not]
|
||||
by_cases hba : r b a
|
||||
· simp only [hba, decide_true, Bool.not_true, Bool.and_false, Bool.false_eq_true,
|
||||
not_false_eq_true, List.filter_cons_of_neg]
|
||||
rw [List.filter_cons_of_pos]
|
||||
rw [dropWhile_sorted_eq_filter_filter]
|
||||
simp only [Bool.decide_and, decide_not, List.cons_append]
|
||||
exact hl.2
|
||||
simp_all
|
||||
· simp only [hba, decide_false, Bool.and_false, Bool.false_eq_true, not_false_eq_true,
|
||||
List.filter_cons_of_neg]
|
||||
have h1 : List.filter (fun c => decide (r a c) && decide (r c a)) l = [] := by
|
||||
rw [@List.filter_eq_nil_iff]
|
||||
intro c hc
|
||||
simp only [Bool.and_eq_true, decide_eq_true_eq, not_and]
|
||||
intro hac hca
|
||||
apply hba
|
||||
apply IsTrans.trans b c a _ hca
|
||||
exact hl.1 c hc
|
||||
rw [h1]
|
||||
rw [dropWhile_sorted_eq_filter_filter]
|
||||
simp only [Bool.decide_and, h1, decide_not, List.nil_append]
|
||||
rw [List.filter_cons_of_pos]
|
||||
simp_all only [List.filter_eq_nil_iff, Bool.and_eq_true, decide_eq_true_eq, not_and,
|
||||
decide_true, decide_false, Bool.not_false, Bool.and_self]
|
||||
exact hl.2
|
||||
|
||||
lemma filter_rel_eq_insertionSort {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] (a : α) : (l : List α) →
|
||||
List.filter (fun c => r a c ∧ r c a) (l.insertionSort r) =
|
||||
List.filter (fun c => r a c ∧ r c a) l
|
||||
| [] => by simp
|
||||
| b :: l => by
|
||||
simp only [List.insertionSort]
|
||||
by_cases h : r a b ∧ r b a
|
||||
· have hl := orderedInsert_filter_of_pos r b (fun c => r a c ∧ r c a) h
|
||||
(List.insertionSort r l) (by exact List.sorted_insertionSort r l)
|
||||
simp only [Bool.decide_and] at hl ⊢
|
||||
rw [hl]
|
||||
rw [List.orderedInsert_eq_take_drop]
|
||||
have ht : List.takeWhile (fun b_1 => decide ¬r b b_1)
|
||||
(List.filter (fun b => decide (r a b) && decide (r b a))
|
||||
(List.insertionSort r l)) = [] := by
|
||||
rw [List.takeWhile_eq_nil_iff]
|
||||
intro hl
|
||||
simp only [List.get_eq_getElem, decide_not, Bool.not_eq_eq_eq_not, Bool.not_true,
|
||||
decide_eq_false_iff_not, Decidable.not_not]
|
||||
have hx := List.getElem_mem hl
|
||||
simp only [List.mem_filter, List.mem_insertionSort, Bool.and_eq_true,
|
||||
decide_eq_true_eq] at hx
|
||||
apply IsTrans.trans b a _ h.2
|
||||
simp_all
|
||||
rw [ht]
|
||||
simp only [decide_not, List.nil_append]
|
||||
rw [List.filter_cons_of_pos]
|
||||
simp only [List.cons.injEq, true_and]
|
||||
have ih := filter_rel_eq_insertionSort r a l
|
||||
simp only [Bool.decide_and] at ih
|
||||
rw [← ih]
|
||||
have htd := List.takeWhile_append_dropWhile (fun b_1 => decide ¬r b b_1)
|
||||
(List.filter (fun b => decide (r a b) && decide (r b a)) (List.insertionSort r l))
|
||||
simp only [decide_not] at htd
|
||||
conv_rhs => rw [← htd]
|
||||
simp only [List.self_eq_append_left, List.takeWhile_eq_nil_iff, List.get_eq_getElem,
|
||||
Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not, Decidable.not_not]
|
||||
intro hl
|
||||
have hx := List.getElem_mem hl
|
||||
simp only [List.mem_filter, List.mem_insertionSort, Bool.and_eq_true, decide_eq_true_eq] at hx
|
||||
apply IsTrans.trans b a _ h.2
|
||||
simp_all only [decide_not, List.takeWhile_eq_nil_iff, List.get_eq_getElem,
|
||||
Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not, Decidable.not_not,
|
||||
List.takeWhile_append_dropWhile]
|
||||
simp_all
|
||||
· have hl := orderedInsert_filter_of_neg r b (fun c => r a c ∧ r c a) h (List.insertionSort r l)
|
||||
simp only [Bool.decide_and] at hl ⊢
|
||||
rw [hl]
|
||||
rw [List.filter_cons_of_neg]
|
||||
have ih := filter_rel_eq_insertionSort r a l
|
||||
simp_all only [not_and, Bool.decide_and]
|
||||
simpa using h
|
||||
|
||||
lemma insertionSort_of_eq_list {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] (a : α) (l1 l l2 : List α)
|
||||
(h : ∀ b ∈ l, r a b ∧ r b a) :
|
||||
List.insertionSort r (l1 ++ l ++ l2) =
|
||||
(List.takeWhile (fun c => ¬ r a c) ((l1 ++ l2).insertionSort r))
|
||||
++ (List.filter (fun c => r a c ∧ r c a) l1)
|
||||
++ l
|
||||
++ (List.filter (fun c => r a c ∧ r c a) l2)
|
||||
++ (List.filter (fun c => r a c ∧ ¬ r c a) ((l1 ++ l2).insertionSort r)) := by
|
||||
have hl : List.insertionSort r (l1 ++ l ++ l2) =
|
||||
List.takeWhile (fun c => ¬ r a c) ((l1 ++ l ++ l2).insertionSort r) ++
|
||||
List.dropWhile (fun c => ¬ r a c) ((l1 ++ l ++ l2).insertionSort r) := by
|
||||
exact (List.takeWhile_append_dropWhile (fun c => decide ¬r a c)
|
||||
(List.insertionSort r (l1 ++ l ++ l2))).symm
|
||||
have hlt : List.takeWhile (fun c => ¬ r a c) ((l1 ++ l ++ l2).insertionSort r)
|
||||
= List.takeWhile (fun c => ¬ r a c) ((l1 ++ l2).insertionSort r) := by
|
||||
rw [takeWhile_sorted_eq_filter, takeWhile_sorted_eq_filter]
|
||||
rw [← insertionSort_filter, ← insertionSort_filter]
|
||||
congr 1
|
||||
simp only [decide_not, List.append_assoc, List.filter_append, List.append_cancel_left_eq,
|
||||
List.append_left_eq_self, List.filter_eq_nil_iff, Bool.not_eq_eq_eq_not, Bool.not_true,
|
||||
decide_eq_false_iff_not, Decidable.not_not]
|
||||
exact fun b hb => (h b hb).1
|
||||
exact List.sorted_insertionSort r (l1 ++ l2)
|
||||
exact List.sorted_insertionSort r (l1 ++ l ++ l2)
|
||||
conv_lhs => rw [hl, hlt]
|
||||
simp only [decide_not, Bool.decide_and]
|
||||
simp only [List.append_assoc, List.append_cancel_left_eq]
|
||||
have h1 := dropWhile_sorted_eq_filter r a (List.insertionSort r (l1 ++ (l ++ l2)))
|
||||
simp only [decide_not] at h1
|
||||
rw [h1]
|
||||
rw [dropWhile_sorted_eq_filter_filter, filter_rel_eq_insertionSort]
|
||||
simp only [Bool.decide_and, List.filter_append, decide_not, List.append_assoc,
|
||||
List.append_cancel_left_eq]
|
||||
congr 1
|
||||
simp only [List.filter_eq_self, Bool.and_eq_true, decide_eq_true_eq]
|
||||
exact fun a a_1 => h a a_1
|
||||
congr 1
|
||||
have h1 := insertionSort_filter r (fun c => decide (r a c) && !decide (r c a)) (l1 ++ (l ++ l2))
|
||||
simp only [Bool.and_eq_true, decide_eq_true_eq, Bool.not_eq_eq_eq_not, Bool.not_true,
|
||||
decide_eq_false_iff_not, Bool.decide_and, decide_not, List.filter_append] at h1
|
||||
rw [← h1]
|
||||
have h2 := insertionSort_filter r (fun c => decide (r a c) && !decide (r c a)) (l1 ++ l2)
|
||||
simp only [Bool.and_eq_true, decide_eq_true_eq, Bool.not_eq_eq_eq_not, Bool.not_true,
|
||||
decide_eq_false_iff_not, Bool.decide_and, decide_not, List.filter_append] at h2
|
||||
rw [← h2]
|
||||
congr
|
||||
have hl : List.filter (fun b => decide (r a b) && !decide (r b a)) l = [] := by
|
||||
rw [@List.filter_eq_nil_iff]
|
||||
intro c hc
|
||||
simp_all
|
||||
rw [hl]
|
||||
simp only [List.nil_append]
|
||||
exact List.sorted_insertionSort r (l1 ++ (l ++ l2))
|
||||
exact List.sorted_insertionSort r (l1 ++ (l ++ l2))
|
||||
|
||||
lemma insertionSort_of_takeWhile_filter {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
[IsTotal α r] [IsTrans α r] (a : α) (l1 l2 : List α) :
|
||||
List.insertionSort r (l1 ++ l2) =
|
||||
(List.takeWhile (fun c => ¬ r a c) ((l1 ++ l2).insertionSort r))
|
||||
++ (List.filter (fun c => r a c ∧ r c a) l1)
|
||||
++ (List.filter (fun c => r a c ∧ r c a) l2)
|
||||
++ (List.filter (fun c => r a c ∧ ¬ r c a) ((l1 ++ l2).insertionSort r)) := by
|
||||
trans List.insertionSort r (l1 ++ [] ++ l2)
|
||||
simp only [List.append_nil]
|
||||
rw [insertionSort_of_eq_list r a l1 [] l2]
|
||||
simp only [decide_not, Bool.decide_and, List.append_nil, List.append_assoc]
|
||||
simp
|
||||
|
||||
end HepLean.List
|
||||
|
|
|
@ -3,8 +3,7 @@ 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.FieldSpecification.NormalOrder
|
||||
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.SuperCommute
|
||||
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
import Mathlib.RingTheory.GradedAlgebra.Basic
|
||||
/-!
|
||||
|
@ -22,9 +21,27 @@ namespace CrAnAlgebra
|
|||
noncomputable section
|
||||
|
||||
/-- The submodule of `CrAnAlgebra` spanned by lists of field statistic `f`. -/
|
||||
def statisticSubmodule (f : FieldStatistic) : Submodule ℂ 𝓕.CrAnAlgebra :=
|
||||
def statisticSubmodule (f : FieldStatistic) : Submodule ℂ 𝓕.CrAnAlgebra :=
|
||||
Submodule.span ℂ {a | ∃ φs, a = ofCrAnList φs ∧ (𝓕 |>ₛ φs) = f}
|
||||
|
||||
lemma ofCrAnList_mem_statisticSubmodule_of (φs : List 𝓕.CrAnStates) (f : FieldStatistic)
|
||||
(h : (𝓕 |>ₛ φs) = f) :
|
||||
ofCrAnList φs ∈ statisticSubmodule f := by
|
||||
refine Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl, h⟩⟩
|
||||
|
||||
lemma ofCrAnList_bosonic_or_fermionic (φs : List 𝓕.CrAnStates) :
|
||||
ofCrAnList φs ∈ statisticSubmodule bosonic ∨ ofCrAnList φs ∈ statisticSubmodule fermionic := by
|
||||
by_cases h : (𝓕 |>ₛ φs) = bosonic
|
||||
· left
|
||||
exact ofCrAnList_mem_statisticSubmodule_of φs bosonic h
|
||||
· right
|
||||
exact ofCrAnList_mem_statisticSubmodule_of φs fermionic (by simpa using h)
|
||||
|
||||
lemma ofCrAnState_bosonic_or_fermionic (φ : 𝓕.CrAnStates) :
|
||||
ofCrAnState φ ∈ statisticSubmodule bosonic ∨ ofCrAnState φ ∈ statisticSubmodule fermionic := by
|
||||
rw [← ofCrAnList_singleton]
|
||||
exact ofCrAnList_bosonic_or_fermionic [φ]
|
||||
|
||||
/-- The projection of an element of `CrAnAlgebra` onto it's bosonic part. -/
|
||||
def bosonicProj : 𝓕.CrAnAlgebra →ₗ[ℂ] statisticSubmodule (𝓕 := 𝓕) bosonic :=
|
||||
Basis.constr ofCrAnListBasis ℂ fun φs =>
|
||||
|
@ -102,7 +119,7 @@ lemma fermionicProj_ofCrAnList (φs : List 𝓕.CrAnStates) :
|
|||
lemma fermionicProj_ofCrAnList_if_bosonic (φs : List 𝓕.CrAnStates) :
|
||||
fermionicProj (ofCrAnList φs) = if h : (𝓕 |>ₛ φs) = bosonic then
|
||||
0 else ⟨ofCrAnList φs, Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl,
|
||||
by simpa using h ⟩⟩⟩ := by
|
||||
by simpa using h⟩⟩⟩ := by
|
||||
rw [fermionicProj_ofCrAnList]
|
||||
by_cases h1 : (𝓕 |>ₛ φs) = fermionic
|
||||
· simp [h1]
|
||||
|
@ -217,7 +234,8 @@ lemma directSum_eq_bosonic_plus_fermionic
|
|||
conv_lhs => rw [hx, hy]
|
||||
abel
|
||||
|
||||
instance : GradedAlgebra (A := 𝓕.CrAnAlgebra) statisticSubmodule where
|
||||
/-- The instance of a graded algebra on `CrAnAlgebra`. -/
|
||||
instance crAnAlgebraGrade : GradedAlgebra (A := 𝓕.CrAnAlgebra) statisticSubmodule where
|
||||
one_mem := by
|
||||
simp only [statisticSubmodule]
|
||||
refine Submodule.mem_span.mpr fun p a => a ?_
|
||||
|
@ -227,7 +245,7 @@ instance : GradedAlgebra (A := 𝓕.CrAnAlgebra) statisticSubmodule where
|
|||
rfl
|
||||
mul_mem f1 f2 a1 a2 h1 h2 := by
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule f2) : Prop :=
|
||||
a1 * a2 ∈ statisticSubmodule (f1 + f2)
|
||||
a1 * a2 ∈ statisticSubmodule (f1 + f2)
|
||||
change p a2 h2
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
|
@ -263,10 +281,10 @@ instance : GradedAlgebra (A := 𝓕.CrAnAlgebra) statisticSubmodule where
|
|||
simp only [Algebra.mul_smul_comm, p]
|
||||
exact Submodule.smul_mem _ _ h1
|
||||
· exact h2
|
||||
decompose' a := DirectSum.of (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) bosonic (bosonicProj a)
|
||||
+ DirectSum.of (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) fermionic (fermionicProj a)
|
||||
decompose' a := DirectSum.of (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) bosonic (bosonicProj a)
|
||||
+ DirectSum.of (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) fermionic (fermionicProj a)
|
||||
left_inv a := by
|
||||
trans a.bosonicProj + fermionicProj a
|
||||
trans a.bosonicProj + fermionicProj a
|
||||
· simp
|
||||
· exact bosonicProj_add_fermionicProj a
|
||||
right_inv a := by
|
||||
|
@ -276,6 +294,117 @@ instance : GradedAlgebra (A := 𝓕.CrAnAlgebra) statisticSubmodule where
|
|||
fermionicProj_of_fermionic_part, zero_add]
|
||||
conv_rhs => rw [directSum_eq_bosonic_plus_fermionic a]
|
||||
|
||||
lemma eq_zero_of_bosonic_and_fermionic {a : 𝓕.CrAnAlgebra}
|
||||
(hb : a ∈ statisticSubmodule bosonic) (hf : a ∈ statisticSubmodule fermionic) : a = 0 := by
|
||||
have ha := bosonicProj_of_mem_bosonic a hb
|
||||
have hb := fermionicProj_of_mem_fermionic a hf
|
||||
have hc := (bosonicProj_add_fermionicProj a)
|
||||
rw [ha, hb] at hc
|
||||
simpa using hc
|
||||
|
||||
lemma bosonicProj_mul (a b : 𝓕.CrAnAlgebra) :
|
||||
(a * b).bosonicProj.1 = a.bosonicProj.1 * b.bosonicProj.1
|
||||
+ a.fermionicProj.1 * b.fermionicProj.1 := by
|
||||
conv_lhs =>
|
||||
rw [← bosonicProj_add_fermionicProj a]
|
||||
rw [← bosonicProj_add_fermionicProj b]
|
||||
simp only [mul_add, add_mul, map_add, Submodule.coe_add]
|
||||
rw [bosonicProj_of_mem_bosonic]
|
||||
conv_lhs =>
|
||||
left
|
||||
right
|
||||
rw [bosonicProj_of_mem_fermionic _
|
||||
(by
|
||||
have h1 : fermionic = fermionic + bosonic := by simp
|
||||
conv_lhs => rw [h1]
|
||||
apply crAnAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
conv_lhs =>
|
||||
right
|
||||
left
|
||||
rw [bosonicProj_of_mem_fermionic _
|
||||
(by
|
||||
have h1 : fermionic = bosonic + fermionic := by simp
|
||||
conv_lhs => rw [h1]
|
||||
apply crAnAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
conv_lhs =>
|
||||
right
|
||||
right
|
||||
rw [bosonicProj_of_mem_bosonic _
|
||||
(by
|
||||
have h1 : bosonic = fermionic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
conv_lhs => rw [h1]
|
||||
apply crAnAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
simp only [ZeroMemClass.coe_zero, add_zero, zero_add]
|
||||
· have h1 : bosonic = bosonic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
conv_lhs => rw [h1]
|
||||
apply crAnAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp
|
||||
|
||||
lemma fermionicProj_mul (a b : 𝓕.CrAnAlgebra) :
|
||||
(a * b).fermionicProj.1 = a.bosonicProj.1 * b.fermionicProj.1
|
||||
+ a.fermionicProj.1 * b.bosonicProj.1 := by
|
||||
conv_lhs =>
|
||||
rw [← bosonicProj_add_fermionicProj a]
|
||||
rw [← bosonicProj_add_fermionicProj b]
|
||||
simp only [mul_add, add_mul, map_add, Submodule.coe_add]
|
||||
conv_lhs =>
|
||||
left
|
||||
left
|
||||
rw [fermionicProj_of_mem_bosonic _
|
||||
(by
|
||||
have h1 : bosonic = bosonic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
conv_lhs => rw [h1]
|
||||
apply crAnAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
conv_lhs =>
|
||||
left
|
||||
right
|
||||
rw [fermionicProj_of_mem_fermionic _
|
||||
(by
|
||||
have h1 : fermionic = fermionic + bosonic := by simp
|
||||
conv_lhs => rw [h1]
|
||||
apply crAnAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
conv_lhs =>
|
||||
right
|
||||
left
|
||||
rw [fermionicProj_of_mem_fermionic _
|
||||
(by
|
||||
have h1 : fermionic = bosonic + fermionic := by simp
|
||||
conv_lhs => rw [h1]
|
||||
apply crAnAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
conv_lhs =>
|
||||
right
|
||||
right
|
||||
rw [fermionicProj_of_mem_bosonic _
|
||||
(by
|
||||
have h1 : bosonic = fermionic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
conv_lhs => rw [h1]
|
||||
apply crAnAlgebraGrade.mul_mem
|
||||
simp only [SetLike.coe_mem]
|
||||
simp)]
|
||||
simp only [ZeroMemClass.coe_zero, zero_add, add_zero]
|
||||
abel
|
||||
|
||||
end
|
||||
|
||||
end CrAnAlgebra
|
||||
|
|
|
@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.Basic
|
||||
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.Grading
|
||||
/-!
|
||||
|
||||
# Super Commute
|
||||
|
@ -439,6 +440,439 @@ lemma superCommute_ofCrAnList_ofStateList_eq_sum (φs : List 𝓕.CrAnStates) :
|
|||
· simp [Finset.mul_sum, smul_smul, ofStateList_cons, mul_assoc,
|
||||
FieldStatistic.ofList_cons_eq_mul, mul_comm]
|
||||
|
||||
lemma summerCommute_jacobi_ofCrAnList (φs1 φs2 φs3 : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs1, [ofCrAnList φs2, ofCrAnList φs3]ₛca]ₛca =
|
||||
𝓢(𝓕 |>ₛ φs1, 𝓕 |>ₛ φs3) •
|
||||
(- 𝓢(𝓕 |>ₛ φs2, 𝓕 |>ₛ φs3) • [ofCrAnList φs3, [ofCrAnList φs1, ofCrAnList φs2]ₛca]ₛca -
|
||||
𝓢(𝓕 |>ₛ φs1, 𝓕 |>ₛ φs2) • [ofCrAnList φs2, [ofCrAnList φs3, ofCrAnList φs1]ₛca]ₛca) := by
|
||||
repeat rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
simp only [instCommGroup, map_sub, map_smul, neg_smul]
|
||||
repeat rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
simp only [instCommGroup.eq_1, ofList_append_eq_mul, List.append_assoc]
|
||||
by_cases h1 : (𝓕 |>ₛ φs1) = bosonic <;>
|
||||
by_cases h2 : (𝓕 |>ₛ φs2) = bosonic <;>
|
||||
by_cases h3 : (𝓕 |>ₛ φs3) = bosonic
|
||||
· simp only [h1, h2, h3, mul_self, bosonic_exchangeSign, one_smul, exchangeSign_bosonic, neg_sub]
|
||||
abel
|
||||
· simp only [h1, h2, bosonic_exchangeSign, one_smul, mul_bosonic, mul_self, map_one,
|
||||
exchangeSign_bosonic, neg_sub]
|
||||
abel
|
||||
· simp only [h1, h3, mul_bosonic, bosonic_exchangeSign, one_smul, exchangeSign_bosonic, neg_sub,
|
||||
mul_self, map_one]
|
||||
abel
|
||||
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
|
||||
simp only [h1, h2, h3, mul_self, bosonic_exchangeSign, one_smul,
|
||||
fermionic_exchangeSign_fermionic, neg_smul, neg_sub, bosonic_mul_fermionic, sub_neg_eq_add,
|
||||
mul_bosonic, smul_add, exchangeSign_bosonic]
|
||||
abel
|
||||
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
|
||||
simp only [h1, h2, h3, mul_self, map_one, one_smul, exchangeSign_bosonic, mul_bosonic,
|
||||
bosonic_exchangeSign, bosonic_mul_fermionic, neg_sub]
|
||||
abel
|
||||
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
|
||||
simp only [h1, h2, h3, bosonic_mul_fermionic, fermionic_exchangeSign_fermionic, neg_smul,
|
||||
one_smul, sub_neg_eq_add, bosonic_exchangeSign, mul_bosonic, smul_add, exchangeSign_bosonic,
|
||||
neg_sub, mul_self]
|
||||
abel
|
||||
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
|
||||
simp only [h1, h2, h3, mul_bosonic, fermionic_exchangeSign_fermionic, neg_smul, one_smul,
|
||||
sub_neg_eq_add, exchangeSign_bosonic, bosonic_mul_fermionic, smul_add, mul_self,
|
||||
bosonic_exchangeSign, neg_sub]
|
||||
abel
|
||||
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
|
||||
simp only [h1, h2, h3, mul_self, map_one, one_smul, fermionic_exchangeSign_fermionic, neg_smul,
|
||||
neg_sub]
|
||||
abel
|
||||
/-!
|
||||
|
||||
## Interaction with grading.
|
||||
|
||||
-/
|
||||
|
||||
lemma superCommute_grade {a b : 𝓕.CrAnAlgebra} {f1 f2 : FieldStatistic}
|
||||
(ha : a ∈ statisticSubmodule f1) (hb : b ∈ statisticSubmodule f2) :
|
||||
[a, b]ₛca ∈ statisticSubmodule (f1 + f2) := by
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule f2) : Prop :=
|
||||
[a, a2]ₛca ∈ statisticSubmodule (f1 + f2)
|
||||
change p b hb
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl, hφs⟩ := hx
|
||||
simp only [add_eq_mul, instCommGroup, p]
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule f1) : Prop :=
|
||||
[a2, ofCrAnList φs]ₛca ∈ statisticSubmodule (f1 + f2)
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [add_eq_mul, instCommGroup, p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
apply Submodule.sub_mem _
|
||||
· apply ofCrAnList_mem_statisticSubmodule_of
|
||||
rw [ofList_append_eq_mul, hφs, hφs']
|
||||
· apply Submodule.smul_mem
|
||||
apply ofCrAnList_mem_statisticSubmodule_of
|
||||
rw [ofList_append_eq_mul, hφs, hφs']
|
||||
rw [mul_comm]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp only [add_eq_mul, instCommGroup, map_add, LinearMap.add_apply, p]
|
||||
exact Submodule.add_mem _ hp1 hp2
|
||||
· intro c x hx hp1
|
||||
simp only [add_eq_mul, instCommGroup, map_smul, LinearMap.smul_apply, p]
|
||||
exact Submodule.smul_mem _ c hp1
|
||||
· exact ha
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp only [add_eq_mul, instCommGroup, map_add, p]
|
||||
exact Submodule.add_mem _ hp1 hp2
|
||||
· intro c x hx hp1
|
||||
simp only [add_eq_mul, instCommGroup, map_smul, p]
|
||||
exact Submodule.smul_mem _ c hp1
|
||||
· exact hb
|
||||
|
||||
lemma superCommute_bosonic_bosonic {a b : 𝓕.CrAnAlgebra}
|
||||
(ha : a ∈ statisticSubmodule bosonic) (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop :=
|
||||
[a, a2]ₛca = a * a2 - a2 * a
|
||||
change p b hb
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl, hφs⟩ := hx
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop :=
|
||||
[a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact ha
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, mul_add, add_mul]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommute_bosonic_fermionic {a b : 𝓕.CrAnAlgebra}
|
||||
(ha : a ∈ statisticSubmodule bosonic) (hb : b ∈ statisticSubmodule fermionic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop :=
|
||||
[a, a2]ₛca = a * a2 - a2 * a
|
||||
change p b hb
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl, hφs⟩ := hx
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop :=
|
||||
[a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, hφs', ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact ha
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, mul_add, add_mul]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommute_fermionic_bonsonic {a b : 𝓕.CrAnAlgebra}
|
||||
(ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop :=
|
||||
[a, a2]ₛca = a * a2 - a2 * a
|
||||
change p b hb
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl, hφs⟩ := hx
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop :=
|
||||
[a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, hφs', ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact ha
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [map_add, mul_add, add_mul, p]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommute_bonsonic {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
rw [← bosonicProj_add_fermionicProj a]
|
||||
simp only [map_add, LinearMap.add_apply]
|
||||
rw [superCommute_bosonic_bosonic (by simp) hb, superCommute_fermionic_bonsonic (by simp) hb]
|
||||
simp only [add_mul, mul_add]
|
||||
abel
|
||||
|
||||
lemma bosonic_superCommute {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = a * b - b * a := by
|
||||
rw [← bosonicProj_add_fermionicProj b]
|
||||
simp only [map_add, LinearMap.add_apply]
|
||||
rw [superCommute_bosonic_bosonic ha (by simp), superCommute_bosonic_fermionic ha (by simp)]
|
||||
simp only [add_mul, mul_add]
|
||||
abel
|
||||
|
||||
lemma superCommute_bonsonic_symm {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = - [b, a]ₛca := by
|
||||
rw [bosonic_superCommute hb, superCommute_bonsonic hb]
|
||||
simp
|
||||
|
||||
lemma bonsonic_superCommute_symm {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) :
|
||||
[a, b]ₛca = - [b, a]ₛca := by
|
||||
rw [bosonic_superCommute ha, superCommute_bonsonic ha]
|
||||
simp
|
||||
|
||||
lemma superCommute_fermionic_fermionic {a b : 𝓕.CrAnAlgebra}
|
||||
(ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule fermionic) :
|
||||
[a, b]ₛca = a * b + b * a := by
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop :=
|
||||
[a, a2]ₛca = a * a2 + a2 * a
|
||||
change p b hb
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl, hφs⟩ := hx
|
||||
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop :=
|
||||
[a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs + ofCrAnList φs * a2
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl, hφs'⟩ := hx
|
||||
simp only [p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
simp [hφs, hφs', ofCrAnList_append]
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact ha
|
||||
· simp [p]
|
||||
· intro x y hx hy hp1 hp2
|
||||
simp_all only [map_add, mul_add, add_mul, p]
|
||||
abel
|
||||
· intro c x hx hp1
|
||||
simp_all [p, smul_sub]
|
||||
· exact hb
|
||||
|
||||
lemma superCommute_fermionic_fermionic_symm {a b : 𝓕.CrAnAlgebra}
|
||||
(ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule fermionic) :
|
||||
[a, b]ₛca = [b, a]ₛca := by
|
||||
rw [superCommute_fermionic_fermionic ha hb]
|
||||
rw [superCommute_fermionic_fermionic hb ha]
|
||||
abel
|
||||
|
||||
lemma superCommute_expand_bosonicProj_fermionicProj (a b : 𝓕.CrAnAlgebra) :
|
||||
[a, b]ₛca = bosonicProj a * bosonicProj b - bosonicProj b * bosonicProj a +
|
||||
bosonicProj a * fermionicProj b - fermionicProj b * bosonicProj a +
|
||||
fermionicProj a * bosonicProj b - bosonicProj b * fermionicProj a +
|
||||
fermionicProj a * fermionicProj b + fermionicProj b * fermionicProj a := by
|
||||
conv_lhs => rw [← bosonicProj_add_fermionicProj a, ← bosonicProj_add_fermionicProj b]
|
||||
simp only [map_add, LinearMap.add_apply]
|
||||
rw [superCommute_bonsonic (by simp),
|
||||
superCommute_fermionic_bonsonic (by simp) (by simp),
|
||||
superCommute_bosonic_fermionic (by simp) (by simp),
|
||||
superCommute_fermionic_fermionic (by simp) (by simp)]
|
||||
abel
|
||||
|
||||
lemma superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List 𝓕.CrAnStates) :
|
||||
[ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule bosonic ∨
|
||||
[ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule fermionic := by
|
||||
by_cases h1 : (𝓕 |>ₛ φs) = bosonic <;> by_cases h2 : (𝓕 |>ₛ φs') = bosonic
|
||||
· left
|
||||
have h : bosonic = bosonic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h1
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h2
|
||||
· right
|
||||
have h : fermionic = bosonic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h1
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h2)
|
||||
· right
|
||||
have h : fermionic = fermionic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h1)
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ h2
|
||||
· left
|
||||
have h : bosonic = fermionic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h1)
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h2)
|
||||
|
||||
lemma superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic (φ φ' : 𝓕.CrAnStates) :
|
||||
[ofCrAnState φ, ofCrAnState φ']ₛca ∈ statisticSubmodule bosonic ∨
|
||||
[ofCrAnState φ, ofCrAnState φ']ₛca ∈ statisticSubmodule fermionic := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
exact superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ] [φ']
|
||||
|
||||
lemma superCommute_superCommute_ofCrAnState_bosonic_or_fermionic (φ1 φ2 φ3 : 𝓕.CrAnStates) :
|
||||
[ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca ∈ statisticSubmodule bosonic ∨
|
||||
[ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca ∈ statisticSubmodule fermionic := by
|
||||
rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φ2 φ3 with hs | hs
|
||||
<;> rcases ofCrAnState_bosonic_or_fermionic φ1 with h1 | h1
|
||||
· left
|
||||
have h : bosonic = bosonic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade h1 hs
|
||||
· right
|
||||
have h : fermionic = fermionic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade h1 hs
|
||||
· right
|
||||
have h : fermionic = bosonic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade h1 hs
|
||||
· left
|
||||
have h : bosonic = fermionic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h]
|
||||
apply superCommute_grade h1 hs
|
||||
|
||||
lemma superCommute_bosonic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List 𝓕.CrAnStates)
|
||||
(ha : a ∈ statisticSubmodule bosonic) :
|
||||
[a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length),
|
||||
ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca *
|
||||
ofCrAnList (φs.drop (n + 1)) := by
|
||||
let p (a : 𝓕.CrAnAlgebra) (ha : a ∈ statisticSubmodule bosonic) : Prop :=
|
||||
[a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length),
|
||||
ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca *
|
||||
ofCrAnList (φs.drop (n + 1))
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro a ha
|
||||
obtain ⟨φs, rfl, hφs⟩ := ha
|
||||
simp only [List.get_eq_getElem, p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList_eq_sum]
|
||||
congr
|
||||
funext n
|
||||
simp [hφs]
|
||||
· simp [p]
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all only [List.get_eq_getElem, map_add, LinearMap.add_apply, p]
|
||||
rw [← Finset.sum_add_distrib]
|
||||
congr
|
||||
funext n
|
||||
simp [mul_add, add_mul]
|
||||
· intro c x hx hpx
|
||||
simp_all [p, Finset.smul_sum]
|
||||
· exact ha
|
||||
|
||||
lemma superCommute_fermionic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List 𝓕.CrAnStates)
|
||||
(ha : a ∈ statisticSubmodule fermionic) :
|
||||
[a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), 𝓢(fermionic, 𝓕 |>ₛ φs.take n) •
|
||||
ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca *
|
||||
ofCrAnList (φs.drop (n + 1)) := by
|
||||
let p (a : 𝓕.CrAnAlgebra) (ha : a ∈ statisticSubmodule fermionic) : Prop :=
|
||||
[a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), 𝓢(fermionic, 𝓕 |>ₛ φs.take n) •
|
||||
ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca *
|
||||
ofCrAnList (φs.drop (n + 1))
|
||||
change p a ha
|
||||
apply Submodule.span_induction (p := p)
|
||||
· intro a ha
|
||||
obtain ⟨φs, rfl, hφs⟩ := ha
|
||||
simp only [instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, p]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList_eq_sum]
|
||||
congr
|
||||
funext n
|
||||
simp [hφs]
|
||||
· simp [p]
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all only [p, instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, map_add,
|
||||
LinearMap.add_apply]
|
||||
rw [← Finset.sum_add_distrib]
|
||||
congr
|
||||
funext n
|
||||
simp [mul_add, add_mul]
|
||||
· intro c x hx hpx
|
||||
simp_all only [p, instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, map_smul,
|
||||
LinearMap.smul_apply, Finset.smul_sum, Algebra.mul_smul_comm]
|
||||
congr
|
||||
funext x
|
||||
simp [smul_smul, mul_comm]
|
||||
· exact ha
|
||||
|
||||
lemma statistic_neq_of_superCommute_fermionic {φs φs' : List 𝓕.CrAnStates}
|
||||
(h : [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule fermionic) :
|
||||
(𝓕 |>ₛ φs) ≠ (𝓕 |>ₛ φs') ∨ [ofCrAnList φs, ofCrAnList φs']ₛca = 0 := by
|
||||
by_cases h0 : [ofCrAnList φs, ofCrAnList φs']ₛca = 0
|
||||
· simp [h0]
|
||||
simp only [ne_eq, h0, or_false]
|
||||
by_contra hn
|
||||
refine h0 (eq_zero_of_bosonic_and_fermionic ?_ h)
|
||||
by_cases hc : (𝓕 |>ₛ φs) = bosonic
|
||||
· have h1 : bosonic = bosonic + bosonic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h1]
|
||||
apply superCommute_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _ hc
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _
|
||||
rw [← hn, hc]
|
||||
· have h1 : bosonic = fermionic + fermionic := by
|
||||
simp only [add_eq_mul, instCommGroup, mul_self]
|
||||
rfl
|
||||
rw [h1]
|
||||
apply superCommute_grade
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _
|
||||
simpa using hc
|
||||
apply ofCrAnList_mem_statisticSubmodule_of _ _
|
||||
rw [← hn]
|
||||
simpa using hc
|
||||
|
||||
end CrAnAlgebra
|
||||
|
||||
end FieldSpecification
|
||||
|
|
|
@ -39,6 +39,66 @@ lemma timeOrder_ofCrAnList (φs : List 𝓕.CrAnStates) :
|
|||
rw [← ofListBasis_eq_ofList]
|
||||
simp only [timeOrder, Basis.constr_basis]
|
||||
|
||||
lemma timeOrder_timeOrder_mid (a b c : 𝓕.CrAnAlgebra) : 𝓣ᶠ(a * b * c) = 𝓣ᶠ(a * 𝓣ᶠ(b) * c) := by
|
||||
let pc (c : 𝓕.CrAnAlgebra) (hc : c ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓣ᶠ(a * b * c) = 𝓣ᶠ(a * 𝓣ᶠ(b) * c)
|
||||
change pc c (Basis.mem_span _ c)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pc]
|
||||
let pb (b : 𝓕.CrAnAlgebra) (hb : b ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓣ᶠ(a * b * ofCrAnList φs) = 𝓣ᶠ(a * 𝓣ᶠ(b) * ofCrAnList φs)
|
||||
change pb b (Basis.mem_span _ b)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pb]
|
||||
let pa (a : 𝓕.CrAnAlgebra) (ha : a ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := 𝓣ᶠ(a * ofCrAnList φs' * ofCrAnList φs) = 𝓣ᶠ(a * 𝓣ᶠ(ofCrAnList φs') * ofCrAnList φs)
|
||||
change pa a (Basis.mem_span _ a)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs'', rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pa]
|
||||
rw [timeOrder_ofCrAnList]
|
||||
simp only [← ofCrAnList_append, Algebra.mul_smul_comm,
|
||||
Algebra.smul_mul_assoc, map_smul]
|
||||
rw [timeOrder_ofCrAnList, timeOrder_ofCrAnList, smul_smul]
|
||||
congr 1
|
||||
· simp only [crAnTimeOrderSign, crAnTimeOrderList]
|
||||
rw [Wick.koszulSign_of_append_eq_insertionSort, mul_comm]
|
||||
· congr 1
|
||||
simp only [crAnTimeOrderList]
|
||||
rw [insertionSort_append_insertionSort_append]
|
||||
· simp [pa]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pa, add_mul]
|
||||
· intro x hx h
|
||||
simp_all [pa]
|
||||
· simp [pb]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pb, mul_add, add_mul]
|
||||
· intro x hx h
|
||||
simp_all [pb]
|
||||
· simp [pc]
|
||||
· intro x y hx hy h1 h2
|
||||
simp_all [pc, mul_add]
|
||||
· intro x hx h hp
|
||||
simp_all [pc]
|
||||
|
||||
lemma timeOrder_timeOrder_right (a b : 𝓕.CrAnAlgebra) : 𝓣ᶠ(a * b) = 𝓣ᶠ(a * 𝓣ᶠ(b)) := by
|
||||
trans 𝓣ᶠ(a * b * 1)
|
||||
· simp
|
||||
· rw [timeOrder_timeOrder_mid]
|
||||
simp
|
||||
|
||||
lemma timeOrder_timeOrder_left (a b : 𝓕.CrAnAlgebra) : 𝓣ᶠ(a * b) = 𝓣ᶠ(𝓣ᶠ(a) * b) := by
|
||||
trans 𝓣ᶠ(1 * a * b)
|
||||
· simp
|
||||
· rw [timeOrder_timeOrder_mid]
|
||||
simp
|
||||
|
||||
lemma timeOrder_ofStateList (φs : List 𝓕.States) :
|
||||
𝓣ᶠ(ofStateList φs) = timeOrderSign φs • ofStateList (timeOrderList φs) := by
|
||||
conv_lhs =>
|
||||
|
@ -100,6 +160,120 @@ lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel
|
|||
· rw [crAnTimeOrderList_pair_ordered]
|
||||
simp_all
|
||||
|
||||
lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right
|
||||
{φ ψ : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ ψ) (a : 𝓕.CrAnAlgebra) :
|
||||
𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca) = 0 := by
|
||||
rw [timeOrder_timeOrder_right,
|
||||
timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel h]
|
||||
simp
|
||||
|
||||
lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left
|
||||
{φ ψ : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ ψ) (a : 𝓕.CrAnAlgebra) :
|
||||
𝓣ᶠ([ofCrAnState φ, ofCrAnState ψ]ₛca * a) = 0 := by
|
||||
rw [timeOrder_timeOrder_left,
|
||||
timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel h]
|
||||
simp
|
||||
|
||||
lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_mid
|
||||
{φ ψ : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ ψ) (a b : 𝓕.CrAnAlgebra) :
|
||||
𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = 0 := by
|
||||
rw [timeOrder_timeOrder_mid,
|
||||
timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel h]
|
||||
simp
|
||||
|
||||
lemma timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel
|
||||
{φ1 φ2 : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ1 φ2) (a : 𝓕.CrAnAlgebra) :
|
||||
𝓣ᶠ([a, [ofCrAnState φ1, ofCrAnState φ2]ₛca]ₛca) = 0 := by
|
||||
rw [← bosonicProj_add_fermionicProj a]
|
||||
simp only [map_add, LinearMap.add_apply]
|
||||
rw [bosonic_superCommute (Submodule.coe_mem (bosonicProj a))]
|
||||
simp only [map_sub]
|
||||
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left h]
|
||||
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right h]
|
||||
simp only [sub_self, zero_add]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ1] [φ2] with h' | h'
|
||||
· rw [superCommute_bonsonic h']
|
||||
simp only [ofCrAnList_singleton, map_sub]
|
||||
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left h]
|
||||
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right h]
|
||||
simp
|
||||
· rw [superCommute_fermionic_fermionic (Submodule.coe_mem (fermionicProj a)) h']
|
||||
simp only [ofCrAnList_singleton, map_add]
|
||||
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left h]
|
||||
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right h]
|
||||
simp
|
||||
|
||||
lemma timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel
|
||||
{φ1 φ2 φ3 : 𝓕.CrAnStates} (h12 : ¬ crAnTimeOrderRel φ1 φ2)
|
||||
(h13 : ¬ crAnTimeOrderRel φ1 φ3) :
|
||||
𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rw [summerCommute_jacobi_ofCrAnList]
|
||||
simp only [instCommGroup.eq_1, ofList_singleton, ofCrAnList_singleton, neg_smul, map_smul,
|
||||
map_sub, map_neg, smul_eq_zero]
|
||||
right
|
||||
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h12]
|
||||
rw [superCommute_ofCrAnState_ofCrAnState_symm φ3]
|
||||
simp only [smul_zero, neg_zero, instCommGroup.eq_1, neg_smul, map_neg, map_smul, smul_neg,
|
||||
sub_neg_eq_add, zero_add, smul_eq_zero]
|
||||
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h13]
|
||||
simp
|
||||
|
||||
lemma timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel'
|
||||
{φ1 φ2 φ3 : 𝓕.CrAnStates} (h12 : ¬ crAnTimeOrderRel φ2 φ1)
|
||||
(h13 : ¬ crAnTimeOrderRel φ3 φ1) :
|
||||
𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rw [summerCommute_jacobi_ofCrAnList]
|
||||
simp only [instCommGroup.eq_1, ofList_singleton, ofCrAnList_singleton, neg_smul, map_smul,
|
||||
map_sub, map_neg, smul_eq_zero]
|
||||
right
|
||||
rw [superCommute_ofCrAnState_ofCrAnState_symm φ1]
|
||||
simp only [instCommGroup.eq_1, neg_smul, map_neg, map_smul, smul_neg, neg_neg]
|
||||
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h12]
|
||||
simp only [smul_zero, zero_sub, neg_eq_zero, smul_eq_zero]
|
||||
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h13]
|
||||
simp
|
||||
|
||||
lemma timeOrder_superCommute_ofCrAnState_superCommute_all_not_crAnTimeOrderRel
|
||||
(φ1 φ2 φ3 : 𝓕.CrAnStates) (h : ¬
|
||||
(crAnTimeOrderRel φ1 φ2 ∧ crAnTimeOrderRel φ1 φ3 ∧
|
||||
crAnTimeOrderRel φ2 φ1 ∧ crAnTimeOrderRel φ2 φ3 ∧
|
||||
crAnTimeOrderRel φ3 φ1 ∧ crAnTimeOrderRel φ3 φ2)) :
|
||||
𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by
|
||||
simp only [not_and] at h
|
||||
by_cases h23 : ¬ crAnTimeOrderRel φ2 φ3
|
||||
· simp_all only [IsEmpty.forall_iff, implies_true]
|
||||
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h23]
|
||||
simp_all only [Decidable.not_not, forall_const]
|
||||
by_cases h32 : ¬ crAnTimeOrderRel φ3 φ2
|
||||
· simp_all only [not_false_eq_true, implies_true]
|
||||
rw [superCommute_ofCrAnState_ofCrAnState_symm]
|
||||
simp only [instCommGroup.eq_1, neg_smul, map_neg, map_smul, neg_eq_zero, smul_eq_zero]
|
||||
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h32]
|
||||
simp
|
||||
simp_all only [imp_false, Decidable.not_not]
|
||||
by_cases h12 : ¬ crAnTimeOrderRel φ1 φ2
|
||||
· have h13 : ¬ crAnTimeOrderRel φ1 φ3 := by
|
||||
intro h13
|
||||
apply h12
|
||||
exact IsTrans.trans φ1 φ3 φ2 h13 h32
|
||||
rw [timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel h12 h13]
|
||||
simp_all only [Decidable.not_not, forall_const]
|
||||
have h13 : crAnTimeOrderRel φ1 φ3 := IsTrans.trans φ1 φ2 φ3 h12 h23
|
||||
simp_all only [forall_const]
|
||||
by_cases h21 : ¬ crAnTimeOrderRel φ2 φ1
|
||||
· simp_all only [IsEmpty.forall_iff]
|
||||
have h31 : ¬ crAnTimeOrderRel φ3 φ1 := by
|
||||
intro h31
|
||||
apply h21
|
||||
exact IsTrans.trans φ2 φ3 φ1 h23 h31
|
||||
rw [timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel' h21 h31]
|
||||
simp_all only [Decidable.not_not, forall_const]
|
||||
refine False.elim (h ?_)
|
||||
exact IsTrans.trans φ3 φ2 φ1 h32 h21
|
||||
|
||||
lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_eq_time
|
||||
{φ ψ : 𝓕.CrAnStates} (h1 : crAnTimeOrderRel φ ψ) (h2 : crAnTimeOrderRel ψ φ) :
|
||||
𝓣ᶠ([ofCrAnState φ, ofCrAnState ψ]ₛca) = [ofCrAnState φ, ofCrAnState ψ]ₛca := by
|
||||
|
|
|
@ -3,10 +3,9 @@ Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
|
|||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.PerturbationTheory.WickContraction.TimeContract
|
||||
import HepLean.Meta.Remark.Basic
|
||||
import Mathlib.RingTheory.TwoSidedIdeal.Operations
|
||||
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.SuperCommute
|
||||
import Mathlib.Algebra.RingQuot
|
||||
import Mathlib.RingTheory.TwoSidedIdeal.Operations
|
||||
/-!
|
||||
|
||||
# Field operator algebra
|
||||
|
@ -15,9 +14,7 @@ import Mathlib.Algebra.RingQuot
|
|||
|
||||
namespace FieldSpecification
|
||||
open CrAnAlgebra
|
||||
open ProtoOperatorAlgebra
|
||||
open HepLean.List
|
||||
open WickContraction
|
||||
open FieldStatistic
|
||||
|
||||
variable (𝓕 : FieldSpecification)
|
||||
|
@ -26,8 +23,8 @@ variable (𝓕 : FieldSpecification)
|
|||
This contains e.g. the super-commutor of two creation operators. -/
|
||||
def fieldOpIdealSet : Set (CrAnAlgebra 𝓕) :=
|
||||
{ x |
|
||||
(∃ (φ ψ : 𝓕.CrAnStates) (a : CrAnAlgebra 𝓕),
|
||||
x = a * [ofCrAnState φ, ofCrAnState ψ]ₛca - [ofCrAnState φ, ofCrAnState ψ]ₛca * a)
|
||||
(∃ (φ1 φ2 φ3 : 𝓕.CrAnStates),
|
||||
x = [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca)
|
||||
∨ (∃ (φc φc' : 𝓕.CrAnStates) (_ : 𝓕 |>ᶜ φc = .create) (_ : 𝓕 |>ᶜ φc' = .create),
|
||||
x = [ofCrAnState φc, ofCrAnState φc']ₛca)
|
||||
∨ (∃ (φa φa' : 𝓕.CrAnStates) (_ : 𝓕 |>ᶜ φa = .annihilate) (_ : 𝓕 |>ᶜ φa' = .annihilate),
|
||||
|
@ -75,19 +72,6 @@ lemma ι_of_mem_fieldOpIdealSet (x : CrAnAlgebra 𝓕) (hx : x ∈ 𝓕.fieldOpI
|
|||
refine RingConGen.Rel.of x 0 ?_
|
||||
simpa using hx
|
||||
|
||||
lemma ι_superCommute_ofCrAnState_ofCrAnState_mem_center (φ ψ : 𝓕.CrAnStates) :
|
||||
ι [ofCrAnState φ, ofCrAnState ψ]ₛca ∈ Subalgebra.center ℂ 𝓕.FieldOpAlgebra := by
|
||||
rw [Subalgebra.mem_center_iff]
|
||||
intro b
|
||||
obtain ⟨b, rfl⟩ := ι_surjective b
|
||||
rw [← map_mul, ← map_mul]
|
||||
rw [LinearMap.sub_mem_ker_iff.mp]
|
||||
simp only [LinearMap.mem_ker]
|
||||
apply ι_of_mem_fieldOpIdealSet
|
||||
simp only [fieldOpIdealSet, exists_prop, exists_and_left, Set.mem_setOf_eq]
|
||||
left
|
||||
use φ, ψ, b
|
||||
|
||||
lemma ι_superCommute_of_create_create (φc φc' : 𝓕.CrAnStates) (hφc : 𝓕 |>ᶜ φc = .create)
|
||||
(hφc' : 𝓕 |>ᶜ φc' = .create) : ι [ofCrAnState φc, ofCrAnState φc']ₛca = 0 := by
|
||||
apply ι_of_mem_fieldOpIdealSet
|
||||
|
@ -108,7 +92,7 @@ lemma ι_superCommute_of_annihilate_annihilate (φa φa' : 𝓕.CrAnStates)
|
|||
left
|
||||
use φa, φa', hφa, hφa'
|
||||
|
||||
lemma ι_superCommute_of_diff_statistic (φ ψ : 𝓕.CrAnStates)
|
||||
lemma ι_superCommute_of_diff_statistic {φ ψ : 𝓕.CrAnStates}
|
||||
(h : (𝓕 |>ₛ φ) ≠ (𝓕 |>ₛ ψ)) : ι [ofCrAnState φ, ofCrAnState ψ]ₛca = 0 := by
|
||||
apply ι_of_mem_fieldOpIdealSet
|
||||
simp only [fieldOpIdealSet, exists_prop, exists_and_left, Set.mem_setOf_eq]
|
||||
|
@ -117,5 +101,326 @@ lemma ι_superCommute_of_diff_statistic (φ ψ : 𝓕.CrAnStates)
|
|||
right
|
||||
use φ, ψ
|
||||
|
||||
lemma ι_superCommute_zero_of_fermionic (φ ψ : 𝓕.CrAnStates)
|
||||
(h : [ofCrAnState φ, ofCrAnState ψ]ₛca ∈ statisticSubmodule fermionic) :
|
||||
ι [ofCrAnState φ, ofCrAnState ψ]ₛca = 0 := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton] at h ⊢
|
||||
rcases statistic_neq_of_superCommute_fermionic h with h | h
|
||||
· simp only [ofCrAnList_singleton]
|
||||
apply ι_superCommute_of_diff_statistic
|
||||
simpa using h
|
||||
· simp [h]
|
||||
|
||||
lemma ι_superCommute_ofCrAnState_ofCrAnState_bosonic_or_zero (φ ψ : 𝓕.CrAnStates) :
|
||||
[ofCrAnState φ, ofCrAnState ψ]ₛca ∈ statisticSubmodule bosonic ∨
|
||||
ι [ofCrAnState φ, ofCrAnState ψ]ₛca = 0 := by
|
||||
rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ] [ψ] with h | h
|
||||
· simp_all [ofCrAnList_singleton]
|
||||
· simp_all only [ofCrAnList_singleton]
|
||||
right
|
||||
exact ι_superCommute_zero_of_fermionic _ _ h
|
||||
|
||||
/-!
|
||||
|
||||
## Super-commutes are in the center
|
||||
|
||||
-/
|
||||
|
||||
@[simp]
|
||||
lemma ι_superCommute_ofCrAnState_superCommute_ofCrAnState_ofCrAnState (φ1 φ2 φ3 : 𝓕.CrAnStates) :
|
||||
ι [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca = 0 := by
|
||||
apply ι_of_mem_fieldOpIdealSet
|
||||
simp only [fieldOpIdealSet, exists_prop, exists_and_left, Set.mem_setOf_eq]
|
||||
left
|
||||
use φ1, φ2, φ3
|
||||
|
||||
lemma ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnState (φ1 φ2 φ3 : 𝓕.CrAnStates) :
|
||||
ι [[ofCrAnState φ1, ofCrAnState φ2]ₛca, ofCrAnState φ3]ₛca = 0 := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ1] [φ2] with h | h
|
||||
· rw [bonsonic_superCommute_symm h]
|
||||
simp [ofCrAnList_singleton]
|
||||
· rcases ofCrAnList_bosonic_or_fermionic [φ3] with h' | h'
|
||||
· rw [superCommute_bonsonic_symm h']
|
||||
simp [ofCrAnList_singleton]
|
||||
· rw [superCommute_fermionic_fermionic_symm h h']
|
||||
simp [ofCrAnList_singleton]
|
||||
|
||||
lemma ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnList (φ1 φ2 : 𝓕.CrAnStates)
|
||||
(φs : List 𝓕.CrAnStates) :
|
||||
ι [[ofCrAnState φ1, ofCrAnState φ2]ₛca, ofCrAnList φs]ₛca = 0 := by
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ1] [φ2] with h | h
|
||||
· rw [superCommute_bosonic_ofCrAnList_eq_sum _ _ h]
|
||||
simp [ofCrAnList_singleton, ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnState]
|
||||
· rw [superCommute_fermionic_ofCrAnList_eq_sum _ _ h]
|
||||
simp [ofCrAnList_singleton, ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnState]
|
||||
|
||||
@[simp]
|
||||
lemma ι_superCommute_superCommute_ofCrAnState_ofCrAnState_crAnAlgebra (φ1 φ2 : 𝓕.CrAnStates)
|
||||
(a : 𝓕.CrAnAlgebra) : ι [[ofCrAnState φ1, ofCrAnState φ2]ₛca, a]ₛca = 0 := by
|
||||
change (ι.toLinearMap ∘ₗ superCommute [ofCrAnState φ1, ofCrAnState φ2]ₛca) a = _
|
||||
have h1 : (ι.toLinearMap ∘ₗ superCommute [ofCrAnState φ1, ofCrAnState φ2]ₛca) = 0 := by
|
||||
apply (ofCrAnListBasis.ext fun l ↦ ?_)
|
||||
simp [ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnList]
|
||||
rw [h1]
|
||||
simp
|
||||
|
||||
lemma ι_commute_crAnAlgebra_superCommute_ofCrAnState_ofCrAnState (φ1 φ2 : 𝓕.CrAnStates)
|
||||
(a : 𝓕.CrAnAlgebra) : ι a * ι [ofCrAnState φ1, ofCrAnState φ2]ₛca -
|
||||
ι [ofCrAnState φ1, ofCrAnState φ2]ₛca * ι a = 0 := by
|
||||
rcases ι_superCommute_ofCrAnState_ofCrAnState_bosonic_or_zero φ1 φ2 with h | h
|
||||
swap
|
||||
· simp [h]
|
||||
trans - ι [[ofCrAnState φ1, ofCrAnState φ2]ₛca, a]ₛca
|
||||
· rw [bosonic_superCommute h]
|
||||
simp
|
||||
· simp
|
||||
|
||||
lemma ι_superCommute_ofCrAnState_ofCrAnState_mem_center (φ ψ : 𝓕.CrAnStates) :
|
||||
ι [ofCrAnState φ, ofCrAnState ψ]ₛca ∈ Subalgebra.center ℂ 𝓕.FieldOpAlgebra := by
|
||||
rw [Subalgebra.mem_center_iff]
|
||||
intro a
|
||||
obtain ⟨a, rfl⟩ := ι_surjective a
|
||||
have h0 := ι_commute_crAnAlgebra_superCommute_ofCrAnState_ofCrAnState φ ψ a
|
||||
trans ι ((superCommute (ofCrAnState φ)) (ofCrAnState ψ)) * ι a + 0
|
||||
swap
|
||||
simp only [add_zero]
|
||||
rw [← h0]
|
||||
abel
|
||||
|
||||
/-!
|
||||
|
||||
## The kernal of ι
|
||||
-/
|
||||
|
||||
lemma ι_eq_zero_iff_mem_ideal (x : CrAnAlgebra 𝓕) :
|
||||
ι x = 0 ↔ x ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet := by
|
||||
rw [ι_apply]
|
||||
change ⟦x⟧ = ⟦0⟧ ↔ _
|
||||
simp only [ringConGen, Quotient.eq]
|
||||
rw [TwoSidedIdeal.mem_iff]
|
||||
simp only
|
||||
rfl
|
||||
|
||||
lemma bosonicProj_mem_fieldOpIdealSet_or_zero (x : CrAnAlgebra 𝓕) (hx : x ∈ 𝓕.fieldOpIdealSet) :
|
||||
x.bosonicProj.1 ∈ 𝓕.fieldOpIdealSet ∨ x.bosonicProj = 0 := by
|
||||
have hx' := hx
|
||||
simp only [fieldOpIdealSet, exists_prop, Set.mem_setOf_eq] at hx
|
||||
rcases hx with ⟨φ1, φ2, φ3, rfl⟩ | ⟨φc, φc', hφc, hφc', rfl⟩ | ⟨φa, φa', hφa, hφa', rfl⟩ |
|
||||
⟨φ, φ', hdiff, rfl⟩
|
||||
· rcases superCommute_superCommute_ofCrAnState_bosonic_or_fermionic φ1 φ2 φ3 with h | h
|
||||
· left
|
||||
rw [bosonicProj_of_mem_bosonic _ h]
|
||||
simpa using hx'
|
||||
· right
|
||||
rw [bosonicProj_of_mem_fermionic _ h]
|
||||
· rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φc φc' with h | h
|
||||
· left
|
||||
rw [bosonicProj_of_mem_bosonic _ h]
|
||||
simpa using hx'
|
||||
· right
|
||||
rw [bosonicProj_of_mem_fermionic _ h]
|
||||
· rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φa φa' with h | h
|
||||
· left
|
||||
rw [bosonicProj_of_mem_bosonic _ h]
|
||||
simpa using hx'
|
||||
· right
|
||||
rw [bosonicProj_of_mem_fermionic _ h]
|
||||
· rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φ φ' with h | h
|
||||
· left
|
||||
rw [bosonicProj_of_mem_bosonic _ h]
|
||||
simpa using hx'
|
||||
· right
|
||||
rw [bosonicProj_of_mem_fermionic _ h]
|
||||
|
||||
lemma fermionicProj_mem_fieldOpIdealSet_or_zero (x : CrAnAlgebra 𝓕) (hx : x ∈ 𝓕.fieldOpIdealSet) :
|
||||
x.fermionicProj.1 ∈ 𝓕.fieldOpIdealSet ∨ x.fermionicProj = 0 := by
|
||||
have hx' := hx
|
||||
simp only [fieldOpIdealSet, exists_prop, Set.mem_setOf_eq] at hx
|
||||
rcases hx with ⟨φ1, φ2, φ3, rfl⟩ | ⟨φc, φc', hφc, hφc', rfl⟩ | ⟨φa, φa', hφa, hφa', rfl⟩ |
|
||||
⟨φ, φ', hdiff, rfl⟩
|
||||
· rcases superCommute_superCommute_ofCrAnState_bosonic_or_fermionic φ1 φ2 φ3 with h | h
|
||||
· right
|
||||
rw [fermionicProj_of_mem_bosonic _ h]
|
||||
· left
|
||||
rw [fermionicProj_of_mem_fermionic _ h]
|
||||
simpa using hx'
|
||||
· rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φc φc' with h | h
|
||||
· right
|
||||
rw [fermionicProj_of_mem_bosonic _ h]
|
||||
· left
|
||||
rw [fermionicProj_of_mem_fermionic _ h]
|
||||
simpa using hx'
|
||||
· rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φa φa' with h | h
|
||||
· right
|
||||
rw [fermionicProj_of_mem_bosonic _ h]
|
||||
· left
|
||||
rw [fermionicProj_of_mem_fermionic _ h]
|
||||
simpa using hx'
|
||||
· rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φ φ' with h | h
|
||||
· right
|
||||
rw [fermionicProj_of_mem_bosonic _ h]
|
||||
· left
|
||||
rw [fermionicProj_of_mem_fermionic _ h]
|
||||
simpa using hx'
|
||||
|
||||
lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet) :
|
||||
x.bosonicProj.1 ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet := by
|
||||
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] at hx
|
||||
let p {k : Set 𝓕.CrAnAlgebra} (a : CrAnAlgebra 𝓕) (h : a ∈ AddSubgroup.closure k) : Prop :=
|
||||
a.bosonicProj.1 ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet
|
||||
change p x hx
|
||||
apply AddSubgroup.closure_induction
|
||||
· intro x hx
|
||||
simp only [p]
|
||||
obtain ⟨a, ha, b, hb, rfl⟩ := Set.mem_mul.mp hx
|
||||
obtain ⟨d, hd, y, hy, rfl⟩ := Set.mem_mul.mp ha
|
||||
rw [bosonicProj_mul, bosonicProj_mul, fermionicProj_mul]
|
||||
simp only [add_mul]
|
||||
rcases fermionicProj_mem_fieldOpIdealSet_or_zero y hy with hfy | hfy
|
||||
<;> rcases bosonicProj_mem_fieldOpIdealSet_or_zero y hy with hby | hby
|
||||
· apply TwoSidedIdeal.add_mem
|
||||
apply TwoSidedIdeal.add_mem
|
||||
· /- boson, boson, boson mem-/
|
||||
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure]
|
||||
refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure
|
||||
apply Set.mem_mul.mpr
|
||||
use ↑(bosonicProj d) * ↑(bosonicProj y)
|
||||
apply And.intro
|
||||
· apply Set.mem_mul.mpr
|
||||
use bosonicProj d
|
||||
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
|
||||
use (bosonicProj y).1
|
||||
simp [hby]
|
||||
· use ↑(bosonicProj b)
|
||||
simp
|
||||
· /- fermion, fermion, boson mem-/
|
||||
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure]
|
||||
refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure
|
||||
apply Set.mem_mul.mpr
|
||||
use ↑(fermionicProj d) * ↑(fermionicProj y)
|
||||
apply And.intro
|
||||
· apply Set.mem_mul.mpr
|
||||
use fermionicProj d
|
||||
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
|
||||
use (fermionicProj y).1
|
||||
simp [hby, hfy]
|
||||
· use ↑(bosonicProj b)
|
||||
simp
|
||||
apply TwoSidedIdeal.add_mem
|
||||
· /- boson, fermion, fermion mem-/
|
||||
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure]
|
||||
refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure
|
||||
apply Set.mem_mul.mpr
|
||||
use ↑(bosonicProj d) * ↑(fermionicProj y)
|
||||
apply And.intro
|
||||
· apply Set.mem_mul.mpr
|
||||
use bosonicProj d
|
||||
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
|
||||
use (fermionicProj y).1
|
||||
simp [hby, hfy]
|
||||
· use ↑(fermionicProj b)
|
||||
simp
|
||||
· /- fermion, boson, fermion mem-/
|
||||
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure]
|
||||
refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure
|
||||
apply Set.mem_mul.mpr
|
||||
use ↑(fermionicProj d) * ↑(bosonicProj y)
|
||||
apply And.intro
|
||||
· apply Set.mem_mul.mpr
|
||||
use fermionicProj d
|
||||
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
|
||||
use (bosonicProj y).1
|
||||
simp [hby, hfy]
|
||||
· use ↑(fermionicProj b)
|
||||
simp
|
||||
· simp only [hby, ZeroMemClass.coe_zero, mul_zero, zero_mul, zero_add, add_zero]
|
||||
apply TwoSidedIdeal.add_mem
|
||||
· /- fermion, fermion, boson mem-/
|
||||
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure]
|
||||
refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure
|
||||
apply Set.mem_mul.mpr
|
||||
use ↑(fermionicProj d) * ↑(fermionicProj y)
|
||||
apply And.intro
|
||||
· apply Set.mem_mul.mpr
|
||||
use fermionicProj d
|
||||
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
|
||||
use (fermionicProj y).1
|
||||
simp [hby, hfy]
|
||||
· use ↑(bosonicProj b)
|
||||
simp
|
||||
· /- boson, fermion, fermion mem-/
|
||||
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure]
|
||||
refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure
|
||||
apply Set.mem_mul.mpr
|
||||
use ↑(bosonicProj d) * ↑(fermionicProj y)
|
||||
apply And.intro
|
||||
· apply Set.mem_mul.mpr
|
||||
use bosonicProj d
|
||||
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
|
||||
use (fermionicProj y).1
|
||||
simp [hby, hfy]
|
||||
· use ↑(fermionicProj b)
|
||||
simp
|
||||
· simp only [hfy, ZeroMemClass.coe_zero, mul_zero, zero_mul, add_zero, zero_add]
|
||||
apply TwoSidedIdeal.add_mem
|
||||
· /- boson, boson, boson mem-/
|
||||
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure]
|
||||
refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure
|
||||
apply Set.mem_mul.mpr
|
||||
use ↑(bosonicProj d) * ↑(bosonicProj y)
|
||||
apply And.intro
|
||||
· apply Set.mem_mul.mpr
|
||||
use bosonicProj d
|
||||
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
|
||||
use (bosonicProj y).1
|
||||
simp [hby]
|
||||
· use ↑(bosonicProj b)
|
||||
simp
|
||||
· /- fermion, boson, fermion mem-/
|
||||
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure]
|
||||
refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure
|
||||
apply Set.mem_mul.mpr
|
||||
use ↑(fermionicProj d) * ↑(bosonicProj y)
|
||||
apply And.intro
|
||||
· apply Set.mem_mul.mpr
|
||||
use fermionicProj d
|
||||
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
|
||||
use (bosonicProj y).1
|
||||
simp [hby, hfy]
|
||||
· use ↑(fermionicProj b)
|
||||
simp
|
||||
· simp [hfy, hby]
|
||||
· simp [p]
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all only [map_add, Submodule.coe_add, p]
|
||||
apply TwoSidedIdeal.add_mem
|
||||
exact hpx
|
||||
exact hpy
|
||||
· intro x hx
|
||||
simp [p]
|
||||
|
||||
lemma fermionicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet) :
|
||||
x.fermionicProj.1 ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet := by
|
||||
have hb := bosonicProj_mem_ideal x hx
|
||||
rw [← ι_eq_zero_iff_mem_ideal] at hx hb ⊢
|
||||
rw [← bosonicProj_add_fermionicProj x] at hx
|
||||
simp only [map_add] at hx
|
||||
simp_all
|
||||
|
||||
lemma ι_eq_zero_iff_ι_bosonicProj_fermonicProj_zero (x : CrAnAlgebra 𝓕) :
|
||||
ι x = 0 ↔ ι x.bosonicProj.1 = 0 ∧ ι x.fermionicProj.1 = 0 := by
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
rw [@ι_eq_zero_iff_mem_ideal] at h ⊢
|
||||
rw [ι_eq_zero_iff_mem_ideal]
|
||||
apply And.intro
|
||||
· exact bosonicProj_mem_ideal x h
|
||||
· exact fermionicProj_mem_ideal x h
|
||||
· intro h
|
||||
rw [← bosonicProj_add_fermionicProj x]
|
||||
simp_all
|
||||
|
||||
end FieldOpAlgebra
|
||||
end FieldSpecification
|
||||
|
|
|
@ -3,6 +3,7 @@ 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.Algebras.FieldOpAlgebra.Basic
|
||||
/-!
|
||||
|
||||
|
@ -12,9 +13,7 @@ import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.Basic
|
|||
|
||||
namespace FieldSpecification
|
||||
open CrAnAlgebra
|
||||
open ProtoOperatorAlgebra
|
||||
open HepLean.List
|
||||
open WickContraction
|
||||
open FieldStatistic
|
||||
|
||||
namespace FieldOpAlgebra
|
||||
|
@ -222,15 +221,13 @@ lemma ι_normalOrder_eq_of_equiv (a b : 𝓕.CrAnAlgebra) (h : a ≈ b) :
|
|||
noncomputable def normalOrder : FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 where
|
||||
toFun := Quotient.lift (ι.toLinearMap ∘ₗ CrAnAlgebra.normalOrder) ι_normalOrder_eq_of_equiv
|
||||
map_add' x y := by
|
||||
obtain ⟨x, hx⟩ := ι_surjective x
|
||||
obtain ⟨y, hy⟩ := ι_surjective y
|
||||
subst hx hy
|
||||
obtain ⟨x, rfl⟩ := ι_surjective x
|
||||
obtain ⟨y, rfl⟩ := ι_surjective y
|
||||
rw [← map_add, ι_apply, ι_apply, ι_apply]
|
||||
rw [Quotient.lift_mk, Quotient.lift_mk, Quotient.lift_mk]
|
||||
simp
|
||||
map_smul' c y := by
|
||||
obtain ⟨y, hy⟩ := ι_surjective y
|
||||
subst hy
|
||||
obtain ⟨y, rfl⟩ := ι_surjective y
|
||||
rw [← map_smul, ι_apply, ι_apply]
|
||||
simp
|
||||
|
||||
|
|
|
@ -0,0 +1,118 @@
|
|||
/-
|
||||
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.TimeOrder
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.Basic
|
||||
/-!
|
||||
|
||||
# SuperCommute on Field operator algebra
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldSpecification
|
||||
open CrAnAlgebra
|
||||
open HepLean.List
|
||||
open FieldStatistic
|
||||
|
||||
namespace FieldOpAlgebra
|
||||
variable {𝓕 : FieldSpecification}
|
||||
|
||||
lemma ι_superCommute_eq_zero_of_ι_right_zero (a b : 𝓕.CrAnAlgebra) (h : ι b = 0) :
|
||||
ι [a, b]ₛca = 0 := by
|
||||
rw [superCommute_expand_bosonicProj_fermionicProj]
|
||||
rw [ι_eq_zero_iff_ι_bosonicProj_fermonicProj_zero] at h
|
||||
simp_all
|
||||
|
||||
lemma ι_superCommute_eq_zero_of_ι_left_zero (a b : 𝓕.CrAnAlgebra) (h : ι a = 0) :
|
||||
ι [a, b]ₛca = 0 := by
|
||||
rw [superCommute_expand_bosonicProj_fermionicProj]
|
||||
rw [ι_eq_zero_iff_ι_bosonicProj_fermonicProj_zero] at h
|
||||
simp_all
|
||||
|
||||
/-!
|
||||
|
||||
## Defining normal order for `FiedOpAlgebra`.
|
||||
|
||||
-/
|
||||
|
||||
lemma ι_superCommute_right_zero_of_mem_ideal (a b : 𝓕.CrAnAlgebra)
|
||||
(h : b ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet) : ι [a, b]ₛca = 0 := by
|
||||
apply ι_superCommute_eq_zero_of_ι_right_zero
|
||||
exact (ι_eq_zero_iff_mem_ideal b).mpr h
|
||||
|
||||
lemma ι_superCommute_eq_of_equiv_right (a b1 b2 : 𝓕.CrAnAlgebra) (h : b1 ≈ b2) :
|
||||
ι [a, b1]ₛca = ι [a, b2]ₛca := by
|
||||
rw [equiv_iff_sub_mem_ideal] at h
|
||||
rw [LinearMap.sub_mem_ker_iff.mp]
|
||||
simp only [LinearMap.mem_ker, ← map_sub]
|
||||
exact ι_superCommute_right_zero_of_mem_ideal a (b1 - b2) h
|
||||
|
||||
/-- The super commutor on the `FieldOpAlgebra` defined as a linear map `[a,_]ₛ`. -/
|
||||
noncomputable def superCommuteRight (a : 𝓕.CrAnAlgebra) :
|
||||
FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 where
|
||||
toFun := Quotient.lift (ι.toLinearMap ∘ₗ CrAnAlgebra.superCommute a)
|
||||
(ι_superCommute_eq_of_equiv_right a)
|
||||
map_add' x y := by
|
||||
obtain ⟨x, hx⟩ := ι_surjective x
|
||||
obtain ⟨y, hy⟩ := ι_surjective y
|
||||
subst hx hy
|
||||
rw [← map_add, ι_apply, ι_apply, ι_apply]
|
||||
rw [Quotient.lift_mk, Quotient.lift_mk, Quotient.lift_mk]
|
||||
simp
|
||||
map_smul' c y := by
|
||||
obtain ⟨y, hy⟩ := ι_surjective y
|
||||
subst hy
|
||||
rw [← map_smul, ι_apply, ι_apply]
|
||||
simp
|
||||
|
||||
lemma superCommuteRight_apply_ι (a b : 𝓕.CrAnAlgebra) :
|
||||
superCommuteRight a (ι b) = ι [a, b]ₛca := by rfl
|
||||
|
||||
lemma superCommuteRight_apply_quot (a b : 𝓕.CrAnAlgebra) :
|
||||
superCommuteRight a ⟦b⟧= ι [a, b]ₛca := by rfl
|
||||
|
||||
lemma superCommuteRight_eq_of_equiv (a1 a2 : 𝓕.CrAnAlgebra) (h : a1 ≈ a2) :
|
||||
superCommuteRight a1 = superCommuteRight a2 := by
|
||||
rw [equiv_iff_sub_mem_ideal] at h
|
||||
ext b
|
||||
obtain ⟨b, rfl⟩ := ι_surjective b
|
||||
have ha1b1 : (superCommuteRight (a1 - a2)) (ι b) = 0 := by
|
||||
rw [superCommuteRight_apply_ι]
|
||||
apply ι_superCommute_eq_zero_of_ι_left_zero
|
||||
exact (ι_eq_zero_iff_mem_ideal (a1 - a2)).mpr h
|
||||
simp_all only [superCommuteRight_apply_ι, map_sub, LinearMap.sub_apply]
|
||||
trans ι ((superCommute a2) b) + 0
|
||||
rw [← ha1b1]
|
||||
simp only [add_sub_cancel]
|
||||
simp
|
||||
|
||||
/-- The super commutor on the `FieldOpAlgebra`. -/
|
||||
noncomputable def superCommute : FieldOpAlgebra 𝓕 →ₗ[ℂ]
|
||||
FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 where
|
||||
toFun := Quotient.lift superCommuteRight superCommuteRight_eq_of_equiv
|
||||
map_add' x y := by
|
||||
obtain ⟨x, rfl⟩ := ι_surjective x
|
||||
obtain ⟨y, rfl⟩ := ι_surjective y
|
||||
ext b
|
||||
obtain ⟨b, rfl⟩ := ι_surjective b
|
||||
rw [← map_add, ι_apply, ι_apply, ι_apply, ι_apply]
|
||||
rw [Quotient.lift_mk, Quotient.lift_mk, Quotient.lift_mk]
|
||||
simp only [LinearMap.add_apply]
|
||||
rw [superCommuteRight_apply_quot, superCommuteRight_apply_quot, superCommuteRight_apply_quot]
|
||||
simp
|
||||
map_smul' c y := by
|
||||
obtain ⟨y, rfl⟩ := ι_surjective y
|
||||
ext b
|
||||
obtain ⟨b, rfl⟩ := ι_surjective b
|
||||
rw [← map_smul, ι_apply, ι_apply, ι_apply]
|
||||
simp only [Quotient.lift_mk, RingHom.id_apply, LinearMap.smul_apply]
|
||||
rw [superCommuteRight_apply_quot, superCommuteRight_apply_quot]
|
||||
simp
|
||||
|
||||
lemma ι_superCommute (a b : 𝓕.CrAnAlgebra) : ι [a, b]ₛca = superCommute (ι a) (ι b) := by
|
||||
rfl
|
||||
|
||||
end FieldOpAlgebra
|
||||
end FieldSpecification
|
|
@ -0,0 +1,385 @@
|
|||
/-
|
||||
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.TimeOrder
|
||||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.Basic
|
||||
/-!
|
||||
|
||||
# Time Ordering on Field operator algebra
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldSpecification
|
||||
open CrAnAlgebra
|
||||
open HepLean.List
|
||||
open FieldStatistic
|
||||
|
||||
namespace FieldOpAlgebra
|
||||
variable {𝓕 : FieldSpecification}
|
||||
|
||||
lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 : 𝓕.CrAnStates}
|
||||
(φs1 φs2 : List 𝓕.CrAnStates) (h :
|
||||
crAnTimeOrderRel φ1 φ2 ∧ crAnTimeOrderRel φ1 φ3 ∧
|
||||
crAnTimeOrderRel φ2 φ1 ∧ crAnTimeOrderRel φ2 φ3 ∧
|
||||
crAnTimeOrderRel φ3 φ1 ∧ crAnTimeOrderRel φ3 φ2) :
|
||||
ι 𝓣ᶠ(ofCrAnList φs1 * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca *
|
||||
ofCrAnList φs2) = 0 := by
|
||||
let l1 :=
|
||||
(List.takeWhile (fun c => ¬ crAnTimeOrderRel φ1 c)
|
||||
((φs1 ++ φs2).insertionSort crAnTimeOrderRel))
|
||||
++ (List.filter (fun c => crAnTimeOrderRel φ1 c ∧ crAnTimeOrderRel c φ1) φs1)
|
||||
let l2 := (List.filter (fun c => crAnTimeOrderRel φ1 c ∧ crAnTimeOrderRel c φ1) φs2)
|
||||
++ (List.filter (fun c => crAnTimeOrderRel φ1 c ∧ ¬ crAnTimeOrderRel c φ1)
|
||||
((φs1 ++ φs2).insertionSort crAnTimeOrderRel))
|
||||
have h123 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ1 :: φ2 :: φ3 :: φs2)) =
|
||||
crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2)
|
||||
• (ι (ofCrAnList l1) * ι (ofCrAnList [φ1, φ2, φ3]) * ι (ofCrAnList l2)) := by
|
||||
have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ1, φ2, φ3] φs2
|
||||
(by simp_all)
|
||||
rw [timeOrder_ofCrAnList, show φs1 ++ φ1 :: φ2 :: φ3 :: φs2 = φs1 ++ [φ1, φ2, φ3] ++ φs2
|
||||
by simp, crAnTimeOrderList, h1]
|
||||
simp only [List.append_assoc, List.singleton_append, decide_not,
|
||||
Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc]
|
||||
have h132 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ1 :: φ3 :: φ2 :: φs2)) =
|
||||
crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2)
|
||||
• (ι (ofCrAnList l1) * ι (ofCrAnList [φ1, φ3, φ2]) * ι (ofCrAnList l2)) := by
|
||||
have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ1, φ3, φ2] φs2
|
||||
(by simp_all)
|
||||
rw [timeOrder_ofCrAnList, show φs1 ++ φ1 :: φ3 :: φ2 :: φs2 = φs1 ++ [φ1, φ3, φ2] ++ φs2
|
||||
by simp, crAnTimeOrderList, h1]
|
||||
simp only [List.singleton_append, decide_not,
|
||||
Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc]
|
||||
congr 1
|
||||
have hp : List.Perm [φ1, φ3, φ2] [φ1, φ2, φ3] := by
|
||||
refine List.Perm.cons φ1 ?_
|
||||
exact List.Perm.swap φ2 φ3 []
|
||||
rw [crAnTimeOrderSign, Wick.koszulSign_perm_eq _ _ φ1 _ _ _ _ _ hp, ← crAnTimeOrderSign]
|
||||
· simp
|
||||
· intro φ4 hφ4
|
||||
simp only [List.mem_cons, List.mem_singleton, List.not_mem_nil, or_false] at hφ4
|
||||
rcases hφ4 with hφ4 | hφ4 | hφ4
|
||||
all_goals
|
||||
subst hφ4
|
||||
simp_all
|
||||
have hp231 : List.Perm [φ2, φ3, φ1] [φ1, φ2, φ3] := by
|
||||
refine List.Perm.trans (l₂ := [φ2, φ1, φ3]) ?_ ?_
|
||||
refine List.Perm.cons φ2 (List.Perm.swap φ1 φ3 [])
|
||||
exact List.Perm.swap φ1 φ2 [φ3]
|
||||
have h231 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ2 :: φ3 :: φ1 :: φs2)) =
|
||||
crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2)
|
||||
• (ι (ofCrAnList l1) * ι (ofCrAnList [φ2, φ3, φ1]) * ι (ofCrAnList l2)) := by
|
||||
have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ2, φ3, φ1] φs2
|
||||
(by simp_all)
|
||||
rw [timeOrder_ofCrAnList, show φs1 ++ φ2 :: φ3 :: φ1 :: φs2 = φs1 ++ [φ2, φ3, φ1] ++ φs2
|
||||
by simp, crAnTimeOrderList, h1]
|
||||
simp only [List.singleton_append, decide_not,
|
||||
Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc]
|
||||
congr 1
|
||||
rw [crAnTimeOrderSign, Wick.koszulSign_perm_eq _ _ φ1 _ _ _ _ _ hp231, ← crAnTimeOrderSign]
|
||||
· simp
|
||||
· intro φ4 hφ4
|
||||
simp only [List.mem_cons, List.mem_singleton, List.not_mem_nil, or_false] at hφ4
|
||||
rcases hφ4 with hφ4 | hφ4 | hφ4
|
||||
all_goals
|
||||
subst hφ4
|
||||
simp_all
|
||||
have h321 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ3 :: φ2 :: φ1 :: φs2)) =
|
||||
crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2)
|
||||
• (ι (ofCrAnList l1) * ι (ofCrAnList [φ3, φ2, φ1]) * ι (ofCrAnList l2)) := by
|
||||
have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ3, φ2, φ1] φs2
|
||||
(by simp_all)
|
||||
rw [timeOrder_ofCrAnList, show φs1 ++ φ3 :: φ2 :: φ1 :: φs2 = φs1 ++ [φ3, φ2, φ1] ++ φs2
|
||||
by simp, crAnTimeOrderList, h1]
|
||||
simp only [List.singleton_append, decide_not,
|
||||
Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc]
|
||||
congr 1
|
||||
have hp : List.Perm [φ3, φ2, φ1] [φ1, φ2, φ3] := by
|
||||
refine List.Perm.trans ?_ hp231
|
||||
exact List.Perm.swap φ2 φ3 [φ1]
|
||||
rw [crAnTimeOrderSign, Wick.koszulSign_perm_eq _ _ φ1 _ _ _ _ _ hp, ← crAnTimeOrderSign]
|
||||
· simp
|
||||
· intro φ4 hφ4
|
||||
simp only [List.mem_cons, List.mem_singleton, List.not_mem_nil, or_false] at hφ4
|
||||
rcases hφ4 with hφ4 | hφ4 | hφ4
|
||||
all_goals
|
||||
subst hφ4
|
||||
simp_all
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
simp only [List.singleton_append, instCommGroup.eq_1, ofList_singleton, map_sub, map_smul]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList, superCommute_ofCrAnList_ofCrAnList]
|
||||
simp only [List.cons_append, List.nil_append, instCommGroup.eq_1, ofList_singleton, mul_sub, ←
|
||||
ofCrAnList_append, Algebra.mul_smul_comm, sub_mul, List.append_assoc, Algebra.smul_mul_assoc,
|
||||
map_sub, map_smul]
|
||||
rw [h123, h132, h231, h321]
|
||||
simp only [smul_smul]
|
||||
rw [mul_comm, ← smul_smul, mul_comm, ← smul_smul]
|
||||
rw [← smul_sub, ← smul_sub, smul_smul, mul_comm, ← smul_smul, ← smul_sub]
|
||||
simp only [smul_eq_zero]
|
||||
right
|
||||
rw [← smul_mul_assoc, ← mul_smul_comm, mul_assoc]
|
||||
rw [← smul_mul_assoc, ← mul_smul_comm]
|
||||
rw [smul_sub]
|
||||
rw [← smul_mul_assoc, ← mul_smul_comm]
|
||||
rw [← smul_mul_assoc, ← mul_smul_comm]
|
||||
repeat rw [mul_assoc]
|
||||
rw [← mul_sub, ← mul_sub, ← mul_sub]
|
||||
rw [← sub_mul, ← sub_mul, ← sub_mul]
|
||||
trans ι (ofCrAnList l1) * ι [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca *
|
||||
ι (ofCrAnList l2)
|
||||
rw [mul_assoc]
|
||||
congr
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList]
|
||||
simp only [List.singleton_append, instCommGroup.eq_1, ofList_singleton, map_sub, map_smul]
|
||||
rw [superCommute_ofCrAnList_ofCrAnList, superCommute_ofCrAnList_ofCrAnList]
|
||||
simp only [List.cons_append, List.nil_append, instCommGroup.eq_1, ofList_singleton, map_sub,
|
||||
map_smul, smul_sub]
|
||||
simp_all
|
||||
|
||||
lemma ι_timeOrder_superCommute_superCommute_ofCrAnList {φ1 φ2 φ3 : 𝓕.CrAnStates}
|
||||
(φs1 φs2 : List 𝓕.CrAnStates) :
|
||||
ι 𝓣ᶠ(ofCrAnList φs1 * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * ofCrAnList φs2)
|
||||
= 0 := by
|
||||
by_cases h :
|
||||
crAnTimeOrderRel φ1 φ2 ∧ crAnTimeOrderRel φ1 φ3 ∧
|
||||
crAnTimeOrderRel φ2 φ1 ∧ crAnTimeOrderRel φ2 φ3 ∧
|
||||
crAnTimeOrderRel φ3 φ1 ∧ crAnTimeOrderRel φ3 φ2
|
||||
· exact ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList φs1 φs2 h
|
||||
· rw [timeOrder_timeOrder_mid]
|
||||
rw [timeOrder_superCommute_ofCrAnState_superCommute_all_not_crAnTimeOrderRel _ _ _ h]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma ι_timeOrder_superCommute_superCommute {φ1 φ2 φ3 : 𝓕.CrAnStates} (a b : 𝓕.CrAnAlgebra) :
|
||||
ι 𝓣ᶠ(a * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * b) = 0 := by
|
||||
let pb (b : 𝓕.CrAnAlgebra) (hc : b ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := ι 𝓣ᶠ(a * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * b) = 0
|
||||
change pb b (Basis.mem_span _ b)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pb]
|
||||
let pa (a : 𝓕.CrAnAlgebra) (hc : a ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := ι 𝓣ᶠ(a * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * ofCrAnList φs) = 0
|
||||
change pa a (Basis.mem_span _ a)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, pa]
|
||||
exact ι_timeOrder_superCommute_superCommute_ofCrAnList φs' φs
|
||||
· simp [pa]
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all [pa,mul_add, add_mul]
|
||||
· intro x hx hpx
|
||||
simp_all [pa, hpx]
|
||||
· simp [pb]
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all [pb,mul_add, add_mul]
|
||||
· intro x hx hpx
|
||||
simp_all [pb, hpx]
|
||||
|
||||
lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates}
|
||||
(hφψ : crAnTimeOrderRel φ ψ) (hψφ : crAnTimeOrderRel ψ φ) (a b : 𝓕.CrAnAlgebra) :
|
||||
ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) =
|
||||
ι ([ofCrAnState φ, ofCrAnState ψ]ₛca * 𝓣ᶠ(a * b)) := by
|
||||
let pb (b : 𝓕.CrAnAlgebra) (hc : b ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) =
|
||||
ι ([ofCrAnState φ, ofCrAnState ψ]ₛca * 𝓣ᶠ(a * b))
|
||||
change pb b (Basis.mem_span _ b)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs, rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, map_mul, pb]
|
||||
let pa (a : 𝓕.CrAnAlgebra) (hc : a ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) :
|
||||
Prop := ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * ofCrAnList φs) =
|
||||
ι ([ofCrAnState φ, ofCrAnState ψ]ₛca * 𝓣ᶠ(a* ofCrAnList φs))
|
||||
change pa a (Basis.mem_span _ a)
|
||||
apply Submodule.span_induction
|
||||
· intro x hx
|
||||
obtain ⟨φs', rfl⟩ := hx
|
||||
simp only [ofListBasis_eq_ofList, map_mul, pa]
|
||||
conv_lhs =>
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
|
||||
simp [mul_sub, sub_mul, ← ofCrAnList_append]
|
||||
rw [timeOrder_ofCrAnList, timeOrder_ofCrAnList]
|
||||
have h1 : crAnTimeOrderSign (φs' ++ φ :: ψ :: φs) =
|
||||
crAnTimeOrderSign (φs' ++ ψ :: φ :: φs) := by
|
||||
trans crAnTimeOrderSign (φs' ++ [φ, ψ] ++ φs)
|
||||
simp only [List.append_assoc, List.cons_append, List.nil_append]
|
||||
rw [crAnTimeOrderSign]
|
||||
have hp : List.Perm [φ,ψ] [ψ,φ] := by exact List.Perm.swap ψ φ []
|
||||
rw [Wick.koszulSign_perm_eq _ _ φ _ _ _ _ _ hp]
|
||||
simp only [List.append_assoc, List.cons_append, List.singleton_append]
|
||||
rfl
|
||||
simp_all
|
||||
rw [h1]
|
||||
simp only [map_smul]
|
||||
have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ φs' [φ, ψ] φs
|
||||
(by simp_all)
|
||||
rw [crAnTimeOrderList, show φs' ++ φ :: ψ :: φs = φs' ++ [φ, ψ] ++ φs by simp, h1]
|
||||
have h2 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ φs' [ψ, φ] φs
|
||||
(by simp_all)
|
||||
rw [crAnTimeOrderList, show φs' ++ ψ :: φ :: φs = φs' ++ [ψ, φ] ++ φs by simp, h2]
|
||||
repeat rw [ofCrAnList_append]
|
||||
rw [smul_smul, mul_comm, ← smul_smul, ← smul_sub]
|
||||
rw [map_mul, map_mul, map_mul, map_mul, map_mul, map_mul, map_mul, map_mul]
|
||||
rw [← mul_smul_comm]
|
||||
rw [mul_assoc, mul_assoc, mul_assoc, mul_assoc, mul_assoc, mul_assoc]
|
||||
rw [← mul_sub, ← mul_sub, mul_smul_comm, mul_smul_comm, ← smul_mul_assoc,
|
||||
← smul_mul_assoc]
|
||||
rw [← sub_mul]
|
||||
have h1 : (ι (ofCrAnList [φ, ψ]) -
|
||||
(exchangeSign (𝓕.crAnStatistics φ)) (𝓕.crAnStatistics ψ) • ι (ofCrAnList [ψ, φ])) =
|
||||
ι [ofCrAnState φ, ofCrAnState ψ]ₛca := by
|
||||
rw [superCommute_ofCrAnState_ofCrAnState]
|
||||
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_append]
|
||||
simp only [instCommGroup.eq_1, List.singleton_append, Algebra.smul_mul_assoc, map_sub,
|
||||
map_smul]
|
||||
rw [← ofCrAnList_append]
|
||||
simp
|
||||
rw [h1]
|
||||
have hc : ι ((superCommute (ofCrAnState φ)) (ofCrAnState ψ)) ∈
|
||||
Subalgebra.center ℂ 𝓕.FieldOpAlgebra := by
|
||||
apply ι_superCommute_ofCrAnState_ofCrAnState_mem_center
|
||||
rw [Subalgebra.mem_center_iff] at hc
|
||||
repeat rw [← mul_assoc]
|
||||
rw [hc]
|
||||
repeat rw [mul_assoc]
|
||||
rw [smul_mul_assoc]
|
||||
rw [← map_mul, ← map_mul, ← map_mul, ← map_mul]
|
||||
rw [← ofCrAnList_append, ← ofCrAnList_append, ← ofCrAnList_append, ← ofCrAnList_append]
|
||||
have h1 := insertionSort_of_takeWhile_filter 𝓕.crAnTimeOrderRel φ φs' φs
|
||||
simp only [decide_not, Bool.decide_and, List.append_assoc, List.cons_append,
|
||||
List.singleton_append, Algebra.mul_smul_comm, map_mul] at h1 ⊢
|
||||
rw [← h1]
|
||||
rw [← crAnTimeOrderList]
|
||||
by_cases hq : (𝓕 |>ₛ φ) ≠ (𝓕 |>ₛ ψ)
|
||||
· rw [ι_superCommute_of_diff_statistic hq]
|
||||
simp
|
||||
· rw [crAnTimeOrderSign, Wick.koszulSign_eq_rel_eq_stat _ _, ← crAnTimeOrderSign]
|
||||
rw [timeOrder_ofCrAnList]
|
||||
simp only [map_smul, Algebra.mul_smul_comm]
|
||||
simp only [List.nil_append]
|
||||
exact hψφ
|
||||
exact hφψ
|
||||
simpa using hq
|
||||
· simp only [map_mul, zero_mul, map_zero, mul_zero, pa]
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all [pa,mul_add, add_mul]
|
||||
· intro x hx hpx
|
||||
simp_all [pa, hpx]
|
||||
· simp only [map_mul, mul_zero, map_zero, pb]
|
||||
· intro x y hx hy hpx hpy
|
||||
simp_all [pb,mul_add, add_mul]
|
||||
· intro x hx hpx
|
||||
simp_all [pb, hpx]
|
||||
|
||||
lemma ι_timeOrder_superCommute_neq_time {φ ψ : 𝓕.CrAnStates}
|
||||
(hφψ : ¬ (crAnTimeOrderRel φ ψ ∧ crAnTimeOrderRel ψ φ)) (a b : 𝓕.CrAnAlgebra) :
|
||||
ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = 0 := by
|
||||
rw [timeOrder_timeOrder_mid]
|
||||
have hφψ : ¬ (crAnTimeOrderRel φ ψ) ∨ ¬ (crAnTimeOrderRel ψ φ) := by
|
||||
exact Decidable.not_and_iff_or_not.mp hφψ
|
||||
rcases hφψ with hφψ | hφψ
|
||||
· rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel]
|
||||
simp_all only [false_and, not_false_eq_true, false_or, mul_zero, zero_mul, map_zero]
|
||||
simp_all
|
||||
· rw [superCommute_ofCrAnState_ofCrAnState_symm]
|
||||
simp only [instCommGroup.eq_1, neg_smul, map_neg, map_smul, mul_neg, Algebra.mul_smul_comm,
|
||||
neg_mul, Algebra.smul_mul_assoc, neg_eq_zero, smul_eq_zero]
|
||||
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel]
|
||||
simp only [mul_zero, zero_mul, map_zero, or_true]
|
||||
simp_all
|
||||
|
||||
/-!
|
||||
|
||||
## Defining normal order for `FiedOpAlgebra`.
|
||||
|
||||
-/
|
||||
|
||||
lemma ι_timeOrder_zero_of_mem_ideal (a : 𝓕.CrAnAlgebra)
|
||||
(h : a ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet) : ι 𝓣ᶠ(a) = 0 := by
|
||||
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] at h
|
||||
let p {k : Set 𝓕.CrAnAlgebra} (a : CrAnAlgebra 𝓕) (h : a ∈ AddSubgroup.closure k) := ι 𝓣ᶠ(a) = 0
|
||||
change p a h
|
||||
apply AddSubgroup.closure_induction
|
||||
· intro x hx
|
||||
obtain ⟨a, ha, b, hb, rfl⟩ := Set.mem_mul.mp hx
|
||||
obtain ⟨a, ha, c, hc, rfl⟩ := ha
|
||||
simp only [p]
|
||||
simp only [fieldOpIdealSet, exists_prop, exists_and_left, Set.mem_setOf_eq] at hc
|
||||
match hc with
|
||||
| Or.inl hc =>
|
||||
obtain ⟨φa, φa', hφa, hφa', rfl⟩ := hc
|
||||
simp only [ι_timeOrder_superCommute_superCommute]
|
||||
| Or.inr (Or.inl hc) =>
|
||||
obtain ⟨φa, hφa, φb, hφb, rfl⟩ := hc
|
||||
by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa)
|
||||
· rw [ι_timeOrder_superCommute_eq_time]
|
||||
simp only [map_mul]
|
||||
rw [ι_superCommute_of_create_create]
|
||||
simp only [zero_mul]
|
||||
· exact hφa
|
||||
· exact hφb
|
||||
· exact heqt.1
|
||||
· exact heqt.2
|
||||
· rw [ι_timeOrder_superCommute_neq_time heqt]
|
||||
| Or.inr (Or.inr (Or.inl hc)) =>
|
||||
obtain ⟨φa, hφa, φb, hφb, rfl⟩ := hc
|
||||
by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa)
|
||||
· rw [ι_timeOrder_superCommute_eq_time]
|
||||
simp only [map_mul]
|
||||
rw [ι_superCommute_of_annihilate_annihilate]
|
||||
simp only [zero_mul]
|
||||
· exact hφa
|
||||
· exact hφb
|
||||
· exact heqt.1
|
||||
· exact heqt.2
|
||||
· rw [ι_timeOrder_superCommute_neq_time heqt]
|
||||
| Or.inr (Or.inr (Or.inr hc)) =>
|
||||
obtain ⟨φa, φb, hdiff, rfl⟩ := hc
|
||||
by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa)
|
||||
· rw [ι_timeOrder_superCommute_eq_time]
|
||||
simp only [map_mul]
|
||||
rw [ι_superCommute_of_diff_statistic]
|
||||
simp only [zero_mul]
|
||||
· exact hdiff
|
||||
· exact heqt.1
|
||||
· exact heqt.2
|
||||
· rw [ι_timeOrder_superCommute_neq_time heqt]
|
||||
· simp [p]
|
||||
· intro x y hx hy
|
||||
simp only [map_add, p]
|
||||
intro h1 h2
|
||||
simp [h1, h2]
|
||||
· intro x hx
|
||||
simp [p]
|
||||
|
||||
lemma ι_timeOrder_eq_of_equiv (a b : 𝓕.CrAnAlgebra) (h : a ≈ b) :
|
||||
ι 𝓣ᶠ(a) = ι 𝓣ᶠ(b) := by
|
||||
rw [equiv_iff_sub_mem_ideal] at h
|
||||
rw [LinearMap.sub_mem_ker_iff.mp]
|
||||
simp only [LinearMap.mem_ker, ← map_sub]
|
||||
exact ι_timeOrder_zero_of_mem_ideal (a - b) h
|
||||
|
||||
/-- Normal ordering on `FieldOpAlgebra`. -/
|
||||
noncomputable def timeOrder : FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 where
|
||||
toFun := Quotient.lift (ι.toLinearMap ∘ₗ CrAnAlgebra.timeOrder) ι_timeOrder_eq_of_equiv
|
||||
map_add' x y := by
|
||||
obtain ⟨x, hx⟩ := ι_surjective x
|
||||
obtain ⟨y, hy⟩ := ι_surjective y
|
||||
subst hx hy
|
||||
rw [← map_add, ι_apply, ι_apply, ι_apply]
|
||||
rw [Quotient.lift_mk, Quotient.lift_mk, Quotient.lift_mk]
|
||||
simp
|
||||
map_smul' c y := by
|
||||
obtain ⟨y, hy⟩ := ι_surjective y
|
||||
subst hy
|
||||
rw [← map_smul, ι_apply, ι_apply]
|
||||
simp
|
||||
|
||||
end FieldOpAlgebra
|
||||
end FieldSpecification
|
|
@ -234,27 +234,10 @@ lemma crAnTimeOrderSign_pair_not_ordered {φ ψ : 𝓕.CrAnStates} (h : ¬ crAnT
|
|||
rw [if_neg h]
|
||||
simp [FieldStatistic.exchangeSign_eq_if]
|
||||
|
||||
lemma crAnTimeOrderSign_swap_eq_time_cons {φ ψ : 𝓕.CrAnStates}
|
||||
(h1 : crAnTimeOrderRel φ ψ) (h2 : crAnTimeOrderRel ψ φ) (φs' : List 𝓕.CrAnStates) :
|
||||
crAnTimeOrderSign (φ :: ψ :: φs') = crAnTimeOrderSign (ψ :: φ :: φs') := by
|
||||
simp only [crAnTimeOrderSign, Wick.koszulSign, ← mul_assoc, mul_eq_mul_right_iff]
|
||||
left
|
||||
rw [mul_comm]
|
||||
simp [Wick.koszulSignInsert, h1, h2]
|
||||
|
||||
lemma crAnTimeOrderSign_swap_eq_time {φ ψ : 𝓕.CrAnStates}
|
||||
(h1 : crAnTimeOrderRel φ ψ) (h2 : crAnTimeOrderRel ψ φ) : (φs φs' : List 𝓕.CrAnStates) →
|
||||
crAnTimeOrderSign (φs ++ φ :: ψ :: φs') = crAnTimeOrderSign (φs ++ ψ :: φ :: φs')
|
||||
| [], φs' => by
|
||||
simp only [crAnTimeOrderSign, List.nil_append]
|
||||
exact crAnTimeOrderSign_swap_eq_time_cons h1 h2 φs'
|
||||
| φ'' :: φs, φs' => by
|
||||
simp only [crAnTimeOrderSign, Wick.koszulSign, List.append_eq]
|
||||
rw [← crAnTimeOrderSign, ← crAnTimeOrderSign]
|
||||
rw [crAnTimeOrderSign_swap_eq_time h1 h2]
|
||||
congr 1
|
||||
apply Wick.koszulSignInsert_eq_perm
|
||||
exact List.Perm.append_left φs (List.Perm.swap ψ φ φs')
|
||||
(h1 : crAnTimeOrderRel φ ψ) (h2 : crAnTimeOrderRel ψ φ) (φs φs' : List 𝓕.CrAnStates) :
|
||||
crAnTimeOrderSign (φs ++ φ :: ψ :: φs') = crAnTimeOrderSign (φs ++ ψ :: φ :: φs') := by
|
||||
exact Wick.koszulSign_swap_eq_rel _ _ h1 h2 _ _
|
||||
|
||||
/-- Sort a list of `CrAnStates` based on `crAnTimeOrderRel`. -/
|
||||
def crAnTimeOrderList (φs : List 𝓕.CrAnStates) : List 𝓕.CrAnStates :=
|
||||
|
|
|
@ -235,6 +235,10 @@ lemma ofList_map_eq_finset_prod (s : 𝓕 → FieldStatistic) :
|
|||
simp only [List.length_cons, List.nodup_cons] at hl
|
||||
exact hl.2
|
||||
|
||||
lemma ofList_pair (s : 𝓕 → FieldStatistic) (φ1 φ2 : 𝓕) :
|
||||
ofList s [φ1, φ2] = s φ1 * s φ2 := by
|
||||
rw [ofList_cons_eq_mul, ofList_singleton]
|
||||
|
||||
/-!
|
||||
|
||||
## ofList and take
|
||||
|
@ -288,11 +292,12 @@ instance : AddMonoid FieldStatistic where
|
|||
add a b := a * b
|
||||
nsmul n a := ∏ (i : Fin n), a
|
||||
zero_add a := by
|
||||
cases a <;> simp <;> rfl
|
||||
cases a <;> simp only [instCommGroup] <;> rfl
|
||||
add_zero a := by
|
||||
cases a <;> simp <;> rfl
|
||||
cases a <;>
|
||||
simp only [instCommGroup] <;> rfl
|
||||
add_assoc a b c := by
|
||||
cases a <;> cases b <;> cases c <;> simp <;> rfl
|
||||
cases a <;> cases b <;> cases c <;> simp only [instCommGroup] <;> rfl
|
||||
nsmul_zero a := by
|
||||
simp only [Finset.univ_eq_empty, Finset.prod_const, instCommGroup, Finset.card_empty, pow_zero]
|
||||
rfl
|
||||
|
@ -300,5 +305,8 @@ instance : AddMonoid FieldStatistic where
|
|||
simp only [instCommGroup, Finset.prod_const, Finset.card_univ, Fintype.card_fin]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma add_eq_mul (a b : FieldStatistic) : a + b = a * b := rfl
|
||||
|
||||
end ofListTake
|
||||
end FieldStatistic
|
||||
|
|
|
@ -71,6 +71,10 @@ lemma exchangeSign_bosonic (a : FieldStatistic) : 𝓢(a, bosonic) = 1 := by
|
|||
lemma bosonic_exchangeSign (a : FieldStatistic) : 𝓢(bosonic, a) = 1 := by
|
||||
rw [exchangeSign_symm, exchangeSign_bosonic]
|
||||
|
||||
@[simp]
|
||||
lemma fermionic_exchangeSign_fermionic : 𝓢(fermionic, fermionic) = - 1 := by
|
||||
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
|
||||
|
|
|
@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSignInsert
|
||||
import HepLean.Mathematics.List.InsertionSort
|
||||
/-!
|
||||
|
||||
# Koszul sign
|
||||
|
@ -259,4 +260,182 @@ lemma koszulSign_eraseIdx_insertionSortMinPos [IsTotal 𝓕 le] [IsTrans 𝓕 le
|
|||
apply Or.inl
|
||||
rfl
|
||||
|
||||
lemma koszulSign_swap_eq_rel_cons {ψ φ : 𝓕}
|
||||
(h1 : le φ ψ) (h2 : le ψ φ) (φs' : List 𝓕) :
|
||||
koszulSign q le (φ :: ψ :: φs') = koszulSign q le (ψ :: φ :: φs') := by
|
||||
simp only [Wick.koszulSign, ← mul_assoc, mul_eq_mul_right_iff]
|
||||
left
|
||||
rw [mul_comm]
|
||||
simp [Wick.koszulSignInsert, h1, h2]
|
||||
|
||||
lemma koszulSign_swap_eq_rel {ψ φ : 𝓕} (h1 : le φ ψ) (h2 : le ψ φ) : (φs φs' : List 𝓕) →
|
||||
koszulSign q le (φs ++ φ :: ψ :: φs') = koszulSign q le (φs ++ ψ :: φ :: φs')
|
||||
| [], φs' => by
|
||||
simp only [List.nil_append]
|
||||
exact koszulSign_swap_eq_rel_cons q le h1 h2 φs'
|
||||
| φ'' :: φs, φs' => by
|
||||
simp only [Wick.koszulSign, List.append_eq]
|
||||
rw [koszulSign_swap_eq_rel h1 h2]
|
||||
congr 1
|
||||
apply Wick.koszulSignInsert_eq_perm
|
||||
exact List.Perm.append_left φs (List.Perm.swap ψ φ φs')
|
||||
|
||||
lemma koszulSign_eq_rel_eq_stat_append {ψ φ : 𝓕} [IsTrans 𝓕 le]
|
||||
(h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) →
|
||||
koszulSign q le (φ :: ψ :: φs) = koszulSign q le φs := by
|
||||
intro φs
|
||||
simp only [koszulSign, ← mul_assoc]
|
||||
trans 1 * koszulSign q le φs
|
||||
swap
|
||||
simp only [one_mul]
|
||||
congr
|
||||
simp only [koszulSignInsert, ite_mul, neg_mul]
|
||||
simp_all only [and_self, ite_true]
|
||||
rw [koszulSignInsert_eq_rel_eq_stat q le h1 h2 hq]
|
||||
simp
|
||||
|
||||
lemma koszulSign_eq_rel_eq_stat {ψ φ : 𝓕} [IsTrans 𝓕 le]
|
||||
(h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs' φs : List 𝓕) →
|
||||
koszulSign q le (φs' ++ φ :: ψ :: φs) = koszulSign q le (φs' ++ φs)
|
||||
| [], φs => by
|
||||
simp only [List.nil_append]
|
||||
exact koszulSign_eq_rel_eq_stat_append q le h1 h2 hq φs
|
||||
| φ'' :: φs', φs => by
|
||||
simp only [koszulSign, List.append_eq]
|
||||
rw [koszulSign_eq_rel_eq_stat h1 h2 hq φs' φs]
|
||||
simp only [mul_eq_mul_right_iff]
|
||||
left
|
||||
trans koszulSignInsert q le φ'' (φ :: ψ :: (φs' ++ φs))
|
||||
apply koszulSignInsert_eq_perm
|
||||
refine List.Perm.symm (List.perm_cons_append_cons φ ?_)
|
||||
exact List.Perm.symm List.perm_middle
|
||||
rw [koszulSignInsert_eq_remove_same_stat_append q le]
|
||||
exact h1
|
||||
exact h2
|
||||
exact hq
|
||||
|
||||
lemma koszulSign_of_sorted : (φs : List 𝓕)
|
||||
→ (hs : List.Sorted le φs) → koszulSign q le φs = 1
|
||||
| [], _ => by
|
||||
simp [koszulSign]
|
||||
| φ :: φs, h => by
|
||||
simp only [koszulSign]
|
||||
simp only [List.sorted_cons] at h
|
||||
rw [koszulSign_of_sorted φs h.2]
|
||||
simp only [mul_one]
|
||||
exact koszulSignInsert_of_le_mem _ _ _ _ h.1
|
||||
|
||||
@[simp]
|
||||
lemma koszulSign_of_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φs : List 𝓕) :
|
||||
koszulSign q le (List.insertionSort le φs) = 1 := by
|
||||
apply koszulSign_of_sorted
|
||||
exact List.sorted_insertionSort le φs
|
||||
|
||||
lemma koszulSign_of_append_eq_insertionSort_left [IsTotal 𝓕 le] [IsTrans 𝓕 le] :
|
||||
(φs φs' : List 𝓕) → koszulSign q le (φs ++ φs') =
|
||||
koszulSign q le (List.insertionSort le φs ++ φs') * koszulSign q le φs
|
||||
| φs, [] => by
|
||||
simp
|
||||
| φs, φ :: φs' => by
|
||||
have h1 : (φs ++ φ :: φs') = List.insertIdx φs.length φ (φs ++ φs') := by
|
||||
rw [insertIdx_length_fst_append]
|
||||
have h2 : (List.insertionSort le φs ++ φ :: φs') =
|
||||
List.insertIdx (List.insertionSort le φs).length φ (List.insertionSort le φs ++ φs') := by
|
||||
rw [insertIdx_length_fst_append]
|
||||
rw [h1, h2]
|
||||
rw [koszulSign_insertIdx]
|
||||
simp only [instCommGroup.eq_1, List.take_left', List.length_insertionSort]
|
||||
rw [koszulSign_insertIdx]
|
||||
simp only [mul_assoc, instCommGroup.eq_1, List.length_insertionSort, List.take_left',
|
||||
ofList_insertionSort, mul_eq_mul_left_iff]
|
||||
left
|
||||
rw [koszulSign_of_append_eq_insertionSort_left φs φs']
|
||||
simp only [mul_assoc, mul_eq_mul_left_iff]
|
||||
left
|
||||
simp only [mul_comm, mul_eq_mul_left_iff]
|
||||
left
|
||||
congr 3
|
||||
· have h2 : (List.insertionSort le φs ++ φ :: φs') =
|
||||
List.insertIdx φs.length φ (List.insertionSort le φs ++ φs') := by
|
||||
rw [← insertIdx_length_fst_append]
|
||||
simp
|
||||
rw [insertionSortEquiv_congr _ _ h2.symm]
|
||||
simp only [Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.cast_mk,
|
||||
Fin.coe_cast]
|
||||
rw [insertionSortEquiv_insertionSort_append]
|
||||
simp only [finCongr_apply, Fin.coe_cast]
|
||||
rw [insertionSortEquiv_congr _ _ h1.symm]
|
||||
simp
|
||||
· rw [insertIdx_length_fst_append]
|
||||
rw [show φs.length = (List.insertionSort le φs).length by simp]
|
||||
rw [insertIdx_length_fst_append]
|
||||
symm
|
||||
apply insertionSort_insertionSort_append
|
||||
· simp
|
||||
· simp
|
||||
|
||||
lemma koszulSign_of_append_eq_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le] : (φs'' φs φs' : List 𝓕) →
|
||||
koszulSign q le (φs'' ++ φs ++ φs') =
|
||||
koszulSign q le (φs'' ++ List.insertionSort le φs ++ φs') * koszulSign q le φs
|
||||
| [], φs, φs'=> by
|
||||
simp only [List.nil_append]
|
||||
exact koszulSign_of_append_eq_insertionSort_left q le φs φs'
|
||||
| φ'' :: φs'', φs, φs' => by
|
||||
simp only [koszulSign, List.append_eq]
|
||||
rw [koszulSign_of_append_eq_insertionSort φs'' φs φs', ← mul_assoc]
|
||||
congr 2
|
||||
apply koszulSignInsert_eq_perm
|
||||
refine (List.perm_append_right_iff φs').mpr ?_
|
||||
refine List.Perm.append_left φs'' ?_
|
||||
exact List.Perm.symm (List.perm_insertionSort le φs)
|
||||
|
||||
/-!
|
||||
|
||||
# koszulSign with permutations
|
||||
|
||||
-/
|
||||
|
||||
lemma koszulSign_perm_eq_append [IsTrans 𝓕 le] (φ : 𝓕) (φs φs' φs2 : List 𝓕)
|
||||
(hp : φs.Perm φs') : (h : ∀ φ' ∈ φs, le φ φ' ∧ le φ' φ) →
|
||||
koszulSign q le (φs ++ φs2) = koszulSign q le (φs' ++ φs2) := by
|
||||
let motive (φs φs' : List 𝓕) (hp : φs.Perm φs') : Prop :=
|
||||
(h : ∀ φ' ∈ φs, le φ φ' ∧ le φ' φ) →
|
||||
koszulSign q le (φs ++ φs2) = koszulSign q le (φs' ++ φs2)
|
||||
change motive φs φs' hp
|
||||
apply List.Perm.recOn
|
||||
· simp [motive]
|
||||
· intro x l1 l2 h ih hxφ
|
||||
simp_all only [List.mem_cons, or_true, and_self, implies_true, nonempty_prop, forall_const,
|
||||
forall_eq_or_imp, List.cons_append, motive]
|
||||
simp only [koszulSign, ih, mul_eq_mul_right_iff]
|
||||
left
|
||||
apply koszulSignInsert_eq_perm
|
||||
exact (List.perm_append_right_iff φs2).mpr h
|
||||
· intro x y l h
|
||||
simp_all only [List.mem_cons, forall_eq_or_imp, List.cons_append]
|
||||
apply Wick.koszulSign_swap_eq_rel_cons
|
||||
exact IsTrans.trans y φ x h.1.2 h.2.1.1
|
||||
exact IsTrans.trans x φ y h.2.1.2 h.1.1
|
||||
· intro l1 l2 l3 h1 h2 ih1 ih2 h
|
||||
simp_all only [and_self, implies_true, nonempty_prop, forall_const, motive]
|
||||
refine (ih2 ?_)
|
||||
intro φ' hφ
|
||||
refine h φ' ?_
|
||||
exact (List.Perm.mem_iff (id (List.Perm.symm h1))).mp hφ
|
||||
|
||||
lemma koszulSign_perm_eq [IsTrans 𝓕 le] (φ : 𝓕) : (φs1 φs φs' φs2 : List 𝓕) →
|
||||
(h : ∀ φ' ∈ φs, le φ φ' ∧ le φ' φ) → (hp : φs.Perm φs') →
|
||||
koszulSign q le (φs1 ++ φs ++ φs2) = koszulSign q le (φs1 ++ φs' ++ φs2)
|
||||
| [], φs, φs', φs2, h, hp => by
|
||||
simp only [List.nil_append]
|
||||
exact koszulSign_perm_eq_append q le φ φs φs' φs2 hp h
|
||||
| φ1 :: φs1, φs, φs', φs2, h, hp => by
|
||||
simp only [koszulSign, List.append_eq]
|
||||
have ih := koszulSign_perm_eq φ φs1 φs φs' φs2 h hp
|
||||
rw [ih]
|
||||
congr 1
|
||||
apply koszulSignInsert_eq_perm
|
||||
refine (List.perm_append_right_iff φs2).mpr ?_
|
||||
exact List.Perm.append_left φs1 hp
|
||||
|
||||
end Wick
|
||||
|
|
|
@ -235,4 +235,52 @@ lemma koszulSignInsert_cons (r0 r1 : 𝓕) (r : List 𝓕) :
|
|||
koszulSignInsert q le r0 r := by
|
||||
simp [koszulSignInsert, koszulSignCons]
|
||||
|
||||
lemma koszulSignInsert_of_le_mem (φ0 : 𝓕) : (φs : List 𝓕) → (h : ∀ b ∈ φs, le φ0 b) →
|
||||
koszulSignInsert q le φ0 φs = 1
|
||||
| [], _ => by
|
||||
simp [koszulSignInsert]
|
||||
| φ1 :: φs, h => by
|
||||
simp only [koszulSignInsert]
|
||||
rw [if_pos]
|
||||
· apply koszulSignInsert_of_le_mem
|
||||
· intro b hb
|
||||
exact h b (List.mem_cons_of_mem _ hb)
|
||||
· exact h φ1 (List.mem_cons_self _ _)
|
||||
|
||||
lemma koszulSignInsert_eq_rel_eq_stat {ψ φ : 𝓕} [IsTrans 𝓕 le]
|
||||
(h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) →
|
||||
koszulSignInsert q le φ φs = koszulSignInsert q le ψ φs
|
||||
| [] => by
|
||||
simp [koszulSignInsert]
|
||||
| φ' :: φs => by
|
||||
simp only [koszulSignInsert]
|
||||
simp_all only
|
||||
by_cases hr : le φ φ'
|
||||
· simp only [hr, ↓reduceIte]
|
||||
have h1' : le ψ φ' := by
|
||||
apply IsTrans.trans ψ φ φ' h2 hr
|
||||
simp only [h1', ↓reduceIte]
|
||||
exact koszulSignInsert_eq_rel_eq_stat h1 h2 hq φs
|
||||
· have hψφ' : ¬ le ψ φ' := by
|
||||
intro hψφ'
|
||||
apply hr
|
||||
apply IsTrans.trans φ ψ φ' h1 hψφ'
|
||||
simp only [hr, ↓reduceIte, hψφ']
|
||||
rw [koszulSignInsert_eq_rel_eq_stat h1 h2 hq φs]
|
||||
|
||||
lemma koszulSignInsert_eq_remove_same_stat_append {ψ φ φ' : 𝓕} [IsTrans 𝓕 le]
|
||||
(h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) →
|
||||
koszulSignInsert q le φ' (φ :: ψ :: φs) = koszulSignInsert q le φ' φs := by
|
||||
intro φs
|
||||
simp_all only [koszulSignInsert, and_self, ite_true, ite_false, ite_self]
|
||||
by_cases hφ'φ : le φ' φ
|
||||
· have hφ'ψ : le φ' ψ := by
|
||||
apply IsTrans.trans φ' φ ψ hφ'φ h1
|
||||
simp [hφ'φ, hφ'ψ]
|
||||
· have hφ'ψ : ¬ le φ' ψ := by
|
||||
intro hφ'ψ
|
||||
apply hφ'φ
|
||||
apply IsTrans.trans φ' ψ φ hφ'ψ h2
|
||||
simp_all [hφ'φ, hφ'ψ]
|
||||
|
||||
end Wick
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue