feat: Static wick's theorem

This commit is contained in:
jstoobysmith 2024-12-17 07:15:47 +00:00
parent dd555b2037
commit dceaab7117
8 changed files with 670 additions and 246 deletions

View file

@ -137,21 +137,59 @@ lemma insertionSortEquiv_get {α : Type} {r : αα → Prop} [DecidableRel
rw [insertEquiv_get (r := r) a (List.insertionSort r l)]
rfl
lemma insertionSort_get_comp_insertionSortEquiv {α : Type} {r : αα → Prop} [DecidableRel r] (l : List α) :
(List.insertionSort r l).get ∘ (insertionSortEquiv r l) = l.get := by
rw [← insertionSortEquiv_get]
funext x
simp
lemma insertionSort_eq_ofFn {α : Type} {r : αα → Prop} [DecidableRel r] (l : List α) :
List.insertionSort r l = List.ofFn (l.get ∘ (insertionSortEquiv r l).symm) := by
rw [insertionSortEquiv_get (r := r)]
exact Eq.symm (List.ofFn_get (List.insertionSort r l))
lemma insertionSort_eraseIdx {α : Type} {r : αα → Prop} [DecidableRel r] :
(l : List α) →
(i : Fin (List.insertionSort r l).length) →
List.eraseIdx (List.insertionSort r l) i =
List.insertionSort r (List.eraseIdx l ((insertionSortEquiv r l).symm i))
| [], i => by
simp
| a :: l, i => by
rw [insertionSortEquiv]
simp
def optionErase {I : Type} (l : List I) (i : Option (Fin l.length)) : List I :=
match i with
| none => l
| some i => List.eraseIdx l i
def optionEraseZ {I : Type} (l : List I) (a : I) (i : Option (Fin l.length)) : List I :=
match i with
| none => a :: l
| some i => List.eraseIdx l i
lemma eraseIdx_length {I : Type} (l : List I) (i : Fin l.length) :
(List.eraseIdx l i).length + 1 = l.length := by
simp [List.length_eraseIdx]
have hi := i.prop
omega
lemma eraseIdx_cons_length {I : Type} (a : I) (l : List I) (i : Fin (a :: l).length) :
(List.eraseIdx (a :: l) i).length= l.length := by
simp [List.length_eraseIdx]
lemma eraseIdx_get {I : Type} (l : List I) (i : Fin l.length) :
(List.eraseIdx l i).get = l.get ∘ (Fin.cast (eraseIdx_length l i)) ∘ i.succAbove := by
ext x
simp only [Function.comp_apply, List.get_eq_getElem, List.eraseIdx, List.getElem_eraseIdx]
simp [Fin.succAbove]
by_cases hi: x.1 < i.val
· have h0 : x.castSucc < ↑↑i := by
simp [Fin.lt_def]
rw [Nat.mod_eq_of_lt]
exact hi
rw [eraseIdx_length]
exact i.prop
simp [h0, hi]
· have h0 : ¬ x.castSucc < ↑↑i := by
simp [Fin.lt_def]
rw [Nat.mod_eq_of_lt]
exact Nat.le_of_not_lt hi
rw [eraseIdx_length]
exact i.prop
simp [h0, hi]
end HepLean.List

View file

@ -23,16 +23,12 @@ namespace Wick
noncomputable section
def optionErase {I : Type} (l : List I) (i : Option (Fin l.length)) : List I :=
match i with
| none => l
| some i => List.eraseIdx l i
open HepLean.List
inductive ContractionsAux {I : Type} : (l : List I) → (aux : List I) → Type
| nil : ContractionsAux [] []
| single {a : I} : ContractionsAux [a] [a]
| cons {l : List I} {aux : List I} {a b: I} (i : Option (Fin (b :: aux).length)) :
ContractionsAux (b :: l) aux → ContractionsAux (a :: b :: l) (optionErase (b :: aux) i)
| cons {l : List I} {aux : List I} {a : I} (i : Option (Fin aux.length)) :
ContractionsAux l aux → ContractionsAux (a :: l) (optionEraseZ aux a i)
def Contractions {I : Type} (l : List I) : Type := Σ aux, ContractionsAux l aux
@ -42,25 +38,6 @@ variable {I : Type} {l : List I} (c : Contractions l)
def normalize : List I := c.1
lemma normalize_length_le : c.normalize.length ≤ l.length := by
cases c
rename_i aux c
induction c with
| nil =>
simp [normalize]
| single =>
simp [normalize]
| cons i c ih =>
simp [normalize, optionErase]
match i with
| none =>
simpa using ih
| some i =>
simp
rw [List.length_eraseIdx]
simp [normalize] at ih
simp
exact Nat.le_add_right_of_le ih
lemma contractions_nil (a : Contractions ([] : List I)) : a = ⟨[], ContractionsAux.nil⟩ := by
cases a
@ -68,21 +45,35 @@ lemma contractions_nil (a : Contractions ([] : List I)) : a = ⟨[], Contraction
cases c
rfl
lemma contractions_single {i : I} (a : Contractions [i]) : a = ⟨[i], ContractionsAux.single⟩ := by
lemma contractions_single {i : I} (a : Contractions [i]) : a =
⟨[i], ContractionsAux.cons none ContractionsAux.nil⟩ := by
cases a
rename_i aux c
cases c
rfl
rename_i aux' c'
cases c'
cases aux'
simp [optionEraseZ]
rename_i x
exact Fin.elim0 x
def consConsEquiv {a b : I} {l : List I} :
Contractions (a :: b :: l) ≃ (c : Contractions (b :: l)) × Option (Fin (b :: c.normalize).length) where
def nilEquiv : Contractions ([] : List I) ≃ Unit where
toFun _ := ()
invFun _ := ⟨[], ContractionsAux.nil⟩
left_inv a := by
exact Eq.symm (contractions_nil a)
right_inv _ := by
rfl
def consEquiv {a : I} {l : List I} :
Contractions (a :: l) ≃ (c : Contractions l) × Option (Fin (c.normalize).length) where
toFun c :=
match c with
| ⟨aux, c⟩ =>
match c with
| ContractionsAux.cons (aux := aux') i c => ⟨⟨aux', c⟩, i⟩
invFun ci :=
⟨(optionErase (b :: ci.fst.normalize) ci.2), ContractionsAux.cons (a := a) ci.2 ci.1.2⟩
⟨(optionEraseZ (ci.fst.normalize) a ci.2), ContractionsAux.cons (a := a) ci.2 ci.1.2⟩
left_inv c := by
match c with
| ⟨aux, c⟩ =>
@ -96,16 +87,11 @@ instance decidable : (l : List I) → DecidableEq (Contractions l)
| ⟨_, a⟩, ⟨_, b⟩ =>
match a, b with
| ContractionsAux.nil , ContractionsAux.nil => isTrue rfl
| _ :: [] => fun a b =>
match a, b with
| ⟨_, a⟩, ⟨_, b⟩ =>
match a, b with
| ContractionsAux.single , ContractionsAux.single => isTrue rfl
| _ :: b :: l =>
haveI : DecidableEq (Contractions (b :: l)) := decidable (b :: l)
haveI : DecidableEq ((c : Contractions (b :: l)) × Option (Fin (b :: c.normalize).length)) :=
| _ :: l =>
haveI : DecidableEq (Contractions l) := decidable l
haveI : DecidableEq ((c : Contractions l) × Option (Fin (c.normalize).length)) :=
Sigma.instDecidableEqSigma
Equiv.decidableEq consConsEquiv
Equiv.decidableEq consEquiv
instance fintype : (l : List I) → Fintype (Contractions l)
| [] => {
@ -114,102 +100,143 @@ instance fintype : (l : List I) → Fintype (Contractions l)
intro a
rw [Finset.mem_singleton]
exact contractions_nil a}
| a :: [] => {
elems := {⟨[a], ContractionsAux.single⟩}
complete := by
intro a
rw [Finset.mem_singleton]
exact contractions_single a}
| a :: b :: l =>
haveI : Fintype (Contractions (b :: l)) := fintype (b :: l)
haveI : Fintype ((c : Contractions (b :: l)) × Option (Fin (b :: c.normalize).length)) :=
| a :: l =>
haveI : Fintype (Contractions l) := fintype l
haveI : Fintype ((c : Contractions l) × Option (Fin (c.normalize).length)) :=
Sigma.instFintype
Fintype.ofEquiv _ consConsEquiv.symm
Fintype.ofEquiv _ consEquiv.symm
-- This definition is not correct.
def superCommuteTermAux {l : List I} {aux : List I} : (c : ContractionsAux l aux) → FreeAlgebra I
| ContractionsAux.nil => 1
| ContractionsAux.single => 1
| ContractionsAux.cons i c => superCommuteTermAux c
def superCommuteTerm {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
{r : List I} (c : Contractions r) : FreeAlgebra (Σ i, f i) :=
freeAlgebraMap f (superCommuteTermAux c.2)
structure Splitting {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1] where
𝓑n : I → (Σ i, f i)
𝓑p : I → (Σ i, f i)
𝓧n : I →
𝓧p : I →
h𝓑 : ∀ i, ofListM f [i] 1 = ofList [𝓑n i] (𝓧n i) + ofList [𝓑p i] (𝓧p i)
h𝓑n : ∀ i j, le1 (𝓑n i) j
h𝓑p : ∀ i j, le1 j (𝓑p i)
lemma superCommuteTerm_center {f : I → Type} [∀ i, Fintype (f i)]
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F] :
F (c.superCommuteTerm) ∈ Subalgebra.center A := by
sorry
def toTerm {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
def toCenterTerm {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
(q : I → Fin 2) {r : List I}
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
(c : Contractions r) : FreeAlgebra (Σ i, f i) :=
c.superCommuteTerm * koszulOrder le1 (fun i => q i.fst) (ofListM f c.normalize 1)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F]
(c : Contractions r) (S : Splitting f le1) : A :=
match c with
| ⟨aux, c⟩ =>
match c with
| .nil => 1
| .cons (a := a) (l := l) (aux := aux') none c => toCenterTerm f q le1 F ⟨aux', c⟩ S
| .cons (a := a) (l := l) (aux := aux') (some n) c =>
toCenterTerm f q le1 F ⟨aux', c⟩ S *
superCommuteCoef q [aux'.get n] (List.take (↑n) aux') •
F (((superCommute fun i => q i.fst) (ofList [S.𝓑p a] (S.𝓧p a))) (ofListM f [aux'.get n] 1))
lemma toTerm_nil {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
: toTerm q le1 (⟨[], ContractionsAux.nil⟩ : Contractions []) = 1 := by
simp [toTerm, normalize]
rw [ofListM_empty]
simp
lemma toCenterTerm_none {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
(q : I → Fin 2) {r : List I}
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F]
(S : Splitting f le1) (a : I) (c : Contractions r) :
toCenterTerm (r := a :: r) f q le1 F (Contractions.consEquiv.symm ⟨c, none⟩) S = toCenterTerm f q le1 F c S := by
rw [consEquiv]
simp [optionErase]
dsimp [toCenterTerm]
rfl
lemma toCenterTerm_center {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
(q : I → Fin 2) {r : List I}
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F]
(c : Contractions r) (S : Splitting f le1) :
(c.toCenterTerm f q le1 F S) ∈ Subalgebra.center A := by
sorry
end Contractions
lemma wick_cons_cons {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (r : List I)
lemma static_wick_nil {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2)
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
(tle : I → I → Prop) [DecidableRel tle]
(i : (Σ i, f i)) (hi : ∀ j, le1 j i)
{A : Type} [Semiring A] [Algebra A] (r : List I) (b a : I)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F]
(bn bp : (Σ i, f i))
(hb : ofListM f [b] 1 = ofList [bn] 1 + ofList [bp] 1)
(ih : F (ofListM f (a :: r) 1) = ∑ c : Contractions (a :: r), F (c.toTerm q le1)) :
F (ofListM f (b :: a :: r) 1) = ∑ c : Contractions (b :: a :: r), F (c.toTerm q le1) := by
rw [ofListM_cons_eq_ofListM, map_mul]
rw [ih]
rw [Finset.mul_sum]
rw [← Contractions.consConsEquiv.symm.sum_comp]
simp
(S : Contractions.Splitting f le1) :
F (ofListM f [] 1) = ∑ c : Contractions [],
c.toCenterTerm f q le1 F S * F (koszulOrder le1 (fun i => q i.fst) (ofListM f c.normalize 1)) := by
rw [← Contractions.nilEquiv.symm.sum_comp]
simp [Contractions.nilEquiv]
dsimp [Contractions.normalize, Contractions.toCenterTerm]
simp [ofListM_empty]
lemma static_wick_cons {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2)
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
{A : Type} [Semiring A] [Algebra A] (r : List I) (a : I)
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F]
(S : Contractions.Splitting f le1)
(ih : F (ofListM f r 1) =
∑ c : Contractions r, c.toCenterTerm f q le1 F S * F (koszulOrder le1 (fun i => q i.fst) (ofListM f c.normalize 1))) :
F (ofListM f (a :: r) 1) = ∑ c : Contractions ( a :: r),
c.toCenterTerm f q le1 F S * F (koszulOrder le1 (fun i => q i.fst) (ofListM f c.normalize 1)) := by
rw [ofListM_cons_eq_ofListM, map_mul, ih, Finset.mul_sum,
← Contractions.consEquiv.symm.sum_comp]
erw [Finset.sum_sigma]
congr
funext c
rw [Contractions.toTerm]
rw [map_mul, ← mul_assoc]
have hi := c.superCommuteTerm_center F
have hb := S.h𝓑 a
rw [← mul_assoc]
have hi := c.toCenterTerm_center f q le1 F S
rw [Subalgebra.mem_center_iff] at hi
rw [hi]
rw [mul_assoc]
rw [← map_mul]
rw [hb]
rw [add_mul]
rw [ofList_singleton, mul_koszulOrder_le, ← ofList_singleton]
rw [map_add]
rw [hi, mul_assoc, ← map_mul, hb, add_mul, map_add]
conv_lhs =>
rhs
rhs
lhs
rw [ofList_eq_smul_one]
rw [Algebra.smul_mul_assoc]
rw [ofList_singleton]
rw [le_all_mul_koszulOrder]
rw [← add_assoc]
rw [← map_add, ← map_add]
rw [mul_koszulOrder_le]
conv_lhs =>
rhs
rw [← map_add]
rw [← add_mul]
rw [← ofList_singleton]
rw [← hb]
rw [map_add]
rw [mul_add]
lhs
rw [← map_smul, ← Algebra.smul_mul_assoc]
rw [← ofList_singleton, ← ofList_eq_smul_one]
conv_lhs =>
rhs
rw [superCommute_ofList_ofListM_sum]
sorry
rhs
rw [ofList_eq_smul_one, Algebra.smul_mul_assoc, map_smul]
rw [le_all_mul_koszulOrder_ofListM_expand]
conv_lhs =>
rhs
rhs
rw [smul_add, Finset.smul_sum]
rw [← map_smul, ← map_smul, ← Algebra.smul_mul_assoc, ← ofList_eq_smul_one]
rhs
rhs
intro n
rw [← Algebra.smul_mul_assoc, smul_comm, ← map_smul, ← LinearMap.map_smul₂, ← ofList_eq_smul_one]
rw [← add_assoc, ← map_add, ← map_add, ← add_mul, ← hb, ← ofListM_cons_eq_ofListM, mul_add]
rw [Fintype.sum_option]
congr 1
rw [Finset.mul_sum]
congr
funext n
rw [← mul_assoc]
rfl
exact S.h𝓑p a
exact S.h𝓑n a
theorem static_wick_theorem {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2)
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
{A : Type} [Semiring A] [Algebra A] (r : List I)
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F]
(S : Contractions.Splitting f le1) :
F (ofListM f r 1) = ∑ c : Contractions r, c.toCenterTerm f q le1 F S *
F (koszulOrder le1 (fun i => q i.fst) (ofListM f c.normalize 1)) := by
induction r with
| nil => exact static_wick_nil q le1 F S
| cons a r ih => exact static_wick_cons q le1 r a F S ih
end
end Wick

View file

@ -21,8 +21,6 @@ import HepLean.PerturbationTheory.Wick.Koszul.Order
namespace Wick
noncomputable section
def grade {I : Type} (q : I → Fin 2) : (l : List I) → Fin 2
| [] => 0
| a :: l => if q a = grade q l then 0 else 1
@ -82,5 +80,75 @@ lemma grade_insertionSort {I : Type} (q : I → Fin 2) (le1 : I → I → Prop)
simp [grade]
rw [ih]
end
def superCommuteCoef {I : Type} (q : I → Fin 2) (la lb : List I) : :=
if grade q la = 1 ∧ grade q lb = 1 then - 1 else 1
lemma superCommuteCoef_empty {I : Type} (q : I → Fin 2) (la : List I) :
superCommuteCoef q la [] = 1 := by
simp only [superCommuteCoef, Fin.isValue, grade_empty, zero_ne_one, and_false, ↓reduceIte]
lemma superCommuteCoef_append {I : Type} (q : I → Fin 2) (la lb lc : List I) :
superCommuteCoef q la (lb ++ lc) = superCommuteCoef q la lb * superCommuteCoef q la lc := by
simp only [superCommuteCoef, Fin.isValue, grade_append, ite_eq_right_iff, zero_ne_one, imp_false,
mul_ite, mul_neg, mul_one]
by_cases hla : grade q la = 1
· by_cases hlb : grade q lb = 1
· by_cases hlc : grade q lc = 1
· simp [hlc, hlb, hla]
· have hc : grade q lc = 0 := by
omega
simp [hc, hlb, hla]
· have hb : grade q lb = 0 := by
omega
by_cases hlc : grade q lc = 1
· simp [hlc, hb]
· have hc : grade q lc = 0 := by
omega
simp [hc, hb]
· have ha : grade q la = 0 := by
omega
simp [ha]
def superCommuteCoefM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (l : List (Σ i, f i)) (r : List I) : :=
(if grade (fun i => q i.fst) l = 1 ∧ grade q r = 1 then -1 else 1)
lemma superCommuteCoefM_empty {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (l : List (Σ i, f i)):
superCommuteCoefM q l [] = 1 := by
simp [superCommuteCoefM]
def test {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
[DecidableRel le1] (n : Fin r.length) : :=
if grade q (List.take n r) = grade q ((List.take (↑((HepLean.List.insertionSortEquiv le1 r) n))
(List.insertionSort le1 r))) then 1 else -1
def superCommuteCoefLE {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
[DecidableRel le1] (i : I) (n : Fin r.length) : :=
koszulSign le1 q r *
superCommuteCoef q [i] (List.take (↑((HepLean.List.insertionSortEquiv le1 r) n))
(List.insertionSort le1 r)) *
koszulSign le1 q (r.eraseIdx ↑n)
lemma superCommuteCoefLE_zero {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) (r : List I)
(a : I)
[DecidableRel le1] (i : I) :
superCommuteCoefLE q le1 (a :: r) i ⟨0, Nat.zero_lt_succ r.length⟩ = 1 := by
simp [superCommuteCoefLE]
simp [koszulSign]
trans koszulSignInsert le1 q a r * (koszulSign le1 q r * koszulSign le1 q r) *
superCommuteCoef q [i]
(List.take (↑((HepLean.List.insertionSortEquiv le1 (a :: r)) ⟨0, Nat.zero_lt_succ r.length⟩))
(List.orderedInsert le1 a (List.insertionSort le1 r)))
· ring_nf
rfl
rw [koszulSign_mul_self]
simp
sorry
lemma superCommuteCoefLE_eq_get {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
[DecidableRel le1] (i : I) (n : Fin r.length) :
superCommuteCoefLE q le1 r i n = superCommuteCoef q [r.get n] (r.take n) := by
sorry
end Wick

View file

@ -76,6 +76,14 @@ lemma koszulOrder_ofList {I : Type} (r : I → I → Prop) [DecidableRel r] (q :
conv_rhs => rw [ofList_eq_smul_one]
rw [smul_smul]
lemma ofList_insertionSort_eq_koszulOrder {I : Type} (r : I → I → Prop) [DecidableRel r] (q : I → Fin 2)
(l : List I) (x : ) :
ofList (List.insertionSort r l) x = (koszulSign r q l) • koszulOrder r q (ofList l x) := by
rw [koszulOrder_ofList]
rw [smul_smul]
rw [koszulSign_mul_self]
simp
def freeAlgebraMap {I : Type} (f : I → Type) [∀ i, Fintype (f i)] :
FreeAlgebra I →ₐ[] FreeAlgebra (Σ i, f i) :=
FreeAlgebra.lift fun i => ∑ (j : f i), FreeAlgebra.ι ⟨i, j⟩
@ -94,6 +102,13 @@ lemma ofListM_empty {I : Type} (f : I → Type) [∀ i, Fintype (f i)] :
rw [ofList_empty]
exact map_one (freeAlgebraMap f)
lemma ofListM_empty_smul {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (x : ) :
ofListM f [] x = x • 1 := by
simp only [ofListM, EmbeddingLike.map_eq_one_iff]
rw [ofList_eq_smul_one]
rw [ofList_empty]
simp
lemma ofListM_cons {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (i : I) (r : List I) (x : ) :
ofListM f (i :: r) x = (∑ j : f i, FreeAlgebra.ι ⟨i, j⟩) * (ofListM f r x) := by
rw [ofListM, ofList_cons_eq_ofList, ofList_singleton, map_mul]
@ -108,6 +123,14 @@ lemma ofListM_singleton {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (i :
rw [freeAlgebraMap_ι]
rw [Finset.smul_sum]
lemma ofListM_singleton_one {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (i : I) :
ofListM f [i] 1 = ∑ j : f i, FreeAlgebra.ι ⟨i, j⟩ := by
simp only [ofListM]
rw [ofList_eq_smul_one, ofList_singleton, map_smul]
rw [freeAlgebraMap_ι]
rw [Finset.smul_sum]
simp
lemma ofListM_cons_eq_ofListM {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (i : I) (r : List I) (x : ) :
ofListM f (i :: r) x = ofListM f [i] 1 * ofListM f r x := by
rw [ofListM_cons, ofListM_singleton]
@ -145,6 +168,10 @@ lemma liftM_get {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (r : List I)
rw [ih]
simp
def liftMCongrEquiv {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (r0 : I) (r : List I) (n : Fin (r0 :: r).length) :
(Π i, f ((r0 :: r).get i)) ≃ f ((r0 :: r).get n) × Π i, f ((r0 :: r).get (n.succAbove i)) :=
(Fin.insertNthEquiv _ _).symm
lemma ofListM_expand {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (x : ) :
(l : List I) → ofListM f l x = ∑ (a : Π i, f (l.get i)), ofList (liftM f l a) x
| [] => by
@ -192,6 +219,7 @@ lemma liftM_grade {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
rw [h0, h0']
rw [h1]
@[simp]
lemma liftM_grade_take {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) : (r : List I) → (a : Π i, f (r.get i)) → (n : ) →
grade (fun i => q i.fst) (List.take n (liftM f r a)) = grade q (List.take n r)
@ -205,6 +233,85 @@ lemma liftM_grade_take {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
refine (liftM_grade_take q r (fun i => a i.succ) n)
rw [ih]
open HepLean.List
def listMEraseEquiv {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) {r0 : I} {r : List I} (n : Fin (r0 :: r).length) :
(Π (i :Fin ((r0 :: r).eraseIdx ↑n).length) , f (((r0 :: r).eraseIdx ↑n).get i))
≃ Π (i : Fin r.length), f ((r0 :: r).get (n.succAbove i)) :=
Equiv.piCongr (Fin.castOrderIso (by rw [eraseIdx_cons_length])).toEquiv
fun x => Equiv.cast (congrArg f (by
rw [HepLean.List.eraseIdx_get]
simp
congr 1
simp [Fin.succAbove]
sorry
))
/-
lemma liftM_eraseIdx {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (r0 : I): (r : List I) → (n : Fin (r0 :: r).length) → (a : Π i, f ((r0 :: r).get i)) →
(liftM f (r0 :: r) a).eraseIdx ↑n = liftM f (List.eraseIdx (r0 :: r) n) ((listMEraseEquiv q n).symm a)
| r, ⟨0, h⟩, a => by
simp [List.eraseIdx]
rfl
| r, ⟨n + 1, h⟩, a => by
have hf : (r.eraseIdx n).length + 1 = r.length := by
rw [List.length_eraseIdx]
simp at h
simp [h]
omega
have hn : n < (r.eraseIdx n).length + 1 := by
simp at h
rw [hf]
exact h
simp [liftM]
apply And.intro
· refine eq_cast_iff_heq.mpr ?left.a
simp [Fin.cast]
rw [Fin.succAbove]
simp
rw [if_pos]
simp
simp
refine Fin.add_one_pos ↑n ?left.a.hc.h
simp at h
rw [Fin.lt_def]
conv_rhs => simp
rw [hf]
simp
rw [Nat.mod_eq_of_modEq rfl (Nat.le.step h)]
exact h
· have hl := liftM_eraseIdx q r ⟨n, Nat.succ_lt_succ_iff.mp h⟩ (fun i => a i.succ)
rw [hl]
congr
funext i
rw [Equiv.apply_eq_iff_eq_symm_apply]
simp
refine eq_cast_iff_heq.mpr ?right.e_a.h.a
congr
rw [Fin.ext_iff]
simp [Fin.succAbove]
simp [Fin.lt_def]
rw [@Fin.val_add_one]
simp [hn]
rw [Nat.mod_eq_of_lt hn]
rw [Nat.mod_eq_of_lt]
have hnot : ¬ ↑n = Fin.last ((r.eraseIdx n).length + 1) := by
rw [Fin.ext_iff]
simp
rw [Nat.mod_eq_of_lt]
omega
exact Nat.lt_add_right 1 hn
simp [hnot]
by_cases hi : i.val < n
· simp [hi]
· simp [hi]
· exact Nat.lt_add_right 1 hn
-/
lemma koszulSignInsert_liftM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]

View file

@ -13,10 +13,13 @@ import HepLean.Meta.Notes.Basic
import Init.Data.List.Sort.Basic
import Mathlib.Data.Fin.Tuple.Take
import HepLean.PerturbationTheory.Wick.Koszul.SuperCommuteM
import Mathlib.Logic.Equiv.Option
/-!
# Koszul signs and ordering for lists and algebras
See e.g.
- https://pcteserver.mi.infn.it/~molinari/NOTES/WICK23.pdf
-/
namespace Wick
@ -101,30 +104,68 @@ lemma superCommuteTakeM_F {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
exact Eq.symm (List.eraseIdx_eq_take_drop_succ r n)
lemma koszulOrder_superCommuteM_le {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
lemma superCommute_koszulOrder_le_ofList {I : Type}
(q : I → Fin 2) (r : List I) (x : )
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
(i : (Σ i, f i)) (hi : ∀ j, le1 j i)
(le1 :I → I → Prop) [DecidableRel le1]
(i : I)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F] :
F (koszulOrder le1 (fun i => q i.fst)
(superCommute (fun i => q i.1) (FreeAlgebra.ι i) (ofListM f r x))) = 0 := by
sorry
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap F] :
F ((superCommute q (FreeAlgebra.ι i) (koszulOrder le1 q (ofList r x)))) =
∑ n : Fin r.length, (superCommuteCoef q [r.get n] (r.take n)) •
(F (((superCommute q) (ofList [i] 1)) (FreeAlgebra.ι (r.get n))) *
F ((koszulOrder le1 q) (ofList (r.eraseIdx ↑n) x))) := by
rw [koszulOrder_ofList]
rw [map_smul, map_smul, ← ofList_singleton]
rw [superCommute_ofList_sum]
rw [map_sum]
rw [← (HepLean.List.insertionSortEquiv le1 r).sum_comp]
conv_lhs =>
rhs
rhs
intro n
rw [superCommuteTake_superCommuteCenterMap]
lhs
rhs
rhs
rhs
change ((List.insertionSort le1 r).get ∘ (HepLean.List.insertionSortEquiv le1 r)) n
rw [HepLean.List.insertionSort_get_comp_insertionSortEquiv]
have hListErase (n : Fin r.length ) : (List.insertionSort le1 r).eraseIdx ↑((HepLean.List.insertionSortEquiv le1 r) n)
= List.insertionSort le1 (r.eraseIdx n) := by sorry
conv_lhs =>
rhs
rhs
intro n
rw [hListErase]
rw [ofList_insertionSort_eq_koszulOrder le1 q]
rw [Finset.smul_sum]
conv_lhs =>
rhs
intro n
rw [map_smul]
rw [smul_smul]
rw [Algebra.mul_smul_comm]
rw [smul_smul]
congr
funext n
congr 1
trans superCommuteCoefLE q le1 r i n
· rw [superCommuteCoefLE]
rw [mul_assoc]
exact superCommuteCoefLE_eq_get q le1 r i n
lemma koszulOrder_of_le_all {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (r : List I) (x : ) (le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
(i : (Σ i, f i)) (hi : ∀ j, le1 j i)
lemma koszulOrder_of_le_all_ofList {I : Type}
(q : I → Fin 2) (r : List I) (x : ) (le1 : I → I → Prop) [DecidableRel le1]
(i : I) (hi : ∀ j, le1 j i)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F] :
F (koszulOrder le1 (fun i => q i.fst)
(ofListM f r x * FreeAlgebra.ι i))
= superCommuteCoefM q [i] r • F (koszulOrder le1 (fun i => q i.fst)
(FreeAlgebra.ι i * ofListM f r x)) := by
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap F] :
F (koszulOrder le1 q (ofList r x * FreeAlgebra.ι i))
= superCommuteCoef q [i] r • F (koszulOrder le1 q (FreeAlgebra.ι i * ofList r x)) := by
conv_lhs =>
rhs
rhs
rw [← ofList_singleton]
rw [ofListM_ofList_superCommute q]
rw [ofListM_ofList_superCommute' q]
rw [map_sub]
rw [sub_eq_add_neg]
rw [map_add]
@ -134,61 +175,176 @@ lemma koszulOrder_of_le_all {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
rw [map_smul]
rw [← neg_smul]
rw [map_smul, map_smul, map_smul]
rw [ofList_singleton, koszulOrder_superCommuteM_le]
· simp
· exact fun j => hi j
sorry
lemma le_all_mul_koszulOrder {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (r : List I) (x : ) (le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
(i : (Σ i, f i)) (hi : ∀ j, le1 j i)
lemma le_all_mul_koszulOrder_ofList {I : Type}
(q : I → Fin 2) (r : List I) (x : ) (le1 : I → I→ Prop) [DecidableRel le1]
(i : I) (hi : ∀ (j : I), le1 j i)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F] :
F (FreeAlgebra.ι i * koszulOrder le1 (fun i => q i.fst)
(ofListM f r x)) = F ((koszulOrder le1 fun i => q i.fst) (FreeAlgebra.ι i * ofListM f r x)) +
F (((superCommute fun i => q i.fst) (ofList [i] 1))
((koszulOrder le1 fun i => q i.fst) (ofListM f r x))) := by
sorry
/-
rw [map_smul]
rw [Algebra.mul_smul_comm, map_smul]
change koszulSign le1 q r • F (FreeAlgebra.ι i * (ofListM f (List.insertionSort le1 r) x)) = _
rw [← ofList_singleton]
rw [ofList_ofListM_superCommute q]
rw [map_add]
rw [smul_add]
rw [← map_smul]
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap F] :
F (FreeAlgebra.ι i * koszulOrder le1 q (ofList r x)) =
F ((koszulOrder le1 q) (FreeAlgebra.ι i * ofList r x)) +
F (((superCommute q) (ofList [i] 1)) ((koszulOrder le1 q) (ofList r x))) := by
rw [koszulOrder_ofList, Algebra.mul_smul_comm, map_smul, ← ofList_singleton,
ofList_ofList_superCommute q, map_add, smul_add, ← map_smul]
conv_lhs =>
lhs
rhs
rw [← Algebra.smul_mul_assoc]
rw [smul_smul, mul_comm, ← smul_smul]
rw [ ofListM, ← map_smul, ← koszulOrder_ofList, ← koszulOrder_ofListM, ofList_singleton]
rw [Algebra.smul_mul_assoc]
rw [koszulOrder_mul_ge]
rw [map_smul]
rw [koszulOrder_of_le_all]
rw [smul_smul]
have h1 : (superCommuteCoefM q [i] (List.insertionSort le1 r) * superCommuteCoefM q [i] r) = 1 := by
simp [superCommuteCoefM]
have ha (a b : Fin 2): (if a = 1 ∧ b = 1 then
-if a = 1 ∧ b = 1 then -1 else 1
else if a = 1 ∧ b = 1 then -1 else (1 : )) = 1 := by
fin_cases a <;> fin_cases b
· rfl
· rfl
· rfl
· simp only [Fin.mk_one, Fin.isValue, and_self, ↓reduceIte, neg_neg]
exact ha _ _
rw [h1]
simp only [one_smul]
rw [← Algebra.smul_mul_assoc, smul_smul, mul_comm, ← smul_smul, ← koszulOrder_ofList,
Algebra.smul_mul_assoc, ofList_singleton]
rw [koszulOrder_mul_ge, map_smul]
congr
· rw [koszulOrder_of_le_all_ofList]
sorry
sorry
· rw [map_smul, map_smul]
· exact fun j => hi j
def superCommuteCenterOrder {I : Type}
(q : I → Fin 2) (r : List I) (i : I)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap F]
(n : Option (Fin r.length)) : A :=
match n with
| none => 1
| some n => superCommuteCoef q [r.get n] (r.take n) • F (((superCommute q) (ofList [i] 1)) (FreeAlgebra.ι (r.get n)))
@[simp]
lemma superCommuteCenterOrder_none {I : Type}
(q : I → Fin 2) (r : List I) (i : I)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap F] :
superCommuteCenterOrder q r i F none = 1 := by
simp [superCommuteCenterOrder]
open HepLean.List
lemma le_all_mul_koszulOrder_ofList_expand {I : Type}
(q : I → Fin 2) (r : List I) (x : ) (le1 : I → I→ Prop) [DecidableRel le1]
(i : I) (hi : ∀ (j : I), le1 j i)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap F] :
F (FreeAlgebra.ι i * koszulOrder le1 q (ofList r x)) =
∑ n, superCommuteCenterOrder q r i F n * F ((koszulOrder le1 q) (ofList (optionEraseZ r i n) x)) := by
rw [le_all_mul_koszulOrder_ofList]
conv_lhs =>
rhs
rw [← map_smul, ← map_smul]
rw [ ofListM, ← map_smul, ← koszulOrder_ofList, ← koszulOrder_ofListM]
rw [ofList_singleton]
rw [superCommute_koszulOrder_le_ofList]
simp only [List.get_eq_getElem, Fintype.sum_option, superCommuteCenterOrder_none, one_mul]
congr
rw [ofList_singleton]
· exact fun j => hi j
· exact fun j => hi j.fst
-/
· rw [← ofList_singleton, ← ofList_pair]
simp only [List.singleton_append, one_mul]
rfl
· funext n
simp only [superCommuteCenterOrder, List.get_eq_getElem, Algebra.smul_mul_assoc]
rfl
exact fun j => hi j
lemma le_all_mul_koszulOrder_ofListM_expand {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (r : List I) (x : ) (le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
(i : (Σ i, f i)) (hi : ∀ (j : (Σ i, f i)), le1 j i)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F] :
F (ofList [i] 1 * koszulOrder le1 (fun i => q i.1) (ofListM f r x)) =
F ((koszulOrder le1 fun i => q i.fst) (ofList [i] 1 * ofListM f r x)) +
∑ n : (Fin r.length), superCommuteCoef q [r.get n] (List.take (↑n) r) •
F (((superCommute fun i => q i.fst) (ofList [i] 1)) (ofListM f [r.get n] 1)) *
F ((koszulOrder le1 fun i => q i.fst) (ofListM f (r.eraseIdx ↑n) x)) := by
match r with
| [] =>
simp only [map_mul, List.length_nil, Finset.univ_eq_empty, List.get_eq_getElem, List.take_nil,
List.eraseIdx_nil, Algebra.smul_mul_assoc, Finset.sum_empty, add_zero]
rw [ofListM_empty_smul]
simp only [map_smul, koszulOrder_one, map_one, Algebra.mul_smul_comm, mul_one]
rw [ofList_singleton, koszulOrder_ι]
| r0 :: r =>
rw [ofListM_expand, map_sum, Finset.mul_sum, map_sum]
let e1 (a : (i : Fin (r0 :: r).length) → f ((r0 :: r).get i)) :
Option (Fin (liftM f (r0 :: r) a).length) ≃ Option (Fin (r0 :: r).length) :=
Equiv.optionCongr (Fin.castOrderIso (liftM_length f (r0 :: r) a)).toEquiv
conv_lhs =>
rhs
intro a
rw [ofList_singleton, le_all_mul_koszulOrder_ofList_expand _ _ _ _ _ hi]
rw [← (e1 a).symm.sum_comp]
rhs
intro n
rw [Finset.sum_comm]
simp only [ Fintype.sum_option]
congr 1
· simp only [List.length_cons, List.get_eq_getElem, superCommuteCenterOrder,
Equiv.optionCongr_symm, OrderIso.toEquiv_symm, Fin.symm_castOrderIso, Equiv.optionCongr_apply,
RelIso.coe_fn_toEquiv, Option.map_none', optionEraseZ, one_mul, e1]
rw [← map_sum, Finset.mul_sum, ← map_sum]
apply congrArg
apply congrArg
congr
funext x
rw [ofList_cons_eq_ofList]
· congr
funext n
rw [← (liftMCongrEquiv _ _ _ n).symm.sum_comp]
simp only [List.get_eq_getElem, List.length_cons, Equiv.optionCongr_symm, OrderIso.toEquiv_symm,
Fin.symm_castOrderIso, Equiv.optionCongr_apply, RelIso.coe_fn_toEquiv, Option.map_some',
Fin.castOrderIso_apply, Algebra.smul_mul_assoc, e1]
erw [Finset.sum_product]
have h1 (a0 : f (r0 :: r)[↑n]) (a : (i : Fin r.length) → f (r0 :: r)[↑(n.succAbove i)]):
superCommuteCenterOrder (fun i => q i.fst) (liftM f (r0 :: r) ((liftMCongrEquiv f r0 r n).symm (a0, a))) i F
(some (Fin.cast (by simp) n)) = superCommuteCoef q [(r0 :: r).get n] (List.take (↑n) (r0 :: r)) •
F (((superCommute fun i => q i.fst) (ofList [i] 1)) (FreeAlgebra.ι ⟨(r0 :: r).get n, a0⟩)) := by
simp only [superCommuteCenterOrder, List.get_eq_getElem, List.length_cons, Fin.coe_cast]
have hx : (liftM f (r0 :: r) ((liftMCongrEquiv f r0 r n).symm (a0, a)))[n.1] =
⟨(r0 :: r).get n, a0⟩ := by
trans (liftM f (r0 :: r) ((liftMCongrEquiv f r0 r n).symm (a0, a))).get (Fin.cast (by simp) n)
· simp only [List.get_eq_getElem, List.length_cons, Fin.coe_cast]
rw [liftM_get]
simp [liftMCongrEquiv]
erw [hx]
have hsc : superCommuteCoef (fun i => q i.fst) [⟨(r0 :: r).get n, a0⟩]
(List.take (↑n) (liftM f (r0 :: r) ((liftMCongrEquiv f r0 r n).symm (a0, a)))) =
superCommuteCoef q [(r0 :: r).get n] (List.take (↑n) ((r0 :: r))) := by
simp only [superCommuteCoef, List.get_eq_getElem, List.length_cons, Fin.isValue,
liftM_grade_take]
rfl
erw [hsc]
rfl
conv_lhs =>
rhs
intro a0
rhs
intro a
erw [h1]
conv_lhs =>
rhs
intro a0
rw [← Finset.mul_sum]
have hl (a0 : f (r0 :: r)[↑n]) (a : (i : Fin r.length) → f (r0 :: r)[↑(n.succAbove i)]):
(ofList (optionEraseZ (liftM f (r0 :: r) ((liftMCongrEquiv f r0 r n).symm (a0, a))) i (some (Fin.cast (by simp ) n))) x)
= ofList ((liftM f ((r0 :: r).eraseIdx ↑n) ((listMEraseEquiv q n).symm a))) x := by
simp only [optionEraseZ, List.get_eq_getElem, List.length_cons, Fin.coe_cast]
simp [liftMCongrEquiv]
sorry
conv_lhs =>
rhs
intro a0
rhs
rhs
intro a
erw [hl a0 a]
rw [← Finset.sum_mul]
conv_lhs =>
lhs
rw [← Finset.smul_sum]
erw [← map_sum, ← map_sum, ← ofListM_singleton_one]
conv_lhs =>
rhs
erw [← (listMEraseEquiv q n).sum_comp]
rw [← map_sum, ← map_sum]
simp only [List.get_eq_getElem, List.length_cons, Equiv.symm_apply_apply,
Algebra.smul_mul_assoc]
erw [← ofListM_expand]
simp only [List.get_eq_getElem, List.length_cons, Algebra.smul_mul_assoc]
end
end Wick

View file

@ -25,7 +25,7 @@ namespace Wick
def koszulSignInsert {I : Type} (r : I → I → Prop) [DecidableRel r] (q : I → Fin 2) (a : I) :
List I →
| [] => 1
| b :: l => if r a b then 1 else
| b :: l => if r a b then koszulSignInsert r q a l else
if q a = 1 ∧ q b = 1 then - koszulSignInsert r q a l else koszulSignInsert r q a l
/-- When inserting a boson the `koszulSignInsert` is always `1`. -/
@ -35,9 +35,49 @@ lemma koszulSignInsert_boson {I : Type} (r : I → I → Prop) [DecidableRel r]
simp [koszulSignInsert]
| b :: l => by
simp only [koszulSignInsert, Fin.isValue, ite_eq_left_iff]
intro _
simp only [ha, Fin.isValue, zero_ne_one, false_and, ↓reduceIte]
exact koszulSignInsert_boson r q a ha l
rw [koszulSignInsert_boson r q a ha l, ha]
simp only [Fin.isValue, zero_ne_one, false_and, ↓reduceIte, ite_self]
@[simp]
lemma koszulSignInsert_mul_self {I : Type} (r : I → I → Prop) [DecidableRel r] (q : I → Fin 2) (a : I) :
(l : List I) → koszulSignInsert r q a l * koszulSignInsert r q a l = 1
| [] => by
simp [koszulSignInsert]
| b :: l => by
simp [koszulSignInsert]
by_cases hr : r a b
· simp [hr]
rw [koszulSignInsert_mul_self]
· simp [hr]
by_cases hq : q a = 1 ∧ q b = 1
· simp [hq]
rw [koszulSignInsert_mul_self]
· simp [hq]
rw [koszulSignInsert_mul_self]
lemma koszulSignInsert_le_forall {I : Type} (r : I → I → Prop) [DecidableRel r] (q : I → Fin 2) (a : I)
(l : List I) (hi : ∀ b, r a b) : koszulSignInsert r q a l = 1 := by
induction l with
| nil => rfl
| cons j l ih =>
simp only [koszulSignInsert, Fin.isValue, ite_eq_left_iff]
rw [ih]
simp only [Fin.isValue, ite_eq_left_iff, ite_eq_right_iff, and_imp]
intro h
exact False.elim (h (hi j))
lemma koszulSignInsert_ge_forall_append {I : Type} (r : I → I → Prop) [DecidableRel r]
(q : I → Fin 2) (l : List I) (j i : I) (hi : ∀ j, r j i) :
koszulSignInsert r q j l = koszulSignInsert r q j (l ++ [i]) := by
induction l with
| nil => simp [koszulSignInsert, hi]
| cons b l ih =>
simp only [koszulSignInsert, Fin.isValue, List.append_eq]
by_cases hr : r j b
· rw [if_pos hr, if_pos hr]
rw [ih]
· rw [if_neg hr, if_neg hr]
rw [ih]
/-- Gives a factor of `- 1` for every fermion-fermion (`q` is `1`) crossing that occurs when sorting
a list of based on `r`. -/
@ -46,6 +86,20 @@ def koszulSign {I : Type} (r : I → I → Prop) [DecidableRel r] (q : I → Fin
| [] => 1
| a :: l => koszulSignInsert r q a l * koszulSign r q l
def natTestQ : → Fin 2 := fun n => if n % 2 = 0 then 0 else 1
lemma koszulSign_mul_self {I : Type} (r : I → I → Prop) [DecidableRel r] (q : I → Fin 2)
(l : List I) : koszulSign r q l * koszulSign r q l = 1 := by
induction l with
| nil => simp [koszulSign]
| cons a l ih =>
simp [koszulSign]
trans (koszulSignInsert r q a l * koszulSignInsert r q a l) * (koszulSign r q l * koszulSign r q l)
ring
rw [ih]
rw [koszulSignInsert_mul_self, mul_one]
@[simp]
lemma koszulSign_freeMonoid_of {I : Type} (r : I → I → Prop) [DecidableRel r] (q : I → Fin 2)
(i : I) : koszulSign r q (FreeMonoid.of i) = 1 := by
@ -209,12 +263,7 @@ lemma mul_koszulOrder_le {I : Type} (r : I → I → Prop) [DecidableRel r] (q :
· congr 1
rw [koszulSign]
have h1 (l : List I) : koszulSignInsert r q i l = 1 := by
induction l with
| nil => rfl
| cons j l ih =>
simp only [koszulSignInsert, Fin.isValue, ite_eq_left_iff]
intro h
exact False.elim (h (hi j))
exact koszulSignInsert_le_forall r q i l hi
rw [h1]
simp
rw [h1]
@ -285,16 +334,7 @@ lemma koszulOrder_mul_ge {I : Type} (r : I → I → Prop) [DecidableRel r] (q :
rw [ih]
simp only [mul_eq_mul_right_iff]
apply Or.inl
have hKI (l : List I) (j : I) : koszulSignInsert r q j l = koszulSignInsert r q j (l ++ [i]) := by
induction l with
| nil => simp [koszulSignInsert, hi]
| cons b l ih =>
simp only [koszulSignInsert, Fin.isValue, List.append_eq]
by_cases hr : r j b
· rw [if_pos hr, if_pos hr]
· rw [if_neg hr, if_neg hr]
rw [ih]
rw [hKI]
rw [koszulSignInsert_ge_forall_append r q l j i hi]
rw [hI]
rfl
rw [h1]

View file

@ -161,35 +161,6 @@ lemma superCommute_one {I : Type} (q : I → Fin 2) (a : FreeAlgebra I) :
rw [hf]
simp
def superCommuteCoef {I : Type} (q : I → Fin 2) (la lb : List I) : :=
if grade q la = 1 ∧ grade q lb = 1 then - 1 else 1
lemma superCommuteCoef_empty {I : Type} (q : I → Fin 2) (la : List I) :
superCommuteCoef q la [] = 1 := by
simp only [superCommuteCoef, Fin.isValue, grade_empty, zero_ne_one, and_false, ↓reduceIte]
lemma superCommuteCoef_append {I : Type} (q : I → Fin 2) (la lb lc : List I) :
superCommuteCoef q la (lb ++ lc) = superCommuteCoef q la lb * superCommuteCoef q la lc := by
simp only [superCommuteCoef, Fin.isValue, grade_append, ite_eq_right_iff, zero_ne_one, imp_false,
mul_ite, mul_neg, mul_one]
by_cases hla : grade q la = 1
· by_cases hlb : grade q lb = 1
· by_cases hlc : grade q lc = 1
· simp [hlc, hlb, hla]
· have hc : grade q lc = 0 := by
omega
simp [hc, hlb, hla]
· have hb : grade q lb = 0 := by
omega
by_cases hlc : grade q lc = 1
· simp [hlc, hb]
· have hc : grade q lc = 0 := by
omega
simp [hc, hb]
· have ha : grade q la = 0 := by
omega
simp [ha]
lemma superCommute_ofList_mul {I : Type} (q : I → Fin 2) (la lb lc : List I) (xa xb xc : ) :
superCommute q (ofList la xa) (ofList lb xb * ofList lc xc) =
(superCommute q (ofList la xa) (ofList lb xb) * ofList lc xc +
@ -279,6 +250,32 @@ lemma superCommute_ofList_sum {I : Type} (q : I → Fin 2) (la lb : List I) (xa
· simp only [← mul_assoc, mul_eq_mul_right_iff]
exact Or.inl (Or.inl (ofList_cons_eq_ofList (List.take (↑n) lb) b 1).symm)
lemma superCommute_ofList_ofList_superCommuteCoef {I : Type} (q : I → Fin 2) (la lb : List I) (xa xb : ) :
superCommute q (ofList la xa) (ofList lb xb) =
ofList la xa * ofList lb xb - superCommuteCoef q la lb • ofList lb xb * ofList la xa := by
rw [superCommute_ofList_ofList, superCommuteCoef]
by_cases hq : grade q la = 1 ∧ grade q lb = 1
· simp [hq, ofList_pair]
· simp [hq, ofList_pair]
abel
lemma ofList_ofList_superCommute {I : Type} (q : I → Fin 2) (la lb : List I) (xa xb : ) :
ofList la xa * ofList lb xb = superCommuteCoef q la lb • ofList lb xb * ofList la xa
+ superCommute q (ofList la xa) (ofList lb xb) := by
rw [superCommute_ofList_ofList_superCommuteCoef]
abel
lemma ofListM_ofList_superCommute' {I : Type}
(q : I → Fin 2) (l : List I) (r : List I) (x y : ) :
ofList r y * ofList l x = superCommuteCoef q l r • (ofList l x * ofList r y)
- superCommuteCoef q l r • superCommute q (ofList l x) (ofList r y) := by
nth_rewrite 2 [ofList_ofList_superCommute q]
rw [superCommuteCoef]
by_cases hq : grade q l = 1 ∧ grade q r = 1
· simp [hq, superCommuteCoef]
· simp [hq]
lemma koszulOrder_superCommute_le {I : Type} (r : I → I → Prop) [DecidableRel r] (q : I → Fin 2)
(i j : I) (hle : r i j) (a1 a2 : FreeAlgebra I) :
koszulOrder r q (a1 * superCommute q (FreeAlgebra.ι i) (FreeAlgebra.ι j) * a2) =

View file

@ -55,15 +55,6 @@ lemma superCommute_ofList_ofListM {I : Type} {f : I → Type} [∀ i, Fintype (
· rw [ofList_pair]
simp only [neg_mul]
def superCommuteCoefM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (l : List (Σ i, f i)) (r : List I) : :=
(if grade (fun i => q i.fst) l = 1 ∧ grade q r = 1 then -1 else 1)
lemma superCommuteCoefM_empty {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (l : List (Σ i, f i)):
superCommuteCoefM q l [] = 1 := by
simp [superCommuteCoefM]
lemma superCommute_ofList_ofListM_superCommuteCoefM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (l : List (Σ i, f i)) (r : List I) (x y : ) :
superCommute (fun i => q i.1) (ofList l x) (ofListM f r y) =