Skip to content

Commit

Permalink
feat: lemmas about dropWhile (#445)
Browse files Browse the repository at this point in the history
  • Loading branch information
kim-em authored Dec 13, 2023
1 parent cc8713e commit 1bd72ed
Showing 1 changed file with 20 additions and 0 deletions.
20 changes: 20 additions & 0 deletions Std/Data/List/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -102,6 +102,11 @@ theorem not_mem_cons_of_ne_of_not_mem {a y : α} {l : List α} : a ≠ y → a
theorem ne_and_not_mem_of_not_mem_cons {a y : α} {l : List α} : a ∉ y::l → a ≠ y ∧ a ∉ l :=
fun p => ⟨ne_of_not_mem_cons p, not_mem_of_not_mem_cons p⟩

/-! ### isEmpty -/

@[simp] theorem isEmpty_nil : ([] : List α).isEmpty = true := rfl
@[simp] theorem isEmpty_cons : (x :: xs : List α).isEmpty = false := rfl

/-! ### append -/

theorem append_eq_append : List.append l₁ l₂ = l₁ ++ l₂ := rfl
Expand Down Expand Up @@ -2062,6 +2067,21 @@ theorem disjoint_take_drop : ∀ {l : List α}, l.Nodup → m ≤ n → Disjoint
| [] => rfl
| x :: xs => by simp [takeWhile, dropWhile]; cases p x <;> simp [takeWhile_append_dropWhile p xs]

@[simp] theorem dropWhile_nil : ([] : List α).dropWhile p = [] := rfl

theorem dropWhile_cons :
(x :: xs : List α).dropWhile p = if p x then xs.dropWhile p else x :: xs := by
split <;> simp_all [dropWhile]

theorem dropWhile_append {xs ys : List α} :
(xs ++ ys).dropWhile p =
if (xs.dropWhile p).isEmpty then ys.dropWhile p else xs.dropWhile p ++ ys := by
induction xs with
| nil => simp
| cons h t ih =>
simp only [cons_append, dropWhile_cons]
split <;> simp_all

/-! ### Chain -/

--Porting note: attribute in Lean3, but not in Lean4 Std so added here instead
Expand Down

0 comments on commit 1bd72ed

Please sign in to comment.