Skip to content

Commit

Permalink
feat: add another List.find?_eq_some lemma (#5974)
Browse files Browse the repository at this point in the history
  • Loading branch information
kim-em authored Nov 6, 2024
1 parent 76d32cb commit 1e98fd7
Show file tree
Hide file tree
Showing 3 changed files with 34 additions and 4 deletions.
8 changes: 6 additions & 2 deletions src/Init/Data/List/Find.lean
Original file line number Diff line number Diff line change
Expand Up @@ -206,7 +206,8 @@ theorem IsInfix.findSome?_eq_none {l₁ l₂ : List α} {f : α → Option β} (
@[simp] theorem find?_eq_none : find? p l = none ↔ ∀ x ∈ l, ¬ p x := by
induction l <;> simp [find?_cons]; split <;> simp [*]

theorem find?_eq_some : xs.find? p = some b ↔ p b ∧ ∃ as bs, xs = as ++ b :: bs ∧ ∀ a ∈ as, !p a := by
theorem find?_eq_some_iff_append :
xs.find? p = some b ↔ p b ∧ ∃ as bs, xs = as ++ b :: bs ∧ ∀ a ∈ as, !p a := by
induction xs with
| nil => simp
| cons x xs ih =>
Expand Down Expand Up @@ -242,6 +243,9 @@ theorem find?_eq_some : xs.find? p = some b ↔ p b ∧ ∃ as bs, xs = as ++ b
cases h₁
simp

@[deprecated find?_eq_some_iff_append (since := "2024-11-06")]
abbrev find?_eq_some := @find?_eq_some_iff_append

@[simp]
theorem find?_cons_eq_some : (a :: xs).find? p = some b ↔ (p a ∧ a = b) ∨ (!p a ∧ xs.find? p = some b) := by
rw [find?_cons]
Expand Down Expand Up @@ -347,7 +351,7 @@ theorem find?_flatten_eq_some {xs : List (List α)} {p : α → Bool} {a : α} :
xs.flatten.find? p = some a ↔
p a ∧ ∃ as ys zs bs, xs = as ++ (ys ++ a :: zs) :: bs ∧
(∀ a ∈ as, ∀ x ∈ a, !p x) ∧ (∀ x ∈ ys, !p x) := by
rw [find?_eq_some]
rw [find?_eq_some_iff_append]
constructor
· rintro ⟨h, ⟨ys, zs, h₁, h₂⟩⟩
refine ⟨h, ?_⟩
Expand Down
26 changes: 26 additions & 0 deletions src/Init/Data/List/Nat/Find.lean
Original file line number Diff line number Diff line change
Expand Up @@ -9,6 +9,32 @@ import Init.Data.List.Find

namespace List

open Nat

theorem find?_eq_some_iff_getElem {xs : List α} {p : α → Bool} {b : α} :
xs.find? p = some b ↔ p b ∧ ∃ i h, xs[i] = b ∧ ∀ j : Nat, (hj : j < i) → !p xs[j] := by
rw [find?_eq_some_iff_append]
simp only [Bool.not_eq_eq_eq_not, Bool.not_true, exists_and_right, and_congr_right_iff]
intro w
constructor
· rintro ⟨as, ⟨bs, rfl⟩, h⟩
refine ⟨as.length, ⟨?_, ?_, ?_⟩⟩
· simp only [length_append, length_cons]
refine Nat.lt_add_of_pos_right (zero_lt_succ bs.length)
· rw [getElem_append_right (Nat.le_refl as.length)]
simp
· intro j h'
rw [getElem_append_left h']
exact h _ (getElem_mem h')
· rintro ⟨i, h, rfl, h'⟩
refine ⟨xs.take i, ⟨xs.drop (i+1), ?_⟩, ?_⟩
· rw [getElem_cons_drop, take_append_drop]
· intro a m
rw [mem_take_iff_getElem] at m
obtain ⟨j, h, rfl⟩ := m
apply h'
omega

theorem findIdx?_eq_some_le_of_findIdx?_eq_some {xs : List α} {p q : α → Bool} (w : ∀ x ∈ xs, p x → q x) {i : Nat}
(h : xs.findIdx? p = some i) : ∃ j, j ≤ i ∧ xs.findIdx? q = some j := by
simp only [findIdx?_eq_findSome?_enum] at h
Expand Down
4 changes: 2 additions & 2 deletions src/Init/Data/List/Nat/Range.lean
Original file line number Diff line number Diff line change
Expand Up @@ -108,7 +108,7 @@ theorem range'_eq_append_iff : range' s n = xs ++ ys ↔ ∃ k, k ≤ n ∧ xs =

@[simp] theorem find?_range'_eq_some {s n : Nat} {i : Nat} {p : Nat → Bool} :
(range' s n).find? p = some i ↔ p i ∧ i ∈ range' s n ∧ ∀ j, s ≤ j → j < i → !p j := by
rw [find?_eq_some]
rw [find?_eq_some_iff_append]
simp only [Bool.not_eq_eq_eq_not, Bool.not_true, exists_and_right, mem_range'_1,
and_congr_right_iff]
simp only [range'_eq_append_iff, eq_comm (a := i :: _), range'_eq_cons_iff]
Expand Down Expand Up @@ -282,7 +282,7 @@ theorem find?_iota_eq_none {n : Nat} {p : Nat → Bool} :

@[simp] theorem find?_iota_eq_some {n : Nat} {i : Nat} {p : Nat → Bool} :
(iota n).find? p = some i ↔ p i ∧ i ∈ iota n ∧ ∀ j, i < j → j ≤ n → !p j := by
rw [find?_eq_some]
rw [find?_eq_some_iff_append]
simp only [iota_eq_reverse_range', reverse_eq_append_iff, reverse_cons, append_assoc, cons_append,
nil_append, Bool.not_eq_eq_eq_not, Bool.not_true, exists_and_right, mem_reverse, mem_range'_1,
and_congr_right_iff]
Expand Down

0 comments on commit 1e98fd7

Please sign in to comment.