refactor: Start filling in sorries

This commit is contained in:
jstoobysmith 2024-12-17 16:35:34 +00:00
parent dceaab7117
commit 3123e831d8
6 changed files with 467 additions and 28 deletions

View file

@ -174,6 +174,7 @@ def finExtractOne {n : } (i : Fin n.succ) : Fin n.succ ≃ Fin 1 ⊕ Fin n :=
(Equiv.sumAssoc (Fin 1) (Fin i) (Fin (n - i))).trans <| (Equiv.sumAssoc (Fin 1) (Fin i) (Fin (n - i))).trans <|
Equiv.sumCongr (Equiv.refl (Fin 1)) (finSumFinEquiv.trans (finCongr (by omega))) Equiv.sumCongr (Equiv.refl (Fin 1)) (finSumFinEquiv.trans (finCongr (by omega)))
@[simp]
lemma finExtractOne_apply_eq {n : } (i : Fin n.succ) : lemma finExtractOne_apply_eq {n : } (i : Fin n.succ) :
finExtractOne i i = Sum.inl 0 := by finExtractOne i i = Sum.inl 0 := by
simp only [Nat.succ_eq_add_one, finExtractOne, Equiv.trans_apply, finCongr_apply, simp only [Nat.succ_eq_add_one, finExtractOne, Equiv.trans_apply, finCongr_apply,
@ -239,12 +240,19 @@ lemma finExtractOne_symm_inl_apply {n : } (i : Fin n.succ) :
(finExtractOne i).symm (Sum.inl 0) = i := by (finExtractOne i).symm (Sum.inl 0) = i := by
rfl rfl
lemma finExtractOne_apply_neq {n : } (i j : Fin n.succ.succ) (hij : i ≠ j) :
finExtractOne i j = Sum.inr (predAboveI i j) := by
symm
apply (Equiv.symm_apply_eq _).mp ?_
simp
exact succsAbove_predAboveI hij
/-- Given an equivalence `Fin n.succ.succ ≃ Fin n.succ.succ`, and an `i : Fin n.succ.succ`, /-- Given an equivalence `Fin n.succ.succ ≃ Fin n.succ.succ`, and an `i : Fin n.succ.succ`,
the map `Fin n.succ → Fin n.succ` obtained by dropping `i` and it's image. -/ the map `Fin n.succ → Fin n.succ` obtained by dropping `i` and it's image. -/
def finExtractOnPermHom (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fin n.succ.succ) : def finExtractOnPermHom {m : } (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fin m.succ.succ) :
Fin n.succ → Fin n.succ := fun x => predAboveI (σ i) (σ ((finExtractOne i).symm (Sum.inr x))) Fin n.succ → Fin m.succ := fun x => predAboveI (σ i) (σ ((finExtractOne i).symm (Sum.inr x)))
lemma finExtractOnPermHom_inv (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fin n.succ.succ) : lemma finExtractOnPermHom_inv {m : } (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fin m.succ.succ) :
(finExtractOnPermHom (σ i) σ.symm) ∘ (finExtractOnPermHom i σ) = id := by (finExtractOnPermHom (σ i) σ.symm) ∘ (finExtractOnPermHom i σ) = id := by
funext x funext x
simp only [Nat.succ_eq_add_one, Function.comp_apply, finExtractOnPermHom, Equiv.symm_apply_apply, simp only [Nat.succ_eq_add_one, Function.comp_apply, finExtractOnPermHom, Equiv.symm_apply_apply,
@ -270,8 +278,8 @@ lemma finExtractOnPermHom_inv (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fi
/-- Given an equivalence `Fin n.succ.succ ≃ Fin n.succ.succ`, and an `i : Fin n.succ.succ`, /-- Given an equivalence `Fin n.succ.succ ≃ Fin n.succ.succ`, and an `i : Fin n.succ.succ`,
the equivalence `Fin n.succ ≃ Fin n.succ` obtained by dropping `i` and it's image. -/ the equivalence `Fin n.succ ≃ Fin n.succ` obtained by dropping `i` and it's image. -/
def finExtractOnePerm (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fin n.succ.succ) : def finExtractOnePerm {m : } (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fin m.succ.succ) :
Fin n.succ ≃ Fin n.succ where Fin n.succ ≃ Fin m.succ where
toFun x := finExtractOnPermHom i σ x toFun x := finExtractOnPermHom i σ x
invFun x := finExtractOnPermHom (σ i) σ.symm x invFun x := finExtractOnPermHom (σ i) σ.symm x
left_inv x := by left_inv x := by
@ -279,6 +287,15 @@ def finExtractOnePerm (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fin n.succ
right_inv x := by right_inv x := by
simpa using congrFun (finExtractOnPermHom_inv (σ i) σ.symm) x simpa using congrFun (finExtractOnPermHom_inv (σ i) σ.symm) x
lemma finExtractOnePerm_equiv {n m : } (e : Fin n.succ.succ ≃ Fin m.succ.succ) (i : Fin n.succ.succ) :
e ∘ i.succAbove = (e i).succAbove ∘ finExtractOnePerm i e := by
simp [finExtractOnePerm]
funext x
simp [finExtractOnPermHom]
rw [succsAbove_predAboveI]
simp
exact Fin.ne_succAbove i x
@[simp] @[simp]
lemma finExtractOnePerm_apply (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fin n.succ.succ) lemma finExtractOnePerm_apply (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fin n.succ.succ)
(x : Fin n.succ) : finExtractOnePerm i σ x = predAboveI (σ i) (x : Fin n.succ) : finExtractOnePerm i σ x = predAboveI (σ i)
@ -376,6 +393,11 @@ def equivCons {n m : } (e : Fin n ≃ Fin m) : Fin n.succ ≃ Fin m.succ wher
subst hj subst hj
simp simp
@[simp]
lemma equivCons_zero {n m : } (e : Fin n ≃ Fin m) :
equivCons e 0 = 0 := by
simp [equivCons]
@[simp] @[simp]
lemma equivCons_trans {n m k : } (e : Fin n ≃ Fin m) (f : Fin m ≃ Fin k) : lemma equivCons_trans {n m k : } (e : Fin n ≃ Fin m) (f : Fin m ≃ Fin k) :
Fin.equivCons (e.trans f) = (Fin.equivCons e).trans (Fin.equivCons f) := by Fin.equivCons (e.trans f) = (Fin.equivCons e).trans (Fin.equivCons f) := by
@ -409,4 +431,15 @@ lemma equivCons_symm_succ {n m : } (e : Fin n ≃ Fin m) (i : ) (hi : i +
rw [Fin.cons_succ] rw [Fin.cons_succ]
simp simp
@[simp]
lemma equivCons_succ {n m : } (e : Fin n ≃ Fin m) (i : ) (hi : i + 1 < n.succ) :
(Fin.equivCons e) ⟨i + 1, hi⟩ = (e ⟨i, Nat.succ_lt_succ_iff.mp hi⟩).succ := by
simp only [Nat.succ_eq_add_one, equivCons, Equiv.toFun_as_coe, Equiv.invFun_as_coe,
Equiv.coe_fn_symm_mk]
have hi : ⟨i + 1, hi⟩ = Fin.succ ⟨i, Nat.succ_lt_succ_iff.mp hi⟩ := by rfl
simp
rw [hi]
rw [Fin.cons_succ]
rfl
end HepLean.Fin end HepLean.Fin

View file

@ -17,6 +17,116 @@ open Fin
open HepLean open HepLean
variable {n : Nat} variable {n : Nat}
def orderedInsertPos {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I) (r0 : I) :
Fin (List.orderedInsert le1 r0 r).length :=
⟨(List.takeWhile (fun b => decide ¬ le1 r0 b) r).length, by
rw [List.orderedInsert_length]
have h1 : (List.takeWhile (fun b => decide ¬le1 r0 b) r).length ≤ r.length :=
List.Sublist.length_le (List.takeWhile_sublist fun b => decide ¬le1 r0 b)
omega⟩
lemma orderedInsertPos_lt_length {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(r0 : I) : orderedInsertPos le1 r r0 < (r0 :: r).length := by
simp only [orderedInsertPos, List.length_cons]
have h1 : (List.takeWhile (fun b => decide ¬le1 r0 b) r).length ≤ r.length :=
List.Sublist.length_le (List.takeWhile_sublist fun b => decide ¬le1 r0 b)
omega
@[simp]
lemma orderedInsert_get_orderedInsertPos {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) :
(List.orderedInsert le1 r0 r)[(orderedInsertPos le1 r r0).val] = r0 := by
simp [orderedInsertPos, List.orderedInsert_eq_take_drop]
rw [List.getElem_append]
simp
@[simp]
lemma orderedInsert_get_lt {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) (i : Fin (List.orderedInsert le1 r0 r).length)
(hi : i.val < orderedInsertPos le1 r r0) :
(List.orderedInsert le1 r0 r)[i] = r.get ⟨i, by
simp only [orderedInsertPos] at hi
have h1 : (List.takeWhile (fun b => decide ¬le1 r0 b) r).length ≤ r.length :=
List.Sublist.length_le (List.takeWhile_sublist fun b => decide ¬le1 r0 b)
omega⟩ := by
simp [orderedInsertPos] at hi
simp [List.orderedInsert_eq_take_drop]
rw [List.getElem_append]
simp [hi]
rw [List.IsPrefix.getElem]
exact List.takeWhile_prefix fun b => !decide (le1 r0 b)
def orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I) (r0 : I) :
Fin (r0 :: r).length ≃ Fin (List.orderedInsert le1 r0 r).length := by
let e2 : Fin (List.orderedInsert le1 r0 r).length ≃ Fin (r0 :: r).length :=
(Fin.castOrderIso (List.orderedInsert_length le1 r r0)).toEquiv
let e3 : Fin (r0 :: r).length ≃ Fin 1 ⊕ Fin (r).length := finExtractOne 0
let e4 : Fin (r0 :: r).length ≃ Fin 1 ⊕ Fin (r).length := finExtractOne ⟨orderedInsertPos le1 r r0, orderedInsertPos_lt_length le1 r r0⟩
exact e3.trans (e4.symm.trans e2.symm)
@[simp]
lemma orderedInsertEquiv_zero {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(r0 : I) : orderedInsertEquiv le1 r r0 ⟨0, by simp⟩ = orderedInsertPos le1 r r0 := by
simp [orderedInsertEquiv]
@[simp]
lemma orderedInsertEquiv_succ {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(r0 : I) (n : ) (hn : Nat.succ n < (r0 :: r).length) :
orderedInsertEquiv le1 r r0 ⟨Nat.succ n, hn⟩ = Fin.cast (List.orderedInsert_length le1 r r0).symm
((Fin.succAbove ⟨(orderedInsertPos le1 r r0), orderedInsertPos_lt_length le1 r r0⟩) ⟨n, Nat.succ_lt_succ_iff.mp hn⟩) := by
simp [orderedInsertEquiv]
match r with
| [] =>
simp
| r1 :: r =>
erw [finExtractOne_apply_neq]
simp [orderedInsertPos]
rfl
exact ne_of_beq_false rfl
lemma get_eq_orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(r0 : I) :
(r0 :: r).get = (List.orderedInsert le1 r0 r).get ∘ (orderedInsertEquiv le1 r r0) := by
funext x
match x with
| ⟨0, h⟩ =>
simp
erw [orderedInsertEquiv_zero]
simp
| ⟨Nat.succ n, h⟩ =>
simp
erw [orderedInsertEquiv_succ]
simp [Fin.succAbove]
by_cases hn : n < ↑(orderedInsertPos le1 r r0)
· simp [hn]
erw [orderedInsert_get_lt le1 r r0 ⟨n, by omega⟩ ]
rfl
simpa using hn
· simp [hn]
simp [List.orderedInsert_eq_take_drop]
rw [List.getElem_append]
have hn' : ¬ n + 1 < (List.takeWhile (fun b => !decide (le1 r0 b)) r).length := by
simp [orderedInsertPos] at hn
omega
simp [hn']
have hnn : n + 1 - (List.takeWhile (fun b => !decide (le1 r0 b)) r).length =
(n - (List.takeWhile (fun b => !decide (le1 r0 b)) r).length) + 1 := by
simp [orderedInsertPos] at hn
omega
simp [hnn]
conv_rhs =>
rw [List.IsSuffix.getElem (List.dropWhile_suffix fun b => !decide (le1 r0 b))]
congr
have hr : r.length = (List.takeWhile (fun b => !decide (le1 r0 b)) r).length + (List.dropWhile (fun b => !decide (le1 r0 b)) r).length := by
rw [← List.length_append]
congr
exact Eq.symm (List.takeWhile_append_dropWhile (fun b => !decide (le1 r0 b)) r)
simp [hr]
omega
/-- The equivalence between `Fin (a :: l).length` and `Fin (List.orderedInsert r a l).length` /-- The equivalence between `Fin (a :: l).length` and `Fin (List.orderedInsert r a l).length`
mapping `0` in the former to the location of `a` in the latter. -/ mapping `0` in the former to the location of `a` in the latter. -/
def insertEquiv {α : Type} (r : αα → Prop) [DecidableRel r] (a : α) : (l : List α) → def insertEquiv {α : Type} (r : αα → Prop) [DecidableRel r] (a : α) : (l : List α) →
@ -172,24 +282,107 @@ lemma eraseIdx_cons_length {I : Type} (a : I) (l : List I) (i : Fin (a :: l).len
lemma eraseIdx_get {I : Type} (l : List I) (i : Fin l.length) : 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 (List.eraseIdx l i).get = l.get ∘ (Fin.cast (eraseIdx_length l i)) ∘ (Fin.cast (eraseIdx_length l i).symm i).succAbove := by
ext x ext x
simp only [Function.comp_apply, List.get_eq_getElem, List.eraseIdx, List.getElem_eraseIdx] simp only [Function.comp_apply, List.get_eq_getElem, List.eraseIdx, List.getElem_eraseIdx]
simp [Fin.succAbove] simp [Fin.succAbove]
by_cases hi: x.1 < i.val by_cases hi: x.castSucc < Fin.cast (by exact Eq.symm (eraseIdx_length l i)) i
· have h0 : x.castSucc < ↑↑i := by · simp [ hi]
simp [Fin.lt_def] intro h
rw [Nat.mod_eq_of_lt] rw [Fin.lt_def] at hi
exact hi simp_all
rw [eraseIdx_length] omega
exact i.prop · simp [ hi]
simp [h0, hi] rw [Fin.lt_def] at hi
· have h0 : ¬ x.castSucc < ↑↑i := by simp at hi
simp [Fin.lt_def] have hn : ¬ x.val < i.val := by omega
rw [Nat.mod_eq_of_lt] simp [hn]
exact Nat.le_of_not_lt hi
rw [eraseIdx_length]
exact i.prop lemma eraseIdx_insertionSort_length {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
simp [h0, hi] (n : Fin r.length ) :
((List.insertionSort le1 r).eraseIdx ↑((HepLean.List.insertionSortEquiv le1 r) n)).length
= (List.insertionSort le1 (r.eraseIdx n)).length := by
rw [List.length_eraseIdx]
have hn := ((insertionSortEquiv le1 r) n).prop
simp [List.length_insertionSort] at hn
simp [hn]
rw [List.length_eraseIdx]
simp
lemma eraseIdx_insertionSort_get {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(n : Fin r.length ) :
((List.insertionSort le1 r).eraseIdx ↑((HepLean.List.insertionSortEquiv le1 r) n)).get
= (List.insertionSort le1 (r.eraseIdx n)).get ∘ Fin.cast (eraseIdx_insertionSort_length le1 r n) := by
rw [eraseIdx_get, ← insertionSortEquiv_get, ← insertionSortEquiv_get,
eraseIdx_get]
@[simp]
lemma eraseIdx_orderedInsert_zero {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(r0 : I) :
(List.orderedInsert le1 r0 r).eraseIdx ↑((insertEquiv le1 r0 r) ⟨0, by simp⟩) = r := by
induction r with
| nil => simp
| cons r1 r ih =>
by_cases hr0 : le1 r0 r1
· simp [hr0, insertEquiv]
· simp [hr0, insertEquiv, equivCons]
simpa using ih
@[simp]
lemma eraseIdx_orderedInsert_succ_succ {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(r0 : I) (n : ) (hn : Nat.succ (Nat.succ n) < (r0 :: r).length) :
(List.orderedInsert le1 r0 r).eraseIdx ↑((insertEquiv le1 r0 r) ⟨Nat.succ (Nat.succ n), hn⟩) =
List.orderedInsert le1 r0 (r.eraseIdx (Nat.succ n)) := by
induction r with
| nil => simp at hn
| cons r1 r ih =>
by_cases hr0 : le1 r0 r1
· simp [hr0, insertEquiv]
· simp [hr0, insertEquiv]
rw [Equiv.swap_apply_of_ne_of_ne]
simp
have h1 (n : ) (hn : n + 1 < (r0 :: r).length) : (List.orderedInsert le1 r0 r).eraseIdx ↑((insertEquiv le1 r0 r) ⟨n + 1, hn⟩) =
List.orderedInsert le1 r0 (r.eraseIdx n) := by
match n with
| 0 => simp
| Nat.succ n =>
exact ih ?_
sorry
exact h1 n (Nat.succ_lt_succ_iff.mp hn)
simp
rw [Fin.ext_iff]
simp
simp
rw [Fin.ext_iff]
simp
lemma eraseIdx_insertionSort_zero {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(h0 : 0 < r.length) :
((List.insertionSort le1 r).eraseIdx ↑((HepLean.List.insertionSortEquiv le1 r) ⟨0, h0⟩))
= (List.insertionSort le1 (r.eraseIdx 0)) := by
induction r with
| nil => simp
| cons r0 r ih =>
simp [insertionSortEquiv]
erw [eraseIdx_orderedInsert_zero]
lemma eraseIdx_insertionSort_succ {I : Type} (le1 : I → I → Prop) [DecidableRel le1] :
(n : ) → (r : List I) → (hn : n < r.length) →
((List.insertionSort le1 r).eraseIdx ↑((HepLean.List.insertionSortEquiv le1 r) ⟨n, hn⟩))
= (List.insertionSort le1 (r.eraseIdx n))
| 0, _, _ => by exact eraseIdx_insertionSort_zero le1 _ _
| Nat.succ n, [], hn => by
simp [insertionSortEquiv]
| Nat.succ 0, (r0 :: r), hn => by
simp [insertionSortEquiv]
| Nat.succ (Nat.succ n), (r0 :: r), hn => by
simp [insertionSortEquiv]
rw [eraseIdx_orderedInsert_succ_succ]
exact eraseIdx_insertionSort_succ n r (Nat.lt_of_succ_lt hn)
end HepLean.List end HepLean.List

View file

@ -118,10 +118,12 @@ lemma superCommuteCoefM_empty {I : Type} {f : I → Type} [∀ i, Fintype (f i)
superCommuteCoefM q l [] = 1 := by superCommuteCoefM q l [] = 1 := by
simp [superCommuteCoefM] simp [superCommuteCoefM]
def test {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I) lemma koszulSign_first_remove {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (l : List I)
[DecidableRel le1] (n : Fin r.length) : := [DecidableRel le1] (a : I) :
if grade q (List.take n r) = grade q ((List.take (↑((HepLean.List.insertionSortEquiv le1 r) n)) let n0 := ((HepLean.List.insertionSortEquiv le1 (a :: l)).symm ⟨0, by
(List.insertionSort le1 r))) then 1 else -1 rw [List.length_insertionSort]; simp⟩)
koszulSign le1 q (a :: l) = superCommuteCoef q [(a :: l).get n0] (List.take n0 (a :: l)) *
koszulSign le1 q ((a :: l).eraseIdx n0) := by sorry
def superCommuteCoefLE {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I) def superCommuteCoefLE {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
[DecidableRel le1] (i : I) (n : Fin r.length) : := [DecidableRel le1] (i : I) (n : Fin r.length) : :=
@ -130,6 +132,22 @@ def superCommuteCoefLE {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r
(List.insertionSort le1 r)) * (List.insertionSort le1 r)) *
koszulSign le1 q (r.eraseIdx ↑n) koszulSign le1 q (r.eraseIdx ↑n)
def natTestQ' : → Fin 2 := fun n => if n % 2 = 0 then 0 else 1
#eval superCommuteCoefLE (natTestQ') (fun i j => i ≤ j) [5, 4, 4, 3, 0] 0 ⟨0, by simp⟩
#eval koszulSign (fun i j => i ≤ j) (natTestQ') [5, 4, 4, 3, 0]
lemma koszulSignInsert_eq_superCommuteCoef{I : Type} (q : I → Fin 2) (le1 : I → I → Prop) (r : List I)
(a : I) [DecidableRel le1] (i : I) :
koszulSignInsert le1 q a r = superCommuteCoef q [i]
(List.take (↑((HepLean.List.insertionSortEquiv le1 (a :: r)) ⟨0, by simp⟩))
(List.orderedInsert le1 a (List.insertionSort le1 r))) := by
sorry
lemma superCommuteCoefLE_zero {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) (r : List I) lemma superCommuteCoefLE_zero {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) (r : List I)
(a : I) (a : I)
[DecidableRel le1] (i : I) : [DecidableRel le1] (i : I) :
@ -146,9 +164,165 @@ lemma superCommuteCoefLE_zero {I : Type} (q : I → Fin 2) (le1 : I → I → Pr
simp simp
sorry sorry
lemma superCommuteCoefLE_eq_get_boson {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
[DecidableRel le1] (i : I) (hi : q i = 0 ) (n : Fin r.length)
(heq : q i = q (r.get n)) :
superCommuteCoefLE q le1 r i n = superCommuteCoef q [r.get n] (r.take n) := by
simp [superCommuteCoefLE, superCommuteCoef]
simp [grade, hi]
simp [hi] at heq
simp [heq]
rw [koszulSign_erase_boson]
rw [koszulSign_mul_self]
exact id (Eq.symm heq)
lemma koszulSignInsert_eq_filter {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1] (r0 : I)
: (r : List I) →
koszulSignInsert le1 q r0 r = koszulSignInsert le1 q r0 (List.filter (fun i => decide (¬ le1 r0 i)) r)
| [] => by
simp [koszulSignInsert]
| r1 :: r => by
dsimp [koszulSignInsert]
simp [List.filter]
by_cases h : le1 r0 r1
· simp [h]
rw [koszulSignInsert_eq_filter]
congr
simp
· simp [h]
dsimp [koszulSignInsert]
simp [h]
rw [koszulSignInsert_eq_filter]
congr
simp
simp
lemma filter_le_orderedInsert {I : Type} (le1 : I → I → Prop)
(le_trans : ∀ {i j k}, le1 i j → le1 j k → le1 i k)
(le_connected : ∀ {i j}, ¬ le1 i j → le1 j i) [DecidableRel le1] (rn r0 : I)
(r : List I) (hr : (List.filter (fun i => decide (¬ le1 rn i)) r) =
List.takeWhile (fun i => decide (¬ le1 rn i)) r) :
(List.filter (fun i => decide (¬ le1 rn i)) (List.orderedInsert le1 r0 r)) =
List.takeWhile (fun i => decide (¬ le1 rn i)) (List.orderedInsert le1 r0 r) := by
induction r with
| nil =>
simp [List.filter, List.takeWhile]
| cons r1 r ih =>
simp
by_cases hr1 : le1 r0 r1
· simp [hr1]
rw [List.filter, List.takeWhile]
by_cases hrn : le1 rn r0
· simp [hrn]
apply And.intro
· exact le_trans hrn hr1
· intro rp hr'
have hrn1 := le_trans hrn hr1
simp [List.filter, List.takeWhile, hrn1] at hr
exact hr rp hr'
· simp [hrn]
simpa using hr
· simp [hr1, List.filter, List.takeWhile]
by_cases hrn : le1 rn r1
· simp [hrn]
apply And.intro
· exact le_trans hrn (le_connected hr1)
· simp [List.filter, List.takeWhile, hrn] at hr
exact hr
· simp [hrn]
have hr' : List.filter (fun i => decide ¬le1 rn i) r = List.takeWhile (fun i => decide ¬le1 rn i) r := by
simp [List.filter, List.takeWhile, hrn] at hr
simpa using hr
simpa only [decide_not] using ih hr'
lemma filter_le_sort {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(le_trans : ∀ {i j k}, le1 i j → le1 j k → le1 i k)
(le_connected : ∀ {i j}, ¬ le1 i j → le1 j i) (r0 : I) (r : List I) :
List.filter (fun i => decide (¬ le1 r0 i)) (List.insertionSort le1 r) =
List.takeWhile (fun i => decide (¬ le1 r0 i)) (List.insertionSort le1 r) := by
induction r with
| nil =>
simp
| cons r1 r ih =>
simp only [ List.insertionSort]
rw [filter_le_orderedInsert]
exact fun {i j k} a a_1 => le_trans a a_1
exact fun {i j} a => le_connected a
exact ih
lemma koszulSignInsert_eq_grade {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1] (r0 : I)
(r : List I) :
koszulSignInsert le1 q r0 r = if grade q [r0] = 1 ∧
grade q (List.filter (fun i => decide (¬ le1 r0 i)) r) = 1 then -1 else 1 := by
induction r with
| nil =>
simp [koszulSignInsert]
| cons r1 r ih =>
rw [koszulSignInsert_eq_filter]
by_cases hr1 : ¬ le1 r0 r1
· rw [List.filter_cons_of_pos]
· dsimp [koszulSignInsert]
rw [if_neg hr1]
dsimp [grade]
simp [grade]
have ha (a b c : Fin 2) : (if a = 1 ∧ b = 1 then -if ¬a = 0 ∧
c = 1 then -1 else (1 : )
else if ¬a = 0 ∧ c = 1 then -1 else 1) =
if ¬a = 0 ∧ ¬b = c then -1 else 1:= by
fin_cases a <;> fin_cases b <;> fin_cases c
any_goals rfl
simp
rw [← ha (q r0) (q r1) (grade q (List.filter (fun a => !decide (le1 r0 a)) r) )]
congr
· rw [koszulSignInsert_eq_filter] at ih
simpa [grade] using ih
· rw [koszulSignInsert_eq_filter] at ih
simpa [grade] using ih
· simp [hr1]
· rw [List.filter_cons_of_neg]
simp
rw [koszulSignInsert_eq_filter] at ih
simpa [grade] using ih
simpa using hr1
lemma koszulSign_erase_fermion {I : Type} (q : I → Fin 2) (le1 :I → I → Prop)
[DecidableRel le1] :
(r : List I) → (n : Fin r.length) → (heq : q (r.get n) = 1) →
koszulSign le1 q (r.eraseIdx n) = koszulSign le1 q r *
if grade q (r.take n) = grade q (List.take (↑((HepLean.List.insertionSortEquiv le1 r) n)) (List.insertionSort le1 r))
then 1 else -1
| [], _ => by
simp
| r0 :: r, ⟨0, h⟩ => by
simp [koszulSign]
intro h
rw [koszulSignInsert_boson]
simp
sorry
sorry
| r0 :: r, ⟨n + 1, h⟩ => by
intro hn
simp only [List.eraseIdx_cons_succ, List.take_succ_cons, List.insertionSort, List.length_cons,
mul_one, mul_neg]
sorry
lemma superCommuteCoefLE_eq_get_fermion {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
[DecidableRel le1] (i : I) (hi : q i = 1 ) (n : Fin r.length)
(heq : q i = q (r.get n)) :
superCommuteCoefLE q le1 r i n = superCommuteCoef q [r.get n] (r.take n) := by
simp [superCommuteCoefLE, superCommuteCoef]
simp [grade, hi]
simp [hi] at heq
simp [← heq]
sorry
lemma superCommuteCoefLE_eq_get {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I) 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) : [DecidableRel le1] (i : I) (n : Fin r.length) (heq : q i = q (r.get n)) :
superCommuteCoefLE q le1 r i n = superCommuteCoef q [r.get n] (r.take n) := by superCommuteCoefLE q le1 r i n = superCommuteCoef q [r.get n] (r.take n) := by
sorry sorry
end Wick end Wick

View file

@ -245,6 +245,7 @@ def listMEraseEquiv {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
simp simp
congr 1 congr 1
simp [Fin.succAbove] simp [Fin.succAbove]
sorry sorry
)) ))

View file

@ -103,7 +103,6 @@ lemma superCommuteTakeM_F {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
congr congr
exact Eq.symm (List.eraseIdx_eq_take_drop_succ r n) exact Eq.symm (List.eraseIdx_eq_take_drop_succ r n)
lemma superCommute_koszulOrder_le_ofList {I : Type} lemma superCommute_koszulOrder_le_ofList {I : Type}
(q : I → Fin 2) (r : List I) (x : ) (q : I → Fin 2) (r : List I) (x : )
(le1 :I → I → Prop) [DecidableRel le1] (le1 :I → I → Prop) [DecidableRel le1]

View file

@ -107,6 +107,45 @@ lemma koszulSign_freeMonoid_of {I : Type} (r : I → I → Prop) [DecidableRel r
simp only [koszulSign, mul_one] simp only [koszulSign, mul_one]
rfl rfl
lemma koszulSignInsert_erase_boson {I : Type} (q : I → Fin 2) (le1 :I → I → Prop)
[DecidableRel le1] (r0 : I) :
(r : List I) → (n : Fin r.length) → (heq : q (r.get n) = 0) →
koszulSignInsert le1 q r0 (r.eraseIdx n) = koszulSignInsert le1 q r0 r
| [], _, _ => by
simp
| r1 :: r, ⟨0, h⟩, hr => by
simp
simp at hr
rw [koszulSignInsert]
simp [hr]
| r1 :: r, ⟨n + 1, h⟩, hr => by
simp
rw [koszulSignInsert, koszulSignInsert]
rw [koszulSignInsert_erase_boson q le1 r0 r ⟨n, Nat.succ_lt_succ_iff.mp h⟩ hr]
lemma koszulSign_erase_boson {I : Type} (q : I → Fin 2) (le1 :I → I → Prop)
[DecidableRel le1] :
(r : List I) → (n : Fin r.length) → (heq : q (r.get n) = 0) →
koszulSign le1 q (r.eraseIdx n) = koszulSign le1 q r
| [], _ => by
simp
| r0 :: r, ⟨0, h⟩ => by
simp [koszulSign]
intro h
rw [koszulSignInsert_boson]
simp
exact h
| r0 :: r, ⟨n + 1, h⟩ => by
simp
intro h'
rw [koszulSign, koszulSign]
rw [koszulSign_erase_boson q le1 r ⟨n, Nat.succ_lt_succ_iff.mp h⟩]
congr 1
rw [koszulSignInsert_erase_boson q le1 r0 r ⟨n, Nat.succ_lt_succ_iff.mp h⟩ h']
exact h'
noncomputable section noncomputable section
/-- Given a relation `r` on `I` sorts elements of `MonoidAlgebra (FreeMonoid I)` by `r` giving it /-- Given a relation `r` on `I` sorts elements of `MonoidAlgebra (FreeMonoid I)` by `r` giving it