Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

feat: fill in proof of Array.data_erase #690

Merged
merged 20 commits into from
Sep 26, 2024
Merged
Show file tree
Hide file tree
Changes from 18 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
57 changes: 56 additions & 1 deletion Batteries/Data/Array/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -115,9 +115,64 @@ theorem mem_join : ∀ {L : Array (Array α)}, a ∈ L.join ↔ ∃ l, l ∈ L
· rintro ⟨s, h₁, h₂⟩
refine ⟨s.data, ⟨⟨s, h₁, rfl⟩, h₂⟩⟩

/-! ### indexOf? -/

Seppel3210 marked this conversation as resolved.
Show resolved Hide resolved
theorem indexOf?_data [BEq α] {a : α} {l : Array α} :
fgdorais marked this conversation as resolved.
Show resolved Hide resolved
l.data.indexOf? a = (l.indexOf? a).map Fin.val := by
simpa using aux l 0
where
aux (l : Array α) (i : Nat) :
((l.data.drop i).indexOf? a).map (·+i) = (indexOfAux l a i).map Fin.val := by
rw [indexOfAux]
if h : i < l.size then
rw [List.drop_eq_getElem_cons h, ←getElem_eq_data_getElem, List.indexOf?_cons]
if h' : l[i] == a then
simp [h, h']
else
simp [h, h', ←aux l (i+1), Function.comp_def, ←Nat.add_assoc, Nat.add_right_comm]
else
have h' : l.size ≤ i := Nat.le_of_not_lt h
simp [h, List.drop_of_length_le h', List.indexOf?]
termination_by l.size - i

/-! ### erase -/

Seppel3210 marked this conversation as resolved.
Show resolved Hide resolved
@[simp] proof_wanted data_erase [BEq α] {l : Array α} {a : α} : (l.erase a).data = l.data.erase a
theorem eraseIdx_data_swap {l : Array α} (i : Nat) (lt : i + 1 < size l) :
(l.swap ⟨i+1, lt⟩ ⟨i, Nat.lt_of_succ_lt lt⟩).data.eraseIdx (i+1) = l.data.eraseIdx i := by
let ⟨xs⟩ := l
induction i generalizing xs <;> let x₀::x₁::xs := xs
case zero => simp [swap, get]
case succ i ih _ =>
have lt' := Nat.lt_of_succ_lt_succ lt
have : (swap ⟨x₀::x₁::xs⟩ ⟨i.succ + 1, lt⟩ ⟨i.succ, Nat.lt_of_succ_lt lt⟩).data
= x₀::(swap ⟨x₁::xs⟩ ⟨i + 1, lt'⟩ ⟨i, Nat.lt_of_succ_lt lt'⟩).data := by
simp [swap_def, getElem_eq_data_getElem]
simp [this, ih]

theorem data_feraseIdx {l : Array α} (i : Fin l.size) :
Seppel3210 marked this conversation as resolved.
Show resolved Hide resolved
(l.feraseIdx i).data = l.data.eraseIdx i := by
induction l, i using feraseIdx.induct with
Seppel3210 marked this conversation as resolved.
Show resolved Hide resolved
| @case1 a i lt a' i' ih =>
rw [feraseIdx]
simp [lt, ih, a', eraseIdx_data_swap i lt]
| case2 a i lt =>
have : i + 1 ≥ a.size := Nat.ge_of_not_lt lt
have last : i + 1 = a.size := Nat.le_antisymm i.is_lt this
simp [feraseIdx, lt, List.dropLast_eq_eraseIdx last]

@[simp] theorem data_erase [BEq α] (l : Array α) (a : α) : (l.erase a).data = l.data.erase a := by
match h : indexOf? l a with
| none =>
simp only [erase, h]
apply Eq.symm
apply List.erase_of_forall_bne
rw [←List.indexOf?_eq_none_iff, indexOf?_data, h, Option.map_none']
| some i =>
simp only [erase, h]
rw [data_feraseIdx, ←List.eraseIdx_indexOf_eq_erase]
congr
rw [List.indexOf_eq_indexOf?, indexOf?_data]
simp [h]

/-! ### shrink -/

Expand Down
56 changes: 56 additions & 0 deletions Batteries/Data/List/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -22,6 +22,22 @@ open Nat
@[simp] theorem next?_nil : @next? α [] = none := rfl
@[simp] theorem next?_cons (a l) : @next? α (a :: l) = some (a, l) := rfl

/-! ### dropLast -/

Seppel3210 marked this conversation as resolved.
Show resolved Hide resolved
theorem dropLast_eq_eraseIdx {xs : List α} {i : Nat} (last_idx : i + 1 = xs.length) :
xs.dropLast = List.eraseIdx xs i := by
induction i generalizing xs with
| zero =>
let [x] := xs
Seppel3210 marked this conversation as resolved.
Show resolved Hide resolved
rfl
| succ n ih =>
let x::xs := xs
simp at last_idx
rw [dropLast, eraseIdx]
congr
exact ih last_idx
exact fun _ => nomatch xs

/-! ### get? -/

@[deprecated getElem_eq_iff (since := "2024-06-12")]
Expand Down Expand Up @@ -210,6 +226,25 @@ theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) :

@[deprecated (since := "2024-04-22")] alias sublist.erase := Sublist.erase

theorem erase_of_forall_bne [BEq α] (a : α) (xs : List α) (h : ∀ (x : α), x ∈ xs → ¬x == a) :
xs.erase a = xs := by
rw [erase_eq_eraseP', eraseP_of_forall_not h]

-- TODO a version of the above theorem with LawfulBEq and ∉

/-! ### findIdx? -/

theorem findIdx_eq_findIdx? (p : α → Bool) (l : List α) :
l.findIdx p = (match l.findIdx? p with | some i => i | none => l.length) := by
Seppel3210 marked this conversation as resolved.
Show resolved Hide resolved
induction l with
| nil => rfl
| cons x xs ih =>
rw [findIdx_cons, findIdx?_cons]
if h : p x then
simp [h]
else
cases h' : findIdx? p xs <;> simp [h, h', ih]

/-! ### replaceF -/

theorem replaceF_nil : [].replaceF p = [] := rfl
Expand Down Expand Up @@ -553,6 +588,14 @@ theorem indexesOf_cons [BEq α] : (x :: xs : List α).indexesOf y =
bif x == y then 0 :: (xs.indexesOf y).map (· + 1) else (xs.indexesOf y).map (· + 1) := by
simp [indexesOf, findIdxs_cons]

@[simp] theorem eraseIdx_indexOf_eq_erase [BEq α] (a : α) (l : List α) :
Seppel3210 marked this conversation as resolved.
Show resolved Hide resolved
l.eraseIdx (l.indexOf a) = l.erase a := by
induction l with
| nil => rfl
| cons x xs ih =>
rw [List.erase, indexOf_cons]
cases x == a <;> simp [ih]

theorem indexOf_mem_indexesOf [BEq α] [LawfulBEq α] {xs : List α} (m : x ∈ xs) :
xs.indexOf x ∈ xs.indexesOf x := by
induction xs with
Expand All @@ -567,6 +610,19 @@ theorem indexOf_mem_indexesOf [BEq α] [LawfulBEq α] {xs : List α} (m : x ∈
specialize ih m
simpa

@[simp] theorem indexOf?_nil [BEq α] : ([] : List α).indexOf? x = none := rfl
Seppel3210 marked this conversation as resolved.
Show resolved Hide resolved
theorem indexOf?_cons [BEq α] :
(x :: xs : List α).indexOf? y = if x == y then some 0 else (xs.indexOf? y).map Nat.succ := by
simp [indexOf?]

theorem indexOf?_eq_none_iff [BEq α] {a : α} {l : List α} :
l.indexOf? a = none ↔ ∀ x ∈ l, ¬x == a := by
simp [indexOf?, findIdx?_eq_none_iff]

theorem indexOf_eq_indexOf? [BEq α] (a : α) (l : List α) :
l.indexOf a = (match l.indexOf? a with | some i => i | none => l.length) := by
simp [indexOf, indexOf?, findIdx_eq_findIdx?]

/-! ### insertP -/

theorem insertP_loop (a : α) (l r : List α) :
Expand Down