Skip to content

Commit

Permalink
refactor: move theorems about lists from mathlib
Browse files Browse the repository at this point in the history
* `List.isEmpty_iff_eq_nil` and `List.modifyHead_modifyHead` are from
  `Mathlib.Data.List.Basic`. I removed the `simp` priority and `nolint`
  attribute from `modifyHead_modifyHead` because we don't need it
  anymore.
* `List.cons_prefix_cons` is from `Mathlib.Data.List.Infix` and was
  renamed `List.cons_prefix_iff`.

We need these theorems to prove `String.splitOn_of_valid`.

Co-authored-by: Kim Morrison <kim@tqft.net>
  • Loading branch information
chabulhwi and semorrison committed May 9, 2024
1 parent 92ad41e commit 3b0a7fc
Showing 1 changed file with 19 additions and 0 deletions.
19 changes: 19 additions & 0 deletions Batteries/Data/List/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -7,6 +7,7 @@ import Batteries.Control.ForInStep.Lemmas
import Batteries.Data.List.Basic
import Batteries.Tactic.Init
import Batteries.Tactic.Alias
import Batteries.Tactic.Lint.Simp

namespace List

Expand Down Expand Up @@ -138,6 +139,8 @@ theorem drop_left' {l₁ l₂ : List α} {n} (h : length l₁ = n) : drop n (l
@[simp] theorem isEmpty_nil : ([] : List α).isEmpty = true := rfl
@[simp] theorem isEmpty_cons : (x :: xs : List α).isEmpty = false := rfl

theorem isEmpty_iff_eq_nil {l : List α} : l.isEmpty ↔ l = [] := by cases l <;> simp [isEmpty]

/-! ### append -/

theorem append_eq_append : List.append l₁ l₂ = l₁ ++ l₂ := rfl
Expand Down Expand Up @@ -1156,6 +1159,11 @@ theorem drop_eq_nil_of_eq_nil : ∀ {as : List α} {i}, as = [] → as.drop i =
theorem ne_nil_of_drop_ne_nil {as : List α} {i : Nat} (h: as.drop i ≠ []) : as ≠ [] :=
mt drop_eq_nil_of_eq_nil h

/-! ### modifyHead -/

@[simp] theorem modifyHead_modifyHead (l : List α) (f g : α → α) :
(l.modifyHead f).modifyHead g = l.modifyHead (g ∘ f) := by cases l <;> simp [modifyHead]

/-! ### modify nth -/

theorem modifyNthTail_id : ∀ n (l : List α), l.modifyNthTail id n = l
Expand Down Expand Up @@ -2356,6 +2364,17 @@ theorem IsInfix.filter (p : α → Bool) ⦃l₁ l₂ : List α⦄ (h : l₁ <:+
obtain ⟨xs, ys, rfl⟩ := h
rw [filter_append, filter_append]; apply infix_append _

theorem cons_prefix_cons : a :: l₁ <+: b :: l₂ ↔ a = b ∧ l₁ <+: l₂ := by
constructor
· rintro ⟨L, hL⟩
simp only [cons_append] at hL
injection hL with hLLeft hLRight
exact ⟨hLLeft, ⟨L, hLRight⟩⟩
· rintro ⟨rfl, h⟩
rwa [prefix_cons_inj]

@[deprecated] alias cons_prefix_iff := cons_prefix_cons

/-! ### drop -/

theorem mem_of_mem_drop {n} {l : List α} (h : a ∈ l.drop n) : a ∈ l := drop_subset _ _ h
Expand Down

0 comments on commit 3b0a7fc

Please sign in to comment.