Skip to content

Commit

Permalink
chore: remove unnecessary case for splitAt.go (#921)
Browse files Browse the repository at this point in the history
  • Loading branch information
kim-em authored Aug 17, 2024
1 parent b1a2ec9 commit 65f464e
Show file tree
Hide file tree
Showing 2 changed files with 17 additions and 3 deletions.
4 changes: 1 addition & 3 deletions Batteries/Data/List/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -117,10 +117,8 @@ splitAt 2 [a, b, c] = ([a, b], [c])
```
-/
def splitAt (n : Nat) (l : List α) : List α × List α := go l n [] where
/-- Auxiliary for `splitAt`: `splitAt.go l n xs acc = (acc.reverse ++ take n xs, drop n xs)`
if `n < length xs`, else `(l, [])`. -/
/-- Auxiliary for `splitAt`: `splitAt.go xs n acc = (acc.reverse ++ take n xs, drop n xs)`. -/
go : List α → Nat → List α → List α × List α
| [], _, _ => (l, [])
| x :: xs, n+1, acc => go xs n (x :: acc)
| xs, _, acc => (acc.reverse, xs)

Expand Down
16 changes: 16 additions & 0 deletions Batteries/Data/List/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -189,6 +189,22 @@ theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) :

@[deprecated (since := "2024-05-06")] alias length_removeNth := length_eraseIdx

/-! ### splitAt -/

theorem splitAt_go (n : Nat) (l acc : List α) :
splitAt.go l n acc = (acc.reverse ++ l.take n, l.drop n) := by
induction l generalizing n acc with
| nil => simp [splitAt.go]
| cons x xs ih =>
cases n with
| zero => simp [splitAt.go]
| succ n =>
rw [splitAt.go, take_succ_cons, drop_succ_cons, ih n (x :: acc),
reverse_cons, append_assoc, singleton_append]

theorem splitAt_eq (n : Nat) (l : List α) : splitAt n l = (l.take n, l.drop n) := by
rw [splitAt, splitAt_go, reverse_nil, nil_append]

/-! ### eraseP -/

@[simp] theorem extractP_eq_find?_eraseP
Expand Down

0 comments on commit 65f464e

Please sign in to comment.