refactor: Style Lint

This commit is contained in:
jstoobysmith 2024-12-19 12:59:14 +00:00
parent aad3afd3a7
commit 63c4cabdf4
9 changed files with 512 additions and 453 deletions

View file

@ -244,7 +244,7 @@ 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
simp only [Nat.succ_eq_add_one, finExtractOne_symm_inr_apply]
exact succsAbove_predAboveI hij
/-- Given an equivalence `Fin n.succ.succ ≃ Fin n.succ.succ`, and an `i : Fin n.succ.succ`,
@ -252,7 +252,8 @@ lemma finExtractOne_apply_neq {n : } (i j : Fin n.succ.succ) (hij : i ≠ j)
def finExtractOnPermHom {m : } (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fin m.succ.succ) :
Fin n.succ → Fin m.succ := fun x => predAboveI (σ i) (σ ((finExtractOne i).symm (Sum.inr x)))
lemma finExtractOnPermHom_inv {m : } (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃ Fin m.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
funext x
simp only [Nat.succ_eq_add_one, Function.comp_apply, finExtractOnPermHom, Equiv.symm_apply_apply,
@ -287,13 +288,15 @@ def finExtractOnePerm {m : } (i : Fin n.succ.succ) (σ : Fin n.succ.succ ≃
right_inv x := by
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]
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 only [Nat.succ_eq_add_one, finExtractOnePerm, Equiv.coe_fn_mk]
funext x
simp [finExtractOnPermHom]
simp only [Function.comp_apply, finExtractOnPermHom, Nat.succ_eq_add_one,
finExtractOne_symm_inr_apply]
rw [succsAbove_predAboveI]
simp
simp only [Nat.succ_eq_add_one, ne_eq, EmbeddingLike.apply_eq_iff_eq]
exact Fin.ne_succAbove i x
@[simp]
@ -437,7 +440,7 @@ lemma equivCons_succ {n m : } (e : Fin n ≃ Fin m) (i : ) (hi : i + 1 < n
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
simp only [Equiv.coe_fn_mk]
rw [hi]
rw [Fin.cons_succ]
rfl

View file

@ -14,78 +14,82 @@ open Fin
open HepLean
variable {n : Nat}
lemma takeWile_eraseIdx {I : Type} (P : I → Prop) [DecidablePred P] :
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
rfl
| a :: [], 0, h => by
simp [List.takeWhile]
simp only [List.takeWhile, List.eraseIdx_zero, List.nil_eq]
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 ]
· rfl
· rfl
| a :: [], Nat.succ n, h => by
simp only [Nat.succ_eq_add_one, List.eraseIdx_cons_succ, List.eraseIdx_nil]
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
simp only [List.length_singleton] at h1
omega
| a :: b :: l, 0, h => by
simp [List.takeWhile]
simp only [List.takeWhile, List.eraseIdx_zero]
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 only [hPb, decide_False, List.nil_eq]
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]
· rfl
· rfl
| a :: b :: l, Nat.succ n, h => by
simp
simp only [Nat.succ_eq_add_one, List.eraseIdx_cons_succ]
by_cases hPa : P a
· dsimp [List.takeWhile]
simp [hPa]
simp only [hPa, decide_True, List.eraseIdx_cons_succ, List.cons.injEq, true_and]
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)
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] :
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
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]
simp only [List.dropWhile, nonpos_iff_eq_zero, List.length_eq_zero, List.takeWhile_eq_nil_iff,
List.length_singleton, zero_lt_one, Fin.zero_eta, Fin.isValue, List.get_eq_getElem,
Fin.val_eq_zero, List.getElem_cons_zero, decide_eq_true_eq, forall_const, zero_le,
Nat.sub_eq_zero_of_le, List.eraseIdx_zero, ite_not, List.nil_eq]
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 only [List.dropWhile, List.eraseIdx_nil, List.takeWhile, Nat.succ_eq_add_one]
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]
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]
simp only [List.dropWhile, List.takeWhile, nonpos_iff_eq_zero, List.length_eq_zero, zero_le,
Nat.sub_eq_zero_of_le, List.eraseIdx_zero]
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 only [hPb, decide_False, nonpos_iff_eq_zero, List.length_eq_zero]
simp_all only [List.length_cons, List.get_eq_getElem]
simp_all only [decide_False, nonpos_iff_eq_zero, List.length_eq_zero]
split
@ -93,29 +97,32 @@ lemma dropWile_eraseIdx {I : Type} (P : I → Prop) [DecidablePred P] :
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]
· rfl
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 => rfl
next x heq => simp_all only [not_true_eq_false]
| a :: b :: l, Nat.succ n, h => by
simp
simp only [Nat.succ_eq_add_one, List.eraseIdx_cons_succ]
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]
simp only [List.dropWhile, hPa, decide_True, List.takeWhile, hPb, List.length_cons,
add_le_add_iff_right, Nat.reduceSubDiff]
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]
intro i j hij hP
simpa using h (Fin.succ i) (Fin.succ j) (by simpa using hij) (by simpa using hP)
· simp only [List.dropWhile, List.takeWhile, hPb, decide_False]
by_cases hPa : P a
· rw [dropWile_eraseIdx]
simp [hPa, List.dropWhile, hPb]
simp only [hPa, decide_True, hPb, decide_False, Bool.false_eq_true, not_false_eq_true,
List.takeWhile_cons_of_neg, List.length_nil, zero_le, ↓reduceIte, List.dropWhile,
tsub_zero, List.length_singleton, le_add_iff_nonneg_left, add_tsub_cancel_right]
intro i j hij hP
simpa using h (Fin.succ i) (Fin.succ j) (by simpa using hij) (by simpa using 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) :
@ -137,49 +144,49 @@ lemma orderedInsertPos_lt_length {I : Type} (le1 : I → I → Prop) [DecidableR
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]
simp only [List.orderedInsert_eq_take_drop, decide_not, orderedInsertPos]
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
(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]
· rfl
lemma orderedInsertPos_cons {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
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]
simp only [List.orderedInsert.eq_2, orderedInsertPos, List.takeWhile, decide_not, Fin.zero_eta,
Fin.succ_mk]
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) :
(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]
simp only [orderedInsertPos, decide_not]
induction l with
| nil =>
simp
| cons a l ih =>
simp [List.takeWhile]
simp only [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]
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_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]
@ -190,10 +197,10 @@ lemma orderedInsert_get_lt {I : Type} (le1 : I → I → Prop) [DecidableRel le1
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]
simp only [orderedInsertPos, decide_not] at hi
simp only [List.orderedInsert_eq_take_drop, decide_not, List.get_eq_getElem]
rw [List.getElem_append]
simp [hi]
simp only [hi, ↓reduceDIte]
rw [List.IsPrefix.getElem]
exact List.takeWhile_prefix fun b => !decide (le1 r0 b)
@ -203,19 +210,18 @@ lemma orderedInsertPos_take_orderedInsert {I : Type} (le1 : I → I → Prop) [D
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 )
refine List.ext_get ?_ ?_
· simp only [List.length_take, Fin.is_le', inf_of_le_left, inf_eq_left]
exact Nat.le_of_lt_succ (orderedInsertPos_lt_length le1 r r0)
· intro n h1 h2
simp
simp only [List.get_eq_getElem, List.getElem_take]
erw [orderedInsert_get_lt le1 r r0 n]
rfl
simp at h1
simp only [List.length_take, lt_inf_iff] at h1
exact h1.1
lemma orderedInsertPos_drop_eq_orderedInsert {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
@ -223,7 +229,8 @@ lemma orderedInsertPos_drop_eq_orderedInsert {I : Type} (le1 : I → I → Prop)
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
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
@ -243,26 +250,27 @@ lemma orderedInsertPos_drop {I : Type} (le1 : I → I → Prop) [DecidableRel le
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]
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
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)
(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
simp only [List.get_eq_getElem, lt_inf_iff, Fin.is_lt, and_true, exists_prop]
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)
[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
@ -278,7 +286,6 @@ lemma gt_orderedInsertPos_rel {I : Type} (le1 : I → I → Prop) [DecidableRel
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)
@ -291,10 +298,10 @@ lemma orderedInsert_eraseIdx_lt_orderedInsertPos {I : Type} (le1 : I → I → P
· rw [takeWile_eraseIdx]
exact hr
· rw [dropWile_eraseIdx]
simp [orderedInsertPos] at hi
simp only [orderedInsertPos, decide_not] at hi
have hi' : ¬ (List.takeWhile (fun b => !decide (le1 r0 b)) r).length ≤ ↑i := by
omega
simp [hi']
simp only [decide_not, hi', ↓reduceIte]
exact fun i j a a_1 => hr i j a a_1
· exact hi
@ -302,15 +309,16 @@ lemma orderedInsert_eraseIdx_orderedInsertPos_le {I : Type} (le1 : I → I → P
(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
(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
simp only [orderedInsertPos, decide_not] at hi
simp only [decide_not]
omega
exact hr
· simp only [Nat.succ_eq_add_one]
@ -328,12 +336,13 @@ lemma orderedInsert_eraseIdx_orderedInsertPos_le {I : Type} (le1 : I → I → P
· simp only [orderedInsertPos] at hi
omega
def orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I) (r0 : I) :
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⟩
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]
@ -344,35 +353,44 @@ lemma orderedInsertEquiv_zero {I : Type} (le1 : I → I → Prop) [DecidableRel
@[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]
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 only [List.length_cons, orderedInsertEquiv, Nat.succ_eq_add_one, OrderIso.toEquiv_symm,
Fin.symm_castOrderIso, Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply]
match r with
| [] =>
simp
| r1 :: r =>
erw [finExtractOne_apply_neq]
simp [orderedInsertPos]
simp only [List.orderedInsert.eq_2, List.length_cons, orderedInsertPos, decide_not,
Nat.succ_eq_add_one, finExtractOne_symm_inr_apply]
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) :
(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]
((Fin.succAbove ⟨(orderedInsertPos le1 r r0), orderedInsertPos_lt_length le1 r r0⟩)
⟨n, n.isLt⟩) := by
simp only [List.length_cons, orderedInsertEquiv, Nat.succ_eq_add_one, OrderIso.toEquiv_symm,
Fin.symm_castOrderIso, Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
Fin.eta]
match r with
| [] =>
simp
| r1 :: r =>
erw [finExtractOne_apply_neq]
simp [orderedInsertPos]
simp only [List.orderedInsert.eq_2, List.length_cons, orderedInsertPos, decide_not,
Nat.succ_eq_add_one, finExtractOne_symm_inr_apply]
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
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
@ -383,35 +401,38 @@ lemma get_eq_orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRe
funext x
match x with
| ⟨0, h⟩ =>
simp
simp only [List.length_cons, Fin.zero_eta, List.get_eq_getElem, Fin.val_zero,
List.getElem_cons_zero, Function.comp_apply]
erw [orderedInsertEquiv_zero]
simp
| ⟨Nat.succ n, h⟩ =>
simp
simp only [List.length_cons, Nat.succ_eq_add_one, List.get_eq_getElem, List.getElem_cons_succ,
Function.comp_apply]
erw [orderedInsertEquiv_succ]
simp [Fin.succAbove]
simp only [Fin.succAbove, Fin.castSucc_mk, Fin.mk_lt_mk, Fin.succ_mk, Fin.coe_cast]
by_cases hn : n < ↑(orderedInsertPos le1 r r0)
· simp [hn]
· simp [hn]
simp [List.orderedInsert_eq_take_drop]
· simp only [hn, ↓reduceIte]
simp only [List.orderedInsert_eq_take_drop, decide_not]
rw [List.getElem_append]
have hn' : ¬ n + 1 < (List.takeWhile (fun b => !decide (le1 r0 b)) r).length := by
simp [orderedInsertPos] at hn
simp only [orderedInsertPos, decide_not, not_lt] at hn
omega
simp [hn']
simp only [hn', ↓reduceDIte]
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
simp only [orderedInsertPos, decide_not, not_lt] at hn
omega
simp [hnn]
simp only [hnn, List.getElem_cons_succ]
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
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]
simp only [hr, add_tsub_cancel_right]
omega
lemma orderedInsertEquiv_get {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
@ -421,14 +442,14 @@ lemma orderedInsertEquiv_get {I : Type} (le1 : I → I → Prop) [DecidableRel l
funext x
simp
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 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)
{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
@ -453,33 +474,34 @@ lemma orderedInsert_eraseIdx_orderedInsertEquiv_fin_succ
(List.orderedInsert le1 r0 (r.eraseIdx n)) := by
have hn : n.succ = ⟨n.val + 1, by omega⟩ := by
rw [Fin.ext_iff]
simp
rfl
rw [hn]
exact orderedInsert_eraseIdx_orderedInsertEquiv_succ le1 r r0 n.val _ hr
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⟩) =
(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
simp only [List.length_cons, Fin.zero_eta, Equiv.trans_apply, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, Fin.cast_zero, Fin.coe_cast]
erw [orderedInsertEquiv_zero, orderedInsertEquiv_zero]
simp [orderedInsertPos_sigma]
| ⟨Nat.succ n, h0⟩ =>
simp
simp only [List.length_cons, Nat.succ_eq_add_one, Equiv.trans_apply, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, Fin.cast_mk, Fin.coe_cast]
erw [orderedInsertEquiv_succ, orderedInsertEquiv_succ]
simp [orderedInsertPos_sigma]
simp only [orderedInsertPos_sigma, Fin.coe_cast]
rw [Fin.succAbove, Fin.succAbove]
simp
simp only [Fin.castSucc_mk, Fin.mk_lt_mk, Fin.succ_mk]
split
next h => simp_all only
next h => simp_all only [not_lt]
· rfl
· rfl
/-- This result is taken from:
https://github.com/leanprover/lean4/blob/master/src/Init/Data/List/Nat/InsertIdx.lean
@ -489,12 +511,14 @@ lemma orderedInsertEquiv_sigma {I : Type} {f : I → Type} [∀ i, Fintype (f i)
Floris van Doorn, Mario Carneiro.
Once HepLean is updated to a more recent version of Lean this result will be removed.
-/
theorem length_insertIdx' : ∀ n as, (List.insertIdx n a as).length = if n ≤ as.length then as.length + 1 else as.length
-/
theorem length_insertIdx' : ∀ n as, (List.insertIdx n a as).length =
if n ≤ as.length then as.length + 1 else as.length
| 0, _ => by simp
| n + 1, [] => by simp
| n + 1, [] => by rfl
| n + 1, a :: as => by
simp only [List.insertIdx_succ_cons, List.length_cons, length_insertIdx', Nat.add_le_add_iff_right]
simp only [List.insertIdx_succ_cons, List.length_cons, length_insertIdx',
Nat.add_le_add_iff_right]
split <;> rfl
/-- This result is taken from:
@ -505,15 +529,17 @@ theorem length_insertIdx' : ∀ n as, (List.insertIdx n a as).length = if n ≤
Floris van Doorn, Mario Carneiro.
Once HepLean is updated to that version of Lean this result will be removed.
-/
-/
theorem _root_.List.getElem_insertIdx_of_ge {l : List α} {x : α} {n k : Nat} (hn : n + 1 ≤ k)
(hk : k < (List.insertIdx n x l).length) :
(List.insertIdx n x l)[k] = l[k - 1]'(by simp [length_insertIdx'] at hk; split at hk <;> omega) := by
(List.insertIdx n x l)[k] =
l[k - 1]'(by simp only [length_insertIdx'] at hk; split at hk <;> omega) := by
induction l generalizing n k with
| nil =>
cases n with
| zero =>
simp [List.insertIdx_zero, List.length_singleton] at hk
simp only [List.insertIdx_zero, List.length_cons, List.length_nil, zero_add,
Nat.lt_one_iff] at hk
omega
| succ n => simp at hk
| cons _ _ ih =>
@ -533,15 +559,14 @@ theorem _root_.List.getElem_insertIdx_of_ge {l : List α} {x : α} {n k : Nat} (
| zero => omega
| succ k => simp
lemma orderedInsert_eq_insertIdx_orderedInsertPos {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
(r : List I) (r0 : I) :
List.orderedInsert le1 r0 r = List.insertIdx (orderedInsertPos le1 r r0).1 r0 r := by
List.orderedInsert le1 r0 r = List.insertIdx (orderedInsertPos le1 r r0).1 r0 r := by
apply List.ext_get
· simp [List.orderedInsert_length]
· simp only [List.orderedInsert_length]
rw [List.length_insertIdx]
have h1 := orderedInsertPos_lt_length le1 r r0
simp at h1
simp only [List.length_cons] at h1
omega
intro n h1 h2
obtain ⟨n', hn'⟩ := (orderedInsertEquiv le1 r r0).surjective ⟨n, h1⟩
@ -549,39 +574,43 @@ lemma orderedInsert_eq_insertIdx_orderedInsertPos {I : Type} (le1 : I → I →
have hn'' : n = ((orderedInsertEquiv le1 r r0) n').val := by rw [hn']
subst hn''
rw [← orderedInsertEquiv_get]
simp
simp only [List.length_cons, Function.comp_apply, Equiv.symm_apply_apply, List.get_eq_getElem]
match n' with
| ⟨0, h0⟩ =>
simp
simp [orderedInsertEquiv]
simp only [List.getElem_cons_zero, orderedInsertEquiv, List.length_cons, Nat.succ_eq_add_one,
OrderIso.toEquiv_symm, Fin.symm_castOrderIso, Fin.zero_eta, Equiv.trans_apply,
finExtractOne_apply_eq, Fin.isValue, finExtractOne_symm_inl_apply, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, Fin.cast_mk, Fin.eta]
rw [List.getElem_insertIdx_self]
exact Nat.le_of_lt_succ (orderedInsertPos_lt_length le1 r r0)
| ⟨Nat.succ n', h0⟩ =>
simp
simp only [Nat.succ_eq_add_one, List.getElem_cons_succ, List.length_cons]
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⟩
trans (List.insertIdx (↑(orderedInsertPos le1 r r0)) r0 r).get
⟨↑((orderedInsertEquiv le1 r r0) ⟨n' +1, h0⟩), h2⟩
swap
rfl
· 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
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
exact h2⟩ := by
rw [Fin.ext_iff]
simp
simp only [List.length_cons]
simpa using hr
rw [hx]
simp [Fin.succAbove]
simp only [Fin.succAbove, Fin.castSucc_mk, Fin.mk_lt_mk, Fin.succ_mk, List.get_eq_getElem]
by_cases hn' : n' < ↑(orderedInsertPos le1 r r0)
· simp [hn']
· simp only [hn', ↓reduceIte]
erw [List.getElem_insertIdx_of_lt]
exact hn'
· simp [hn']
· simp only [hn', ↓reduceIte]
rw [List.getElem_insertIdx_of_ge]
simp
omega
· rfl
· omega
/-- The equivalence between `Fin l.length ≃ Fin (List.insertionSort r l).length` induced by the
sorting algorithm. -/
@ -593,8 +622,7 @@ def insertionSortEquiv {α : Type} (r : αα → Prop) [DecidableRel r] : (
lemma insertionSortEquiv_get {α : Type} {r : αα → Prop} [DecidableRel r] : (l : List α) →
l.get ∘ (insertionSortEquiv r l).symm = (List.insertionSort r l).get
| [] => by
simp [insertionSortEquiv]
| [] => by rfl
| a :: l => by
rw [insertionSortEquiv]
change ((a :: l).get ∘ ((Fin.equivCons (insertionSortEquiv r l))).symm) ∘
@ -617,8 +645,8 @@ lemma insertionSortEquiv_congr {α : Type} {r : αα → Prop} [DecidableRe
((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
lemma insertionSort_get_comp_insertionSortEquiv {α : Type} {r : αα → Prop} [DecidableRel r]
(l : List α) : (List.insertionSort r l).get ∘ (insertionSortEquiv r l) = l.get := by
rw [← insertionSortEquiv_get]
funext x
simp
@ -628,8 +656,6 @@ lemma insertionSort_eq_ofFn {α : Type} {r : αα → Prop} [DecidableRel r
rw [insertionSortEquiv_get (r := r)]
exact Eq.symm (List.ofFn_get (List.insertionSort r l))
def optionErase {I : Type} (l : List I) (i : Option (Fin l.length)) : List I :=
match i with
| none => l
@ -642,7 +668,7 @@ def optionEraseZ {I : Type} (l : List I) (a : I) (i : Option (Fin l.length)) : L
lemma eraseIdx_length {I : Type} (l : List I) (i : Fin l.length) :
(List.eraseIdx l i).length + 1 = l.length := by
simp [List.length_eraseIdx]
simp only [List.length_eraseIdx, Fin.is_lt, ↓reduceIte]
have hi := i.prop
omega
@ -650,21 +676,21 @@ lemma eraseIdx_cons_length {I : Type} (a : I) (l : List I) (i : Fin (a :: l).len
(List.eraseIdx (a :: l) i).length= l.length := by
simp [List.length_eraseIdx]
lemma eraseIdx_get {I : Type} (l : List I) (i : Fin l.length) :
(List.eraseIdx l i).get = l.get ∘ (Fin.cast (eraseIdx_length l i)) ∘ (Fin.cast (eraseIdx_length l i).symm 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
simp only [Function.comp_apply, List.get_eq_getElem, List.eraseIdx, List.getElem_eraseIdx]
simp [Fin.succAbove]
simp only [Fin.succAbove, Fin.coe_cast]
by_cases hi: x.castSucc < Fin.cast (by exact Eq.symm (eraseIdx_length l i)) i
· simp [ hi]
· simp only [hi, ↓reduceIte, Fin.coe_castSucc, dite_eq_left_iff, not_lt]
intro h
rw [Fin.lt_def] at hi
simp_all
simp_all only [Fin.coe_castSucc, Fin.coe_cast]
omega
· simp [ hi]
· simp only [hi, ↓reduceIte, Fin.val_succ]
rw [Fin.lt_def] at hi
simp at hi
simp only [Fin.coe_castSucc, Fin.coe_cast, not_lt] at hi
have hn : ¬ x.val < i.val := by omega
simp [hn]
@ -673,18 +699,17 @@ lemma eraseIdx_insertionSort {I : Type} (le1 : I → I → Prop) [DecidableRel l
(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, [], _ => by rfl
| 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 n, [], hn => by rfl
| Nat.succ n, (r0 :: r), hn => by
simp [insertionSortEquiv]
simp only [List.insertionSort, List.length_cons, insertionSortEquiv, Nat.succ_eq_add_one,
Equiv.trans_apply, equivCons_succ]
have hOr := orderedInsert_eraseIdx_orderedInsertEquiv_fin_succ le1
(List.insertionSort le1 r) r0 ((insertionSortEquiv le1 r) ⟨n, by simpa using hn⟩)
erw [hOr]
@ -703,12 +728,9 @@ lemma eraseIdx_insertionSort {I : Type} (le1 : I → I → Prop) [DecidableRel l
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) :
[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)
end HepLean.List