feat: Sorry free version

This commit is contained in:
jstoobysmith 2024-12-19 11:23:49 +00:00
parent ab7da149c6
commit 681ffbeafd
7 changed files with 208 additions and 474 deletions

View file

@ -484,6 +484,59 @@ lemma orderedInsertEquiv_sigma {I : Type} {f : I → Type} [∀ i, Fintype (f i)
next h => simp_all only [not_lt]
/-- This result is taken from:
https://github.com/leanprover/lean4/blob/master/src/Init/Data/List/Nat/InsertIdx.lean
with simple modification here to make it run.
The file it was taken from is licensed under the Apache License, Version 2.0.
and written by Parikshit Khanna, Jeremy Avigad, Leonardo de Moura,
Floris van Doorn, Mario Carneiro.
Once HepLean is updated to a more recent version of Lean this result will be removed.
-/
theorem length_insertIdx' : ∀ n as, (List.insertIdx n a as).length = if n ≤ as.length then as.length + 1 else as.length
| 0, _ => by simp
| n + 1, [] => by simp
| n + 1, a :: as => by
simp only [List.insertIdx_succ_cons, List.length_cons, length_insertIdx', Nat.add_le_add_iff_right]
split <;> rfl
/-- This result is taken from:
https://github.com/leanprover/lean4/blob/master/src/Init/Data/List/Nat/InsertIdx.lean
with simple modification here to make it run.
The file it was taken from is licensed under the Apache License, Version 2.0.
and written by Parikshit Khanna, Jeremy Avigad, Leonardo de Moura,
Floris van Doorn, Mario Carneiro.
Once HepLean is updated to that version of Lean this result will be removed.
-/
theorem _root_.List.getElem_insertIdx_of_ge {l : List α} {x : α} {n k : Nat} (hn : n + 1 ≤ k)
(hk : k < (List.insertIdx n x l).length) :
(List.insertIdx n x l)[k] = l[k - 1]'(by simp [length_insertIdx'] at hk; split at hk <;> omega) := by
induction l generalizing n k with
| nil =>
cases n with
| zero =>
simp [List.insertIdx_zero, List.length_singleton] at hk
omega
| succ n => simp at hk
| cons _ _ ih =>
cases n with
| zero =>
simp only [List.insertIdx_zero] at hk
cases k with
| zero => omega
| succ k => simp
| succ n =>
cases k with
| zero => simp
| succ k =>
simp only [List.insertIdx_succ_cons, List.getElem_cons_succ]
rw [ih (by omega)]
cases k with
| zero => omega
| succ k => simp
lemma orderedInsert_eq_insertIdx_orderedInsertPos {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) :
List.orderedInsert le1 r0 r = List.insertIdx (orderedInsertPos le1 r r0).1 r0 r := by
@ -529,7 +582,9 @@ lemma orderedInsert_eq_insertIdx_orderedInsertPos {I : Type} (le1 : I → I →
erw [List.getElem_insertIdx_of_lt]
exact hn'
· simp [hn']
sorry
rw [List.getElem_insertIdx_of_ge]
simp
omega
/-- The equivalence between `Fin l.length ≃ Fin (List.insertionSort r l).length` induced by the
sorting algorithm. -/

View file

@ -118,18 +118,14 @@ structure Splitting {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
h𝓑p : ∀ i j, le1 j (𝓑p i)
def toCenterTerm {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
(q : I → Fin 2) {r : List I}
(q : I → Fin 2)
(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) : 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 *
(F : FreeAlgebra (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le1 F]
: {r : List I} → (c : Contractions r) → (S : Splitting f le1) → A
| [], ⟨[], .nil⟩, _ => 1
| _ :: _, ⟨_, .cons (aux := aux') none c⟩, S => toCenterTerm f q le1 F ⟨aux', c⟩ S
| a :: _, ⟨_, .cons (aux := aux') (some n) c⟩, S => 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))
@ -137,7 +133,7 @@ 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]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le1 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]
@ -146,13 +142,30 @@ lemma toCenterTerm_none {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
rfl
lemma toCenterTerm_center {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
(q : I → Fin 2) {r : List I}
(q : I → Fin 2)
(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
(F : FreeAlgebra (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le1 F]
: {r : List I} → (c : Contractions r) → (S : Splitting f le1) →
(c.toCenterTerm f q le1 F S) ∈ Subalgebra.center A
| [], ⟨[], .nil⟩, _ => by
dsimp [toCenterTerm]
exact Subalgebra.one_mem (Subalgebra.center A)
| _ :: _, ⟨_, .cons (aux := aux') none c⟩, S => by
dsimp [toCenterTerm]
exact toCenterTerm_center f q le1 F ⟨aux', c⟩ S
| a :: _, ⟨_, .cons (aux := aux') (some n) c⟩, S => by
dsimp [toCenterTerm]
refine Subalgebra.mul_mem (Subalgebra.center A) ?hx ?hy
exact toCenterTerm_center f q le1 F ⟨aux', c⟩ S
apply Subalgebra.smul_mem
rw [ofListM_expand]
rw [map_sum, map_sum]
refine Subalgebra.sum_mem (Subalgebra.center A) ?hy.hx.h
intro x _
simp [CreatAnnilateSect.toList]
rw [ofList_singleton]
exact OperatorMap.superCommute_ofList_singleton_ι_center (q := fun i => q i.1) (le1 := le1) F (S.𝓑p a) ⟨aux'[↑n], x.head⟩
end Contractions
@ -160,7 +173,7 @@ 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]
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le1 F]
(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
@ -172,8 +185,9 @@ lemma static_wick_nil {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
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]
[IsTrans ((i : I) × f i) le1] [IsTotal ((i : I) × f i) le1]
{A : Type} [Semiring A] [Algebra A] (r : List I) (a : I)
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le1 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))) :
@ -228,9 +242,9 @@ lemma static_wick_cons {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
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]
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1] [IsTrans ((i : I) × f i) le1] [IsTotal ((i : I) × f i) le1]
{A : Type} [Semiring A] [Algebra A] (r : List I)
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap F]
(F : FreeAlgebra (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le1 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

View file

@ -213,6 +213,20 @@ lemma toList_grade (q : I → Fin 2) :
rw [h0, h0']
rw [h1]
@[simp]
lemma toList_grade_take {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) : (r : List I) → (a : CreatAnnilateSect f r) → (n : ) →
grade (fun i => q i.fst) (List.take n a.toList) = grade q (List.take n r)
| [], _, _ => by
simp [toList]
| i :: r, a, 0 => by
simp
| i :: r, a, Nat.succ n => by
simp only [grade, Fin.isValue]
rw [toList_grade_take q r a.tail n]
def extractEquiv {I : Type} {f : I → Type} [(i : I) → Fintype (f i)] {l : List I} (n : Fin l.length) : CreatAnnilateSect f l ≃
f (l.get n) × CreatAnnilateSect f (l.eraseIdx n) := by
match l with
@ -246,6 +260,25 @@ def extractEquiv {I : Type} {f : I → Type} [(i : I) → Fintype (f i)] {l : L
next h_1 => simp_all only [not_lt, Fin.val_succ, Fin.coe_cast]))
exact (Fin.insertNthEquiv _ _).symm.trans (Equiv.prodCongr (Equiv.refl _) e1.symm)
lemma extractEquiv_symm_toList_get_same {I : Type} {f : I → Type} [(i : I) → Fintype (f i)]
{l : List I} (n : Fin l.length) (a0 : f (l.get n)) (a : CreatAnnilateSect f (l.eraseIdx n)) :
((extractEquiv n).symm (a0, a)).toList[n] = ⟨l[n], a0⟩ := by
match l with
| [] => exact Fin.elim0 n
| l0 :: l =>
trans (((CreatAnnilateSect.extractEquiv n).symm (a0, a)).toList).get (Fin.cast (by simp) n)
· simp only [List.length_cons, List.get_eq_getElem, Fin.coe_cast]
rfl
rw [CreatAnnilateSect.toList_get]
simp only [List.get_eq_getElem, List.length_cons, extractEquiv, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, Equiv.symm_trans_apply, Equiv.symm_symm, Equiv.prodCongr_symm,
Equiv.refl_symm, Equiv.prodCongr_apply, Equiv.coe_refl, Prod.map_apply, id_eq,
Function.comp_apply, Fin.cast_trans, Fin.cast_eq_self, Sigma.mk.inj_iff, heq_eq_eq]
apply And.intro
· rfl
erw [Fin.insertNthEquiv_apply]
simp only [Fin.insertNth_apply_same]
def eraseIdx (n : Fin l.length) : CreatAnnilateSect f (l.eraseIdx n) :=
(extractEquiv n a).2
@ -334,6 +367,15 @@ lemma eraseIdx_toList : {l : List I} → {n : Fin l.length} → (a : CreatAnnil
· conv_rhs => rw [← eraseIdx_toList (l := r) (n := ⟨n, Nat.succ_lt_succ_iff.mp h⟩) a.tail]
rw [eraseIdx_succ_tail]
lemma extractEquiv_symm_eraseIdx {I : Type} {f : I → Type} [(i : I) → Fintype (f i)]
{l : List I} (n : Fin l.length) (a0 : f l[↑n]) (a : CreatAnnilateSect f (l.eraseIdx n)) :
((extractEquiv n).symm (a0, a)).eraseIdx n = a := by
match l with
| [] => exact Fin.elim0 n
| l0 :: l =>
rw [eraseIdx, extractEquiv]
simp
lemma toList_koszulSignInsert {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
(l : List I) (a : CreatAnnilateSect f l) (x : (i : I) × f i):
@ -524,375 +566,5 @@ lemma koszulOrder_ofListM_eq_ofListM {I : Type} {f : I → Type} [∀ i, Fintype
rw [koszulOrder_ofListM, koszulOrder_ofList, map_smul]
rfl
def liftM {I : Type} (f : I → Type) [∀ i, Fintype (f i)] :
(l : List I) → (a : Π i, f (l.get i)) → List (Σ i, f i)
| [], _ => []
| i :: l, a => ⟨i, a ⟨0, Nat.zero_lt_succ l.length⟩⟩ :: liftM f l (fun i => a (Fin.succ i))
@[simp]
lemma liftM_length {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (r : List I) (a : Π i, f (r.get i)) :
(liftM f r a).length = r.length := by
induction r with
| nil => rfl
| cons i r ih =>
simp only [liftM, List.length_cons, Fin.zero_eta, add_left_inj]
rw [ih]
lemma liftM_cons {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (r0 : I) (r : List I) (a : Π i, f ((r0 :: r).get i)) :
liftM f (r0 :: r) a = ⟨r0, a ⟨0, Nat.zero_lt_succ r.length⟩⟩ :: liftM f r (fun i => a (Fin.succ i)) := by
simp [liftM, List.length_cons, Fin.zero_eta]
lemma liftM_get {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (r : List I) (a : Π i, f (r.get i)) :
(liftM f r a).get = (fun i => ⟨r.get i, a i⟩) ∘ Fin.cast (by simp) := by
induction r with
| nil =>
funext i
exact Fin.elim0 i
| cons i l ih =>
simp only [liftM, List.length_cons, Fin.zero_eta, List.get_eq_getElem]
funext x
match x with
| ⟨0, h⟩ => rfl
| ⟨x + 1, h⟩ =>
simp only [List.length_cons, List.get_eq_getElem, Prod.mk.eta, List.getElem_cons_succ,
Function.comp_apply, Fin.cast_mk]
change (liftM f _ _).get _ = _
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 liftMCongrEquiv_symm_succAbove {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (r0 : I) (r : List I)
(n : Fin (r0 :: r).length) (a0 : f ((r0 :: r).get n) ) (a : Π i, f ((r0 :: r).get (n.succAbove i)))
(i : Fin r.length) :
(liftMCongrEquiv f r0 r n).symm (a0, a) (n.succAbove i) = a i := by
simp [liftMCongrEquiv]
@[simp]
lemma liftMCongrEquiv_symm_zero_succ {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (r0 : I) (r : List I)
(a0 : f ((r0 :: r).get ⟨0, by simp⟩) ) (a : Π i, f ((r0 :: r).get ( i.succ)))
(i : Fin r.length) :
(liftMCongrEquiv f r0 r ⟨0, by simp⟩).symm (a0, a) i.succ = a i := by
trans (liftMCongrEquiv f r0 r ⟨0, by simp⟩).symm (a0, a)
((⟨0, by simp⟩ : Fin (r0 :: r).length).succAbove i)
rfl
rw [liftMCongrEquiv_symm_succAbove]
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
simp only [ofListM, List.length_nil, List.get_eq_getElem, Finset.univ_unique, liftM,
Finset.sum_const, Finset.card_singleton, one_smul]
rw [ofList_eq_smul_one, map_smul, ofList_empty, ofList_eq_smul_one, ofList_empty, map_one]
| i :: l => by
rw [ofListM_cons, ofListM_expand f x l]
let e1 : f i × (Π j, f (l.get j)) ≃ (Π j, f ((i :: l).get j)) :=
(Fin.insertNthEquiv (fun j => f ((i :: l).get j)) 0)
rw [← e1.sum_comp (α := FreeAlgebra _)]
erw [Finset.sum_product]
rw [Finset.sum_mul]
conv_lhs =>
rhs
intro n
rw [Finset.mul_sum]
congr
funext j
congr
funext n
rw [← ofList_singleton, ← ofList_pair, one_mul]
rfl
@[simp]
lemma liftM_grade {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (r : List I) (a : Π i, f (r.get i)) :
grade (fun i => q i.fst) (liftM f r a) = 1 ↔ grade q r = 1 := by
induction r with
| nil =>
simp [liftM]
| cons i r ih =>
simp only [grade, Fin.isValue, ite_eq_right_iff, zero_ne_one, imp_false]
have ih' := ih (fun i => a i.succ)
have h1 : grade (fun i => q i.fst) (liftM f r fun i => a i.succ) = grade q r := by
by_cases h : grade q r = 1
· simp_all
· have h0 : grade q r = 0 := by
omega
rw [h0] at ih'
simp only [Fin.isValue, zero_ne_one, iff_false] at ih'
have h0' : grade (fun i => q i.fst) (liftM f r fun i => a i.succ) = 0 := by
omega
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)
| [], _, _ => by
simp [liftM]
| i :: r, a, 0 => by
simp
| i :: r, a, Nat.succ n => by
simp only [grade, Fin.isValue]
have ih : grade (fun i => q i.fst) (List.take n (liftM f r fun i => a i.succ)) = grade q (List.take n r) := by
refine (liftM_grade_take q r (fun i => a i.succ) n)
rw [ih]
def listMEraseEquiv {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
{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]
split
next h =>
simp_all only [Fin.coe_castSucc]
split
next h_1 => simp_all only [Fin.coe_castSucc, Fin.coe_cast]
next h_1 =>
simp_all only [not_lt, Fin.val_succ, Fin.coe_cast, self_eq_add_right, one_ne_zero]
simp [Fin.le_def] at h_1
simp [Fin.lt_def] at h
omega
next h =>
simp_all only [not_lt, Fin.val_succ]
split
next h_1 =>
simp_all only [Fin.coe_castSucc, Fin.coe_cast, add_right_eq_self, one_ne_zero]
simp [Fin.lt_def] at h_1
simp [Fin.le_def] at h
omega
next h_1 => simp_all only [not_lt, Fin.val_succ, Fin.coe_cast]))
lemma liftM_eraseIdx {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(r0 : I) :
(r : List I) → (n : Fin (r0 :: r).length) →
(a0 : f (r0 :: r)[↑n]) → (a : (i : Fin r.length) → f (r0 :: r)[↑(n.succAbove i)]) →
(liftM f (r0 :: r) ((liftMCongrEquiv f r0 r n).symm (a0, a))).eraseIdx ↑n =
liftM f ((r0 :: r).eraseIdx ↑n) ((listMEraseEquiv n).symm a) := by
intro r n a0 a
match n with
| ⟨0, h0⟩ =>
simp
rw [liftM_cons]
simp
conv_lhs =>
rhs
intro n
erw [liftMCongrEquiv_symm_zero_succ]
simp [listMEraseEquiv]
| ⟨n + 1, hn⟩ =>
simp
rw [liftM_cons, liftM_cons]
simp
apply And.intro
· 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]
(l : List I) (a : (j : Fin l.length) → f (l.get j)) (x : (i : I) × f i):
koszulSignInsert (fun i j => le1 i.fst j.fst) (fun i => q i.fst) x
(liftM f l a) =
koszulSignInsert le1 q x.1 l := by
induction l with
| nil => simp [koszulSignInsert]
| cons b l ih =>
simp [koszulSignInsert]
rw [ih]
lemma koszulSign_liftM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
(l : List I) (a : (i : Fin l.length) → f (l.get i)) :
koszulSign (fun i j => le1 i.fst j.fst) (fun i => q i.fst) (liftM f l a) =
koszulSign le1 q l := by
induction l with
| nil => simp [koszulSign]
| cons i l ih =>
simp [koszulSign, liftM]
rw [ih]
congr 1
rw [koszulSignInsert_liftM]
lemma insertionSortEquiv_liftM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(le1 : I → I → Prop) [DecidableRel le1](l : List I) (a : (i : Fin l.length) → f (l.get i)) :
(HepLean.List.insertionSortEquiv (fun i j => le1 i.fst j.fst) (liftM f l a)) =
(Fin.castOrderIso (by simp)).toEquiv.trans ((HepLean.List.insertionSortEquiv le1 l).trans
(Fin.castOrderIso (by simp)).toEquiv) := by
induction l with
| nil =>
simp [liftM, HepLean.List.insertionSortEquiv]
| cons i l ih =>
simp only [liftM, List.length_cons, Fin.zero_eta, List.insertionSort]
conv_lhs => simp [HepLean.List.insertionSortEquiv]
erw [orderedInsertEquiv_sigma]
rw [ih]
simp only [HepLean.Fin.equivCons_trans, Nat.succ_eq_add_one,
HepLean.Fin.equivCons_castOrderIso, List.length_cons, Nat.add_zero, Nat.zero_eq,
Fin.zero_eta]
ext x
conv_rhs => simp [HepLean.List.insertionSortEquiv]
simp only [Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.cast_trans,
Fin.coe_cast]
have h2' (i : Σ i, f i) (l' : List ( Σ i, f i)) :
List.map (fun i => i.1) (List.orderedInsert (fun i j => le1 i.fst j.fst) i l') =
List.orderedInsert le1 i.1 (List.map (fun i => i.1) l') := by
induction l' with
| nil =>
simp [HepLean.List.orderedInsertEquiv]
| cons j l' ih' =>
by_cases hij : (fun i j => le1 i.fst j.fst) i j
· rw [List.orderedInsert_of_le]
· erw [List.orderedInsert_of_le]
· simp
· exact hij
· exact hij
· simp only [List.orderedInsert, hij, ↓reduceIte, List.unzip_snd, List.map_cons]
have hn : ¬ le1 i.1 j.1 := hij
simp only [hn, ↓reduceIte, List.cons.injEq, true_and]
simpa using ih'
have h2 (l' : List ( Σ i, f i)) :
List.map (fun i => i.1) (List.insertionSort (fun i j => le1 i.fst j.fst) l') =
List.insertionSort le1 (List.map (fun i => i.1) l') := by
induction l' with
| nil =>
simp [HepLean.List.orderedInsertEquiv]
| cons i l' ih' =>
simp only [List.insertionSort, List.unzip_snd]
simp only [List.unzip_snd] at h2'
rw [h2']
congr
rw [HepLean.List.orderedInsertEquiv_congr _ _ _ (h2 _)]
simp only [List.length_cons, Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
Fin.cast_trans, Fin.coe_cast]
have h3 : (List.insertionSort le1 (List.map (fun i => i.1) (liftM f l (fun i => a i.succ)))) =
List.insertionSort le1 l := by
congr
have h3' (l : List I) (a : Π (i : Fin l.length), f (l.get i)) :
List.map (fun i => i.1) (liftM f l a) = l := by
induction l with
| nil => rfl
| cons i l ih' =>
simp only [liftM, List.length_cons, Fin.zero_eta, Prod.mk.eta,
List.unzip_snd, List.map_cons, List.cons.injEq, true_and]
simpa using ih' _
rw [h3']
rw [HepLean.List.orderedInsertEquiv_congr _ _ _ h3]
simp only [List.length_cons, Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
Fin.cast_trans, Fin.cast_eq_self, Fin.coe_cast]
lemma insertionSort_liftM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(le1 : I → I → Prop) [DecidableRel le1](l : List I) (a : (i : Fin l.length) → f (l.get i))
:
List.insertionSort (fun i j => le1 i.fst j.fst) (liftM f l a) =
liftM f (List.insertionSort le1 l)
(Equiv.piCongr (HepLean.List.insertionSortEquiv le1 l) (fun i => (Equiv.cast (by
congr 1
rw [← HepLean.List.insertionSortEquiv_get]
simp))) a) := by
let l1 := List.insertionSort (fun i j => le1 i.fst j.fst) (liftM f l a)
let l2 := liftM f (List.insertionSort le1 l)
(Equiv.piCongr (HepLean.List.insertionSortEquiv le1 l) (fun i => (Equiv.cast (by
congr 1
rw [← HepLean.List.insertionSortEquiv_get]
simp))) a)
change l1 = l2
have hlen : l1.length = l2.length := by
simp [l1, l2]
have hget : l1.get = l2.get ∘ Fin.cast hlen := by
rw [← HepLean.List.insertionSortEquiv_get]
rw [liftM_get, liftM_get]
funext i
rw [insertionSortEquiv_liftM]
simp only [ Function.comp_apply, Equiv.symm_trans_apply,
OrderIso.toEquiv_symm, Fin.symm_castOrderIso, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
Fin.cast_trans, Fin.cast_eq_self, id_eq, eq_mpr_eq_cast, Fin.coe_cast, Sigma.mk.inj_iff]
apply And.intro
· have h1 := congrFun (HepLean.List.insertionSortEquiv_get (r := le1) l) (Fin.cast (by simp) i)
rw [← h1]
simp
· simp [Equiv.piCongr]
exact (cast_heq _ _).symm
apply List.ext_get hlen
rw [hget]
simp
end
end Wick

View file

@ -26,19 +26,23 @@ namespace Wick
noncomputable section
class SuperCommuteCenterMap {A : Type} [Semiring A] [Algebra A]
(q : I → Fin 2) (F : FreeAlgebra I →ₐ[] A) : Prop where
prop : ∀ i j, F (superCommute q (FreeAlgebra.ι i) (FreeAlgebra.ι j)) ∈ Subalgebra.center A
dif_grade : ∀ i j, q i ≠ q j → F (superCommute q (FreeAlgebra.ι i) (FreeAlgebra.ι j)) = 0
namespace SuperCommuteCenterMap
variable {I : Type} {A : Type} [Semiring A] [Algebra A]
(f : FreeAlgebra I →ₐ[] A) (q : I → Fin 2) [SuperCommuteCenterMap q f]
class OperatorMap {A : Type} [Semiring A] [Algebra A]
(q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1] (F : FreeAlgebra I →ₐ[] A) : Prop where
superCommute_mem_center : ∀ i j, F (superCommute q (FreeAlgebra.ι i) (FreeAlgebra.ι j)) ∈ Subalgebra.center A
superCommute_diff_grade_zero : ∀ i j, q i ≠ q j → F (superCommute q (FreeAlgebra.ι i) (FreeAlgebra.ι j)) = 0
superCommute_ordered_zero : ∀ i j, ∀ a b,
F (koszulOrder le1 q (a * superCommute q (FreeAlgebra.ι i) (FreeAlgebra.ι j) * b)) = 0
lemma ofList_fst (i j : I) :
f (superCommute q (ofList [i] xa) (FreeAlgebra.ι j)) ∈ Subalgebra.center A := by
have h1 : f (superCommute q (ofList [i] xa) (FreeAlgebra.ι j)) =
xa • f (superCommute q (FreeAlgebra.ι i) (FreeAlgebra.ι j)) := by
namespace OperatorMap
variable {A : Type} [Semiring A] [Algebra A] {q : I → Fin 2} {le1 : I → I → Prop}
[DecidableRel le1] (F : FreeAlgebra I →ₐ[] A)
lemma superCommute_ofList_singleton_ι_center [OperatorMap q le1 F] (i j :I) :
F (superCommute q (ofList [i] xa) (FreeAlgebra.ι j)) ∈ Subalgebra.center A := by
have h1 : F (superCommute q (ofList [i] xa) (FreeAlgebra.ι j)) =
xa • F (superCommute q (FreeAlgebra.ι i) (FreeAlgebra.ι j)) := by
rw [← map_smul]
congr
rw [ofList_eq_smul_one, ofList_singleton]
@ -46,31 +50,23 @@ lemma ofList_fst (i j : I) :
rfl
rw [h1]
refine Subalgebra.smul_mem (Subalgebra.center A) ?_ xa
exact prop i j
exact superCommute_mem_center (le1 := le1) i j
lemma ofList_freeAlgebraMap {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (c : (Σ i, f i)) (x : )
{A : Type} [Semiring A] [Algebra A] (F : FreeAlgebra (Σ i, f i) →ₐ[] A)
[SuperCommuteCenterMap (fun i => q i.1) F] (b : I) :
F ((superCommute fun i => q i.fst) (ofList [c] x) ((freeAlgebraMap f) (FreeAlgebra.ι b)))
∈ Subalgebra.center A := by
rw [freeAlgebraMap_ι]
rw [map_sum, map_sum]
refine Subalgebra.sum_mem (Subalgebra.center A) ?h
intro n hn
exact ofList_fst F (fun i => q i.fst) c ⟨b, n⟩
end SuperCommuteCenterMap
end OperatorMap
lemma superCommuteTake_superCommuteCenterMap {I : Type} (q : I → Fin 2) (lb : List I) (xa xb : ) (n : )
lemma superCommuteTake_operatorMap {I : Type} (q : I → Fin 2)
(le1 : I → I → Prop) [DecidableRel le1]
(lb : List I) (xa xb : ) (n : )
(hn : n < lb.length) {A : Type} [Semiring A] [Algebra A] (f : FreeAlgebra I →ₐ[] A)
[SuperCommuteCenterMap q f] (i : I) :
[OperatorMap q le1 f] (i : I) :
f (superCommuteTake q [i] lb xa xb n hn) =
f (superCommute q (ofList [i] xa) (FreeAlgebra.ι (lb.get ⟨n, hn⟩)))
* (superCommuteCoef q [i] (List.take n lb) •
f (ofList (List.eraseIdx lb n) xb)) := by
have hn : f ((superCommute q) (ofList [i] xa) (FreeAlgebra.ι (lb.get ⟨n, hn⟩))) ∈
Subalgebra.center A := SuperCommuteCenterMap.ofList_fst f q i (lb.get ⟨n, hn⟩)
Subalgebra.center A := OperatorMap.superCommute_ofList_singleton_ι_center (le1 := le1) f i (lb.get ⟨n, hn⟩)
rw [Subalgebra.mem_center_iff] at hn
rw [superCommuteTake, map_mul, map_mul, map_smul, hn, mul_assoc, smul_mul_assoc,
← map_mul, ← ofList_pair]
@ -79,11 +75,12 @@ lemma superCommuteTake_superCommuteCenterMap {I : Type} (q : I → Fin 2) (lb :
· exact one_mul xb
lemma superCommuteTakeM_F {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
lemma superCommuteTakeM_operatorMap {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (c : (Σ i, f i)) (r : List I) (x y : ) (n : )
(hn : n < r.length)
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
{A : Type} [Semiring A] [Algebra A] (F : FreeAlgebra (Σ i, f i) →ₐ[] A)
[SuperCommuteCenterMap (fun i => q i.1) F] :
[OperatorMap (fun i => q i.1) le1 F] :
F (superCommuteTakeM q [c] r x y n hn) = superCommuteCoefM q [c] (List.take n r) •
(F (superCommute (fun i => q i.1) (ofList [c] x) (freeAlgebraMap f (FreeAlgebra.ι (r.get ⟨n, hn⟩))))
* F (ofListM f (List.eraseIdx r n) y)) := by
@ -92,8 +89,12 @@ lemma superCommuteTakeM_F {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
congr
rw [map_mul, map_mul]
have h1 : F ((superCommute fun i => q i.fst) (ofList [c] x) ((freeAlgebraMap f) (FreeAlgebra.ι (r.get ⟨n, hn⟩))))
∈ Subalgebra.center A :=
SuperCommuteCenterMap.ofList_freeAlgebraMap q c x F (r.get ⟨n, hn⟩)
∈ Subalgebra.center A := by
rw [freeAlgebraMap_ι]
rw [map_sum, map_sum]
refine Subalgebra.sum_mem _ ?_
intro n
exact fun a => OperatorMap.superCommute_ofList_singleton_ι_center (le1 := le1) F c _
rw [Subalgebra.mem_center_iff] at h1
rw [h1, mul_assoc, ← map_mul]
congr
@ -108,7 +109,7 @@ lemma superCommute_koszulOrder_le_ofList {I : Type}
(le1 :I → I → Prop) [DecidableRel le1] [IsTotal I le1] [IsTrans I le1]
(i : I)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap q F] :
(F : FreeAlgebra I →ₐ A) [OperatorMap q le1 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))) *
@ -118,7 +119,7 @@ lemma superCommute_koszulOrder_le_ofList {I : Type}
conv_lhs =>
enter [2, 2]
intro n
rw [superCommuteTake_superCommuteCenterMap]
rw [superCommuteTake_operatorMap (le1 := le1)]
enter [1, 2, 2, 2]
change ((List.insertionSort le1 r).get ∘ (HepLean.List.insertionSortEquiv le1 r)) n
rw [HepLean.List.insertionSort_get_comp_insertionSortEquiv]
@ -135,7 +136,7 @@ lemma superCommute_koszulOrder_le_ofList {I : Type}
congr
funext n
by_cases hq : q i ≠ q (r.get n)
· have hn := SuperCommuteCenterMap.dif_grade (q := q) (F := F) i (r.get n) hq
· have hn := OperatorMap.superCommute_diff_grade_zero (q := q) (F := F) le1 i (r.get n) hq
conv_lhs =>
enter [2, 1]
rw [ofList_singleton, hn]
@ -151,9 +152,9 @@ lemma superCommute_koszulOrder_le_ofList {I : Type}
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)
(i : I)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap q F] :
(F : FreeAlgebra I →ₐ A) [OperatorMap q le1 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 =>
@ -168,14 +169,26 @@ lemma koszulOrder_of_le_all_ofList {I : Type}
rw [map_smul]
rw [← neg_smul]
rw [map_smul, map_smul, map_smul]
sorry
conv_lhs =>
rhs
rhs
rw [superCommute_ofList_sum]
rw [map_sum, map_sum]
dsimp [superCommuteTake]
rw [ofList_singleton]
rhs
intro n
rw [Algebra.smul_mul_assoc, Algebra.smul_mul_assoc]
rw [map_smul, map_smul]
rw [OperatorMap.superCommute_ordered_zero ]
simp
rw [ofList_singleton]
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 →ₐ A) [SuperCommuteCenterMap q F] :
(F : FreeAlgebra I →ₐ A) [OperatorMap q le1 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
@ -193,14 +206,13 @@ lemma le_all_mul_koszulOrder_ofList {I : Type}
rw [smul_smul]
rw [superCommuteCoef_mul_self]
simp [ofList_singleton]
exact fun j => hi j
· 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 q F]
(F : FreeAlgebra I →ₐ[] A)
(n : Option (Fin r.length)) : A :=
match n with
| none => 1
@ -210,7 +222,7 @@ def superCommuteCenterOrder {I : Type}
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 q F] :
(F : FreeAlgebra I →ₐ[] A) :
superCommuteCenterOrder q r i F none = 1 := by
simp [superCommuteCenterOrder]
@ -221,7 +233,7 @@ lemma le_all_mul_koszulOrder_ofList_expand {I : Type}
[IsTotal I le1] [IsTrans I le1]
(i : I) (hi : ∀ (j : I), le1 j i)
{A : Type} [Semiring A] [Algebra A]
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap q F] :
(F : FreeAlgebra I →ₐ[] A) [OperatorMap q le1 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]
@ -244,7 +256,7 @@ lemma le_all_mul_koszulOrder_ofListM_expand {I : Type} {f : I → Type} [∀ i,
[IsTotal (Σ i, f i) le1] [IsTrans (Σ i, f i) 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 (fun i => q i.1) F] :
(F : FreeAlgebra (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le1 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) •
@ -259,9 +271,9 @@ lemma le_all_mul_koszulOrder_ofListM_expand {I : Type} {f : I → Type} [∀ i,
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
let e1 (a : CreatAnnilateSect f (r0 :: r)) :
Option (Fin a.toList.length) ≃ Option (Fin (r0 :: r).length) :=
Equiv.optionCongr (Fin.castOrderIso (CreatAnnilateSect.toList_length a)).toEquiv
conv_lhs =>
rhs
intro a
@ -283,28 +295,22 @@ lemma le_all_mul_koszulOrder_ofListM_expand {I : Type} {f : I → Type} [∀ i,
rw [ofList_cons_eq_ofList]
· congr
funext n
rw [← (liftMCongrEquiv _ _ _ n).symm.sum_comp]
rw [← (CreatAnnilateSect.extractEquiv n).symm.sum_comp]
simp only [List.get_eq_getElem, List.length_cons, Equiv.optionCongr_symm, OrderIso.toEquiv_symm,
Fin.symm_castOrderIso, Equiv.optionCongr_apply, RelIso.coe_fn_toEquiv, Option.map_some',
Fin.castOrderIso_apply, Algebra.smul_mul_assoc, e1]
erw [Finset.sum_product]
have h1 (a0 : f (r0 :: r)[↑n]) (a : (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
have h1 (a0 : f (r0 :: r)[↑n]) (a : CreatAnnilateSect f ((r0 :: r).eraseIdx ↑n)):
superCommuteCenterOrder (fun i => q i.fst) ((CreatAnnilateSect.extractEquiv n).symm (a0, a)).toList i F
(some (Fin.cast (by simp) n)) = superCommuteCoef q [(r0 :: r).get n] (List.take (↑n) (r0 :: r)) •
F (((superCommute fun i => q i.fst) (ofList [i] 1)) (FreeAlgebra.ι ⟨(r0 :: r).get n, a0⟩)) := by
simp only [superCommuteCenterOrder, List.get_eq_getElem, List.length_cons, Fin.coe_cast]
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]
erw [CreatAnnilateSect.extractEquiv_symm_toList_get_same]
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)))) =
(List.take (↑n) ((CreatAnnilateSect.extractEquiv n).symm (a0, a)).toList) =
superCommuteCoef q [(r0 :: r).get n] (List.take (↑n) ((r0 :: r))) := by
simp only [superCommuteCoef, List.get_eq_getElem, List.length_cons, Fin.isValue,
liftM_grade_take]
CreatAnnilateSect.toList_grade_take]
rfl
erw [hsc]
rfl
@ -318,19 +324,18 @@ lemma le_all_mul_koszulOrder_ofListM_expand {I : Type} {f : I → Type} [∀ i,
rhs
intro a0
rw [← Finset.mul_sum]
have hl (n : Fin (r0 :: r).length) (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]
congr
sorry
conv_lhs =>
rhs
intro a0
enter [2, 2]
intro a
erw [hl n a0 a]
simp [optionEraseZ]
rhs
rhs
lhs
rw [← CreatAnnilateSect.eraseIdx_toList]
erw [CreatAnnilateSect.extractEquiv_symm_eraseIdx]
rw [← Finset.sum_mul]
conv_lhs =>
lhs
@ -338,7 +343,6 @@ lemma le_all_mul_koszulOrder_ofListM_expand {I : Type} {f : I → Type} [∀ i,
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]

View file

@ -87,10 +87,6 @@ def koszulSign {I : Type} (r : I → I → Prop) [DecidableRel r] (q : I → Fin
| 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
def natTest3 : × × → Fin 2 := fun ⟨a, b, c⟩ => if a % 2 = 0 then 0 else 1
#eval List.insertionSort (fun i j => i.2 ≤ j.2) [(1, 1, 0), (1, 0, 3)]
#eval koszulSign (fun i j => i.2 ≤ j.2) natTest3 [ (0, 0, 2), (1, 1, 0), (1, 1, 3)]
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

View file

@ -275,12 +275,5 @@ lemma ofListM_ofList_superCommute' {I : Type}
· 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) =
0 := by
sorry
end
end Wick

View file

@ -48,10 +48,10 @@ lemma superCommute_ofList_ofListM {I : Type} {f : I → Type} [∀ i, Fintype (
funext a
rw [superCommute_ofList_ofList]
congr 1
· exact ofList_pair l (liftM f r a) x y
· exact ofList_pair l a.toList x y
congr 1
· simp
· exact ofList_pair (liftM f r a) l y x
· exact ofList_pair a.toList l y x
· rw [ofList_pair]
simp only [neg_mul]
@ -124,7 +124,7 @@ lemma superCommuteM_ofList_cons {I : Type} {f : I → Type} [∀ i, Fintype (f i
rw [ofListM_expand]
rw [Finset.mul_sum]
rw [map_sum]
trans ∑ n, ∑ j : f b1, ((superCommute fun i => q i.fst) (ofList l x)) (( FreeAlgebra.ι ⟨b1, j⟩) * ofList (liftM f r n) y)
trans ∑ (n : CreatAnnilateSect f r), ∑ j : f b1, ((superCommute fun i => q i.fst) (ofList l x)) (( FreeAlgebra.ι ⟨b1, j⟩) * ofList n.toList y)
· apply congrArg
funext n
rw [← map_sum]
@ -152,7 +152,7 @@ lemma superCommuteM_ofList_cons {I : Type} {f : I → Type} [∀ i, Fintype (f i
rw [← Finset.sum_add_distrib]
congr
funext b
trans ((superCommute fun i => q i.fst) (ofList l x)) (ofList (⟨b1, b⟩ :: liftM f r n) y)
trans ((superCommute fun i => q i.fst) (ofList l x)) (ofList (⟨b1, b⟩ :: n.toList) y)
· congr
rw [ofList_cons_eq_ofList]
rw [ofList_singleton]