Documentation

Batteries.Data.Array.Lemmas

@[simp]
theorem getElem_fin {Cont : Type u_1} {Elem : Type u_2} {Dom : ContNatProp} {n : Nat} [GetElem Cont Nat Elem Dom] (a : Cont) (i : Fin n) (h : Dom a i) :
a[i] = a[i]
@[simp]
theorem getElem?_fin {Cont : Type u_1} {Elem : Type u_2} {Dom : ContNatProp} {n : Nat} [GetElem Cont Nat Elem Dom] (a : Cont) (i : Fin n) [Decidable (Dom a i)] :
a[i]? = a[i]?
@[simp]
theorem getElem!_fin {Cont : Type u_1} {Elem : Type u_2} {Dom : ContNatProp} {n : Nat} [GetElem Cont Nat Elem Dom] (a : Cont) (i : Fin n) [Decidable (Dom a i)] [Inhabited Elem] :
a[i]! = a[i]!
@[simp]
theorem mkArray_data {α : Type u_1} (n : Nat) (v : α) :
(mkArray n v).data = List.replicate n v
@[simp]
theorem getElem_mkArray {α : Type u_1} {i : Nat} (n : Nat) (v : α) (h : i < (mkArray n v).size) :
(mkArray n v)[i] = v
@[simp]
theorem Array.singleton_def {α : Type u_1} (v : α) :
@[simp]
theorem Array.toArray_data {α : Type u_1} (a : Array α) :
List.toArray a.data = a
@[simp]
theorem Array.data_length {α : Type u_1} {l : Array α} :
l.data.length = l.size
theorem Array.mem_data {α : Type u_1} {a : α} {l : Array α} :
a l.data a l

mem #

theorem Array.not_mem_nil {α : Type u_1} (a : α) :
¬a #[]
theorem Array.getElem?_mem {α : Type u_1} {l : Array α} {i : Fin l.size} :
l[i] l

get lemmas #

