feat: Fill in sorries

This commit is contained in:
jstoobysmith 2024-12-19 09:48:35 +00:00
parent 3123e831d8
commit ab7da149c6
5 changed files with 1432 additions and 454 deletions

View file

@ -17,6 +17,110 @@ open Fin
open HepLean
variable {n : Nat}
lemma takeWile_eraseIdx {I : Type} (P : I → Prop) [DecidablePred P] :
(l : List I) → (i : ) → (hi : ∀ (i j : Fin l.length), i < j → P (l.get j) → P (l.get i)) →
List.takeWhile P (List.eraseIdx l i) = (List.takeWhile P l).eraseIdx i
| [], _, h => by
simp
| a :: [], 0, h => by
simp [List.takeWhile]
split
next x heq => simp_all only [decide_eq_true_eq, List.tail_cons]
next x heq => simp_all only [decide_eq_false_iff_not, List.tail_nil]
| a :: [], Nat.succ n , h => by
simp
rw [List.eraseIdx_of_length_le ]
have h1 : (List.takeWhile P [a]).length ≤ [a].length :=
List.Sublist.length_le (List.takeWhile_sublist _)
simp at h1
omega
| a :: b :: l, 0, h => by
simp [List.takeWhile]
by_cases hPb : P b
· have hPa : P a := by
simpa using h ⟨0, by simp⟩ ⟨1, by simp⟩ (by simp [Fin.lt_def]) (by simpa using hPb)
simp [hPb, hPa]
· simp [hPb]
simp_all only [List.length_cons, List.get_eq_getElem]
split
next x heq => simp_all only [decide_eq_true_eq, List.tail_cons]
next x heq => simp_all only [decide_eq_false_iff_not, List.tail_nil]
| a :: b :: l, Nat.succ n, h => by
simp
by_cases hPa : P a
· dsimp [List.takeWhile]
simp [hPa]
rw [takeWile_eraseIdx]
rfl
intro i j hij hP
simpa using h (Fin.succ i) (Fin.succ j) (by simpa using hij) (by simpa using hP)
· simp [hPa]
lemma dropWile_eraseIdx {I : Type} (P : I → Prop) [DecidablePred P] :
(l : List I) → (i : ) → (hi : ∀ (i j : Fin l.length), i < j → P (l.get j) → P (l.get i)) →
List.dropWhile P (List.eraseIdx l i) =
if (List.takeWhile P l).length ≤ i then
(List.dropWhile P l).eraseIdx (i - (List.takeWhile P l).length)
else (List.dropWhile P l)
| [], _, h => by
simp
| a :: [], 0, h => by
simp [List.dropWhile]
simp_all only [List.length_singleton, List.get_eq_getElem, Fin.val_eq_zero, List.getElem_cons_zero, implies_true,
decide_True, decide_False, List.tail_cons, ite_self]
| a :: [], Nat.succ n , h => by
simp [List.dropWhile, List.takeWhile]
rw [List.eraseIdx_of_length_le ]
simp_all only [List.length_singleton, List.get_eq_getElem, Fin.val_eq_zero, List.getElem_cons_zero, implies_true,
ite_self]
simp_all only [List.length_singleton, List.get_eq_getElem, Fin.val_eq_zero, List.getElem_cons_zero, implies_true]
split
next x heq =>
simp_all only [decide_eq_true_eq, List.length_nil, List.length_singleton, add_tsub_cancel_right, zero_le]
next x heq =>
simp_all only [decide_eq_false_iff_not, List.length_singleton, List.length_nil, tsub_zero,
le_add_iff_nonneg_left, zero_le]
| a :: b :: l, 0, h => by
simp [List.takeWhile, List.dropWhile]
by_cases hPb : P b
· have hPa : P a := by
simpa using h ⟨0, by simp⟩ ⟨1, by simp⟩ (by simp [Fin.lt_def]) (by simpa using hPb)
simp [hPb, hPa]
· simp [hPb]
simp_all only [List.length_cons, List.get_eq_getElem]
simp_all only [decide_False, nonpos_iff_eq_zero, List.length_eq_zero]
split
next h_1 =>
simp_all only [nonpos_iff_eq_zero, List.length_eq_zero]
split
next x heq => simp_all only [List.cons_ne_self]
next x heq => simp_all only [decide_eq_false_iff_not, List.tail_cons]
next h_1 =>
simp_all only [nonpos_iff_eq_zero, List.length_eq_zero]
split
next x heq => simp_all only [List.cons_ne_self, not_false_eq_true, decide_eq_true_eq]
next x heq => simp_all only [not_true_eq_false]
| a :: b :: l, Nat.succ n, h => by
simp
by_cases hPb : P b
· have hPa : P a := by
simpa using h ⟨0, by simp⟩ ⟨1, by simp⟩ (by simp [Fin.lt_def]) (by simpa using hPb)
simp [List.takeWhile, List.dropWhile, hPb, hPa]
rw [dropWile_eraseIdx]
simp_all only [List.length_cons, List.get_eq_getElem, decide_True, List.takeWhile_cons_of_pos,
List.dropWhile_cons_of_pos]
intro i j hij hP
simpa using h (Fin.succ i) (Fin.succ j) (by simpa using hij) (by simpa using hP)
· simp [List.takeWhile, List.dropWhile, hPb]
by_cases hPa : P a
· rw [dropWile_eraseIdx]
simp [hPa, List.dropWhile, hPb]
intro i j hij hP
simpa using h (Fin.succ i) (Fin.succ j) (by simpa using hij) (by simpa using hP)
· simp [hPa]
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
@ -40,11 +144,50 @@ lemma orderedInsert_get_orderedInsertPos {I : Type} (le1 : I → I → Prop) [De
rw [List.getElem_append]
simp
@[simp]
lemma orderedInsert_eraseIdx_orderedInsertPos {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) :
(List.orderedInsert le1 r0 r).eraseIdx ↑(orderedInsertPos le1 r r0) = r := by
simp only [List.orderedInsert_eq_take_drop]
rw [List.eraseIdx_append_of_length_le]
· simp [orderedInsertPos]
· simp [orderedInsertPos]
lemma orderedInsertPos_cons {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 r1 : I) :
(orderedInsertPos le1 (r1 ::r) r0).val =
if le1 r0 r1 then ⟨0, by simp⟩ else (Fin.succ (orderedInsertPos le1 r r0)) := by
simp [orderedInsertPos, List.takeWhile]
by_cases h : le1 r0 r1
· simp [h]
· simp [h]
lemma orderedInsertPos_sigma {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(le1 : I → I → Prop) [DecidableRel le1] (l : List (Σ i, f i))
(k : I) (a : f k) :
(orderedInsertPos (fun (i j : Σ i, f i) => le1 i.1 j.1) l ⟨k, a⟩).1 =
(orderedInsertPos le1 (List.map (fun (i : Σ i, f i) => i.1) l) k).1 := by
simp [orderedInsertPos]
induction l with
| nil =>
simp
| cons a l ih =>
simp [List.takeWhile]
obtain ⟨fst, snd⟩ := a
simp_all only
split
next x heq =>
simp_all only [Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not, List.length_cons, decide_False,
Bool.not_false]
next x heq =>
simp_all only [Bool.not_eq_eq_eq_not, Bool.not_false, decide_eq_true_eq, List.length_nil, decide_True,
Bool.not_true]
@[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) :
(r : List I) (r0 : I) (i : )
(hi : i < 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 :=
@ -57,6 +200,137 @@ lemma orderedInsert_get_lt {I : Type} (le1 : I → I → Prop) [DecidableRel le1
rw [List.IsPrefix.getElem]
exact List.takeWhile_prefix fun b => !decide (le1 r0 b)
lemma orderedInsertPos_take_orderedInsert {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) :
(List.take (orderedInsertPos le1 r r0) (List.orderedInsert le1 r0 r)) =
List.takeWhile (fun b => decide ¬le1 r0 b) r := by
simp [orderedInsertPos, List.orderedInsert_eq_take_drop]
lemma orderedInsertPos_take_eq_orderedInsert {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) :
List.take (orderedInsertPos le1 r r0) r =
List.take (orderedInsertPos le1 r r0) (List.orderedInsert le1 r0 r) := by
refine List.ext_get ?_ ?_
· simp
exact Nat.le_of_lt_succ (orderedInsertPos_lt_length le1 r r0 )
· intro n h1 h2
simp
erw [orderedInsert_get_lt le1 r r0 n]
rfl
simp at h1
exact h1.1
lemma orderedInsertPos_drop_eq_orderedInsert {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) :
List.drop (orderedInsertPos le1 r r0) r =
List.drop (orderedInsertPos le1 r r0).succ (List.orderedInsert le1 r0 r) := by
conv_rhs => simp [orderedInsertPos, List.orderedInsert_eq_take_drop]
have hr : r = List.takeWhile (fun b => !decide (le1 r0 b)) r ++ List.dropWhile (fun b => !decide (le1 r0 b)) r := by
exact Eq.symm (List.takeWhile_append_dropWhile (fun b => !decide (le1 r0 b)) r)
conv_lhs =>
rhs
rw [hr]
rw [List.drop_append_eq_append_drop]
simp [orderedInsertPos]
lemma orderedInsertPos_take {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) :
List.take (orderedInsertPos le1 r r0) r = List.takeWhile (fun b => decide ¬le1 r0 b) r := by
rw [orderedInsertPos_take_eq_orderedInsert]
rw [orderedInsertPos_take_orderedInsert]
lemma orderedInsertPos_drop {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) :
List.drop (orderedInsertPos le1 r r0) r = List.dropWhile (fun b => decide ¬le1 r0 b) r := by
rw [orderedInsertPos_drop_eq_orderedInsert]
simp [orderedInsertPos, List.orderedInsert_eq_take_drop]
lemma orderedInsertPos_succ_take_orderedInsert {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) :
(List.take (orderedInsertPos le1 r r0).succ (List.orderedInsert le1 r0 r)) =
List.takeWhile (fun b => decide ¬le1 r0 b) r ++ [r0] := by
simp [orderedInsertPos, List.orderedInsert_eq_take_drop, List.take_append_eq_append_take]
lemma lt_orderedInsertPos_rel {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r0 : I) (r : List I) (n : Fin r.length)
(hn : n.val < (orderedInsertPos le1 r r0).val) : ¬ le1 r0 (r.get n) := by
have htake : r.get n ∈ List.take (orderedInsertPos le1 r r0) r := by
rw [@List.mem_take_iff_getElem]
use n
simp
exact hn
rw [orderedInsertPos_take] at htake
have htake' := List.mem_takeWhile_imp htake
simpa using htake'
lemma gt_orderedInsertPos_rel {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
[IsTotal I le1] [IsTrans I le1] (r0 : I) (r : List I) (hs : List.Sorted le1 r) (n : Fin r.length)
(hn : ¬ n.val < (orderedInsertPos le1 r r0).val) : le1 r0 (r.get n) := by
have hrsSorted : List.Sorted le1 (List.orderedInsert le1 r0 r) :=
List.Sorted.orderedInsert r0 r hs
apply List.Sorted.rel_of_mem_take_of_mem_drop (k := (orderedInsertPos le1 r r0).succ) hrsSorted
· rw [orderedInsertPos_succ_take_orderedInsert]
simp
· rw [← orderedInsertPos_drop_eq_orderedInsert]
refine List.mem_drop_iff_getElem.mpr ?hy.a
use n - (orderedInsertPos le1 r r0).val
have hn : ↑n - ↑(orderedInsertPos le1 r r0) + ↑(orderedInsertPos le1 r r0) < r.length := by
omega
use hn
congr
omega
lemma orderedInsert_eraseIdx_lt_orderedInsertPos {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) (i : )
(hi : i < orderedInsertPos le1 r r0)
(hr : ∀ (i j : Fin r.length), i < j → ¬le1 r0 (r.get j) → ¬le1 r0 (r.get i)) :
(List.orderedInsert le1 r0 r).eraseIdx i = List.orderedInsert le1 r0 (r.eraseIdx i) := by
conv_lhs => simp only [List.orderedInsert_eq_take_drop]
rw [List.eraseIdx_append_of_lt_length]
· simp only [List.orderedInsert_eq_take_drop]
congr 1
· rw [takeWile_eraseIdx]
exact hr
· rw [dropWile_eraseIdx]
simp [orderedInsertPos] at hi
have hi' : ¬ (List.takeWhile (fun b => !decide (le1 r0 b)) r).length ≤ ↑i := by
omega
simp [hi']
exact fun i j a a_1 => hr i j a a_1
· exact hi
lemma orderedInsert_eraseIdx_orderedInsertPos_le {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) (i : )
(hi : orderedInsertPos le1 r r0 ≤ i)
(hr : ∀ (i j : Fin r.length), i < j → ¬le1 r0 (r.get j) → ¬le1 r0 (r.get i)) :
(List.orderedInsert le1 r0 r).eraseIdx (Nat.succ i) = List.orderedInsert le1 r0 (r.eraseIdx i) := by
conv_lhs => simp only [List.orderedInsert_eq_take_drop]
rw [List.eraseIdx_append_of_length_le]
· simp only [List.orderedInsert_eq_take_drop]
congr 1
· rw [takeWile_eraseIdx]
rw [List.eraseIdx_of_length_le]
simp [orderedInsertPos] at hi
simp
omega
exact hr
· simp only [Nat.succ_eq_add_one]
have hn : (i + 1 - (List.takeWhile (fun b => (decide (¬ le1 r0 b))) r).length)
= (i - (List.takeWhile (fun b => decide (¬ le1 r0 b)) r).length) + 1 := by
simp only [orderedInsertPos] at hi
omega
rw [hn]
simp only [List.eraseIdx_cons_succ, List.cons.injEq, true_and]
rw [dropWile_eraseIdx]
rw [if_pos]
· simp only [orderedInsertPos] at hi
omega
· exact hr
· simp only [orderedInsertPos] at hi
omega
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 :=
@ -85,6 +359,26 @@ lemma orderedInsertEquiv_succ {I : Type} (le1 : I → I → Prop) [DecidableRel
rfl
exact ne_of_beq_false rfl
@[simp]
lemma orderedInsertEquiv_fin_succ {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(r0 : I) (n : Fin r.length) :
orderedInsertEquiv le1 r r0 n.succ = Fin.cast (List.orderedInsert_length le1 r r0).symm
((Fin.succAbove ⟨(orderedInsertPos le1 r r0), orderedInsertPos_lt_length le1 r r0⟩) ⟨n, n.isLt⟩) := by
simp [orderedInsertEquiv]
match r with
| [] =>
simp
| r1 :: r =>
erw [finExtractOne_apply_neq]
simp [orderedInsertPos]
rfl
exact ne_of_beq_false rfl
lemma orderedInsertEquiv_congr {α : Type} {r : αα → Prop} [DecidableRel r] (a : α) (l l' : List α)
(h : l = l') : orderedInsertEquiv r l a = (Fin.castOrderIso (by simp [h])).toEquiv.trans
((orderedInsertEquiv r l' a).trans (Fin.castOrderIso (by simp [h])).toEquiv) := by
subst h
rfl
lemma get_eq_orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(r0 : I) :
@ -101,9 +395,6 @@ lemma get_eq_orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRe
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]
@ -126,97 +417,119 @@ lemma get_eq_orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRe
simp [hr]
omega
lemma orderedInsertEquiv_get {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(r0 : I) :
(r0 :: r).get ∘ (orderedInsertEquiv le1 r r0).symm = (List.orderedInsert le1 r0 r).get := by
rw [get_eq_orderedInsertEquiv le1]
funext x
simp
/-- 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. -/
def insertEquiv {α : Type} (r : αα → Prop) [DecidableRel r] (a : α) : (l : List α) →
Fin (a :: l).length ≃ Fin (List.orderedInsert r a l).length
| [] => Equiv.refl _
| b :: l => by
if r a b then
exact (Fin.castOrderIso (List.orderedInsert_length r (b :: l) a).symm).toEquiv
else
let e := insertEquiv (r := r) a l
let e2 : Fin (a :: b :: l).length ≃ Fin (b :: a :: l).length :=
Equiv.swap ⟨0, Nat.zero_lt_succ (b :: l).length⟩ ⟨1, Nat.one_lt_succ_succ l.length⟩
let e3 : Fin (b :: a :: l).length ≃ Fin (b :: List.orderedInsert r a l).length :=
Fin.equivCons e
let e4 : Fin (b :: List.orderedInsert r a l).length ≃
Fin (List.orderedInsert r a (b :: l)).length :=
(Fin.castOrderIso (by
rw [List.orderedInsert_length]
simpa using List.orderedInsert_length r l a)).toEquiv
exact e2.trans (e3.trans e4)
lemma insertEquiv_congr {α : Type} {r : αα → Prop} [DecidableRel r] (a : α) (l l' : List α)
(h : l = l') : insertEquiv r a l = (Fin.castOrderIso (by simp [h])).toEquiv.trans
((insertEquiv r a l').trans (Fin.castOrderIso (by simp [h])).toEquiv) := by
subst h
rfl
lemma orderedInsert_eraseIdx_orderedInsertEquiv_zero
{I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I) (r0 : I) :
(List.orderedInsert le1 r0 r).eraseIdx (orderedInsertEquiv le1 r r0 ⟨0, by simp⟩) = r := by
simp [orderedInsertEquiv]
lemma insertEquiv_cons_pos {α : Type} {r : αα → Prop} [DecidableRel r] (a b : α) (hab : r a b)
(l : List α) : insertEquiv r a (b :: l) =
(Fin.castOrderIso (List.orderedInsert_length r (b :: l) a).symm).toEquiv := by
simp [insertEquiv, hab]
lemma orderedInsert_eraseIdx_orderedInsertEquiv_succ
{I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I) (r0 : I) (n : ) (hn : Nat.succ n < (r0 :: r).length)
(hr : ∀ (i j : Fin r.length), i < j → ¬le1 r0 (r.get j) → ¬le1 r0 (r.get i)) :
(List.orderedInsert le1 r0 r).eraseIdx (orderedInsertEquiv le1 r r0 ⟨Nat.succ n, hn⟩) =
(List.orderedInsert le1 r0 (r.eraseIdx n)) := by
induction r with
| nil =>
simp at hn
| cons r1 r ih =>
rw [orderedInsertEquiv_succ]
simp only [List.length_cons, Fin.succAbove,
Fin.castSucc_mk, Fin.mk_lt_mk, Fin.succ_mk, Fin.coe_cast]
by_cases hn' : n < (orderedInsertPos le1 (r1 :: r) r0)
· simp only [hn', ↓reduceIte]
rw [orderedInsert_eraseIdx_lt_orderedInsertPos le1 (r1 :: r) r0 n hn' hr]
· simp only [hn', ↓reduceIte]
rw [orderedInsert_eraseIdx_orderedInsertPos_le le1 (r1 :: r) r0 n _ hr]
omega
lemma insertEquiv_cons_neg {α : Type} {r : αα → Prop} [DecidableRel r] (a b : α) (hab : ¬ r a b)
(l : List α) : insertEquiv r a (b :: l) =
let e := insertEquiv r a l
let e2 : Fin (a :: b :: l).length ≃ Fin (b :: a :: l).length :=
Equiv.swap ⟨0, Nat.zero_lt_succ (b :: l).length⟩ ⟨1, Nat.one_lt_succ_succ l.length⟩
let e3 : Fin (b :: a :: l).length ≃ Fin (b :: List.orderedInsert r a l).length :=
Fin.equivCons e
let e4 : Fin (b :: List.orderedInsert r a l).length ≃
Fin (List.orderedInsert r a (b :: l)).length :=
(Fin.castOrderIso (by
rw [List.orderedInsert_length]
simpa using List.orderedInsert_length r l a)).toEquiv
e2.trans (e3.trans e4) := by
simp [insertEquiv, hab]
lemma orderedInsert_eraseIdx_orderedInsertEquiv_fin_succ
{I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I) (r0 : I) (n : Fin r.length)
(hr : ∀ (i j : Fin r.length), i < j → ¬le1 r0 (r.get j) → ¬le1 r0 (r.get i)) :
(List.orderedInsert le1 r0 r).eraseIdx (orderedInsertEquiv le1 r r0 n.succ) =
(List.orderedInsert le1 r0 (r.eraseIdx n)) := by
have hn : n.succ = ⟨n.val + 1, by omega⟩ := by
rw [Fin.ext_iff]
simp
rw [hn]
exact orderedInsert_eraseIdx_orderedInsertEquiv_succ le1 r r0 n.val _ hr
lemma insertEquiv_get {α : Type} {r : αα → Prop} [DecidableRel r] (a : α) : (l : List α) →
(a :: l).get ∘ (insertEquiv r a l).symm = (List.orderedInsert r a l).get
| [] => by
simp [insertEquiv]
| b :: l => by
by_cases hr : r a b
· rw [insertEquiv_cons_pos a b hr l]
simp_all only [List.orderedInsert.eq_2, List.length_cons, OrderIso.toEquiv_symm,
Fin.symm_castOrderIso, RelIso.coe_fn_toEquiv]
ext x : 1
simp_all only [Function.comp_apply, Fin.castOrderIso_apply, List.get_eq_getElem,
List.length_cons, Fin.coe_cast, ↓reduceIte]
· rw [insertEquiv_cons_neg a b hr l]
trans (b :: List.orderedInsert r a l).get ∘ Fin.cast (by
rw [List.orderedInsert_length]
simp [List.orderedInsert_length])
· simp only [List.orderedInsert.eq_2, List.length_cons, Fin.zero_eta, Fin.mk_one]
ext x
match x with
| ⟨0, h⟩ => rfl
| ⟨Nat.succ x, h⟩ =>
simp only [Nat.succ_eq_add_one, Function.comp_apply, Equiv.symm_trans_apply,
Equiv.symm_swap, OrderIso.toEquiv_symm, Fin.symm_castOrderIso, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, Fin.cast_mk, equivCons_symm_succ, List.get_eq_getElem,
List.length_cons, List.getElem_cons_succ]
have hswap (n : Fin (b :: a :: l).length) :
(a :: b :: l).get (Equiv.swap ⟨0, by simp⟩ ⟨1, by simp⟩ n) = (b :: a :: l).get n := by
match n with
| ⟨0, h⟩ => rfl
| ⟨1, h⟩ => rfl
| ⟨Nat.succ (Nat.succ x), h⟩ => rfl
trans (a :: b :: l).get (Equiv.swap ⟨0, by simp⟩ ⟨1, by simp⟩
((insertEquiv r a l).symm ⟨x, by simpa [List.orderedInsert_length, hr] using h⟩).succ)
· simp
· rw [hswap]
simp only [List.length_cons, List.get_eq_getElem, Fin.val_succ, List.getElem_cons_succ]
change _ = (List.orderedInsert r a l).get _
rw [← insertEquiv_get (r := r) a l]
simp
· simp_all only [List.orderedInsert.eq_2, List.length_cons]
ext x : 1
simp_all only [Function.comp_apply, List.get_eq_getElem, List.length_cons, Fin.coe_cast,
↓reduceIte]
lemma orderedInsertEquiv_sigma {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(le1 : I → I → Prop) [DecidableRel le1] (l : List (Σ i, f i))
(i : I) (a : f i) :
(orderedInsertEquiv (fun i j => le1 i.fst j.fst) l ⟨i, a⟩) =
(Fin.castOrderIso (by simp)).toEquiv.trans
((orderedInsertEquiv le1 (List.map (fun i => i.1) l) i).trans
(Fin.castOrderIso (by simp [List.orderedInsert_length])).toEquiv) := by
ext x
match x with
| ⟨0, h0⟩ =>
simp
erw [orderedInsertEquiv_zero, orderedInsertEquiv_zero]
simp [orderedInsertPos_sigma]
| ⟨Nat.succ n, h0⟩ =>
simp
erw [orderedInsertEquiv_succ, orderedInsertEquiv_succ]
simp [orderedInsertPos_sigma]
rw [Fin.succAbove, Fin.succAbove]
simp
split
next h => simp_all only
next h => simp_all only [not_lt]
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
apply List.ext_get
· simp [List.orderedInsert_length]
rw [List.length_insertIdx]
have h1 := orderedInsertPos_lt_length le1 r r0
simp at h1
omega
intro n h1 h2
obtain ⟨n', hn'⟩ := (orderedInsertEquiv le1 r r0).surjective ⟨n, h1⟩
rw [← hn']
have hn'' : n = ((orderedInsertEquiv le1 r r0) n').val := by rw [hn']
subst hn''
rw [← orderedInsertEquiv_get]
simp
match n' with
| ⟨0, h0⟩ =>
simp
simp [orderedInsertEquiv]
rw [List.getElem_insertIdx_self]
exact Nat.le_of_lt_succ (orderedInsertPos_lt_length le1 r r0)
| ⟨Nat.succ n', h0⟩ =>
simp
have hr := orderedInsertEquiv_succ le1 r r0 n' h0
trans (List.insertIdx (↑(orderedInsertPos le1 r r0)) r0 r).get ⟨↑((orderedInsertEquiv le1 r r0) ⟨n' +1, h0⟩), h2⟩
swap
rfl
rw [Fin.ext_iff] at hr
have hx : (⟨↑((orderedInsertEquiv le1 r r0) ⟨n' +1, h0⟩), h2⟩ :Fin (List.insertIdx (↑(orderedInsertPos le1 r r0)) r0 r).length) =
⟨(
(⟨↑(orderedInsertPos le1 r r0), orderedInsertPos_lt_length le1 r r0⟩ : Fin ((r).length + 1))).succAbove ⟨n', Nat.succ_lt_succ_iff.mp h0⟩, by
erw [← hr]
exact h2
⟩ := by
rw [Fin.ext_iff]
simp
simpa using hr
rw [hx]
simp [Fin.succAbove]
by_cases hn' : n' < ↑(orderedInsertPos le1 r r0)
· simp [hn']
erw [List.getElem_insertIdx_of_lt]
exact hn'
· simp [hn']
sorry
/-- The equivalence between `Fin l.length ≃ Fin (List.insertionSort r l).length` induced by the
sorting algorithm. -/
@ -224,7 +537,7 @@ def insertionSortEquiv {α : Type} (r : αα → Prop) [DecidableRel r] : (
Fin l.length ≃ Fin (List.insertionSort r l).length
| [] => Equiv.refl _
| a :: l =>
(Fin.equivCons (insertionSortEquiv r l)).trans (insertEquiv r a (List.insertionSort r l))
(Fin.equivCons (insertionSortEquiv r l)).trans (orderedInsertEquiv r (List.insertionSort r l) a)
lemma insertionSortEquiv_get {α : Type} {r : αα → Prop} [DecidableRel r] : (l : List α) →
l.get ∘ (insertionSortEquiv r l).symm = (List.insertionSort r l).get
@ -233,7 +546,7 @@ lemma insertionSortEquiv_get {α : Type} {r : αα → Prop} [DecidableRel
| a :: l => by
rw [insertionSortEquiv]
change ((a :: l).get ∘ ((Fin.equivCons (insertionSortEquiv r l))).symm) ∘
(insertEquiv r a (List.insertionSort r l)).symm = _
(orderedInsertEquiv r (List.insertionSort r l) a).symm = _
have hl : (a :: l).get ∘ ((Fin.equivCons (insertionSortEquiv r l))).symm =
(a :: List.insertionSort r l).get := by
ext x
@ -244,9 +557,14 @@ lemma insertionSortEquiv_get {α : Type} {r : αα → Prop} [DecidableRel
rw [← insertionSortEquiv_get (r := r) l]
rfl
rw [hl]
rw [insertEquiv_get (r := r) a (List.insertionSort r l)]
rw [orderedInsertEquiv_get r (List.insertionSort r l) a]
rfl
lemma insertionSortEquiv_congr {α : Type} {r : αα → Prop} [DecidableRel r] (l l' : List α)
(h : l = l') : insertionSortEquiv r l = (Fin.castOrderIso (by simp [h])).toEquiv.trans
((insertionSortEquiv r l').trans (Fin.castOrderIso (by simp [h])).toEquiv) := by
subst h
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]
@ -298,91 +616,47 @@ lemma eraseIdx_get {I : Type} (l : List I) (i : Fin l.length) :
have hn : ¬ x.val < i.val := by omega
simp [hn]
lemma eraseIdx_insertionSort_length {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)).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 _ _
lemma eraseIdx_insertionSort {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
[IsTotal I le1] [IsTrans I le1] :
(n : ) → (r : List I) → (hn : n < r.length) →
(List.insertionSort le1 r).eraseIdx ↑((HepLean.List.insertionSortEquiv le1 r) ⟨n, hn⟩)
= List.insertionSort le1 (r.eraseIdx n)
| 0, [], _ => by
simp
| 0, (r0 :: r), hn => by
simp only [List.insertionSort, List.insertionSort.eq_2, List.length_cons, insertionSortEquiv,
Nat.succ_eq_add_one, Fin.zero_eta, Equiv.trans_apply, equivCons_zero, List.eraseIdx_zero,
List.tail_cons]
erw [orderedInsertEquiv_zero]
simp
| Nat.succ n, [], hn => by
simp [insertionSortEquiv]
| Nat.succ 0, (r0 :: r), hn => by
| Nat.succ n, (r0 :: r), hn => by
simp [insertionSortEquiv]
have hOr := orderedInsert_eraseIdx_orderedInsertEquiv_fin_succ le1
(List.insertionSort le1 r) r0 ((insertionSortEquiv le1 r) ⟨n, by simpa using hn⟩)
erw [hOr]
congr
refine eraseIdx_insertionSort le1 n r _
intro i j hij hn
have hx := List.Sorted.rel_get_of_lt (r := le1) (l := (List.insertionSort le1 r))
(List.sorted_insertionSort le1 r) hij
have hr : le1 ((List.insertionSort le1 r).get j) r0 := by
have hn := IsTotal.total (r := le1) ((List.insertionSort le1 r).get j) r0
simp_all only [List.get_eq_getElem, List.length_cons, or_false]
have ht (i j k : I) (hij : le1 i j) (hjk : ¬ le1 k j) : ¬ le1 k i := by
intro hik
have ht := IsTrans.trans (r := le1) k i j hik hij
exact hjk ht
exact ht ((List.insertionSort le1 r).get i) ((List.insertionSort le1 r).get j) r0 hx hn
lemma eraseIdx_insertionSort_fin {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
[IsTotal I le1] [IsTrans I le1] (r : List I) (n : Fin r.length) :
(List.insertionSort le1 r).eraseIdx ↑((HepLean.List.insertionSortEquiv le1 r) n)
= List.insertionSort le1 (r.eraseIdx n) :=
eraseIdx_insertionSort le1 n.val r (Fin.prop n)
| 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

View file

@ -20,6 +20,7 @@ import HepLean.PerturbationTheory.Wick.Koszul.Order
-/
namespace Wick
open HepLean.List
def grade {I : Type} (q : I → Fin 2) : (l : List I) → Fin 2
| [] => 0
@ -58,13 +59,13 @@ lemma grade_orderedInsert {I : Type} (q : I → Fin 2) (le1 : I → I → Prop)
induction l with
| nil => simp
| cons j l ih =>
simp
simp only [List.orderedInsert]
by_cases hij : le1 i j
· simp [hij]
· simp [hij]
· simp only [hij, ↓reduceIte]
rw [grade]
rw [ih]
simp [grade]
simp only [grade, Fin.isValue]
have h1 (a b c : Fin 2) : (if a = if b = c then 0 else 1 then (0 : Fin 2) else 1) = if b = if a = c then 0 else 1 then 0 else 1 := by
fin_cases a <;> fin_cases b <;> fin_cases c <;> rfl
exact h1 _ _ _
@ -75,14 +76,75 @@ lemma grade_insertionSort {I : Type} (q : I → Fin 2) (le1 : I → I → Prop)
induction l with
| nil => simp
| cons j l ih =>
simp [grade]
simp only [List.insertionSort, grade, Fin.isValue]
rw [grade_orderedInsert]
simp [grade]
simp only [grade, Fin.isValue]
rw [ih]
lemma grade_count {I : Type} (q : I → Fin 2) (l : List I) :
grade q l = if Nat.mod (List.countP (fun i => decide (q i = 1)) l) 2 = 0 then 0 else 1 := by
induction l with
| nil => simp
| cons r0 r ih =>
simp only [grade, Fin.isValue]
rw [List.countP_cons]
simp only [Fin.isValue, decide_eq_true_eq]
rw [ih]
by_cases h: q r0 = 1
· simp only [h, Fin.isValue, ↓reduceIte]
split
next h1 =>
simp_all only [Fin.isValue, ↓reduceIte, one_ne_zero]
split
next h2 =>
simp_all only [Fin.isValue, one_ne_zero]
have ha (a : ) (ha : a % 2 = 0) : (a + 1) % 2 ≠ 0 := by
omega
exact ha (List.countP (fun i => decide (q i = 1)) r) h1 h2
next h2 => simp_all only [Fin.isValue]
next h1 =>
simp_all only [Fin.isValue, ↓reduceIte]
split
next h2 => simp_all only [Fin.isValue]
next h2 =>
simp_all only [Fin.isValue, zero_ne_one]
have ha (a : ) (ha : ¬ a % 2 = 0) : (a + 1) % 2 = 0 := by
omega
exact h2 (ha (List.countP (fun i => decide (q i = 1)) r) h1)
· have h0 : q r0 = 0 := by omega
simp only [h0, Fin.isValue, zero_ne_one, ↓reduceIte, add_zero]
by_cases hn : (List.countP (fun i => decide (q i = 1)) r).mod 2 = 0
· simp [hn]
· simp [hn]
lemma grade_perm {I : Type} (q : I → Fin 2) {l l' : List I} (h : l.Perm l') :
grade q l = grade q l' := by
rw [grade_count, grade_count, h.countP_eq]
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_comm {I : Type} (q : I → Fin 2) (la lb : List I) :
superCommuteCoef q la lb = superCommuteCoef q lb la := by
simp only [superCommuteCoef, Fin.isValue]
congr 1
exact Eq.propIntro (fun a => id (And.symm a)) fun a => id (And.symm a)
lemma superCommuteCoef_perm_snd {I : Type} (q : I → Fin 2) (la lb lb' : List I)
(h : lb.Perm lb') :
superCommuteCoef q la lb = superCommuteCoef q la lb' := by
rw [superCommuteCoef, superCommuteCoef, grade_perm q h ]
lemma superCommuteCoef_mul_self {I : Type} (q : I → Fin 2) (l lb : List I) :
superCommuteCoef q l lb * superCommuteCoef q l lb = 1 := by
simp only [superCommuteCoef, Fin.isValue, mul_ite, mul_neg, mul_one]
have ha (a b : 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
any_goals rfl
simp
exact ha (grade q l) (grade q lb)
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]
@ -109,6 +171,12 @@ lemma superCommuteCoef_append {I : Type} (q : I → Fin 2) (la lb lc : List I)
omega
simp [ha]
lemma superCommuteCoef_cons {I : Type} (q : I → Fin 2) (i : I) (la lb : List I) :
superCommuteCoef q la (i :: lb) = superCommuteCoef q la [i] * superCommuteCoef q la lb := by
trans superCommuteCoef q la ([i] ++ lb)
simp only [List.singleton_append]
rw [superCommuteCoef_append]
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)
@ -118,13 +186,6 @@ lemma superCommuteCoefM_empty {I : Type} {f : I → Type} [∀ i, Fintype (f i)
superCommuteCoefM q l [] = 1 := by
simp [superCommuteCoefM]
lemma koszulSign_first_remove {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (l : List I)
[DecidableRel le1] (a : I) :
let n0 := ((HepLean.List.insertionSortEquiv le1 (a :: l)).symm ⟨0, by
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)
[DecidableRel le1] (i : I) (n : Fin r.length) : :=
koszulSign le1 q r *
@ -132,124 +193,37 @@ def superCommuteCoefLE {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r
(List.insertionSort le1 r)) *
koszulSign le1 q (r.eraseIdx ↑n)
lemma superCommuteCoefLE_eq_q {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
[DecidableRel le1] (i : I) (n : Fin r.length)
(hq : q i = q (r.get n)) :
superCommuteCoefLE q le1 r i n =
koszulSign le1 q r *
superCommuteCoef q [r.get n] (List.take (↑(insertionSortEquiv le1 r n))
(List.insertionSort le1 r)) *
koszulSign le1 q (r.eraseIdx ↑n) := by
simp [superCommuteCoefLE, superCommuteCoef, grade, hq]
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)
(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_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]
dsimp only [koszulSignInsert, Fin.isValue]
simp only [Fin.isValue, List.filter, decide_not]
by_cases h : le1 r0 r1
· simp [h]
· simp only [h, ↓reduceIte, decide_True, Bool.not_true]
rw [koszulSignInsert_eq_filter]
congr
simp
· simp [h]
dsimp [koszulSignInsert]
simp [h]
· simp only [h, ↓reduceIte, Fin.isValue, decide_False, Bool.not_false]
dsimp only [Fin.isValue, koszulSignInsert]
simp only [Fin.isValue, h, ↓reduceIte]
rw [koszulSignInsert_eq_filter]
congr
simp only [decide_not]
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) :
@ -262,10 +236,10 @@ lemma koszulSignInsert_eq_grade {I : Type} (q : I → Fin 2) (le1 : I → I →
rw [koszulSignInsert_eq_filter]
by_cases hr1 : ¬ le1 r0 r1
· rw [List.filter_cons_of_pos]
· dsimp [koszulSignInsert]
· dsimp only [koszulSignInsert, Fin.isValue, decide_not]
rw [if_neg hr1]
dsimp [grade]
simp [grade]
dsimp only [Fin.isValue, grade, ite_eq_right_iff, zero_ne_one, imp_false, decide_not]
simp only [Fin.isValue, decide_not, ite_eq_right_iff, zero_ne_one, imp_false]
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) =
@ -281,48 +255,392 @@ lemma koszulSignInsert_eq_grade {I : Type} (q : I → Fin 2) (le1 : I → I →
simpa [grade] using ih
· simp [hr1]
· rw [List.filter_cons_of_neg]
simp
simp only [decide_not, Fin.isValue]
rw [koszulSignInsert_eq_filter] at ih
simpa [grade] using ih
simpa using hr1
lemma koszulSignInsert_eq_perm {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) (r r' : List I)
(a : I) [DecidableRel le1] (h : r.Perm r') :
koszulSignInsert le1 q a r = koszulSignInsert le1 q a r' := by
rw [koszulSignInsert_eq_grade]
rw [koszulSignInsert_eq_grade]
congr 1
simp only [Fin.isValue, decide_not, eq_iff_iff, and_congr_right_iff]
intro h'
have hg : grade q (List.filter (fun i => !decide (le1 a i)) r) =
grade q (List.filter (fun i => !decide (le1 a i)) r') := by
rw [grade_count, grade_count]
rw [List.countP_filter, List.countP_filter]
rw [h.countP_eq]
rw [hg]
lemma koszulSignInsert_eq_sort {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) (r : List I)
(a : I) [DecidableRel le1] :
koszulSignInsert le1 q a r = koszulSignInsert le1 q a (List.insertionSort le1 r) := by
apply koszulSignInsert_eq_perm
exact List.Perm.symm (List.perm_insertionSort le1 r)
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
lemma koszulSignInsert_eq_cons {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
[IsTotal I le1] [IsTrans I le1] (r0 : I)
(r : List I) :
koszulSignInsert le1 q r0 r = koszulSignInsert le1 q r0 (r0 :: r):= by
simp only [koszulSignInsert, Fin.isValue, and_self]
have h1 : le1 r0 r0 := by
simpa using IsTotal.total (r := le1) r0 r0
simp [h1]
def insertSign {I : Type} (q : I → Fin 2) (n : )
(r0 : I) (r : List I) : :=
superCommuteCoef q [r0] (List.take n r)
lemma take_insert_same {I : Type} (i : I) :
(n : ) → (r : List I) →
List.take n (List.insertIdx n i r) = List.take n r
| 0, _ => by simp
| _+1, [] => by simp
| n+1, a::as => by
simp only [List.insertIdx_succ_cons, List.take_succ_cons, List.cons.injEq, true_and]
exact take_insert_same i n as
lemma insertSign_insert {I : Type} (q : I → Fin 2) (n : )
(r0 : I) (r : List I) : insertSign q n r0 r = insertSign q n r0 (List.insertIdx n r0 r) := by
simp only [insertSign]
congr 1
rw [take_insert_same]
lemma take_eraseIdx_same {I : Type} :
(n : ) → (r : List I) →
List.take n (List.eraseIdx r n) = List.take n r
| 0, _ => by simp
| _+1, [] => by simp
| n+1, a::as => by
simp only [List.eraseIdx_cons_succ, List.take_succ_cons, List.cons.injEq, true_and]
exact take_eraseIdx_same n as
lemma insertSign_eraseIdx {I : Type} (q : I → Fin 2) (n : )
(r0 : I) (r : List I) : insertSign q n r0 (r.eraseIdx n) = insertSign q n r0 r := by
simp only [insertSign]
congr 1
rw [take_eraseIdx_same]
lemma insertSign_zero {I : Type} (q : I → Fin 2) (r0 : I) (r : List I) :
insertSign q 0 r0 r = 1 := by
simp [insertSign, superCommuteCoef]
lemma insertSign_succ_cons {I : Type} (q : I → Fin 2) (n : )
(r0 r1 : I) (r : List I) : insertSign q (n + 1) r0 (r1 :: r) =
superCommuteCoef q [r0] [r1] * insertSign q n r0 r := by
simp only [insertSign, List.take_succ_cons]
rw [superCommuteCoef_cons]
lemma take_insert_gt {I : Type} (i : I) :
(n m : ) → (h : n < m) → (r : List I) →
List.take n (List.insertIdx m i r) = List.take n r
| 0, 0, _, _ => by simp
| 0, m + 1, _, _ => by simp
| n+1, m + 1, _, [] => by simp
| n+1, m + 1, h, a::as => by
simp only [List.insertIdx_succ_cons, List.take_succ_cons, List.cons.injEq, true_and]
refine take_insert_gt i n m (Nat.succ_lt_succ_iff.mp h) as
lemma insertSign_insert_gt {I : Type} (q : I → Fin 2) (n m : )
(r0 r1 : I) (r : List I) (hn : n < m) :
insertSign q n r0 (List.insertIdx m r1 r) = insertSign q n r0 r := by
rw [insertSign, insertSign]
congr 1
exact take_insert_gt r1 n m hn r
lemma take_insert_let {I : Type} (i : I) :
(n m : ) → (h : m ≤ n) → (r : List I) → (hm : m ≤ r.length) →
(List.take (n + 1) (List.insertIdx m i r)).Perm (i :: List.take n r)
| 0, 0, h, _, _ => by simp
| m + 1, 0, h, r, _ => by simp
| n + 1, m + 1, h, [], hm => by
simp at hm
| n + 1, m + 1, h, a::as, hm => by
simp only [List.insertIdx_succ_cons, List.take_succ_cons]
have hp : (i :: a :: List.take n as).Perm (a :: i :: List.take n as) := by
exact List.Perm.swap a i (List.take n as)
refine List.Perm.trans ?_ hp.symm
refine List.Perm.cons a ?_
exact take_insert_let i n m (Nat.le_of_succ_le_succ h) as (Nat.le_of_succ_le_succ hm)
lemma insertSign_insert_lt_eq_insertSort {I : Type} (q : I → Fin 2) (n m : )
(r0 r1 : I) (r : List I) (hn : m ≤ n) (hm : m ≤ r.length):
insertSign q (n + 1) r0 (List.insertIdx m r1 r) = insertSign q (n + 1) r0 (r1 :: r) := by
rw [insertSign, insertSign]
apply superCommuteCoef_perm_snd
simp only [List.take_succ_cons]
refine take_insert_let r1 n m hn r hm
lemma insertSign_insert_lt {I : Type} (q : I → Fin 2) (n m : )
(r0 r1 : I) (r : List I) (hn : m ≤ n) (hm : m ≤ r.length):
insertSign q (n + 1) r0 (List.insertIdx m r1 r) = superCommuteCoef q [r0] [r1] * insertSign q n r0 r := by
rw [insertSign_insert_lt_eq_insertSort, insertSign_succ_cons]
exact hn
exact hm
def koszulSignCons {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1] (r0 r1 : I) :
:=
if le1 r0 r1 then 1 else
if q r0 = 1 ∧ q r1 = 1 then -1 else 1
lemma koszulSignCons_eq_superComuteCoef {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
(r0 r1 : I) : koszulSignCons q le1 r0 r1 =
if le1 r0 r1 then 1 else superCommuteCoef q [r0] [r1] := by
simp only [koszulSignCons, Fin.isValue, superCommuteCoef, grade, ite_eq_right_iff, zero_ne_one,
imp_false]
congr 1
by_cases h0 : q r0 = 1
· by_cases h1 : q r1 = 1
· simp [h0, h1]
· have h1 : q r1 = 0 := by omega
simp [h0, h1]
· have h0 : q r0 = 0 := by omega
by_cases h1 : q r1 = 1
· simp [h0, h1]
· have h1 : q r1 = 0 := by omega
simp [h0, h1]
lemma koszulSignInsert_cons {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
[IsTotal I le1] [IsTrans I le1] (r0 r1 : I) (r : List I) :
koszulSignInsert le1 q r0 (r1 :: r) = koszulSignCons q le1 r0 r1 * koszulSignInsert le1 q r0 r := by
simp [koszulSignInsert, koszulSignCons]
lemma koszulSignInsert_eq_insertSign {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
[IsTotal I le1] [IsTrans I le1] (r0 : I) (r : List I) :
koszulSignInsert le1 q r0 r = insertSign q (orderedInsertPos le1 (List.insertionSort le1 r) r0)
r0 (List.insertionSort le1 r) := by
rw [koszulSignInsert_eq_cons, koszulSignInsert_eq_sort, koszulSignInsert_eq_filter,
koszulSignInsert_eq_grade, insertSign, superCommuteCoef]
congr
simp only [List.filter_filter, Bool.and_self]
rw [List.insertionSort]
nth_rewrite 1 [List.orderedInsert_eq_take_drop]
rw [List.filter_append]
have h1 : List.filter (fun a => decide ¬le1 r0 a) (List.takeWhile (fun b => decide ¬le1 r0 b) (List.insertionSort le1 r))
= (List.takeWhile (fun b => decide ¬le1 r0 b) (List.insertionSort le1 r)) := by
induction r with
| nil => simp
| cons r1 r ih =>
simp only [decide_not, List.insertionSort, List.filter_eq_self, Bool.not_eq_eq_eq_not,
Bool.not_true, decide_eq_false_iff_not]
intro a ha
have ha' := List.mem_takeWhile_imp ha
simp_all
rw [h1]
rw [List.filter_cons]
simp only [decide_not, (IsTotal.to_isRefl le1).refl r0, not_true_eq_false, decide_False,
Bool.false_eq_true, ↓reduceIte]
rw [orderedInsertPos_take]
simp only [decide_not, List.append_right_eq_self, List.filter_eq_nil_iff, Bool.not_eq_eq_eq_not,
Bool.not_true, decide_eq_false_iff_not, Decidable.not_not]
intro a ha
refine List.Sorted.rel_of_mem_take_of_mem_drop
(k := (orderedInsertPos le1 (List.insertionSort le1 r) r0).1 + 1 )
(List.sorted_insertionSort le1 (r0 :: r)) ?_ ?_
· simp only [List.insertionSort, List.orderedInsert_eq_take_drop, decide_not]
rw [List.take_append_eq_append_take]
rw [List.take_of_length_le]
· simp [orderedInsertPos]
· simp [orderedInsertPos]
· simp only [List.insertionSort, List.orderedInsert_eq_take_drop, decide_not]
rw [List.drop_append_eq_append_drop]
rw [List.drop_of_length_le]
· simpa [orderedInsertPos] using ha
· simp [orderedInsertPos]
lemma koszulSignInsert_insertIdx {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
(i j : I) [IsTotal I le1] [IsTrans I le1] (r : List I) (n : ) (hn : n ≤ r.length) :
koszulSignInsert le1 q j (List.insertIdx n i r) = koszulSignInsert le1 q j (i :: r) := by
apply koszulSignInsert_eq_perm
exact List.perm_insertIdx i r hn
lemma take_insertIdx {I : Type} (i : I) : (r : List I) → (n : ) → (hn : n ≤ r.length) →
List.take n (List.insertIdx n i r) = List.take n r
| [], 0, h => by
simp
| r0 :: r, ⟨0, h⟩ => by
simp [koszulSign]
intro h
rw [koszulSignInsert_boson]
| [], n + 1, h => by
simp at h
| r0 :: r, 0, h => by
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
| r0 :: r, n + 1, h => by
simp only [List.insertIdx_succ_cons, List.take_succ_cons, List.cons.injEq, true_and]
exact take_insertIdx i r n (Nat.le_of_lt_succ h)
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 koszulSign_insertIdx {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
(i : I) [IsTotal I le1] [IsTrans I le1] : (r : List I) → (n : ) → (hn : n ≤ r.length) →
koszulSign le1 q (List.insertIdx n i r) = insertSign q n i r
* koszulSign le1 q r
* insertSign q (insertionSortEquiv le1 (List.insertIdx n i r) ⟨n, by
rw [List.length_insertIdx _ _ hn]
omega⟩) i
(List.insertionSort le1 (List.insertIdx n i r))
| [], 0, h => by
simp [koszulSign, insertSign, superCommuteCoef, koszulSignInsert]
| [], n + 1, h => by
simp at h
| r0 :: r, 0, h => by
simp only [List.insertIdx_zero, List.insertionSort, List.length_cons, Fin.zero_eta]
rw [koszulSign]
trans koszulSign le1 q (r0 :: r) * koszulSignInsert le1 q i (r0 :: r)
ring
simp only [insertionSortEquiv, List.length_cons, Nat.succ_eq_add_one, List.insertionSort,
orderedInsertEquiv, OrderIso.toEquiv_symm, Fin.symm_castOrderIso, HepLean.Fin.equivCons_trans,
Equiv.trans_apply, HepLean.Fin.equivCons_zero, HepLean.Fin.finExtractOne_apply_eq,
Fin.isValue, HepLean.Fin.finExtractOne_symm_inl_apply, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, Fin.cast_mk, Fin.eta]
conv_rhs =>
rhs
rhs
rw [orderedInsert_eq_insertIdx_orderedInsertPos]
conv_rhs =>
rhs
rw [← insertSign_insert]
change insertSign q (↑(orderedInsertPos le1 ((List.insertionSort le1 (r0 :: r))) i)) i
(List.insertionSort le1 (r0 :: r))
rw [← koszulSignInsert_eq_insertSign q le1]
rw [insertSign_zero]
simp
| r0 :: r, n + 1, h => by
conv_lhs =>
rw [List.insertIdx_succ_cons]
rw [koszulSign]
rw [koszulSign_insertIdx]
conv_rhs =>
rhs
simp only [List.insertIdx_succ_cons]
simp only [List.insertionSort, List.length_cons, insertionSortEquiv, Nat.succ_eq_add_one,
Equiv.trans_apply, HepLean.Fin.equivCons_succ]
erw [orderedInsertEquiv_fin_succ]
simp only [Fin.eta, Fin.coe_cast]
rhs
rw [orderedInsert_eq_insertIdx_orderedInsertPos]
conv_rhs =>
lhs
rw [insertSign_succ_cons, koszulSign]
ring_nf
conv_lhs =>
lhs
rw [mul_assoc, mul_comm]
rw [mul_assoc]
conv_rhs =>
rw [mul_assoc, mul_assoc]
congr 1
let rs := (List.insertionSort le1 (List.insertIdx n i r))
have hnsL : n < (List.insertIdx n i r).length := by
rw [List.length_insertIdx _ _]
simp only [List.length_cons, add_le_add_iff_right] at h
omega
exact Nat.le_of_lt_succ h
let ni : Fin rs.length := (insertionSortEquiv le1 (List.insertIdx n i r))
⟨n, hnsL⟩
let nro : Fin (rs.length + 1) := ⟨↑(orderedInsertPos le1 rs r0), orderedInsertPos_lt_length le1 rs r0⟩
rw [koszulSignInsert_insertIdx, koszulSignInsert_cons]
trans koszulSignInsert le1 q r0 r * (koszulSignCons q le1 r0 i *insertSign q ni i rs)
· simp only [rs, ni]
ring
trans koszulSignInsert le1 q r0 r * (superCommuteCoef q [i] [r0] *
insertSign q (nro.succAbove ni) i (List.insertIdx nro r0 rs))
swap
· simp only [rs, nro, ni]
ring
congr 1
simp only [Fin.succAbove]
have hns : rs.get ni = i := by
simp only [Fin.eta, rs]
rw [← insertionSortEquiv_get]
simp only [Function.comp_apply, Equiv.symm_apply_apply, List.get_eq_getElem, ni]
simp_all only [List.length_cons, add_le_add_iff_right, List.getElem_insertIdx_self]
have hms : (List.orderedInsert le1 r0 rs).get ⟨nro, by simp⟩ = r0 := by
simp [nro]
have hc1 : ni.castSucc < nro → ¬ le1 r0 i := by
intro hninro
rw [← hns]
exact lt_orderedInsertPos_rel le1 r0 rs ni hninro
have hc2 : ¬ ni.castSucc < nro → le1 r0 i := by
intro hninro
rw [← hns]
refine gt_orderedInsertPos_rel le1 r0 rs ?_ ni hninro
exact List.sorted_insertionSort le1 (List.insertIdx n i r)
by_cases hn : ni.castSucc < nro
· simp only [hn, ↓reduceIte, Fin.coe_castSucc]
rw [insertSign_insert_gt]
swap
· exact hn
congr 1
rw [koszulSignCons_eq_superComuteCoef]
simp only [hc1 hn, ↓reduceIte]
rw [superCommuteCoef_comm]
· simp only [hn, ↓reduceIte, Fin.val_succ]
rw [insertSign_insert_lt]
rw [← mul_assoc]
congr 1
rw [superCommuteCoef_mul_self]
rw [koszulSignCons]
simp only [hc2 hn, ↓reduceIte]
exact Nat.le_of_not_lt hn
exact Nat.le_of_lt_succ (orderedInsertPos_lt_length le1 rs r0)
· exact Nat.le_of_lt_succ h
· exact Nat.le_of_lt_succ h
lemma insertIdx_eraseIdx {I : Type} :
(n : ) → (r : List I) → (hn : n < r.length) →
List.insertIdx n (r.get ⟨n, hn⟩) (r.eraseIdx n) = r
| n, [], hn => by
simp at hn
| 0, r0 :: r, hn => by
simp
| n + 1, r0 :: r, hn => by
simp only [List.length_cons, List.get_eq_getElem, List.getElem_cons_succ,
List.eraseIdx_cons_succ, List.insertIdx_succ_cons, List.cons.injEq, true_and]
exact insertIdx_eraseIdx n r _
lemma superCommuteCoefLE_eq_get {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
[DecidableRel le1] (i : I) (n : Fin r.length) (heq : q i = q (r.get n)) :
[DecidableRel le1] [IsTotal I le1] [IsTrans I 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
sorry
rw [superCommuteCoefLE_eq_q]
let r' := r.eraseIdx ↑n
have hr : List.insertIdx n (r.get n) (r.eraseIdx n) = r := by
exact insertIdx_eraseIdx n.1 r n.prop
conv_lhs =>
lhs
lhs
rw [← hr]
rw [koszulSign_insertIdx q le1 (r.get n) ((r.eraseIdx ↑n)) n (by
rw [List.length_eraseIdx]
simp only [Fin.is_lt, ↓reduceIte]
omega)]
rhs
rhs
rw [hr]
conv_lhs =>
lhs
lhs
rhs
enter [2, 1, 1]
rw [insertionSortEquiv_congr _ _ hr]
simp only [List.get_eq_getElem, Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
Fin.cast_mk, Fin.eta, Fin.coe_cast]
conv_lhs =>
lhs
rw [mul_assoc]
rhs
rw [insertSign]
rw [superCommuteCoef_mul_self]
simp only [mul_one]
rw [mul_assoc]
rw [koszulSign_mul_self]
simp only [mul_one]
rw [insertSign_eraseIdx]
rfl
exact heq
end Wick

View file

@ -20,6 +20,7 @@ import HepLean.PerturbationTheory.Wick.Koszul.Grade
-/
namespace Wick
open HepLean.List
noncomputable section
@ -136,6 +137,393 @@ lemma ofListM_cons_eq_ofListM {I : Type} (f : I → Type) [∀ i, Fintype (f i)]
rw [ofListM_cons, ofListM_singleton]
simp only [one_smul]
def CreatAnnilateSect {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (l : List I) : Type :=
Π i, f (l.get i)
namespace CreatAnnilateSect
variable {I : Type} {f : I → Type} [∀ i, Fintype (f i)] {l : List I} (a : CreatAnnilateSect f l)
instance fintype : Fintype (CreatAnnilateSect f l) := Pi.fintype
def tail : {l : List I} → (a : CreatAnnilateSect f l) → CreatAnnilateSect f l.tail
| [], a => a
| _ :: _, a => fun i => a (Fin.succ i)
def head {i : I} (a : CreatAnnilateSect f (i :: l)) : f i := a ⟨0, Nat.zero_lt_succ l.length⟩
def toList : {l : List I} → (a : CreatAnnilateSect f l) → List (Σ i, f i)
| [], _ => []
| i :: _, a => ⟨i, a.head⟩ :: toList a.tail
@[simp]
lemma toList_length : (toList a).length = l.length := by
induction l with
| nil => rfl
| cons i l ih =>
simp only [toList, List.length_cons, Fin.zero_eta]
rw [ih]
lemma toList_tail : {l : List I} → (a : CreatAnnilateSect f l) → toList a.tail = (toList a).tail
| [], _ => rfl
| i :: l, a => by
simp [toList]
lemma toList_cons {i : I} (a : CreatAnnilateSect f (i :: l)) :
(toList a) = ⟨i, a.head⟩ :: toList a.tail := by
rfl
lemma toList_get (a : CreatAnnilateSect f l) :
(toList a).get = (fun i => ⟨l.get i, a i⟩) ∘ Fin.cast (by simp) := by
induction l with
| nil =>
funext i
exact Fin.elim0 i
| cons i l ih =>
simp only [toList_cons, List.get_eq_getElem, Fin.zero_eta, List.getElem_cons_succ,
Function.comp_apply, Fin.cast_mk]
funext x
match x with
| ⟨0, h⟩ => rfl
| ⟨x + 1, h⟩ =>
simp only [List.get_eq_getElem, Prod.mk.eta, List.getElem_cons_succ, Function.comp_apply]
change (toList a.tail).get _ = _
rw [ih]
simp [tail]
@[simp]
lemma toList_grade (q : I → Fin 2) :
grade (fun i => q i.fst) a.toList = 1 ↔ grade q l = 1 := by
induction l with
| nil =>
simp [toList]
| 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) a.tail.toList = 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) a.tail.toList = 0 := by
simp [tail]
omega
rw [h0, h0']
rw [h1]
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
| [] => exact Fin.elim0 n
| l0 :: l =>
let e1 : CreatAnnilateSect f ((l0 :: l).eraseIdx n) ≃ Π i, f ((l0 :: l).get (n.succAbove i)) :=
Equiv.piCongr (Fin.castOrderIso (by rw [eraseIdx_cons_length])).toEquiv
fun x => Equiv.cast (congrArg f (by
rw [HepLean.List.eraseIdx_get]
simp
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]))
exact (Fin.insertNthEquiv _ _).symm.trans (Equiv.prodCongr (Equiv.refl _) e1.symm)
def eraseIdx (n : Fin l.length) : CreatAnnilateSect f (l.eraseIdx n) :=
(extractEquiv n a).2
@[simp]
lemma eraseIdx_zero_tail {i : I} {l : List I} (a : CreatAnnilateSect f (i :: l)) :
(eraseIdx a (@OfNat.ofNat (Fin (l.length + 1)) 0 Fin.instOfNat : Fin (l.length + 1))) =
a.tail := by
simp [eraseIdx, extractEquiv]
rfl
lemma eraseIdx_succ_head {i : I} {l : List I} (n : ) (hn : n + 1 < (i :: l).length) (a : CreatAnnilateSect f (i :: l)) :
(eraseIdx a ⟨n + 1, hn⟩).head = a.head := by
rw [eraseIdx, extractEquiv]
simp
conv_lhs =>
rhs
rhs
rhs
erw [Fin.insertNthEquiv_symm_apply]
simp [head, Equiv.piCongr, Equiv.piCongrRight, Equiv.piCongrLeft, Equiv.piCongrLeft']
simp [Fin.removeNth, Fin.succAbove]
refine cast_eq_iff_heq.mpr ?_
congr
simp [Fin.ext_iff]
lemma eraseIdx_succ_tail {i : I} {l : List I} (n : ) (hn : n + 1 < (i :: l).length) (a : CreatAnnilateSect f (i :: l)) :
(eraseIdx a ⟨n + 1, hn⟩).tail = eraseIdx a.tail ⟨n , Nat.succ_lt_succ_iff.mp hn⟩ := by
match l with
| [] =>
simp at hn
| r0 :: r =>
rw [eraseIdx, extractEquiv]
simp
conv_lhs =>
rhs
rhs
rhs
erw [Fin.insertNthEquiv_symm_apply]
rw [eraseIdx]
conv_rhs =>
rhs
rw [extractEquiv]
simp
erw [Fin.insertNthEquiv_symm_apply]
simp [tail, Equiv.piCongr, Equiv.piCongrRight, Equiv.piCongrLeft, Equiv.piCongrLeft']
funext i
simp
have hcast {α β : Type} (h : α = β) (a : α) (b : β) : cast h a = b ↔ a = cast (Eq.symm h) b := by
cases h
simp
rw [hcast]
simp
refine eq_cast_iff_heq.mpr ?_
simp [Fin.removeNth, Fin.succAbove]
congr
simp [Fin.ext_iff]
split
next h =>
simp_all only [Fin.coe_castSucc, Fin.val_succ, Fin.coe_cast, add_left_inj]
split
next h_1 => simp_all only [Fin.coe_castSucc, Fin.coe_cast]
next h_1 =>
simp_all only [not_lt, Fin.val_succ, Fin.coe_cast, self_eq_add_right, one_ne_zero]
simp [Fin.lt_def] at h
simp [Fin.le_def] at h_1
omega
next h =>
simp_all only [not_lt, Fin.val_succ, Fin.coe_cast, add_left_inj]
split
next h_1 =>
simp_all only [Fin.coe_castSucc, Fin.coe_cast, add_right_eq_self, one_ne_zero]
simp [Fin.le_def] at h
simp [Fin.lt_def] at h_1
omega
next h_1 => simp_all only [not_lt, Fin.val_succ, Fin.coe_cast]
lemma eraseIdx_toList : {l : List I} → {n : Fin l.length} → (a : CreatAnnilateSect f l) →
(eraseIdx a n).toList = a.toList.eraseIdx n
| [], n, _ => Fin.elim0 n
| r0 :: r, ⟨0, h⟩, a => by
simp [toList_tail]
| r0 :: r, ⟨n + 1, h⟩, a => by
simp [toList]
apply And.intro
· rw [eraseIdx_succ_head]
· conv_rhs => rw [← eraseIdx_toList (l := r) (n := ⟨n, Nat.succ_lt_succ_iff.mp h⟩) a.tail]
rw [eraseIdx_succ_tail]
lemma 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):
koszulSignInsert (fun i j => le1 i.fst j.fst) (fun i => q i.fst) x a.toList =
koszulSignInsert le1 q x.1 l := by
induction l with
| nil => simp [koszulSignInsert]
| cons b l ih =>
simp [koszulSignInsert]
rw [ih]
lemma toList_koszulSign {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) :
koszulSign (fun i j => le1 i.fst j.fst) (fun i => q i.fst) a.toList =
koszulSign le1 q l := by
induction l with
| nil => simp [koszulSign]
| cons i l ih =>
simp [koszulSign, liftM]
rw [ih]
congr 1
rw [toList_koszulSignInsert]
lemma insertionSortEquiv_toList {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(le1 : I → I → Prop) [DecidableRel le1](l : List I)
(a : CreatAnnilateSect f l) :
insertionSortEquiv (fun i j => le1 i.fst j.fst) a.toList =
(Fin.castOrderIso (by simp)).toEquiv.trans ((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) a.tail.toList)) =
List.insertionSort le1 l := by
congr
have h3' (l : List I) (a : CreatAnnilateSect f l) :
List.map (fun i => i.1) a.toList = l := by
induction l with
| nil => rfl
| cons i l ih' =>
simp only [toList, List.length_cons, Fin.zero_eta, Prod.mk.eta,
List.unzip_snd, List.map_cons, List.cons.injEq, true_and]
simpa using ih' _
rw [h3']
rfl
rw [HepLean.List.orderedInsertEquiv_congr _ _ _ h3]
simp only [List.length_cons, Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
Fin.cast_trans, Fin.cast_eq_self, Fin.coe_cast]
rfl
def sort (le1 : I → I → Prop) [DecidableRel le1] : CreatAnnilateSect 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
lemma sort_toList {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(le1 : I → I → Prop) [DecidableRel le1](l : List I) (a : CreatAnnilateSect f l) :
(a.sort le1).toList = List.insertionSort (fun i j => le1 i.fst j.fst) a.toList := by
let l1 := List.insertionSort (fun i j => le1 i.fst j.fst) a.toList
let l2 := (a.sort le1).toList
symm
change l1 = l2
have hlen : l1.length = l2.length := by
simp [l1, l2]
have hget : l1.get = l2.get ∘ Fin.cast hlen := by
rw [← HepLean.List.insertionSortEquiv_get]
rw [toList_get, toList_get]
funext i
rw [insertionSortEquiv_toList]
simp only [ Function.comp_apply, Equiv.symm_trans_apply,
OrderIso.toEquiv_symm, Fin.symm_castOrderIso, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
Fin.cast_trans, Fin.cast_eq_self, id_eq, eq_mpr_eq_cast, Fin.coe_cast, Sigma.mk.inj_iff]
apply And.intro
· have h1 := congrFun (HepLean.List.insertionSortEquiv_get (r := le1) l) (Fin.cast (by simp) i)
rw [← h1]
simp
· simp [Equiv.piCongr, sort]
exact (cast_heq _ _).symm
apply List.ext_get hlen
rw [hget]
simp
end CreatAnnilateSect
lemma ofListM_expand {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (x : ) :
(l : List I) → ofListM f l x = ∑ (a : CreatAnnilateSect f l), ofList a.toList x
| [] => by
simp only [ofListM, CreatAnnilateSect, List.length_nil, List.get_eq_getElem, Finset.univ_unique,
CreatAnnilateSect.toList, Finset.sum_const, Finset.card_singleton, one_smul]
rw [ofList_eq_smul_one, map_smul, ofList_empty, ofList_eq_smul_one, ofList_empty, map_one]
| i :: l => by
rw [ofListM_cons, ofListM_expand f x l]
conv_rhs => rw [← (CreatAnnilateSect.extractEquiv
⟨0, by exact Nat.zero_lt_succ l.length⟩).symm.sum_comp (α := FreeAlgebra _)]
erw [Finset.sum_product]
rw [Finset.sum_mul]
conv_lhs =>
rhs
intro n
rw [Finset.mul_sum]
congr
funext j
congr
funext n
rw [← ofList_singleton, ← ofList_pair, one_mul]
rfl
lemma koszulOrder_ofListM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
(l : List I) (x : ) : koszulOrder (fun i j => le1 i.1 j.1) (fun i => q i.fst) (ofListM f l x) =
freeAlgebraMap f (koszulOrder le1 q (ofList l x)) := by
rw [koszulOrder_ofList]
rw [map_smul]
change _ = _ • ofListM _ _ _
rw [ofListM_expand]
rw [map_sum]
conv_lhs =>
rhs
intro a
rw [koszulOrder_ofList]
rw [CreatAnnilateSect.toList_koszulSign]
rw [← Finset.smul_sum]
apply congrArg
conv_lhs =>
rhs
intro n
rw [← CreatAnnilateSect.sort_toList]
rw [ofListM_expand]
refine Fintype.sum_equiv ((HepLean.List.insertionSortEquiv le1 l).piCongr fun i => Equiv.cast ?_) _ _ ?_
congr 1
· rw [← HepLean.List.insertionSortEquiv_get]
simp
· intro x
rfl
lemma koszulOrder_ofListM_eq_ofListM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
(l : List I) (x : ) : koszulOrder (fun i j => le1 i.1 j.1) (fun i => q i.fst) (ofListM f l x) =
koszulSign le1 q l • ofListM f (List.insertionSort le1 l) x := by
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)
| [], _ => []
@ -150,6 +538,10 @@ lemma liftM_length {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (r : List
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
@ -172,6 +564,22 @@ def liftMCongrEquiv {I : Type} (f : I → Type) [∀ i, Fintype (f i)] (r0 : I)
(Π 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
@ -233,11 +641,10 @@ 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))
{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
@ -245,10 +652,54 @@ def listMEraseEquiv {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
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
·
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)) →
@ -350,42 +801,7 @@ lemma insertionSortEquiv_liftM {I : Type} {f : I → Type} [∀ i, Fintype (f i)
| cons i l ih =>
simp only [liftM, List.length_cons, Fin.zero_eta, List.insertionSort]
conv_lhs => simp [HepLean.List.insertionSortEquiv]
have h1 (l' : List (Σ i, f i)) :
(HepLean.List.insertEquiv (fun i j => le1 i.fst j.fst) ⟨i, a ⟨0, by simp⟩⟩ l') =
(Fin.castOrderIso (by simp)).toEquiv.trans
((HepLean.List.insertEquiv le1 i (List.map (fun i => i.1) l')).trans
(Fin.castOrderIso (by simp [List.orderedInsert_length])).toEquiv) := by
induction l' with
| nil =>
simp only [List.length_cons, Nat.add_zero, Nat.zero_eq, Fin.zero_eta, List.length_singleton,
List.orderedInsert, HepLean.List.insertEquiv, Fin.castOrderIso_refl,
OrderIso.refl_toEquiv, Equiv.trans_refl]
rfl
| cons j l' ih' =>
by_cases hr : (fun (i j : Σ i, f i) => le1 i.fst j.fst) ⟨i, a ⟨0, by simp⟩⟩ j
· rw [HepLean.List.insertEquiv_cons_pos]
· erw [HepLean.List.insertEquiv_cons_pos]
· rfl
· exact hr
· exact hr
· rw [HepLean.List.insertEquiv_cons_neg]
· erw [HepLean.List.insertEquiv_cons_neg]
· simp only [List.length_cons, Nat.add_zero, Nat.zero_eq, Fin.zero_eta,
List.orderedInsert, Prod.mk.eta, Fin.mk_one]
erw [ih']
ext x
simp only [Prod.mk.eta, List.length_cons, Nat.add_zero, Nat.zero_eq, Fin.zero_eta,
HepLean.Fin.equivCons_trans, Nat.succ_eq_add_one,
HepLean.Fin.equivCons_castOrderIso, Equiv.trans_apply, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, Fin.cast_trans, Fin.coe_cast]
congr 2
match x with
| ⟨0, h⟩ => rfl
| ⟨1, h⟩ => rfl
| ⟨Nat.succ (Nat.succ x), h⟩ => rfl
· exact hr
· exact hr
erw [h1]
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,
@ -399,7 +815,7 @@ lemma insertionSortEquiv_liftM {I : Type} {f : I → Type} [∀ i, Fintype (f i)
List.orderedInsert le1 i.1 (List.map (fun i => i.1) l') := by
induction l' with
| nil =>
simp [HepLean.List.insertEquiv]
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]
@ -416,13 +832,13 @@ lemma insertionSortEquiv_liftM {I : Type} {f : I → Type} [∀ i, Fintype (f i)
List.insertionSort le1 (List.map (fun i => i.1) l') := by
induction l' with
| nil =>
simp [HepLean.List.insertEquiv]
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.insertEquiv_congr _ _ _ (h2 _)]
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)))) =
@ -437,7 +853,7 @@ lemma insertionSortEquiv_liftM {I : Type} {f : I → Type} [∀ i, Fintype (f i)
List.unzip_snd, List.map_cons, List.cons.injEq, true_and]
simpa using ih' _
rw [h3']
rw [HepLean.List.insertEquiv_congr _ _ _ 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]
@ -477,40 +893,6 @@ lemma insertionSort_liftM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
rw [hget]
simp
lemma koszulOrder_ofListM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
(l : List I) (x : ) : koszulOrder (fun i j => le1 i.1 j.1) (fun i => q i.fst) (ofListM f l x) =
freeAlgebraMap f (koszulOrder le1 q (ofList l x)) := by
rw [koszulOrder_ofList]
rw [map_smul]
change _ = _ • ofListM _ _ _
rw [ofListM_expand]
rw [map_sum]
conv_lhs =>
rhs
intro a
rw [koszulOrder_ofList]
rw [koszulSign_liftM]
rw [← Finset.smul_sum]
apply congrArg
conv_lhs =>
rhs
intro n
rw [insertionSort_liftM]
rw [ofListM_expand]
refine Fintype.sum_equiv ((HepLean.List.insertionSortEquiv le1 l).piCongr fun i => Equiv.cast ?_) _ _ ?_
congr 1
· rw [← HepLean.List.insertionSortEquiv_get]
simp
· intro x
rfl
lemma koszulOrder_ofListM_eq_ofListM {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
(l : List I) (x : ) : koszulOrder (fun i j => le1 i.1 j.1) (fun i => q i.fst) (ofListM f l x) =
koszulSign le1 q l • ofListM f (List.insertionSort le1 l) x := by
rw [koszulOrder_ofListM, koszulOrder_ofList, map_smul]
rfl
end
end Wick

View file

@ -27,15 +27,15 @@ namespace Wick
noncomputable section
class SuperCommuteCenterMap {A : Type} [Semiring A] [Algebra A]
(f : FreeAlgebra I →ₐ[] A) : Prop where
prop : ∀ i j, f (superCommute q (FreeAlgebra.ι i) (FreeAlgebra.ι j)) ∈ Subalgebra.center 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) [SuperCommuteCenterMap f]
(f : FreeAlgebra I →ₐ[] A) (q : I → Fin 2) [SuperCommuteCenterMap q f]
lemma ofList_fst (q : I → Fin 2) (i j : I) :
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
@ -51,7 +51,7 @@ lemma ofList_fst (q : I → Fin 2) (i j : I) :
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 F] (b : I) :
[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_ι]
@ -64,7 +64,7 @@ end SuperCommuteCenterMap
lemma superCommuteTake_superCommuteCenterMap {I : Type} (q : I → Fin 2) (lb : List I) (xa xb : ) (n : )
(hn : n < lb.length) {A : Type} [Semiring A] [Algebra A] (f : FreeAlgebra I →ₐ[] A)
[SuperCommuteCenterMap f] (i : I) :
[SuperCommuteCenterMap q 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) •
@ -83,7 +83,7 @@ lemma superCommuteTakeM_F {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)
{A : Type} [Semiring A] [Algebra A] (F : FreeAlgebra (Σ i, f i) →ₐ[] A)
[SuperCommuteCenterMap F] :
[SuperCommuteCenterMap (fun i => q i.1) 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
@ -105,104 +105,102 @@ lemma superCommuteTakeM_F {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 → I → Prop) [DecidableRel le1]
(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 F] :
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap q 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]
rw [koszulOrder_ofList, map_smul, map_smul, ← ofList_singleton, superCommute_ofList_sum]
rw [map_sum, ← (HepLean.List.insertionSortEquiv le1 r).sum_comp]
conv_lhs =>
rhs
rhs
enter [2, 2]
intro n
rw [superCommuteTake_superCommuteCenterMap]
lhs
rhs
rhs
rhs
enter [1, 2, 2, 2]
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
enter [2, 2]
intro n
rw [hListErase]
rw [HepLean.List.eraseIdx_insertionSort_fin le1 r n]
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]
rw [map_smul, smul_smul, Algebra.mul_smul_comm, 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
by_cases hq : q i ≠ q (r.get n)
· have hn := SuperCommuteCenterMap.dif_grade (q := q) (F := F) i (r.get n) hq
conv_lhs =>
enter [2, 1]
rw [ofList_singleton, hn]
conv_rhs =>
enter [2, 1]
rw [ofList_singleton, hn]
simp
· congr 1
trans superCommuteCoefLE q le1 r i n
· rw [superCommuteCoefLE, mul_assoc]
refine superCommuteCoefLE_eq_get q le1 r i n ?_
simpa using hq
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 →ₐ A) [SuperCommuteCenterMap F] :
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap q 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
enter [2, 2]
rw [← ofList_singleton]
rw [ofListM_ofList_superCommute' q]
rw [map_sub]
rw [sub_eq_add_neg]
rw [map_add]
conv_lhs =>
rhs
rhs
enter [2, 2]
rw [map_smul]
rw [← neg_smul]
rw [map_smul, map_smul, map_smul]
sorry
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 F] :
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap q 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
enter [1, 2]
rw [← Algebra.smul_mul_assoc, smul_smul, mul_comm, ← smul_smul, ← koszulOrder_ofList,
Algebra.smul_mul_assoc, ofList_singleton]
rw [koszulOrder_mul_ge, map_smul]
congr
· rw [koszulOrder_of_le_all_ofList]
sorry
sorry
rw [superCommuteCoef_perm_snd q [i] (List.insertionSort le1 r) r
(List.perm_insertionSort le1 r)]
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 F]
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap q F]
(n : Option (Fin r.length)) : A :=
match n with
| none => 1
@ -212,7 +210,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 F] :
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap q F] :
superCommuteCenterOrder q r i F none = 1 := by
simp [superCommuteCenterOrder]
@ -220,9 +218,10 @@ 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]
[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 F] :
(F : FreeAlgebra I →ₐ A) [SuperCommuteCenterMap q 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]
@ -242,9 +241,10 @@ lemma le_all_mul_koszulOrder_ofList_expand {I : Type}
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]
[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 F] :
(F : FreeAlgebra (Σ i, f i) →ₐ A) [SuperCommuteCenterMap (fun i => q i.1) 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) •
@ -318,19 +318,19 @@ lemma le_all_mul_koszulOrder_ofListM_expand {I : Type} {f : I → Type} [∀ i,
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)]):
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
rhs
rhs
enter [2, 2]
intro a
erw [hl a0 a]
erw [hl n a0 a]
rw [← Finset.sum_mul]
conv_lhs =>
lhs

View file

@ -87,6 +87,10 @@ 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