Skip to content

Commit

Permalink
Browse files Browse the repository at this point in the history
  • Loading branch information
kim-em committed Dec 9, 2024
1 parent b063652 commit f2f45b2
Show file tree
Hide file tree
Showing 2 changed files with 1 addition and 5 deletions.
4 changes: 0 additions & 4 deletions Mathlib/Data/List/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -487,10 +487,6 @@ theorem forall_mem_iff_getElem {l : List α} {p : α → Prop} :
(∀ x ∈ l, p x) ↔ ∀ (i : ℕ) (_ : i < l.length), p l[i] := by
simp [mem_iff_getElem, @forall_swap α]

theorem getElem_cons {l : List α} {a : α} {n : ℕ} (h : n < (a :: l).length) :
(a :: l)[n] = if hn : n = 0 then a else l[n - 1]'(by rw [length_cons] at h; omega) := by
cases n <;> simp

theorem get_tail (l : List α) (i) (h : i < l.tail.length)
(h' : i + 1 < l.length := (by simp only [length_tail] at h; omega)) :
l.tail.get ⟨i, h⟩ = l.get ⟨i + 1, h'⟩ := by
Expand Down
2 changes: 1 addition & 1 deletion lake-manifest.json
Original file line number Diff line number Diff line change
Expand Up @@ -65,7 +65,7 @@
"type": "git",
"subDir": null,
"scope": "leanprover-community",
"rev": "fb26177ca3f6862cb9b57ea4083bf163f5199c6e",
"rev": "c792d1da5f40fe2df83ae19cdacba457040e1a55",
"name": "batteries",
"manifestFile": "lake-manifest.json",
"inputRev": "lean-pr-testing-6342",
Expand Down

0 comments on commit f2f45b2

Please sign in to comment.