Skip to content

Commit

Permalink
feat: duplicate List.attach/attachWith/pmap API for Array (#6132)
Browse files Browse the repository at this point in the history
This PR duplicates the verification API for
`List.attach`/`attachWith`/`pmap` over to `Array`.
  • Loading branch information
kim-em authored Nov 20, 2024
1 parent 3c75551 commit c8b4f6b
Show file tree
Hide file tree
Showing 8 changed files with 502 additions and 50 deletions.
378 changes: 373 additions & 5 deletions src/Init/Data/Array/Attach.lean

Large diffs are not rendered by default.

6 changes: 6 additions & 0 deletions src/Init/Data/Array/Find.lean
Original file line number Diff line number Diff line change
Expand Up @@ -272,4 +272,10 @@ theorem find?_mkArray_eq_none {n : Nat} {a : α} {p : α → Bool} :
((mkArray n a).find? p).get h = a := by
simp [mkArray_eq_toArray_replicate]

theorem find?_pmap {P : α → Prop} (f : (a : α) → P a → β) (xs : Array α)
(H : ∀ (a : α), a ∈ xs → P a) (p : β → Bool) :
(xs.pmap f H).find? p = (xs.attach.find? (fun ⟨a, m⟩ => p (f a (H a m)))).map fun ⟨a, m⟩ => f a (H a m) := by
simp only [pmap_eq_map_attach, find?_map]
rfl

end Array
101 changes: 81 additions & 20 deletions src/Init/Data/Array/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -23,6 +23,9 @@ import Init.TacticsExtra

namespace Array

@[simp] theorem mem_toArray {a : α} {l : List α} : a ∈ l.toArray ↔ a ∈ l := by
simp [mem_def]

@[simp] theorem getElem_mk {xs : List α} {i : Nat} (h : i < xs.length) : (Array.mk xs)[i] = xs[i] := rfl

theorem getElem_eq_getElem_toList {a : Array α} (h : i < a.size) : a[i] = a.toList[i] := rfl
Expand All @@ -36,12 +39,21 @@ theorem getElem?_eq_getElem {a : Array α} {i : Nat} (h : i < a.size) : a[i]? =
· rw [getElem?_neg a i h]
simp_all

@[simp] theorem none_eq_getElem?_iff {a : Array α} {i : Nat} : none = a[i]? ↔ a.size ≤ i := by
simp [eq_comm (a := none)]

theorem getElem?_eq {a : Array α} {i : Nat} :
a[i]? = if h : i < a.size then some a[i] else none := by
split
· simp_all [getElem?_eq_getElem]
· simp_all

theorem getElem?_eq_some_iff {a : Array α} : a[i]? = some b ↔ ∃ h : i < a.size, a[i] = b := by
simp [getElem?_eq]

theorem some_eq_getElem?_iff {a : Array α} : some b = a[i]? ↔ ∃ h : i < a.size, a[i] = b := by
rw [eq_comm, getElem?_eq_some_iff]

theorem getElem?_eq_getElem?_toList (a : Array α) (i : Nat) : a[i]? = a.toList[i]? := by
rw [getElem?_eq]
split <;> simp_all
Expand All @@ -66,6 +78,35 @@ theorem getElem_push (a : Array α) (x : α) (i : Nat) (h : i < (a.push x).size)
@[deprecated getElem_push_lt (since := "2024-10-21")] abbrev get_push_lt := @getElem_push_lt
@[deprecated getElem_push_eq (since := "2024-10-21")] abbrev get_push_eq := @getElem_push_eq

@[simp] theorem mem_push {a : Array α} {x y : α} : x ∈ a.push y ↔ x ∈ a ∨ x = y := by
simp [mem_def]

theorem mem_push_self {a : Array α} {x : α} : x ∈ a.push x :=
mem_push.2 (Or.inr rfl)

theorem mem_push_of_mem {a : Array α} {x : α} (y : α) (h : x ∈ a) : x ∈ a.push y :=
mem_push.2 (Or.inl h)

theorem getElem_of_mem {a} {l : Array α} (h : a ∈ l) : ∃ (n : Nat) (h : n < l.size), l[n]'h = a := by
cases l
simp [List.getElem_of_mem (by simpa using h)]

theorem getElem?_of_mem {a} {l : Array α} (h : a ∈ l) : ∃ n : Nat, l[n]? = some a :=
let ⟨n, _, e⟩ := getElem_of_mem h; ⟨n, e ▸ getElem?_eq_getElem _⟩

theorem mem_of_getElem? {l : Array α} {n : Nat} {a : α} (e : l[n]? = some a) : a ∈ l :=
let ⟨_, e⟩ := getElem?_eq_some_iff.1 e; e ▸ getElem_mem ..

theorem mem_iff_getElem {a} {l : Array α} : a ∈ l ↔ ∃ (n : Nat) (h : n < l.size), l[n]'h = a :=
⟨getElem_of_mem, fun ⟨_, _, e⟩ => e ▸ getElem_mem ..⟩

theorem mem_iff_getElem? {a} {l : Array α} : a ∈ l ↔ ∃ n : Nat, l[n]? = some a := by
simp [getElem?_eq_some_iff, mem_iff_getElem]

theorem forall_getElem {l : Array α} {p : α → Prop} :
(∀ (n : Nat) h, p (l[n]'h)) ↔ ∀ a, a ∈ l → p a := by
cases l; simp [List.forall_getElem]

@[simp] theorem get!_eq_getElem! [Inhabited α] (a : Array α) (i : Nat) : a.get! i = a[i]! := by
simp [getElem!_def, get!, getD]
split <;> rename_i h
Expand Down Expand Up @@ -93,9 +134,6 @@ We prefer to pull `List.toArray` outwards.
(a.toArrayAux b).size = b.size + a.length := by
simp [size]

@[simp] theorem mem_toArray {a : α} {l : List α} : a ∈ l.toArray ↔ a ∈ l := by
simp [mem_def]

@[simp] theorem push_toArray (l : List α) (a : α) : l.toArray.push a = (l ++ [a]).toArray := by
apply ext'
simp
Expand Down Expand Up @@ -605,19 +643,6 @@ theorem getElem?_mkArray (n : Nat) (v : α) (i : Nat) :

theorem not_mem_nil (a : α) : ¬ a ∈ #[] := nofun

theorem getElem_of_mem {a : α} {as : Array α} :
a ∈ as → (∃ (n : Nat) (h : n < as.size), as[n]'h = a) := by
intro ha
rcases List.getElem_of_mem ha.val with ⟨i, hbound, hi⟩
exists i
exists hbound

theorem getElem?_of_mem {a : α} {as : Array α} :
a ∈ as → ∃ (n : Nat), as[n]? = some a := by
intro ha
rcases List.getElem?_of_mem ha.val with ⟨i, hi⟩
exists i

@[simp] theorem mem_dite_empty_left {x : α} [Decidable p] {l : ¬ p → Array α} :
(x ∈ if h : p then #[] else l h) ↔ ∃ h : ¬ p, x ∈ l h := by
split <;> simp_all
Expand Down Expand Up @@ -659,10 +684,6 @@ theorem get?_eq_get?_toList (a : Array α) (i : Nat) : a.get? i = a.toList.get?
theorem get!_eq_get? [Inhabited α] (a : Array α) : a.get! n = (a.get? n).getD default := by
simp only [get!_eq_getElem?, get?_eq_getElem?]

theorem getElem?_eq_some_iff {as : Array α} : as[n]? = some a ↔ ∃ h : n < as.size, as[n] = a := by
cases as
simp [List.getElem?_eq_some_iff]

theorem back!_eq_back? [Inhabited α] (a : Array α) : a.back! = a.back?.getD default := by
simp only [back!, get!_eq_getElem?, get?_eq_getElem?, back?]

Expand All @@ -672,6 +693,10 @@ theorem back!_eq_back? [Inhabited α] (a : Array α) : a.back! = a.back?.getD de
@[simp] theorem back!_push [Inhabited α] (a : Array α) : (a.push x).back! = x := by
simp [back!_eq_back?]

theorem mem_of_back?_eq_some {xs : Array α} {a : α} (h : xs.back? = some a) : a ∈ xs := by
cases xs
simpa using List.mem_of_getLast?_eq_some (by simpa using h)

theorem getElem?_push_lt (a : Array α) (x : α) (i : Nat) (h : i < a.size) :
(a.push x)[i]? = some a[i] := by
rw [getElem?_pos, getElem_push_lt]
Expand Down Expand Up @@ -1025,6 +1050,10 @@ theorem foldr_congr {as bs : Array α} (h₀ : as = bs) {f g : α → β → β}
@[simp] theorem mem_map {f : α → β} {l : Array α} : b ∈ l.map f ↔ ∃ a, a ∈ l ∧ f a = b := by
simp only [mem_def, toList_map, List.mem_map]

theorem exists_of_mem_map (h : b ∈ map f l) : ∃ a, a ∈ l ∧ f a = b := mem_map.1 h

theorem mem_map_of_mem (f : α → β) (h : a ∈ l) : f a ∈ map f l := mem_map.2 ⟨_, h, rfl⟩

theorem mapM_eq_mapM_toList [Monad m] [LawfulMonad m] (f : α → m β) (arr : Array α) :
arr.mapM f = List.toArray <$> (arr.toList.mapM f) := by
rw [mapM_eq_foldlM, ← foldlM_toList, ← List.foldrM_reverse]
Expand Down Expand Up @@ -1215,6 +1244,12 @@ theorem push_eq_append_singleton (as : Array α) (x) : as.push x = as ++ #[x] :=
@[simp] theorem mem_append {a : α} {s t : Array α} : a ∈ s ++ t ↔ a ∈ s ∨ a ∈ t := by
simp only [mem_def, toList_append, List.mem_append]

theorem mem_append_left {a : α} {l₁ : Array α} (l₂ : Array α) (h : a ∈ l₁) : a ∈ l₁ ++ l₂ :=
mem_append.2 (Or.inl h)

theorem mem_append_right {a : α} (l₁ : Array α) {l₂ : Array α} (h : a ∈ l₂) : a ∈ l₁ ++ l₂ :=
mem_append.2 (Or.inr h)

@[simp] theorem size_append (as bs : Array α) : (as ++ bs).size = as.size + bs.size := by
simp only [size, toList_append, List.length_append]

Expand Down Expand Up @@ -1914,6 +1949,26 @@ theorem array_array_induction (P : Array (Array α) → Prop) (h : ∀ (xss : Li
specialize h (ass.toList.map toList)
simpa [← toList_map, Function.comp_def, map_id] using h

theorem foldl_map (f : β₁ → β₂) (g : α → β₂ → α) (l : Array β₁) (init : α) :
(l.map f).foldl g init = l.foldl (fun x y => g x (f y)) init := by
cases l; simp [List.foldl_map]

theorem foldr_map (f : α₁ → α₂) (g : α₂ → β → β) (l : Array α₁) (init : β) :
(l.map f).foldr g init = l.foldr (fun x y => g (f x) y) init := by
cases l; simp [List.foldr_map]

theorem foldl_filterMap (f : α → Option β) (g : γ → β → γ) (l : Array α) (init : γ) :
(l.filterMap f).foldl g init = l.foldl (fun x y => match f y with | some b => g x b | none => x) init := by
cases l
simp [List.foldl_filterMap]
rfl

theorem foldr_filterMap (f : α → Option β) (g : β → γ → γ) (l : Array α) (init : γ) :
(l.filterMap f).foldr g init = l.foldr (fun x y => match f x with | some b => g b y | none => y) init := by
cases l
simp [List.foldr_filterMap]
rfl

/-! ### flatten -/

@[simp] theorem flatten_empty : flatten (#[] : Array (Array α)) = #[] := rfl
Expand All @@ -1928,6 +1983,12 @@ theorem array_array_induction (P : Array (Array α) → Prop) (h : ∀ (xss : Li
| nil => simp
| cons xs xss ih => simp [ih]

/-! ### reverse -/

@[simp] theorem mem_reverse {x : α} {as : Array α} : x ∈ as.reverse ↔ x ∈ as := by
cases as
simp

/-! ### findSomeRevM?, findRevM?, findSomeRev?, findRev? -/

@[simp] theorem findSomeRevM?_eq_findSomeM?_reverse
Expand Down
38 changes: 27 additions & 11 deletions src/Init/Data/List/Attach.lean
Original file line number Diff line number Diff line change
Expand Up @@ -13,7 +13,7 @@ namespace List
`a : α` satisfying `P`, then `pmap f l h` is essentially the same as `map f l`
but is defined only when all members of `l` satisfy `P`, using the proof
to apply `f`. -/
@[simp] def pmap {P : α → Prop} (f : ∀ a, P a → β) : ∀ l : List α, (H : ∀ a ∈ l, P a) → List β
def pmap {P : α → Prop} (f : ∀ a, P a → β) : ∀ l : List α, (H : ∀ a ∈ l, P a) → List β
| [], _ => []
| a :: l, H => f a (forall_mem_cons.1 H).1 :: pmap f l (forall_mem_cons.1 H).2

Expand Down Expand Up @@ -46,6 +46,11 @@ Unsafe implementation of `attachWith`, taking advantage of the fact that the rep
| cons _ L', hL' => congrArg _ <| go L' fun _ hx => hL' (.tail _ hx)
exact go L h'

@[simp] theorem pmap_nil {P : α → Prop} (f : ∀ a, P a → β) : pmap f [] (by simp) = [] := rfl

@[simp] theorem pmap_cons {P : α → Prop} (f : ∀ a, P a → β) (a : α) (l : List α) (h : ∀ b ∈ a :: l, P b) :
pmap f (a :: l) h = f a (forall_mem_cons.1 h).1 :: pmap f l (forall_mem_cons.1 h).2 := rfl

@[simp] theorem attach_nil : ([] : List α).attach = [] := rfl

@[simp] theorem attachWith_nil : ([] : List α).attachWith P H = [] := rfl
Expand Down Expand Up @@ -148,7 +153,7 @@ theorem mem_pmap_of_mem {p : α → Prop} {f : ∀ a, p a → β} {l H} {a} (h :
exact ⟨a, h, rfl⟩

@[simp]
theorem length_pmap {p : α → Prop} {f : ∀ a, p a → β} {l H} : length (pmap f l H) = length l := by
theorem length_pmap {p : α → Prop} {f : ∀ a, p a → β} {l H} : (pmap f l H).length = l.length := by
induction l
· rfl
· simp only [*, pmap, length]
Expand Down Expand Up @@ -199,7 +204,7 @@ theorem attachWith_ne_nil_iff {l : List α} {P : α → Prop} {H : ∀ a ∈ l,

@[simp]
theorem getElem?_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) (n : Nat) :
(pmap f l h)[n]? = Option.pmap f l[n]? fun x H => h x (getElem?_mem H) := by
(pmap f l h)[n]? = Option.pmap f l[n]? fun x H => h x (mem_of_getElem? H) := by
induction l generalizing n with
| nil => simp
| cons hd tl hl =>
Expand All @@ -215,7 +220,7 @@ theorem getElem?_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h
· simp_all

theorem get?_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) (n : Nat) :
get? (pmap f l h) n = Option.pmap f (get? l n) fun x H => h x (get?_mem H) := by
get? (pmap f l h) n = Option.pmap f (get? l n) fun x H => h x (mem_of_get? H) := by
simp only [get?_eq_getElem?]
simp [getElem?_pmap, h]

Expand All @@ -238,18 +243,18 @@ theorem get_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h :
(hn : n < (pmap f l h).length) :
get (pmap f l h) ⟨n, hn⟩ =
f (get l ⟨n, @length_pmap _ _ p f l h ▸ hn⟩)
(h _ (get_mem l ⟨n, @length_pmap _ _ p f l h ▸ hn)) := by
(h _ (getElem_mem (@length_pmap _ _ p f l h ▸ hn))) := by
simp only [get_eq_getElem]
simp [getElem_pmap]

@[simp]
theorem getElem?_attachWith {xs : List α} {i : Nat} {P : α → Prop} {H : ∀ a ∈ xs, P a} :
(xs.attachWith P H)[i]? = xs[i]?.pmap Subtype.mk (fun _ a => H _ (getElem?_mem a)) :=
(xs.attachWith P H)[i]? = xs[i]?.pmap Subtype.mk (fun _ a => H _ (mem_of_getElem? a)) :=
getElem?_pmap ..

@[simp]
theorem getElem?_attach {xs : List α} {i : Nat} :
xs.attach[i]? = xs[i]?.pmap Subtype.mk (fun _ a => getElem?_mem a) :=
xs.attach[i]? = xs[i]?.pmap Subtype.mk (fun _ a => mem_of_getElem? a) :=
getElem?_attachWith

@[simp]
Expand Down Expand Up @@ -333,6 +338,7 @@ This is useful when we need to use `attach` to show termination.
Unfortunately this can't be applied by `simp` because of the higher order unification problem,
and even when rewriting we need to specify the function explicitly.
See however `foldl_subtype` below.
-/
theorem foldl_attach (l : List α) (f : β → α → β) (b : β) :
l.attach.foldl (fun acc t => f acc t.1) b = l.foldl f b := by
Expand All @@ -348,6 +354,7 @@ This is useful when we need to use `attach` to show termination.
Unfortunately this can't be applied by `simp` because of the higher order unification problem,
and even when rewriting we need to specify the function explicitly.
See however `foldr_subtype` below.
-/
theorem foldr_attach (l : List α) (f : α → β → β) (b : β) :
l.attach.foldr (fun t acc => f t.1 acc) b = l.foldr f b := by
Expand Down Expand Up @@ -452,16 +459,16 @@ theorem pmap_append' {p : α → Prop} (f : ∀ a : α, p a → β) (l₁ l₂ :
pmap_append f l₁ l₂ _

@[simp] theorem attach_append (xs ys : List α) :
(xs ++ ys).attach = xs.attach.map (fun ⟨x, h⟩ => ⟨x, mem_append_of_mem_left ys h⟩) ++
ys.attach.map fun ⟨x, h⟩ => ⟨x, mem_append_of_mem_right xs h⟩ := by
(xs ++ ys).attach = xs.attach.map (fun ⟨x, h⟩ => ⟨x, mem_append_left ys h⟩) ++
ys.attach.map fun ⟨x, h⟩ => ⟨x, mem_append_right xs h⟩ := by
simp only [attach, attachWith, pmap, map_pmap, pmap_append]
congr 1 <;>
exact pmap_congr_left _ fun _ _ _ _ => rfl

@[simp] theorem attachWith_append {P : α → Prop} {xs ys : List α}
{H : ∀ (a : α), a ∈ xs ++ ys → P a} :
(xs ++ ys).attachWith P H = xs.attachWith P (fun a h => H a (mem_append_of_mem_left ys h)) ++
ys.attachWith P (fun a h => H a (mem_append_of_mem_right xs h)) := by
(xs ++ ys).attachWith P H = xs.attachWith P (fun a h => H a (mem_append_left ys h)) ++
ys.attachWith P (fun a h => H a (mem_append_right xs h)) := by
simp only [attachWith, attach_append, map_pmap, pmap_append]

@[simp] theorem pmap_reverse {P : α → Prop} (f : (a : α) → P a → β) (xs : List α)
Expand Down Expand Up @@ -598,6 +605,15 @@ def unattach {α : Type _} {p : α → Prop} (l : List { x // p x }) := l.map (
| nil => simp
| cons a l ih => simp [ih, Function.comp_def]

@[simp] theorem getElem?_unattach {p : α → Prop} {l : List { x // p x }} (i : Nat) :
l.unattach[i]? = l[i]?.map Subtype.val := by
simp [unattach]

@[simp] theorem getElem_unattach
{p : α → Prop} {l : List { x // p x }} (i : Nat) (h : i < l.unattach.length) :
l.unattach[i] = (l[i]'(by simpa using h)).1 := by
simp [unattach]

/-! ### Recognizing higher order functions on subtypes using a function that only depends on the value. -/

/--
Expand Down
4 changes: 2 additions & 2 deletions src/Init/Data/List/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -726,13 +726,13 @@ theorem elem_eq_true_of_mem [BEq α] [LawfulBEq α] {a : α} {as : List α} (h :
instance [BEq α] [LawfulBEq α] (a : α) (as : List α) : Decidable (a ∈ as) :=
decidable_of_decidable_of_iff (Iff.intro mem_of_elem_eq_true elem_eq_true_of_mem)

theorem mem_append_of_mem_left {a : α} {as : List α} (bs : List α) : a ∈ as → a ∈ as ++ bs := by
theorem mem_append_left {a : α} {as : List α} (bs : List α) : a ∈ as → a ∈ as ++ bs := by
intro h
induction h with
| head => apply Mem.head
| tail => apply Mem.tail; assumption

theorem mem_append_of_mem_right {b : α} {bs : List α} (as : List α) : b ∈ bs → b ∈ as ++ bs := by
theorem mem_append_right {b : α} {bs : List α} (as : List α) : b ∈ bs → b ∈ as ++ bs := by
intro h
induction as with
| nil => simp [h]
Expand Down
2 changes: 1 addition & 1 deletion src/Init/Data/List/Control.lean
Original file line number Diff line number Diff line change
Expand Up @@ -256,7 +256,7 @@ theorem findM?_eq_findSomeM? [Monad m] [LawfulMonad m] (p : α → m Bool) (as :
have : a ∈ as := by
have ⟨bs, h⟩ := h
subst h
exact mem_append_of_mem_right _ (Mem.head ..)
exact mem_append_right _ (Mem.head ..)
match (← f a this b) with
| ForInStep.done b => pure b
| ForInStep.yield b =>
Expand Down
19 changes: 10 additions & 9 deletions src/Init/Data/List/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -394,9 +394,9 @@ theorem get?_concat_length (l : List α) (a : α) : (l ++ [a]).get? l.length = s
theorem mem_cons_self (a : α) (l : List α) : a ∈ a :: l := .head ..

theorem mem_concat_self (xs : List α) (a : α) : a ∈ xs ++ [a] :=
mem_append_of_mem_right xs (mem_cons_self a _)
mem_append_right xs (mem_cons_self a _)

theorem mem_append_cons_self : a ∈ xs ++ a :: ys := mem_append_of_mem_right _ (mem_cons_self _ _)
theorem mem_append_cons_self : a ∈ xs ++ a :: ys := mem_append_right _ (mem_cons_self _ _)

theorem eq_append_cons_of_mem {a : α} {xs : List α} (h : a ∈ xs) :
∃ as bs, xs = as ++ a :: bs ∧ a ∉ as := by
Expand Down Expand Up @@ -507,12 +507,16 @@ theorem get_mem : ∀ (l : List α) n, get l n ∈ l
| _ :: _, ⟨0, _⟩ => .head ..
| _ :: l, ⟨_+1, _⟩ => .tail _ (get_mem l ..)

theorem getElem?_mem {l : List α} {n : Nat} {a : α} (e : l[n]? = some a) : a ∈ l :=
theorem mem_of_getElem? {l : List α} {n : Nat} {a : α} (e : l[n]? = some a) : a ∈ l :=
let ⟨_, e⟩ := getElem?_eq_some_iff.1 e; e ▸ getElem_mem ..

theorem get?_mem {l : List α} {n a} (e : l.get? n = some a) : a ∈ l :=
@[deprecated mem_of_getElem? (since := "2024-09-06")] abbrev getElem?_mem := @mem_of_getElem?

theorem mem_of_get? {l : List α} {n a} (e : l.get? n = some a) : a ∈ l :=
let ⟨_, e⟩ := get?_eq_some.1 e; e ▸ get_mem ..

@[deprecated mem_of_get? (since := "2024-09-06")] abbrev get?_mem := @mem_of_get?

theorem mem_iff_getElem {a} {l : List α} : a ∈ l ↔ ∃ (n : Nat) (h : n < l.length), l[n]'h = a :=
⟨getElem_of_mem, fun ⟨_, _, e⟩ => e ▸ getElem_mem ..⟩

Expand Down Expand Up @@ -1997,11 +2001,8 @@ theorem not_mem_append {a : α} {s t : List α} (h₁ : a ∉ s) (h₂ : a ∉ t
theorem mem_append_eq (a : α) (s t : List α) : (a ∈ s ++ t) = (a ∈ s ∨ a ∈ t) :=
propext mem_append

theorem mem_append_left {a : α} {l₁ : List α} (l₂ : List α) (h : a ∈ l₁) : a ∈ l₁ ++ l₂ :=
mem_append.2 (Or.inl h)

theorem mem_append_right {a : α} (l₁ : List α) {l₂ : List α} (h : a ∈ l₂) : a ∈ l₁ ++ l₂ :=
mem_append.2 (Or.inr h)
@[deprecated mem_append_left (since := "2024-11-20")] abbrev mem_append_of_mem_left := @mem_append_left
@[deprecated mem_append_right (since := "2024-11-20")] abbrev mem_append_of_mem_right := @mem_append_right

theorem mem_iff_append {a : α} {l : List α} : a ∈ l ↔ ∃ s t : List α, l = s ++ a :: t :=
⟨append_of_mem, fun ⟨s, t, e⟩ => e ▸ by simp⟩
Expand Down
Loading

0 comments on commit c8b4f6b

Please sign in to comment.