theorem Array.getElem_fin_eq_data_get {α : Type u_1} (a : Array α) (i : Fin a.data.length) :
a[i] = a.data.get i
@[simp]
theorem Array.ugetElem_eq_getElem {α : Type u_1} (a : Array α) {i : USize} (h : i.toNat < a.size) :
a[i] = a[i.toNat]
theorem Array.getElem?_eq_getElem {α : Type u_1} (a : Array α) (i : Nat) (h : i < a.size) :
a[i]? = some a[i]
theorem Array.get?_len_le {α : Type u_1} (a : Array α) (i : Nat) (h : a.size i) :
a[i]? = none
theorem Array.getElem_mem_data {α : Type u_1} {i : Nat} (a : Array α) (h : i < a.size) :
a[i] a.data
theorem Array.getElem?_eq_data_get? {α : Type u_1} (a : Array α) (i : Nat) :
a[i]? = a.data.get? i
theorem Array.get?_eq_data_get? {α : Type u_1} (a : Array α) (i : Nat) :
a.get? i = a.data.get? i
theorem Array.get!_eq_get? {α : Type u_1} {n : Nat} [Inhabited α] (a : Array α) :
a.get! n = (a.get? n).getD default
@[simp]
theorem Array.back_eq_back? {α : Type u_1} [Inhabited α] (a : Array α) :
a.back = a.back?.getD default
@[simp]
theorem Array.back?_push {α : Type u_1} {x : α} (a : Array α) :
(a.push x).back? = some x
theorem Array.back_push {α : Type u_1} {x : α} [Inhabited α] (a : Array α) :
(a.push x).back = x
theorem Array.get?_push_lt {α : Type u_1} (a : Array α) (x : α) (i : Nat) (h : i < a.size) :
(a.push x)[i]? = some a[i]
theorem Array.get?_push_eq {α : Type u_1} (a : Array α) (x : α) :
(a.push x)[a.size]? = some x
theorem Array.get?_push {α : Type u_1} {x : α} {i : Nat} {a : Array α} :
(a.push x)[i]? = if i = a.size then some x else a[i]?
@[simp]
theorem Array.get?_size {α : Type u_1} {a : Array α} :
a[a.size]? = none
@[simp]
theorem Array.data_set {α : Type u_1} (a : Array α) (i : Fin a.size) (v : α) :
(a.set i v).data = a.data.set (i) v
theorem Array.get_set_eq {α : Type u_1} (a : Array α) (i : Fin a.size) (v : α) :
(a.set i v)[i] = v
theorem Array.get?_set_eq {α : Type u_1} (a : Array α) (i : Fin a.size) (v : α) :
(a.set i v)[i]? = some v
@[simp]
theorem Array.get?_set_ne {α : Type u_1} (a : Array α) (i : Fin a.size) {j : Nat} (v : α) (h : i j) :
(a.set i v)[j]? = a[j]?
theorem Array.get?_set {α : Type u_1} (a : Array α) (i : Fin a.size) (j : Nat) (v : α) :
(a.set i v)[j]? = if i = j then some v else a[j]?
theorem Array.get_set {α : Type u_1} (a : Array α) (i : Fin a.size) (j : Nat) (hj : j < a.size) (v : α) :
(a.set i v)[j] = if i = j then v else a[j]
@[simp]
theorem Array.get_set_ne {α : Type u_1} (a : Array α) (i : Fin a.size) {j : Nat} (v : α) (hj : j < a.size) (h : i j) :
(a.set i v)[j] = a[j]
theorem Array.getElem_setD {α : Type u_1} (a : Array α) (i : Nat) (v : α) (h : i < (a.setD i v).size) :
(a.setD i v)[i] = v
theorem Array.set_set {α : Type u_1} (a : Array α) (i : Fin a.size) (v : α) (v' : α) :
(a.set i v).set i, v' = a.set i v'
theorem Array.swap_def {α : Type u_1} (a : Array α) (i : Fin a.size) (j : Fin a.size) :
a.swap i j = (a.set i (a.get j)).set j, (a.get i)
theorem Array.data_swap {α : Type u_1} (a : Array α) (i : Fin a.size) (j : Fin a.size) :
(a.swap i j).data = (a.data.set (i) (a.get j)).set (j) (a.get i)
theorem Array.get?_swap {α : Type u_1} (a : Array α) (i : Fin a.size) (j : Fin a.size) (k : Nat) :
(a.swap i j)[k]? = if j = k then some a[i] else if i = k then some a[j] else a[k]?
@[simp]
theorem Array.swapAt_def {α : Type u_1} (a : Array α) (i : Fin a.size) (v : α) :
a.swapAt i v = (a[i], a.set i v)
theorem Array.swapAt!_def {α : Type u_1} (a : Array α) (i : Nat) (v : α) (h : i < a.size) :
a.swapAt! i v = (a[i], a.set i, h v)
@[simp]
theorem Array.data_pop {α : Type u_1} (a : Array α) :
a.pop.data = a.data.dropLast
@[simp]
theorem Array.pop_empty {α : Type u_1} :
#[].pop = #[]
@[simp]
theorem Array.pop_push {α : Type u_1} {x : α} (a : Array α) :
(a.push x).pop = a
@[simp]
theorem Array.getElem_pop {α : Type u_1} (a : Array α) (i : Nat) (hi : i < a.pop.size) :
a.pop[i] = a[i]
theorem Array.eq_empty_of_size_eq_zero {α : Type u_1} {as : Array α} (h : as.size = 0) :
as = #[]
theorem Array.eq_push_pop_back_of_size_ne_zero {α : Type u_1} [Inhabited α] {as : Array α} (h : as.size 0) :
as = as.pop.push as.back
theorem Array.eq_push_of_size_ne_zero {α : Type u_1} {as : Array α} (h : as.size 0) :
∃ (bs : Array α), ∃ (c : α), as = bs.push c
theorem Array.size_eq_length_data {α : Type u_1} (as : Array α) :
as.size = as.data.length
@[simp]
theorem Array.size_swap! {α : Type u_1} (a : Array α) (i : Nat) (j : Nat) :
(a.swap! i j).size = a.size
@[simp]
theorem Array.size_reverse {α : Type u_1} (a : Array α) :
a.reverse.size = a.size
theorem Array.size_reverse.go {α : Type u_1} (as : Array α) (i : Nat) (j : Fin as.size) :
(Array.reverse.loop as i j).size = as.size
@[simp]
theorem Array.size_range {n : Nat} :
(Array.range n).size = n
@[simp]
theorem Array.reverse_data {α : Type u_1} (a : Array α) :
a.reverse.data = a.data.reverse
theorem Array.reverse_data.go {α : Type u_1} (a : Array α) (as : Array α) (i : Nat) (j : Nat) (hj : j < as.size) (h : i + j + 1 = a.size) (h₂ : as.size = a.size) (H : ∀ (k : Nat), as.data.get? k = if i k k j then a.data.get? k else a.data.reverse.get? k) (k : Nat) :
(Array.reverse.loop as i j, hj).data.get? k = a.data.reverse.get? k
theorem Array.forIn_eq_data_forIn {m : Type u_1 → Type u_2} {α : Type u_3} {β : Type u_1} [Monad m] (as : Array α) (b : β) (f : αβm (ForInStep β)) :
forIn as b f = forIn as.data b f
theorem Array.forIn_eq_data_forIn.loop {m : Type u_1 → Type u_2} {α : Type u_3} {β : Type u_1} [Monad m] (as : Array α) (f : αβm (ForInStep β)) {i : Nat} {h : i as.size} {b : β} {j : Nat} :
j + i = as.sizeArray.forIn.loop as f i h b = forIn (List.drop j as.data) b f

foldl / foldr #

theorem Array.foldl_induction {α : Type u_1} {β : Type u_2} {as : Array α} (motive : NatβProp) {init : β} (h0 : motive 0 init) {f : βαβ} (hf : ∀ (i : Fin as.size) (b : β), motive (i) bmotive (i + 1) (f b as[i])) :
motive as.size (Array.foldl f init as 0)
theorem Array.foldl_induction.go {α : Type u_1} {β : Type u_2} {as : Array α} (motive : NatβProp) {f : βαβ} (hf : ∀ (i : Fin as.size) (b : β), motive (i) bmotive (i + 1) (f b as[i])) {i : Nat} {j : Nat} {b : β} (h₁ : j as.size) (h₂ : as.size i + j) (H : motive j b) :
motive as.size (Array.foldlM.loop f as as.size i j b)
theorem Array.foldr_induction {α : Type u_1} {β : Type u_2} {as : Array α} (motive : NatβProp) {init : β} (h0 : motive as.size init) {f : αββ} (hf : ∀ (i : Fin as.size) (b : β), motive (i + 1) bmotive (i) (f as[i] b)) :
motive 0 (Array.foldr f init as as.size)
theorem Array.foldr_induction.go {α : Type u_1} {β : Type u_2} {as : Array α} (motive : NatβProp) {f : αββ} (hf : ∀ (i : Fin as.size) (b : β), motive (i + 1) bmotive (i) (f as[i] b)) {i : Nat} {b : β} (hi : i as.size) (H : motive i b) :
motive 0 (Array.foldrM.fold f as 0 i hi b)

zipWith / zip #

theorem Array.zipWith_eq_zipWith_data {α : Type u_1} {β : Type u_2} {γ : Type u_3} (f : αβγ) (as : Array α) (bs : Array β) :
(as.zipWith bs f).data = List.zipWith f as.data bs.data
theorem Array.zipWith_eq_zipWith_data.loop {α : Type u_1} {β : Type u_2} {γ : Type u_3} (f : αβγ) (as : Array α) (bs : Array β) (i : Nat) (cs : Array γ) :
i as.sizei bs.size(Array.zipWithAux f as bs i cs).data = cs.data ++ List.zipWith f (List.drop i as.data) (List.drop i bs.data)
theorem Array.size_zipWith {α : Type u_1} {β : Type u_2} {γ : Type u_3} (as : Array α) (bs : Array β) (f : αβγ) :
(as.zipWith bs f).size = min as.size bs.size
theorem Array.zip_eq_zip_data {α : Type u_1} {β : Type u_2} (as : Array α) (bs : Array β) :
(as.zip bs).data = as.data.zip bs.data
theorem Array.size_zip {α : Type u_1} {β : Type u_2} (as : Array α) (bs : Array β) :
(as.zip bs).size = min as.size bs.size

map #

@[simp]
theorem Array.mem_map {α : Type u_1} {β : Type u_2} {b : β} {f : αβ} {l : Array α} :
b Array.map f l ∃ (a : α), a l f a = b
theorem Array.mapM_eq_mapM_data {m : Type u_1 → Type u_2} {α : Type u_3} {β : Type u_1} [Monad m] [LawfulMonad m] (f : αm β) (arr : Array α) :
Array.mapM f arr = do let __do_liftList.mapM f arr.data pure { data := __do_lift }
theorem Array.mapM_map_eq_foldl {α : Type u_1} {β : Type u_2} {b : Array β} (as : Array α) (f : αβ) (i : Nat) :
Array.mapM.map f as i b = Array.foldl (fun (r : Array β) (a : α) => r.push (f a)) b as i
theorem Array.map_eq_foldl {α : Type u_1} {β : Type u_2} (as : Array α) (f : αβ) :
Array.map f as = Array.foldl (fun (r : Array β) (a : α) => r.push (f a)) #[] as 0
theorem Array.map_induction {α : Type u_1} {β : Type u_2} (as : Array α) (f : αβ) (motive : NatProp) (h0 : motive 0) (p : Fin as.sizeβProp) (hs : ∀ (i : Fin as.size), motive ip i (f as[i]) motive (i + 1)) :
motive as.size ∃ (eq : (Array.map f as).size = as.size), ∀ (i : Nat) (h : i < as.size), p i, h (Array.map f as)[i]
theorem Array.map_spec {α : Type u_1} {β : Type u_2} (as : Array α) (f : αβ) (p : Fin as.sizeβProp) (hs : ∀ (i : Fin as.size), p i (f as[i])) :
∃ (eq : (Array.map f as).size = as.size), ∀ (i : Nat) (h : i < as.size), p i, h (Array.map f as)[i]
@[simp]
theorem Array.getElem_map {α : Type u_1} {β : Type u_2} (f : αβ) (as : Array α) (i : Nat) (h : i < (Array.map f as).size) :
(Array.map f as)[i] = f as[i]

mapIdx #

theorem Array.mapIdx_induction {α : Type u_1} {β : Type u_2} (as : Array α) (f : Fin as.sizeαβ) (motive : NatProp) (h0 : motive 0) (p : Fin as.sizeβProp) (hs : ∀ (i : Fin as.size), motive ip i (f i as[i]) motive (i + 1)) :
motive as.size ∃ (eq : (as.mapIdx f).size = as.size), ∀ (i : Nat) (h : i < as.size), p i, h (as.mapIdx f)[i]
theorem Array.mapIdx_induction.go {α : Type u_1} {β : Type u_2} (as : Array α) (f : Fin as.sizeαβ) (motive : NatProp) (p : Fin as.sizeβProp) (hs : ∀ (i : Fin as.size), motive ip i (f i as[i]) motive (i + 1)) {bs : Array β} {i : Nat} {j : Nat} {h : i + j = as.size} (h₁ : j = bs.size) (h₂ : ∀ (i : Nat) (h : i < as.size) (h' : i < bs.size), p i, h bs[i]) (hm : motive j) :
let arr := Array.mapIdxM.map as f i j h bs; motive as.size ∃ (eq : arr.size = as.size), ∀ (i_1 : Nat) (h : i_1 < as.size), p i_1, h arr[i_1]
theorem Array.mapIdx_spec {α : Type u_1} {β : Type u_2} (as : Array α) (f : Fin as.sizeαβ) (p : Fin as.sizeβProp) (hs : ∀ (i : Fin as.size), p i (f i as[i])) :
∃ (eq : (as.mapIdx f).size = as.size), ∀ (i : Nat) (h : i < as.size), p i, h (as.mapIdx f)[i]
@[simp]
theorem Array.size_mapIdx {α : Type u_1} {β : Type u_2} (a : Array α) (f : Fin a.sizeαβ) :
(a.mapIdx f).size = a.size
@[simp]
theorem Array.size_zipWithIndex {α : Type u_1} (as : Array α) :
as.zipWithIndex.size = as.size
@[simp]
theorem Array.getElem_mapIdx {α : Type u_1} {β : Type u_2} (a : Array α) (f : Fin a.sizeαβ) (i : Nat) (h : i < (a.mapIdx f).size) :
(a.mapIdx f)[i] = f i, a[i]

modify #

@[simp]
theorem Array.size_modify {α : Type u_1} (a : Array α) (i : Nat) (f : αα) :
(a.modify i f).size = a.size
theorem Array.get_modify {α : Type u_1} {f : αα} {arr : Array α} {x : Nat} {i : Nat} (h : i < arr.size) :
(arr.modify x f).get i, = if x = i then f (arr.get i, h) else arr.get i, h

filter #

@[simp]
theorem Array.filter_data {α : Type u_1} (p : αBool) (l : Array α) :
(Array.filter p l 0).data = List.filter p l.data
@[simp]
theorem Array.filter_filter {α : Type u_1} {p : αBool} (q : αBool) (l : Array α) :
Array.filter p (Array.filter q l 0) 0 = Array.filter (fun (a : α) => decide (p a = true q a = true)) l 0
theorem Array.size_filter_le {α : Type u_1} (p : αBool) (l : Array α) :
(Array.filter p l 0).size l.size
@[simp]
theorem Array.mem_filter :
∀ {α : Type u_1} {x : α} {p : αBool} {as : Array α}, x Array.filter p as 0 x as p x = true
theorem Array.mem_of_mem_filter {α : Type u_1} {p : αBool} {a : α} {l : Array α} (h : a Array.filter p l 0) :
a l

filterMap #

@[simp]
theorem Array.filterMap_data {α : Type u_1} {β : Type u_2} (f : αOption β) (l : Array α) :
(Array.filterMap f l 0).data = List.filterMap f l.data
@[simp]
theorem Array.mem_filterMap {α : Type u_1} {β : Type u_2} (f : αOption β) (l : Array α) {b : β} :
b Array.filterMap f l 0 ∃ (a : α), a l f a = some b

join #

@[simp]
theorem Array.join_data {α : Type u_1} {l : Array (Array α)} :
l.join.data = (List.map Array.data l.data).join
theorem Array.mem_join {α : Type u_1} {a : α} {L : Array (Array α)} :
a L.join ∃ (l : Array α), l L a l

empty #

theorem Array.size_empty {α : Type u_1} :
#[].size = 0
theorem Array.empty_data {α : Type u_1} :
#[].data = []

append #

theorem Array.push_eq_append_singleton {α : Type u_1} (as : Array α) (x : α) :
as.push x = as ++ #[x]
@[simp]
theorem Array.mem_append {α : Type u_1} {a : α} {s : Array α} {t : Array α} :
a s ++ t a s a t
theorem Array.size_append {α : Type u_1} (as : Array α) (bs : Array α) :
(as ++ bs).size = as.size + bs.size
theorem Array.get_append_left {α : Type u_1} {i : Nat} {as : Array α} {bs : Array α} {h : i < (as ++ bs).size} (hlt : i < as.size) :
(as ++ bs)[i] = as[i]
theorem Array.get_append_right {α : Type u_1} {i : Nat} {as : Array α} {bs : Array α} {h : i < (as ++ bs).size} (hle : as.size i) (hlt : optParam (i - as.size < bs.size) ) :
(as ++ bs)[i] = bs[i - as.size]
@[simp]
theorem Array.append_nil {α : Type u_1} (as : Array α) :
as ++ #[] = as
@[simp]
theorem Array.nil_append {α : Type u_1} (as : Array α) :
#[] ++ as = as
theorem Array.append_assoc {α : Type u_1} (as : Array α) (bs : Array α) (cs : Array α) :
as ++ bs ++ cs = as ++ (bs ++ cs)

extract #

theorem Array.extract_loop_zero {α : Type u_1} (as : Array α) (bs : Array α) (start : Nat) :
Array.extract.loop as 0 start bs = bs
theorem Array.extract_loop_succ {α : Type u_1} (as : Array α) (bs : Array α) (size : Nat) (start : Nat) (h : start < as.size) :
Array.extract.loop as (size + 1) start bs = Array.extract.loop as size (start + 1) (bs.push as[start])
theorem Array.extract_loop_of_ge {α : Type u_1} (as : Array α) (bs : Array α) (size : Nat) (start : Nat) (h : start as.size) :
Array.extract.loop as size start bs = bs
theorem Array.extract_loop_eq_aux {α : Type u_1} (as : Array α) (bs : Array α) (size : Nat) (start : Nat) :
Array.extract.loop as size start bs = bs ++ Array.extract.loop as size start #[]
theorem Array.extract_loop_eq {α : Type u_1} (as : Array α) (bs : Array α) (size : Nat) (start : Nat) (h : start + size as.size) :
Array.extract.loop as size start bs = bs ++ as.extract start (start + size)
theorem Array.size_extract_loop {α : Type u_1} (as : Array α) (bs : Array α) (size : Nat) (start : Nat) :
(Array.extract.loop as size start bs).size = bs.size + min size (as.size - start)
@[simp]
theorem Array.size_extract {α : Type u_1} (as : Array α) (start : Nat) (stop : Nat) :
(as.extract start stop).size = min stop as.size - start
theorem Array.get_extract_loop_lt_aux {α : Type u_1} {i : Nat} (as : Array α) (bs : Array α) (size : Nat) (start : Nat) (hlt : i < bs.size) :
i < (Array.extract.loop as size start bs).size
theorem Array.get_extract_loop_lt {α : Type u_1} {i : Nat} (as : Array α) (bs : Array α) (size : Nat) (start : Nat) (hlt : i < bs.size) (h : optParam (i < (Array.extract.loop as size start bs).size) ) :
(Array.extract.loop as size start bs)[i] = bs[i]
theorem Array.get_extract_loop_ge_aux {α : Type u_1} {i : Nat} (as : Array α) (bs : Array α) (size : Nat) (start : Nat) (hge : i bs.size) (h : i < (Array.extract.loop as size start bs).size) :
start + i - bs.size < as.size
theorem Array.get_extract_loop_ge {α : Type u_1} {i : Nat} (as : Array α) (bs : Array α) (size : Nat) (start : Nat) (hge : i bs.size) (h : i < (Array.extract.loop as size start bs).size) (h' : optParam (start + i - bs.size < as.size) ) :
(Array.extract.loop as size start bs)[i] = as[start + i - bs.size]
theorem Array.get_extract_aux {α : Type u_1} {i : Nat} {as : Array α} {start : Nat} {stop : Nat} (h : i < (as.extract start stop).size) :
start + i < as.size
@[simp]
theorem Array.get_extract {α : Type u_1} {i : Nat} {as : Array α} {start : Nat} {stop : Nat} (h : i < (as.extract start stop).size) :
(as.extract start stop)[i] = as[start + i]
@[simp]
theorem Array.extract_all {α : Type u_1} (as : Array α) :
as.extract 0 as.size = as
theorem Array.extract_empty_of_stop_le_start {α : Type u_1} (as : Array α) {start : Nat} {stop : Nat} (h : stop start) :
as.extract start stop = #[]
theorem Array.extract_empty_of_size_le_start {α : Type u_1} (as : Array α) {start : Nat} {stop : Nat} (h : as.size start) :
as.extract start stop = #[]
@[simp]
theorem Array.extract_empty {α : Type u_1} (start : Nat) (stop : Nat) :
#[].extract start stop = #[]

any #

theorem Array.anyM_loop_iff_exists {α : Type u_1} (p : αBool) (as : Array α) (start : Nat) (stop : Nat) (h : stop as.size) :
Array.anyM.loop p as stop h start = true ∃ (i : Fin as.size), start i i < stop p as[i] = true
theorem Array.any_iff_exists {α : Type u_1} (p : αBool) (as : Array α) (start : Nat) (stop : Nat) :
as.any p start stop = true ∃ (i : Fin as.size), start i i < stop p as[i] = true
theorem Array.any_eq_true {α : Type u_1} (p : αBool) (as : Array α) :
as.any p 0 = true ∃ (i : Fin as.size), p as[i] = true
theorem Array.any_def {α : Type u_1} {p : αBool} (as : Array α) :
as.any p 0 = as.data.any p

all #

theorem Array.all_eq_not_any_not {α : Type u_1} (p : αBool) (as : Array α) (start : Nat) (stop : Nat) :
as.all p start stop = !as.any (fun (x : α) => !p x) start stop
theorem Array.all_iff_forall {α : Type u_1} (p : αBool) (as : Array α) (start : Nat) (stop : Nat) :
as.all p start stop = true ∀ (i : Fin as.size), start i i < stopp as[i] = true
theorem Array.all_eq_true {α : Type u_1} (p : αBool) (as : Array α) :
as.all p 0 = true ∀ (i : Fin as.size), p as[i] = true
theorem Array.all_def {α : Type u_1} {p : αBool} (as : Array α) :
as.all p 0 = as.data.all p
theorem Array.all_eq_true_iff_forall_mem {α : Type u_1} {p : αBool} {l : Array α} :
l.all p 0 = true ∀ (x : α), x lp x = true

contains #

theorem Array.contains_def {α : Type u_1} [DecidableEq α] {a : α} {as : Array α} :
as.contains a = true a as

erase #

swap #

@[simp]
theorem Array.get_swap_right {α : Type u_1} (a : Array α) {i : Fin a.size} {j : Fin a.size} :
(a.swap i j)[j] = a[i]
@[simp]
theorem Array.get_swap_left {α : Type u_1} (a : Array α) {i : Fin a.size} {j : Fin a.size} :
(a.swap i j)[i] = a[j]
@[simp]
theorem Array.get_swap_of_ne {α : Type u_1} {p : Nat} (a : Array α) {i : Fin a.size} {j : Fin a.size} (hp : p < a.size) (hi : p i) (hj : p j) :
(a.swap i j)[p] = a[p]
theorem Array.get_swap {α : Type u_1} (a : Array α) (i : Fin a.size) (j : Fin a.size) (k : Nat) (hk : k < a.size) :
(a.swap i j)[k] = if k = i then a[j] else if k = j then a[i] else a[k]
theorem Array.get_swap' {α : Type u_1} (a : Array α) (i : Fin a.size) (j : Fin a.size) (k : Nat) (hk' : k < (a.swap i j).size) :
(a.swap i j)[k] = if k = i then a[j] else if k = j then a[i] else a[k]
@[simp]
theorem Array.swap_swap {α : Type u_1} (a : Array α) {i : Fin a.size} {j : Fin a.size} :
(a.swap i j).swap i, j, = a
theorem Array.swap_comm {α : Type u_1} (a : Array α) {i : Fin a.size} {j : Fin a.size} :
a.swap i j = a.swap j i