Documentation

Mathlib.Data.List.InsertNth

insertNth #

Proves various lemmas about List.insertNth.

@[simp]
theorem List.insertNth_zero {α : Type u} (s : List α) (x : α) :
List.insertNth 0 x s = x :: s
@[simp]
theorem List.insertNth_succ_nil {α : Type u} (n : ) (a : α) :
List.insertNth (n + 1) a [] = []
@[simp]
theorem List.insertNth_succ_cons {α : Type u} (s : List α) (hd : α) (x : α) (n : ) :
List.insertNth (n + 1) x (hd :: s) = hd :: List.insertNth n x s
theorem List.length_insertNth {α : Type u} {a : α} (n : ) (as : List α) :
n as.length(List.insertNth n a as).length = as.length + 1
theorem List.eraseIdx_insertNth {α : Type u} {a : α} (n : ) (l : List α) :
(List.insertNth n a l).eraseIdx n = l
@[deprecated List.eraseIdx_insertNth]
theorem List.removeNth_insertNth {α : Type u} {a : α} (n : ) (l : List α) :
(List.insertNth n a l).eraseIdx n = l

Alias of List.eraseIdx_insertNth.

theorem List.insertNth_eraseIdx_of_ge {α : Type u} {a : α} (n : ) (m : ) (as : List α) :
n < as.lengthn mList.insertNth m a (as.eraseIdx n) = (List.insertNth (m + 1) a as).eraseIdx n
@[deprecated List.insertNth_eraseIdx_of_ge]
theorem List.insertNth_removeNth_of_ge {α : Type u} {a : α} (n : ) (m : ) (as : List α) :
n < as.lengthn mList.insertNth m a (as.eraseIdx n) = (List.insertNth (m + 1) a as).eraseIdx n

Alias of List.insertNth_eraseIdx_of_ge.

theorem List.insertNth_eraseIdx_of_le {α : Type u} {a : α} (n : ) (m : ) (as : List α) :
n < as.lengthm nList.insertNth m a (as.eraseIdx n) = (List.insertNth m a as).eraseIdx (n + 1)
@[deprecated List.insertNth_eraseIdx_of_le]
theorem List.insertNth_removeNth_of_le {α : Type u} {a : α} (n : ) (m : ) (as : List α) :
n < as.lengthm nList.insertNth m a (as.eraseIdx n) = (List.insertNth m a as).eraseIdx (n + 1)

Alias of List.insertNth_eraseIdx_of_le.

theorem List.insertNth_comm {α : Type u} (a : α) (b : α) (i : ) (j : ) (l : List α) :
i jj l.lengthList.insertNth (j + 1) b (List.insertNth i a l) = List.insertNth i a (List.insertNth j b l)
theorem List.mem_insertNth {α : Type u} {a : α} {b : α} {n : } {l : List α} :
n l.length(a List.insertNth n b l a = b a l)
theorem List.insertNth_of_length_lt {α : Type u} (l : List α) (x : α) (n : ) (h : l.length < n) :
@[simp]
theorem List.insertNth_length_self {α : Type u} (l : List α) (x : α) :
List.insertNth l.length x l = l ++ [x]
theorem List.length_le_length_insertNth {α : Type u} (l : List α) (x : α) (n : ) :
l.length (List.insertNth n x l).length
theorem List.length_insertNth_le_succ {α : Type u} (l : List α) (x : α) (n : ) :
(List.insertNth n x l).length l.length + 1
theorem List.get_insertNth_of_lt {α : Type u} (l : List α) (x : α) (n : ) (k : ) (hn : k < n) (hk : k < l.length) (hk' : optParam (k < (List.insertNth n x l).length) ) :
(List.insertNth n x l).get k, hk' = l.get k, hk
@[deprecated List.get_insertNth_of_lt]
theorem List.nthLe_insertNth_of_lt {α : Type u} (l : List α) (x : α) (n : ) (k : ) :
k < n∀ (hk : k < l.length) (hk' : optParam (k < (List.insertNth n x l).length) ), (List.insertNth n x l).nthLe k hk' = l.nthLe k hk
@[simp]
theorem List.get_insertNth_self {α : Type u} (l : List α) (x : α) (n : ) (hn : n l.length) (hn' : optParam (n < (List.insertNth n x l).length) ) :
(List.insertNth n x l).get n, hn' = x
@[simp, deprecated List.get_insertNth_self]
theorem List.nthLe_insertNth_self {α : Type u} (l : List α) (x : α) (n : ) (hn : n l.length) (hn' : optParam (n < (List.insertNth n x l).length) ) :
(List.insertNth n x l).nthLe n hn' = x
theorem List.get_insertNth_add_succ {α : Type u} (l : List α) (x : α) (n : ) (k : ) (hk' : n + k < l.length) (hk : optParam (n + k + 1 < (List.insertNth n x l).length) ) :
(List.insertNth n x l).get n + k + 1, hk = l.get n + k, hk'
@[deprecated List.get_insertNth_add_succ]
theorem List.nthLe_insertNth_add_succ {α : Type u} (l : List α) (x : α) (n : ) (k : ) (hk' : n + k < l.length) (hk : optParam (n + k + 1 < (List.insertNth n x l).length) ) :
(List.insertNth n x l).nthLe (n + k + 1) hk = l.nthLe (n + k) hk'