insertIdx #
Proves various lemmas about List.insertIdx
.
@[simp]
@[simp]
@[simp]
theorem
List.insertIdx_succ_cons
{α : Type u}
(s : List α)
(hd x : α)
(n : Nat)
:
List.insertIdx (n + 1) x (hd :: s) = hd :: List.insertIdx n x s
theorem
List.length_insertIdx
{α : Type u}
{a : α}
(n : Nat)
(as : List α)
:
(List.insertIdx n a as).length = if n ≤ as.length then as.length + 1 else as.length
theorem
List.length_insertIdx_of_le_length
{α : Type u}
{a : α}
{n : Nat}
{as : List α}
(h : n ≤ as.length)
:
(List.insertIdx n a as).length = as.length + 1
theorem
List.length_insertIdx_of_length_lt
{α : Type u}
{a : α}
{as : List α}
{n : Nat}
(h : as.length < n)
:
(List.insertIdx n a as).length = as.length
theorem
List.eraseIdx_insertIdx
{α : Type u}
{a : α}
(n : Nat)
(l : List α)
:
(List.insertIdx n a l).eraseIdx n = l
theorem
List.insertIdx_eraseIdx_of_ge
{α : Type u}
{a : α}
(n m : Nat)
(as : List α)
:
n < as.length → n ≤ m → List.insertIdx m a (as.eraseIdx n) = (List.insertIdx (m + 1) a as).eraseIdx n
theorem
List.insertIdx_eraseIdx_of_le
{α : Type u}
{a : α}
(n m : Nat)
(as : List α)
:
n < as.length → m ≤ n → List.insertIdx m a (as.eraseIdx n) = (List.insertIdx m a as).eraseIdx (n + 1)
theorem
List.insertIdx_comm
{α : Type u}
(a b : α)
(i j : Nat)
(l : List α)
:
i ≤ j → j ≤ l.length → List.insertIdx (j + 1) b (List.insertIdx i a l) = List.insertIdx i a (List.insertIdx j b l)
theorem
List.insertIdx_of_length_lt
{α : Type u}
(l : List α)
(x : α)
(n : Nat)
(h : l.length < n)
:
List.insertIdx n x l = l
@[simp]
theorem
List.insertIdx_length_self
{α : Type u}
(l : List α)
(x : α)
:
List.insertIdx l.length x l = l ++ [x]
theorem
List.length_le_length_insertIdx
{α : Type u}
(l : List α)
(x : α)
(n : Nat)
:
l.length ≤ (List.insertIdx n x l).length
theorem
List.length_insertIdx_le_succ
{α : Type u}
(l : List α)
(x : α)
(n : Nat)
:
(List.insertIdx n x l).length ≤ l.length + 1
theorem
List.getElem_insertIdx_of_lt
{α : Type u}
{l : List α}
{x : α}
{n k : Nat}
(hn : k < n)
(hk : k < (List.insertIdx n x l).length)
:
(List.insertIdx n x l)[k] = l[k]
@[simp]
theorem
List.getElem_insertIdx_self
{α : Type u}
{l : List α}
{x : α}
{n : Nat}
(hn : n < (List.insertIdx n x l).length)
:
(List.insertIdx n x l)[n] = x
theorem
List.getElem_insertIdx_of_ge
{α : Type u}
{l : List α}
{x : α}
{n k : Nat}
(hn : n + 1 ≤ k)
(hk : k < (List.insertIdx n x l).length)
:
(List.insertIdx n x l)[k] = l[k - 1]
theorem
List.getElem_insertIdx
{α : Type u}
{l : List α}
{x : α}
{n k : Nat}
(h : k < (List.insertIdx n x l).length)
:
(List.insertIdx n x l)[k] = if h₁ : k < n then l[k] else if h₂ : k = n then x else l[k - 1]
theorem
List.getElem?_insertIdx_of_lt
{α : Type u}
{l : List α}
{x : α}
{n k : Nat}
(h : k < n)
:
(List.insertIdx n x l)[k]? = l[k]?
theorem
List.getElem?_insertIdx_self
{α : Type u}
{l : List α}
{x : α}
{n : Nat}
:
(List.insertIdx n x l)[n]? = if n ≤ l.length then some x else none
theorem
List.getElem?_insertIdx_of_ge
{α : Type u}
{l : List α}
{x : α}
{n k : Nat}
(h : n + 1 ≤ k)
:
(List.insertIdx n x l)[k]? = l[k - 1]?