doc-next-gen

Init.Data.List.Attach

Module docstring

{"## unattach

List.unattach is the (one-sided) inverse of List.attach. It is a synonym for List.map Subtype.val.

We use it by providing a simp lemma l.attach.unattach = l, and simp lemmas which recognize higher order functions applied to l : List { x // p x } which only depend on the value, not the predicate, and rewrite these in terms of a simpler function applied to l.unattach.

Further, we provide simp lemmas that push unattach inwards. ","### Recognizing higher order functions on subtypes using a function that only depends on the value. ","### Simp lemmas pushing unattach inwards. ","### Well-founded recursion preprocessing setup "}

List.pmap definition
{P : α → Prop} (f : ∀ a, P a → β) : ∀ l : List α, (H : ∀ a ∈ l, P a) → List β
Full source
/--
Maps a partially defined function (defined on those terms of `α` that satisfy a predicate `P`) over
a list `l : List α`, given a proof that every element of `l` in fact satisfies `P`.

`O(|l|)`. `List.pmap`, named for “partial map,” is the equivalent of `List.map` for such partial
functions.
-/
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
Partial map over a list with a predicate
Informal description
Given a predicate $P$ on elements of type $\alpha$ and a function $f$ defined on elements $a$ of $\alpha$ that satisfy $P(a)$, the function `List.pmap` applies $f$ to each element of a list $l : \text{List } \alpha$ for which $P$ holds, producing a new list of type $\text{List } \beta$. The function requires a proof that every element in $l$ satisfies $P$.
List.attachWith definition
(l : List α) (P : α → Prop) (H : ∀ x ∈ l, P x) : List { x // P x }
Full source
/--
“Attaches” individual proofs to a list of values that satisfy a predicate `P`, returning a list of
elements in the corresponding subtype `{ x // P x }`.

`O(1)`.
-/
@[implemented_by attachWithImpl] def attachWith
    (l : List α) (P : α → Prop) (H : ∀ x ∈ l, P x) : List {x // P x} := pmap Subtype.mk l H
List attachment with predicate proofs
Informal description
Given a list `l` of elements of type `α` and a predicate `P` on `α`, the function `List.attachWith` constructs a new list where each element `x` of `l` is paired with a proof that `P x` holds (provided by `H : ∀ x ∈ l, P x`). The resulting list consists of elements of the subtype `{x // P x}`.
List.attach definition
(l : List α) : List { x // x ∈ l }
Full source
/--
“Attaches” the proof that the elements of `l` are in fact elements of `l`, producing a new list with
the same elements but in the subtype `{ x // x ∈ l }`.

`O(1)`.

This function is primarily used to allow definitions by [well-founded
recursion](lean-manual://section/well-founded-recursion) that use higher-order functions (such as
`List.map`) to prove that an value taken from a list is smaller than the list. This allows the
well-founded recursion mechanism to prove that the function terminates.
-/
@[inline] def attach (l : List α) : List {x // x ∈ l} := attachWith l _ fun _ => id
List attachment with membership proofs
Informal description
Given a list `l` of elements of type `α`, the function `List.attach` constructs a new list where each element `x` of `l` is paired with a proof that `x` is an element of `l`. The resulting list consists of elements of the subtype `{x // x ∈ l}`. This operation has constant time complexity (O(1)). This function is primarily used to facilitate well-founded recursion proofs involving higher-order functions (like `List.map`) by ensuring that values taken from the list can be shown to be "smaller" than the list itself, thereby proving termination.
List.pmap_nil theorem
{P : α → Prop} {f : ∀ a, P a → β} : pmap f [] (by simp) = []
Full source
@[simp] theorem pmap_nil {P : α → Prop} {f : ∀ a, P a → β} : pmap f [] (by simp) = [] := rfl
Partial Map of Empty List is Empty
Informal description
For any predicate $P$ on elements of type $\alpha$ and any function $f$ defined on elements $a$ of $\alpha$ that satisfy $P(a)$, the partial map of $f$ over the empty list is equal to the empty list. That is, $\text{pmap}\,f\,[]\,h = []$, where $h$ is a proof that all elements of the empty list satisfy $P$.
List.pmap_cons theorem
{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
Full source
@[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
Partial Map over Cons List Decomposition
Informal description
For any predicate $P : \alpha \to \text{Prop}$, function $f : \forall a, P a \to \beta$, element $a : \alpha$, and list $l : \text{List}\ \alpha$, if $h$ is a proof that $\forall b \in a :: l, P b$ holds, then the partial map of $f$ over the list $a :: l$ with proof $h$ is equal to $f a h_1$ consed onto the partial map of $f$ over $l$ with proof $h_2$, where $h_1$ and $h_2$ are the components of the decomposed proof $h$ via `forall_mem_cons`.
List.attach_nil theorem
: ([] : List α).attach = []
Full source
@[simp] theorem attach_nil : ([] : List α).attach = [] := rfl
Empty List Attachment Yields Empty List
Informal description
For any type $\alpha$, the attachment of the empty list is equal to the empty list, i.e., $\text{attach}([]) = []$.
List.attachWith_nil theorem
: ([] : List α).attachWith P H = []
Full source
@[simp] theorem attachWith_nil : ([] : List α).attachWith P H = [] := rfl
Empty List Attachment with Predicate Yields Empty List
Informal description
For any type $\alpha$, predicate $P$ on $\alpha$, and proof $H$ that all elements of the empty list satisfy $P$, the result of attaching $P$ to the empty list is the empty list. In other words, $[].\text{attachWith}\ P\ H = []$.
List.pmap_eq_map theorem
{p : α → Prop} {f : α → β} {l : List α} (H) : @pmap _ _ p (fun a _ => f a) l H = map f l
Full source
@[simp]
theorem pmap_eq_map {p : α → Prop} {f : α → β} {l : List α} (H) :
    @pmap _ _ p (fun a _ => f a) l H = map f l := by
  induction l
  · rfl
  · simp only [*, pmap, map]
Equality of Partial Map and Regular Map When Ignoring Predicate Proofs
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $f : \alpha \to \beta$, and any list $l : \text{List } \alpha$, the partial map operation `pmap` with $f$ (ignoring the proof argument) applied to $l$ is equal to the regular map operation `map` with $f$ applied to $l$. In mathematical notation, this can be expressed as: $$\text{pmap}_{p} (\lambda a \_. f a)\ l\ H = \text{map}\ f\ l$$ where $H$ is a proof that all elements of $l$ satisfy $p$.
List.pmap_congr_left theorem
{p q : α → Prop} {f : ∀ a, p a → β} {g : ∀ a, q a → β} (l : List α) {H₁ H₂} (h : ∀ a ∈ l, ∀ (h₁ h₂), f a h₁ = g a h₂) : pmap f l H₁ = pmap g l H₂
Full source
theorem pmap_congr_left {p q : α → Prop} {f : ∀ a, p a → β} {g : ∀ a, q a → β} (l : List α) {H₁ H₂}
    (h : ∀ a ∈ l, ∀ (h₁ h₂), f a h₁ = g a h₂) : pmap f l H₁ = pmap g l H₂ := by
  induction l with
  | nil => rfl
  | cons x l ih =>
    rw [pmap, pmap, h _ mem_cons_self, ih fun a ha => h a (mem_cons_of_mem _ ha)]
Equality of Partial Maps under Pointwise Function Equality
Informal description
For any predicates $p$ and $q$ on elements of type $\alpha$, and any functions $f : \forall a, p a \to \beta$ and $g : \forall a, q a \to \beta$, given a list $l : \text{List } \alpha$ and proofs $H_1 : \forall a \in l, p a$ and $H_2 : \forall a \in l, q a$, if for every $a \in l$ and any proofs $h_1 : p a$ and $h_2 : q a$ we have $f a h_1 = g a h_2$, then the partial maps $\text{pmap}\ f\ l\ H_1$ and $\text{pmap}\ g\ l\ H_2$ are equal.
List.pmap_congr abbrev
Full source
@[deprecated pmap_congr_left (since := "2024-09-06")] abbrev pmap_congr := @pmap_congr_left
Equality of Partial Maps under Pointwise Function Equality
Informal description
For any predicates $p$ and $q$ on elements of type $\alpha$, and any functions $f : \forall a, p a \to \beta$ and $g : \forall a, q a \to \beta$, given a list $l : \text{List } \alpha$ and proofs $H_1 : \forall a \in l, p a$ and $H_2 : \forall a \in l, q a$, if for every $a \in l$ and any proofs $h_1 : p a$ and $h_2 : q a$ we have $f a h_1 = g a h_2$, then the partial maps $\text{pmap}\ f\ l\ H_1$ and $\text{pmap}\ g\ l\ H_2$ are equal.
List.map_pmap theorem
{p : α → Prop} {g : β → γ} {f : ∀ a, p a → β} {l : List α} (H) : map g (pmap f l H) = pmap (fun a h => g (f a h)) l H
Full source
theorem map_pmap {p : α → Prop} {g : β → γ} {f : ∀ a, p a → β} {l : List α} (H) :
    map g (pmap f l H) = pmap (fun a h => g (f a h)) l H := by
  induction l
  · rfl
  · simp only [*, pmap, map]
Commutativity of Map and Partial Map on Lists: $\text{map } g \circ \text{pmap } f = \text{pmap } (g \circ f)$
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $g : \beta \to \gamma$, any function $f$ defined on elements $a$ of $\alpha$ that satisfy $p(a)$, and any list $l$ of type $\text{List } \alpha$, the following equality holds: \[ \text{map } g (\text{pmap } f l H) = \text{pmap } (\lambda a h, g (f a h)) l H \] where $H$ is a proof that all elements in $l$ satisfy $p$.
List.pmap_map theorem
{p : β → Prop} {g : ∀ b, p b → γ} {f : α → β} {l : List α} (H) : pmap g (map f l) H = pmap (fun a h => g (f a) h) l fun _ h => H _ (mem_map_of_mem h)
Full source
theorem pmap_map {p : β → Prop} {g : ∀ b, p b → γ} {f : α → β} {l : List α} (H) :
    pmap g (map f l) H = pmap (fun a h => g (f a) h) l fun _ h => H _ (mem_map_of_mem h) := by
  induction l
  · rfl
  · simp only [*, pmap, map]
Commutativity of Partial Map and Map on Lists: $\text{pmap } g \circ \text{map } f = \text{pmap } (g \circ f)$
Informal description
For any predicate $p$ on elements of type $\beta$, any function $g : \forall b, p(b) \to \gamma$, any function $f : \alpha \to \beta$, and any list $l$ of type $\text{List } \alpha$, given a proof $H$ that all elements in $\text{map } f\, l$ satisfy $p$, the following equality holds: \[ \text{pmap } g (\text{map } f\, l) H = \text{pmap } (\lambda a h, g (f a) h) l (\lambda \_ h, H (f a) (\text{mem\_map\_of\_mem } h)) \]
List.attach_congr theorem
{l₁ l₂ : List α} (h : l₁ = l₂) : l₁.attach = l₂.attach.map (fun x => ⟨x.1, h ▸ x.2⟩)
Full source
theorem attach_congr {l₁ l₂ : List α} (h : l₁ = l₂) :
    l₁.attach = l₂.attach.map (fun x => ⟨x.1, h ▸ x.2⟩) := by
  subst h
  simp
Congruence of List Attachment Under List Equality: $l_1 = l_2 \Rightarrow l_1.\text{attach} = l_2.\text{attach}.\text{map}(\langle x, h \rangle \mapsto \langle x, h \text{ transported along } l_1 = l_2 \rangle)$
Informal description
For any two lists $l_1$ and $l_2$ of elements of type $\alpha$ such that $l_1 = l_2$, the attached list $l_1.\text{attach}$ is equal to the list obtained by mapping each element $\langle x, h \rangle$ in $l_2.\text{attach}$ to $\langle x, h' \rangle$, where $h'$ is the proof that $x \in l_1$ obtained by transporting the proof $h$ of $x \in l_2$ along the equality $h : l_1 = l_2$.
List.attachWith_congr theorem
{l₁ l₂ : List α} (w : l₁ = l₂) {P : α → Prop} {H : ∀ x ∈ l₁, P x} : l₁.attachWith P H = l₂.attachWith P fun _ h => H _ (w ▸ h)
Full source
theorem attachWith_congr {l₁ l₂ : List α} (w : l₁ = l₂) {P : α → Prop} {H : ∀ x ∈ l₁, P x} :
    l₁.attachWith P H = l₂.attachWith P fun _ h => H _ (w ▸ h) := by
  subst w
  simp
Equality of Attached Lists Under Predicate Congruence
Informal description
Given two lists $l_1$ and $l_2$ of elements of type $\alpha$ such that $l_1 = l_2$, and a predicate $P$ on $\alpha$ with proofs $H : \forall x \in l_1, P(x)$, then the attached lists with the predicate $P$ are equal: $$l_1.\text{attachWith}\, P\, H = l_2.\text{attachWith}\, P\, (\lambda \_ h, H \_ (w \mapsto h))$$ where the right-hand side uses the equality $w : l_1 = l_2$ to adjust the membership proofs.
List.attach_cons theorem
{x : α} {xs : List α} : (x :: xs).attach = ⟨x, mem_cons_self⟩ :: xs.attach.map fun ⟨y, h⟩ => ⟨y, mem_cons_of_mem x h⟩
Full source
@[simp] theorem attach_cons {x : α} {xs : List α} :
    (x :: xs).attach =
      ⟨x, mem_cons_self⟩⟨x, mem_cons_self⟩ :: xs.attach.map fun ⟨y, h⟩ => ⟨y, mem_cons_of_mem x h⟩ := by
  simp only [attach, attachWith, pmap, map_pmap, cons.injEq, true_and]
  apply pmap_congr_left
  intros a _ m' _
  rfl
Attached List Construction via Cons: $(x :: xs).\text{attach} = \langle x, \text{mem\_cons\_self}\rangle :: \text{map}\ (\lambda \langle y, h\rangle, \langle y, \text{mem\_cons\_of\_mem}\ x\ h\rangle)\ xs.\text{attach}$
Informal description
For any element $x$ of type $\alpha$ and any list $xs$ of elements of type $\alpha$, the attached list of $x :: xs$ is equal to the list obtained by prepending the pair $\langle x, \text{mem\_cons\_self}\rangle$ to the result of mapping each element $\langle y, h\rangle$ in the attached list of $xs$ to $\langle y, \text{mem\_cons\_of\_mem}\ x\ h\rangle$.
List.attachWith_cons theorem
{x : α} {xs : List α} {p : α → Prop} (h : ∀ a ∈ x :: xs, p a) : (x :: xs).attachWith p h = ⟨x, h x (mem_cons_self)⟩ :: xs.attachWith p (fun a ha ↦ h a (mem_cons_of_mem x ha))
Full source
@[simp]
theorem attachWith_cons {x : α} {xs : List α} {p : α → Prop} (h : ∀ a ∈ x :: xs, p a) :
    (x :: xs).attachWith p h = ⟨x, h x (mem_cons_self)⟩⟨x, h x (mem_cons_self)⟩ ::
      xs.attachWith p (fun a ha ↦ h a (mem_cons_of_mem x ha)) :=
  rfl
Cons Structure of Attached Lists with Predicate Proofs
Informal description
Let $\alpha$ be a type and $p : \alpha \to \text{Prop}$ a predicate. For any element $x \in \alpha$, list $xs$ of elements of $\alpha$, and proof $h$ that all elements in $x :: xs$ satisfy $p$, the attached list with proofs satisfies: $$(x :: xs).\text{attachWith}\ p\ h = \langle x, h_x\rangle :: xs.\text{attachWith}\ p\ h'$$ where $\langle x, h_x\rangle$ is the element $x$ paired with its proof $h_x = h\ x\ (\text{mem\_cons\_self})$, and $h'$ is the proof that all elements in $xs$ satisfy $p$ obtained from $h$ via $\text{mem\_cons\_of\_mem}$.
List.pmap_eq_map_attach theorem
{p : α → Prop} {f : ∀ a, p a → β} {l : List α} (H) : pmap f l H = l.attach.map fun x => f x.1 (H _ x.2)
Full source
theorem pmap_eq_map_attach {p : α → Prop} {f : ∀ a, p a → β} {l : List α} (H) :
    pmap f l H = l.attach.map fun x => f x.1 (H _ x.2) := by
  rw [attach, attachWith, map_pmap]; exact pmap_congr_left l fun _ _ _ _ => rfl
Partial Map Equals Map over Attached List with Predicate Proofs
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $f$ defined on elements $a$ of $\alpha$ that satisfy $p(a)$, and any list $l$ of elements of type $\alpha$ with a proof $H$ that all elements in $l$ satisfy $p$, the partial map of $f$ over $l$ is equal to mapping $f$ over the attached list of $l$ (where each element is paired with its membership proof). More precisely: \[ \text{pmap}\ f\ l\ H = \text{map}\ (\lambda x, f\ x.1\ (H\ x.1\ x.2))\ l.\text{attach} \] where $x.1$ is the value and $x.2$ is the proof that $x.1 \in l$.
List.pmap_eq_attachWith theorem
{p q : α → Prop} {f : ∀ a, p a → q a} {l : List α} (H) : pmap (fun a h => ⟨a, f a h⟩) l H = l.attachWith q (fun x h => f x (H x h))
Full source
@[simp]
theorem pmap_eq_attachWith {p q : α → Prop} {f : ∀ a, p a → q a} {l : List α} (H) :
    pmap (fun a h => ⟨a, f a h⟩) l H = l.attachWith q (fun x h => f x (H x h)) := by
  induction l with
  | nil => rfl
  | cons a l ih =>
    simp [pmap, attachWith, ih]
Equality of Partial Map and AttachWith for Lists with Predicates
Informal description
For any predicates $p$ and $q$ on a type $\alpha$, any function $f$ that maps elements satisfying $p$ to elements satisfying $q$, and any list $l$ of elements of type $\alpha$ with a proof $H$ that all elements in $l$ satisfy $p$, the partial map of $f$ over $l$ is equal to attaching the elements of $l$ with proofs that they satisfy $q$ via $f$. More precisely, we have: \[ \text{pmap} \, (\lambda a \, h, \langle a, f \, a \, h \rangle) \, l \, H = \text{attachWith} \, q \, l \, (\lambda x \, h, f \, x \, (H \, x \, h)) \]
List.attach_map_val theorem
{l : List α} {f : α → β} : (l.attach.map fun (i : { i // i ∈ l }) => f i) = l.map f
Full source
theorem attach_map_val {l : List α} {f : α → β} :
    (l.attach.map fun (i : {i // i ∈ l}) => f i) = l.map f := by
  rw [attach, attachWith, map_pmap]; exact pmap_eq_map _
Equality of Mapped Attached List and Directly Mapped List: $\text{map } f \circ \text{attach} = \text{map } f$
Informal description
For any list $l$ of elements of type $\alpha$ and any function $f : \alpha \to \beta$, mapping $f$ over the attached list $l.\text{attach}$ (where each element is paired with a membership proof) is equivalent to mapping $f$ directly over $l$. In mathematical notation: \[ \text{map } f (l.\text{attach}) = \text{map } f l \]
List.attach_map_coe abbrev
Full source
@[deprecated attach_map_val (since := "2025-02-17")]
abbrev attach_map_coe := @attach_map_val
Coercion Mapping of Attached List Equals Original List
Informal description
For any list $l$ of elements of type $\alpha$, the mapping of the coercion function (which extracts the underlying value from a subtype) over the attached list $l.\text{attach}$ is equal to the original list $l$. In mathematical notation: \[ \text{map } \text{Subtype.val } (l.\text{attach}) = l \]
List.attach_map_subtype_val theorem
(l : List α) : l.attach.map Subtype.val = l
Full source
theorem attach_map_subtype_val (l : List α) : l.attach.map Subtype.val = l :=
  attach_map_val.trans (List.map_id _)
Projection of Attached List Yields Original List: $\text{map val} \circ \text{attach} = \text{id}$
Informal description
For any list $l$ of elements of type $\alpha$, mapping the projection function $\text{Subtype.val}$ (which extracts the underlying value from a subtype) over the attached list $l.\text{attach}$ yields the original list $l$. In mathematical notation: \[ \text{map } \text{Subtype.val } (l.\text{attach}) = l \]
List.attachWith_map_val theorem
{p : α → Prop} {f : α → β} {l : List α} (H : ∀ a ∈ l, p a) : ((l.attachWith p H).map fun (i : { i // p i }) => f i) = l.map f
Full source
theorem attachWith_map_val {p : α → Prop} {f : α → β} {l : List α} (H : ∀ a ∈ l, p a) :
    ((l.attachWith p H).map fun (i : { i // p i}) => f i) = l.map f := by
  rw [attachWith, map_pmap]; exact pmap_eq_map _
Equality of Mapping after Attaching Predicate Proofs: $\text{map } f \circ \text{attachWith } p = \text{map } f$
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $f : \alpha \to \beta$, and any list $l$ of elements of type $\alpha$ where every element satisfies $p$, the following equality holds: \[ \text{map } f \circ \text{attachWith } p\ l\ H = \text{map } f\ l \] where $\text{attachWith } p\ l\ H$ attaches proofs of $p$ to each element of $l$, and $\text{map } f$ applies $f$ to each element of the resulting list of subtypes.
List.attachWith_map_coe abbrev
Full source
@[deprecated attachWith_map_val (since := "2025-02-17")]
abbrev attachWith_map_coe := @attachWith_map_val
Identity Mapping after Attaching Predicate Proofs Preserves Original List
Informal description
Given a list $l$ of elements of type $\alpha$, a predicate $p$ on $\alpha$, and a proof $H$ that every element in $l$ satisfies $p$, the mapping of the coercion function (which extracts the underlying value from a subtype) over the list obtained by attaching proofs of $p$ to each element of $l$ yields the original list $l$. In other words: \[ \text{map } (\lambda x \Rightarrow x) \circ \text{attachWith } p\ l\ H = l \] where $\text{attachWith } p\ l\ H$ attaches proofs of $p$ to each element of $l$, and $\text{map } (\lambda x \Rightarrow x)$ is the identity mapping on the resulting list of subtypes.
List.attachWith_map_subtype_val theorem
{p : α → Prop} {l : List α} (H : ∀ a ∈ l, p a) : (l.attachWith p H).map Subtype.val = l
Full source
theorem attachWith_map_subtype_val {p : α → Prop} {l : List α} (H : ∀ a ∈ l, p a) :
    (l.attachWith p H).map Subtype.val = l :=
  (attachWith_map_val _).trans (List.map_id _)
Mapping Subtype Values Recovers Original List: $\text{map val} \circ \text{attachWith } p = \text{id}$
Informal description
For any predicate $p$ on elements of type $\alpha$ and any list $l$ of elements of type $\alpha$ where every element satisfies $p$, mapping the subtype projection function $\text{Subtype.val}$ over the list obtained by attaching proofs of $p$ to elements of $l$ (via $\text{attachWith}$) yields the original list $l$. In other words: \[ \text{map } \text{Subtype.val } (\text{attachWith } p\ l\ H) = l \] where $\text{Subtype.val}$ extracts the underlying value from a subtype element.
List.mem_attach theorem
(l : List α) : ∀ x, x ∈ l.attach
Full source
@[simp]
theorem mem_attach (l : List α) : ∀ x, x ∈ l.attach
  | ⟨a, h⟩ => by
    have := mem_map.1 (by rw [attach_map_subtype_val]; exact h)
    rcases this with ⟨⟨_, _⟩, m, rfl⟩
    exact m
Membership in Attached List is Universal
Informal description
For any list $l$ of elements of type $\alpha$ and any element $x$ in the attached list $l.\text{attach}$, $x$ is a member of $l.\text{attach}$. In other words, every element in the list obtained by attaching membership proofs to $l$ is trivially a member of that list.
List.mem_attachWith theorem
{l : List α} {q : α → Prop} (H) (x : { x // q x }) : x ∈ l.attachWith q H ↔ x.1 ∈ l
Full source
@[simp]
theorem mem_attachWith {l : List α} {q : α → Prop} (H) (x : {x // q x}) :
    x ∈ l.attachWith q Hx ∈ l.attachWith q H ↔ x.1 ∈ l := by
  induction l with
  | nil => simp
  | cons a l ih =>
    simp [ih]
    constructor
    · rintro (_ | _) <;> simp_all
    · rintro (h | h)
      · simp [← h]
      · simp_all
Membership in Attached List with Predicate Proofs
Informal description
For a list $l$ of elements of type $\alpha$ and a predicate $q$ on $\alpha$, an element $x$ of the subtype $\{x \mid q(x)\}$ is in the list obtained by attaching proofs of $q$ to elements of $l$ (via `List.attachWith`) if and only if the underlying value $x.1$ is in the original list $l$.
List.mem_pmap theorem
{p : α → Prop} {f : ∀ a, p a → β} {l H b} : b ∈ pmap f l H ↔ ∃ (a : _) (h : a ∈ l), f a (H a h) = b
Full source
@[simp]
theorem mem_pmap {p : α → Prop} {f : ∀ a, p a → β} {l H b} :
    b ∈ pmap f l Hb ∈ pmap f l H ↔ ∃ (a : _) (h : a ∈ l), f a (H a h) = b := by
  simp only [pmap_eq_map_attach, mem_map, mem_attach, true_and, Subtype.exists, eq_comm]
Membership in Partial Map: $b \in \operatorname{pmap} f l H \leftrightarrow \exists (a \in l), f a (H a h) = b$
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $f$ defined on elements $a$ of $\alpha$ that satisfy $p(a)$, any list $l$ of elements of type $\alpha$, any proof $H$ that all elements in $l$ satisfy $p$, and any element $b$ of type $\beta$, the following equivalence holds: \[ b \in \operatorname{pmap} f l H \leftrightarrow \exists (a \in l), f a (H a h) = b \] where $h$ is a proof that $a \in l$.
List.mem_pmap_of_mem theorem
{p : α → Prop} {f : ∀ a, p a → β} {l H} {a} (h : a ∈ l) : f a (H a h) ∈ pmap f l H
Full source
theorem mem_pmap_of_mem {p : α → Prop} {f : ∀ a, p a → β} {l H} {a} (h : a ∈ l) :
    f a (H a h) ∈ pmap f l H := by
  rw [mem_pmap]
  exact ⟨a, h, rfl⟩
Element in Partial Map if in Original List
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $f$ defined on elements $a$ of $\alpha$ that satisfy $p(a)$, any list $l$ of elements of type $\alpha$, and any proof $H$ that all elements in $l$ satisfy $p$, if an element $a$ is in $l$, then $f(a, H(a, h))$ is in the partial map $\operatorname{pmap} f l H$, where $h$ is a proof that $a \in l$.
List.length_pmap theorem
{p : α → Prop} {f : ∀ a, p a → β} {l H} : (pmap f l H).length = l.length
Full source
@[simp]
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]
Length Preservation under Partial Mapping: $\operatorname{length}(\operatorname{pmap} f l H) = \operatorname{length}(l)$
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $f$ defined on elements $a$ of $\alpha$ that satisfy $p(a)$, any list $l$ of elements of type $\alpha$, and any proof $H$ that all elements in $l$ satisfy $p$, the length of the partially mapped list $\operatorname{pmap} f l H$ is equal to the length of $l$.
List.length_attach theorem
{l : List α} : l.attach.length = l.length
Full source
@[simp]
theorem length_attach {l : List α} : l.attach.length = l.length :=
  length_pmap
Length Preservation under List Attachment: $\operatorname{length}(l.\text{attach}) = \operatorname{length}(l)$
Informal description
For any list $l$ of elements of type $\alpha$, the length of the attached list $l.\text{attach}$ is equal to the length of $l$.
List.length_attachWith theorem
{p : α → Prop} {l H} : length (l.attachWith p H) = length l
Full source
@[simp]
theorem length_attachWith {p : α → Prop} {l H} : length (l.attachWith p H) = length l :=
  length_pmap
Length Preservation under Attaching Predicate Proofs: $\operatorname{length}(l.\operatorname{attachWith} p H) = \operatorname{length}(l)$
Informal description
For any predicate $p$ on elements of type $\alpha$, any list $l$ of elements of type $\alpha$, and any proof $H$ that all elements in $l$ satisfy $p$, the length of the list obtained by attaching proofs to each element via `attachWith` is equal to the length of the original list $l$. That is, $\operatorname{length}(l.\operatorname{attachWith} p H) = \operatorname{length}(l)$.
List.pmap_eq_nil_iff theorem
{p : α → Prop} {f : ∀ a, p a → β} {l H} : pmap f l H = [] ↔ l = []
Full source
@[simp]
theorem pmap_eq_nil_iff {p : α → Prop} {f : ∀ a, p a → β} {l H} : pmappmap f l H = [] ↔ l = [] := by
  rw [← length_eq_zero_iff, length_pmap, length_eq_zero_iff]
Empty List Characterization under Partial Mapping: $\operatorname{pmap} f l H = [] \leftrightarrow l = []$
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $f$ defined on elements $a$ of $\alpha$ that satisfy $p(a)$, any list $l$ of elements of type $\alpha$, and any proof $H$ that all elements in $l$ satisfy $p$, the partially mapped list $\operatorname{pmap} f l H$ is empty if and only if the original list $l$ is empty.
List.pmap_ne_nil_iff theorem
{P : α → Prop} (f : (a : α) → P a → β) {xs : List α} (H : ∀ (a : α), a ∈ xs → P a) : xs.pmap f H ≠ [] ↔ xs ≠ []
Full source
theorem pmap_ne_nil_iff {P : α → Prop} (f : (a : α) → P a → β) {xs : List α}
    (H : ∀ (a : α), a ∈ xs → P a) : xs.pmap f H ≠ []xs.pmap f H ≠ [] ↔ xs ≠ [] := by
  simp
Non-emptiness of Partial Map ↔ Non-emptiness of Original List
Informal description
For a list `xs` of elements of type `α`, a predicate `P` on `α`, and a function `f` that maps elements `a` of `α` satisfying `P a` to type `β`, the partial map `xs.pmap f H` is non-empty if and only if the original list `xs` is non-empty. Here, `H` is a proof that every element in `xs` satisfies `P`.
List.pmap_eq_self theorem
{l : List α} {p : α → Prop} {hp : ∀ (a : α), a ∈ l → p a} {f : (a : α) → p a → α} : l.pmap f hp = l ↔ ∀ a (h : a ∈ l), f a (hp a h) = a
Full source
theorem pmap_eq_self {l : List α} {p : α → Prop} {hp : ∀ (a : α), a ∈ l → p a}
    {f : (a : α) → p a → α} : l.pmap f hp = l ↔ ∀ a (h : a ∈ l), f a (hp a h) = a := by
  rw [pmap_eq_map_attach]
  conv => lhs; rhs; rw [← attach_map_subtype_val l]
  rw [map_inj_left]
  simp
Partial Map Identity Condition: $\text{pmap}\ f\ l\ hp = l \leftrightarrow \forall a \in l, f(a, hp(a, h)) = a$
Informal description
For a list $l$ of elements of type $\alpha$, a predicate $p$ on $\alpha$, a proof $hp$ that every element in $l$ satisfies $p$, and a function $f$ that maps elements of $\alpha$ satisfying $p$ back to $\alpha$, the partial map of $f$ over $l$ equals $l$ if and only if for every element $a$ in $l$ and its membership proof $h$, $f(a, hp(a, h)) = a$. In mathematical notation: \[ \text{pmap}\ f\ l\ hp = l \leftrightarrow \forall a \in l, f(a, hp(a, h)) = a \]
List.attach_eq_nil_iff theorem
{l : List α} : l.attach = [] ↔ l = []
Full source
@[simp]
theorem attach_eq_nil_iff {l : List α} : l.attach = [] ↔ l = [] :=
  pmap_eq_nil_iff
Empty List Characterization under Attachment: $\text{attach}(l) = [] \leftrightarrow l = []$
Informal description
For any list $l$ of elements of type $\alpha$, the attached list $l.\text{attach}$ is empty if and only if the original list $l$ is empty.
List.attach_ne_nil_iff theorem
{l : List α} : l.attach ≠ [] ↔ l ≠ []
Full source
theorem attach_ne_nil_iff {l : List α} : l.attach ≠ []l.attach ≠ [] ↔ l ≠ [] :=
  pmap_ne_nil_iff _ _
Non-emptiness of Attached List ↔ Non-emptiness of Original List
Informal description
For any list $l$ of elements of type $\alpha$, the attached list $l.\text{attach}$ is non-empty if and only if the original list $l$ is non-empty.
List.attachWith_eq_nil_iff theorem
{l : List α} {P : α → Prop} {H : ∀ a ∈ l, P a} : l.attachWith P H = [] ↔ l = []
Full source
@[simp]
theorem attachWith_eq_nil_iff {l : List α} {P : α → Prop} {H : ∀ a ∈ l, P a} :
    l.attachWith P H = [] ↔ l = [] :=
  pmap_eq_nil_iff
Empty List Characterization for Attached List with Predicate: $l.\text{attachWith}~P~H = [] \leftrightarrow l = []$
Informal description
For any list $l$ of elements of type $\alpha$ and any predicate $P$ on $\alpha$, given a proof $H$ that all elements in $l$ satisfy $P$, the attached list $l.\text{attachWith}~P~H$ is empty if and only if the original list $l$ is empty.
List.attachWith_ne_nil_iff theorem
{l : List α} {P : α → Prop} {H : ∀ a ∈ l, P a} : l.attachWith P H ≠ [] ↔ l ≠ []
Full source
theorem attachWith_ne_nil_iff {l : List α} {P : α → Prop} {H : ∀ a ∈ l, P a} :
    l.attachWith P H ≠ []l.attachWith P H ≠ [] ↔ l ≠ [] :=
  pmap_ne_nil_iff _ _
Non-emptiness of Attached List with Predicate ↔ Non-emptiness of Original List
Informal description
For a list $l$ of elements of type $\alpha$ and a predicate $P$ on $\alpha$, the list $l.\text{attachWith}~P~H$ is non-empty if and only if the original list $l$ is non-empty, where $H$ is a proof that every element in $l$ satisfies $P$.
List.pmap_eq_nil abbrev
Full source
@[deprecated pmap_eq_nil_iff (since := "2024-09-06")] abbrev pmap_eq_nil := @pmap_eq_nil_iff
Empty List Characterization under Partial Mapping: $\operatorname{pmap}~f~l~H = [] \leftrightarrow l = []$
Informal description
For a list $l$ of elements of type $\alpha$, a predicate $P$ on $\alpha$, a function $f$ defined on elements $a$ of $\alpha$ that satisfy $P(a)$, and a proof $H$ that all elements in $l$ satisfy $P$, the partially mapped list $\operatorname{pmap}~f~l~H$ is empty if and only if the original list $l$ is empty.
List.pmap_ne_nil abbrev
Full source
@[deprecated pmap_ne_nil_iff (since := "2024-09-06")] abbrev pmap_ne_nil := @pmap_ne_nil_iff
Non-emptiness of Partial Map ↔ Non-emptiness of Original List
Informal description
For a list `xs` of elements of type `α`, a predicate `P` on `α`, and a function `f` that maps elements `a` of `α` satisfying `P a` to type `β`, the partial map `xs.pmap f H` is non-empty if and only if the original list `xs` is non-empty. Here, `H` is a proof that every element in `xs` satisfies `P`.
List.attach_eq_nil abbrev
Full source
@[deprecated attach_eq_nil_iff (since := "2024-09-06")] abbrev attach_eq_nil := @attach_eq_nil_iff
Empty List Characterization under Attachment: $\text{attach}(l) = [] \leftrightarrow l = []$
Informal description
For any list $l$ of elements of type $\alpha$, the attached list $l.\text{attach}$ is empty if and only if the original list $l$ is empty.
List.attach_ne_nil abbrev
Full source
@[deprecated attach_ne_nil_iff (since := "2024-09-06")] abbrev attach_ne_nil := @attach_ne_nil_iff
Non-emptiness of Attached List ↔ Non-emptiness of Original List
Informal description
For any list $l$ of elements of type $\alpha$, the attached list $l.\text{attach}$ is non-empty if and only if the original list $l$ is non-empty.
List.getElem?_pmap theorem
{p : α → Prop} {f : ∀ a, p a → β} {l : List α} (h : ∀ a ∈ l, p a) (i : Nat) : (pmap f l h)[i]? = Option.pmap f l[i]? fun x H => h x (mem_of_getElem? H)
Full source
@[simp]
theorem getElem?_pmap {p : α → Prop} {f : ∀ a, p a → β} {l : List α} (h : ∀ a ∈ l, p a) (i : Nat) :
    (pmap f l h)[i]? = Option.pmap f l[i]? fun x H => h x (mem_of_getElem? H) := by
  induction l generalizing i with
  | nil => simp
  | cons hd tl hl =>
    rcases i with ⟨i⟩
    · simp only [Option.pmap]
      split <;> simp_all
    · simp only [pmap, getElem?_cons_succ, hl, Option.pmap]
Index Preservation under Partial Mapping: $(\operatorname{pmap} f l h)[i]? = \operatorname{Option.pmap} f (l[i]?)$
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $f$ defined on elements $a$ of $\alpha$ that satisfy $p(a)$, any list $l$ of elements of type $\alpha$, and any proof $h$ that all elements in $l$ satisfy $p$, the optional indexing operation on the partially mapped list $\operatorname{pmap} f l h$ at index $i$ is equal to the partial map of $f$ applied to the optional indexing of $l$ at index $i$. In symbols: $$(\operatorname{pmap} f l h)[i]? = \operatorname{Option.pmap} f (l[i]?) \ (\lambda x H, h x (\operatorname{mem\_of\_getElem?} H))$$
List.get?_pmap theorem
{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 (mem_of_get? H)
Full source
@[deprecated List.getElem?_pmap (since := "2025-02-12")]
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 (mem_of_get? H) := by
  simp only [get?_eq_getElem?]
  simp [getElem?_pmap, h]
Index Preservation under Partial Mapping: $\operatorname{get?} (\operatorname{pmap} f l h) n = \operatorname{Option.pmap} f (\operatorname{get?} l n)$
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $f$ defined on elements $a$ of $\alpha$ that satisfy $p(a)$, any list $l$ of elements of type $\alpha$, and any proof $h$ that all elements in $l$ satisfy $p$, the optional element access operation on the partially mapped list $\operatorname{pmap} f l h$ at index $n$ is equal to the partial map of $f$ applied to the optional element access of $l$ at index $n$. In symbols: $$\operatorname{get?} (\operatorname{pmap} f l h) n = \operatorname{Option.pmap} f (\operatorname{get?} l n) (\lambda x H, h x (\operatorname{mem\_of\_get?} H))$$
List.getElem_pmap theorem
{p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) {i : Nat} (hn : i < (pmap f l h).length) : (pmap f l h)[i] = f (l[i]'(@length_pmap _ _ p f l h ▸ hn)) (h _ (getElem_mem (@length_pmap _ _ p f l h ▸ hn)))
Full source
@[simp]
theorem getElem_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) {i : Nat}
    (hn : i < (pmap f l h).length) :
    (pmap f l h)[i] =
      f (l[i]'(@length_pmap _ _ p f l h ▸ hn))
        (h _ (getElem_mem (@length_pmap _ _ p f l h ▸ hn))) := by
  induction l generalizing i with
  | nil =>
    simp only [length, pmap] at hn
    exact absurd hn (Nat.not_lt_of_le i.zero_le)
  | cons hd tl hl =>
    cases i
    · simp
    · simp [hl]
Index Preservation under Partial Mapping: $(\operatorname{pmap} f l h)[i] = f(l[i])$
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $f$ defined on elements $a$ of $\alpha$ that satisfy $p(a)$, any list $l$ of elements of type $\alpha$, and any proof $h$ that all elements in $l$ satisfy $p$, if $i$ is a valid index for the partially mapped list $\operatorname{pmap} f l h$ (i.e., $i < \operatorname{length}(\operatorname{pmap} f l h)$), then the $i$-th element of $\operatorname{pmap} f l h$ is equal to $f$ applied to the $i$-th element of $l$ and the proof that this element satisfies $p$. In symbols: $$(\operatorname{pmap} f l h)[i] = f(l[i]) (h(l[i]) (\text{proof that } l[i] \in l))$$
List.get_pmap theorem
{p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) {n : Nat} (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 _ (getElem_mem (@length_pmap _ _ p f l h ▸ hn)))
Full source
@[deprecated getElem_pmap (since := "2025-02-13")]
theorem get_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) {n : Nat}
    (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 _ (getElem_mem (@length_pmap _ _ p f l h ▸ hn))) := by
  simp only [get_eq_getElem]
  simp [getElem_pmap]
Element-wise Preservation under Partial Mapping: $\text{pmap}\,f\,l\,h[n] = f(l[n])$
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $f$ defined on elements $a$ of $\alpha$ that satisfy $p(a)$, any list $l$ of elements of type $\alpha$, and any proof $h$ that all elements in $l$ satisfy $p$, if $n$ is a natural number such that $n < \text{length}(\text{pmap}\,f\,l\,h)$, then the $n$-th element of $\text{pmap}\,f\,l\,h$ is equal to $f$ applied to the $n$-th element of $l$ and the proof that this element satisfies $p$. In symbols: $$\text{get}\,(\text{pmap}\,f\,l\,h)\,\langle n, \text{hn}\rangle = f\,(\text{get}\,l\,\langle n, \text{hn}'\rangle)\,(h\,(\text{get}\,l\,\langle n, \text{hn}'\rangle)\,(\text{getElem\_mem}\,\text{hn}'))$$ where $\text{hn}'$ is derived from $\text{hn}$ via the length preservation property of $\text{pmap}$.
List.getElem?_attachWith theorem
{xs : List α} {i : Nat} {P : α → Prop} {H : ∀ a ∈ xs, P a} : (xs.attachWith P H)[i]? = xs[i]?.pmap Subtype.mk (fun _ a => H _ (mem_of_getElem? a))
Full source
@[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 _ (mem_of_getElem? a)) :=
  getElem?_pmap ..
Optional Indexing of Attached List with Predicate Proofs
Informal description
For any list `xs` of elements of type `α`, any natural number index `i`, any predicate `P` on `α`, and any proof `H` that all elements in `xs` satisfy `P`, the optional indexing operation on the attached list `xs.attachWith P H` at index `i` is equal to the partial map of `Subtype.mk` applied to the optional indexing of `xs` at index `i`. In symbols: $$(\text{xs.attachWith}\,P\,H)[i]? = \text{xs}[i]?.\text{pmap}\,\text{Subtype.mk}\,(\lambda \_\,a, H\,(\text{mem\_of\_getElem?}\,a))$$
List.getElem?_attach theorem
{xs : List α} {i : Nat} : xs.attach[i]? = xs[i]?.pmap Subtype.mk (fun _ a => mem_of_getElem? a)
Full source
@[simp]
theorem getElem?_attach {xs : List α} {i : Nat} :
    xs.attach[i]? = xs[i]?.pmap Subtype.mk (fun _ a => mem_of_getElem? a) :=
  getElem?_attachWith
Optional Indexing of Attached List Equals Partial Map of Optional Indexing
Informal description
For any list `xs` of elements of type `α` and any natural number index `i`, the optional indexing operation on the attached list `xs.attach` at index `i` is equal to the partial map of `Subtype.mk` applied to the optional indexing of `xs` at index `i`. In symbols: $$(\text{xs.attach})[i]? = \text{xs}[i]?.\text{pmap}\,\text{Subtype.mk}\,(\lambda \_\,a, \text{mem\_of\_getElem?}\,a)$$
List.getElem_attachWith theorem
{xs : List α} {P : α → Prop} {H : ∀ a ∈ xs, P a} {i : Nat} (h : i < (xs.attachWith P H).length) : (xs.attachWith P H)[i] = ⟨xs[i]'(by simpa using h), H _ (getElem_mem (by simpa using h))⟩
Full source
@[simp]
theorem getElem_attachWith {xs : List α} {P : α → Prop} {H : ∀ a ∈ xs, P a}
    {i : Nat} (h : i < (xs.attachWith P H).length) :
    (xs.attachWith P H)[i] = ⟨xs[i]'(by simpa using h), H _ (getElem_mem (by simpa using h))⟩ :=
  getElem_pmap ..
Index Preservation in Attached List with Predicate Proofs: $(\text{attachWith}\,P\,H)[i] = \langle \text{xs}[i], H(\text{xs}[i]) \rangle$
Informal description
For any list `xs` of elements of type `α`, any predicate `P` on `α`, any proof `H` that all elements in `xs` satisfy `P`, and any natural number index `i` such that `i` is less than the length of `xs.attachWith P H`, the `i`-th element of `xs.attachWith P H` is equal to the pair consisting of the `i`-th element of `xs` and the proof that this element satisfies `P`. In symbols: $$(\text{xs.attachWith}\,P\,H)[i] = \langle \text{xs}[i], H(\text{xs}[i]) \rangle$$
List.getElem_attach theorem
{xs : List α} {i : Nat} (h : i < xs.attach.length) : xs.attach[i] = ⟨xs[i]'(by simpa using h), getElem_mem (by simpa using h)⟩
Full source
@[simp]
theorem getElem_attach {xs : List α} {i : Nat} (h : i < xs.attach.length) :
    xs.attach[i] = ⟨xs[i]'(by simpa using h), getElem_mem (by simpa using h)⟩ :=
  getElem_attachWith h
Index Preservation in Attached List: $\text{attach}(xs)[i] = \langle xs[i], \text{proof}\rangle$
Informal description
For any list $xs$ of elements of type $\alpha$ and any natural number index $i$ such that $i < \text{length}(xs.\text{attach})$, the $i$-th element of the attached list $xs.\text{attach}$ is equal to the pair consisting of the $i$-th element of $xs$ and a proof that this element is in $xs$. In symbols: $$xs.\text{attach}[i] = \langle xs[i], \text{proof that } xs[i] \in xs \rangle$$
List.pmap_attach theorem
{l : List α} {p : { x // x ∈ l } → Prop} {f : ∀ a, p a → β} (H) : pmap f l.attach H = l.pmap (P := fun a => ∃ h : a ∈ l, p ⟨a, h⟩) (fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨h, H ⟨a, h⟩ (by simp)⟩)
Full source
@[simp] theorem pmap_attach {l : List α} {p : {x // x ∈ l} → Prop} {f : ∀ a, p a → β} (H) :
    pmap f l.attach H =
      l.pmap (P := fun a => ∃ h : a ∈ l, p ⟨a, h⟩)
        (fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨h, H ⟨a, h⟩ (by simp)⟩) := by
  apply ext_getElem <;> simp
Equality of Partial Maps on Attached List and Original List via Subtype Predicate
Informal description
For any list $l$ of elements of type $\alpha$, any predicate $p$ on the subtype $\{x \mid x \in l\}$, and any function $f$ defined on elements $\langle a, h \rangle$ of $\{x \mid x \in l\}$ that satisfy $p \langle a, h \rangle$, the following equality holds between two partially mapped lists: \[ \operatorname{pmap} f (l.\operatorname{attach}) H = l.\operatorname{pmap} (\lambda a h', f \langle a, h'.1 \rangle h'.2) (\lambda a h', \langle h', H \langle a, h' \rangle \rangle) \] Here, $H$ is a proof that all elements in $l.\operatorname{attach}$ satisfy $p$.
List.pmap_attachWith theorem
{l : List α} {p : { x // q x } → Prop} {f : ∀ a, p a → β} (H₁ H₂) : pmap f (l.attachWith q H₁) H₂ = l.pmap (P := fun a => ∃ h : q a, p ⟨a, h⟩) (fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨H₁ _ h, H₂ ⟨a, H₁ _ h⟩ (by simpa)⟩)
Full source
@[simp] theorem pmap_attachWith {l : List α} {p : {x // q x} → Prop} {f : ∀ a, p a → β} (H₁ H₂) :
    pmap f (l.attachWith q H₁) H₂ =
      l.pmap (P := fun a => ∃ h : q a, p ⟨a, h⟩)
        (fun a h => f ⟨a, h.1⟩ h.2) (fun a h => ⟨H₁ _ h, H₂ ⟨a, H₁ _ h⟩ (by simpa)⟩) := by
  apply ext_getElem <;> simp
Equality of Partial Maps on Attached List and Original List
Informal description
For any list $l$ of elements of type $\alpha$, any predicate $q$ on $\alpha$, any predicate $p$ on the subtype $\{x \mid q x\}$, and any function $f$ defined on elements $\langle a, h \rangle$ of $\{x \mid q x\}$ that satisfy $p \langle a, h \rangle$, the following equality holds between two partially mapped lists: \[ \operatorname{pmap} f (l.\operatorname{attachWith} q H_1) H_2 = l.\operatorname{pmap} (\lambda a h', f \langle a, h'.1 \rangle h'.2) (\lambda a h', \langle H_1 a h', H_2 \langle a, H_1 a h' \rangle \rangle) \] Here, $H_1$ is a proof that all elements in $l$ satisfy $q$, and $H_2$ is a proof that all elements in $l.\operatorname{attachWith} q H_1$ satisfy $p$.
List.head?_pmap theorem
{P : α → Prop} {f : (a : α) → P a → β} {xs : List α} (H : ∀ (a : α), a ∈ xs → P a) : (xs.pmap f H).head? = xs.attach.head?.map fun ⟨a, m⟩ => f a (H a m)
Full source
@[simp] theorem head?_pmap {P : α → Prop} {f : (a : α) → P a → β} {xs : List α}
    (H : ∀ (a : α), a ∈ xs → P a) :
    (xs.pmap f H).head? = xs.attach.head?.map fun ⟨a, m⟩ => f a (H a m) := by
  induction xs with
  | nil => simp
  | cons x xs ih =>
    simp at ih
    simp [head?_pmap, ih]
Optional Head of Partial Map Equals Mapped Optional Head of Attached List
Informal description
For any predicate $P$ on elements of type $\alpha$, any function $f$ that maps elements $a \in \alpha$ satisfying $P(a)$ to type $\beta$, and any list $xs$ of elements of type $\alpha$ where every element satisfies $P$, the optional head of the partially mapped list $\text{pmap}\ f\ xs\ H$ is equal to the optional head of the attached list $xs.\text{attach}$ mapped under the function $\lambda \langle a, m\rangle, f\ a\ (H\ a\ m)$. More precisely, if $H : \forall a \in xs, P(a)$, then: \[ (\text{pmap}\ f\ xs\ H).\text{head?} = \text{map}\ (\lambda \langle a, m\rangle, f\ a\ (H\ a\ m))\ (xs.\text{attach}.\text{head?}) \]
List.head_pmap theorem
{P : α → Prop} {f : (a : α) → P a → β} {xs : List α} (H : ∀ (a : α), a ∈ xs → P a) (h : xs.pmap f H ≠ []) : (xs.pmap f H).head h = f (xs.head (by simpa using h)) (H _ (head_mem _))
Full source
@[simp] theorem head_pmap {P : α → Prop} {f : (a : α) → P a → β} {xs : List α}
    (H : ∀ (a : α), a ∈ xs → P a) (h : xs.pmap f H ≠ []) :
    (xs.pmap f H).head h = f (xs.head (by simpa using h)) (H _ (head_mem _)) := by
  induction xs with
  | nil => simp at h
  | cons x xs ih => simp [head_pmap, ih]
Head of Partial Map Equals Function Applied to Head
Informal description
For any predicate $P$ on elements of type $\alpha$, any function $f$ that maps elements $a \in \alpha$ satisfying $P(a)$ to type $\beta$, and any non-empty list $xs$ of elements of type $\alpha$ where every element satisfies $P$, the head of the mapped list $\text{pmap}\ f\ xs\ H$ is equal to $f$ applied to the head of $xs$ and its corresponding proof that $P$ holds for it. More formally, if $xs$ is a non-empty list such that $\forall a \in xs, P(a)$, then: \[ \text{head}\ (\text{pmap}\ f\ xs\ H)\ h = f\ (\text{head}\ xs\ h')\ (H\ (\text{head}\ xs\ h')\ (\text{head\_mem}\ xs\ h')) \] where $h$ is a proof that $\text{pmap}\ f\ xs\ H$ is non-empty, and $h'$ is derived from $h$.
List.head?_attachWith theorem
{P : α → Prop} {xs : List α} (H : ∀ (a : α), a ∈ xs → P a) : (xs.attachWith P H).head? = xs.head?.pbind (fun a h => some ⟨a, H _ (mem_of_mem_head? h)⟩)
Full source
@[simp] theorem head?_attachWith {P : α → Prop} {xs : List α}
    (H : ∀ (a : α), a ∈ xs → P a) :
    (xs.attachWith P H).head? = xs.head?.pbind (fun a h => some ⟨a, H _ (mem_of_mem_head? h)⟩) := by
  cases xs <;> simp_all
Optional Head of Attached List Equals Partial Bind of Optional Head with Proof Pairing
Informal description
For any predicate $P$ on elements of type $\alpha$ and any list $xs$ of elements of type $\alpha$ where every element satisfies $P$, the optional head of the attached list $xs.\text{attachWith}\ P\ H$ is equal to the partial bind of the optional head of $xs$ with a function that pairs each element $a$ with its proof that $P(a)$ holds (using the fact that $a$ is in $xs$). More precisely, if $H : \forall a \in xs, P(a)$, then: \[ (xs.\text{attachWith}\ P\ H).\text{head?} = xs.\text{head?}.\text{pbind}\ (\lambda a\ h, \text{some}\ \langle a, H\ a\ (\text{mem\_of\_mem\_head?}\ h) \rangle) \]
List.head_attachWith theorem
{P : α → Prop} {xs : List α} {H : ∀ (a : α), a ∈ xs → P a} (h : xs.attachWith P H ≠ []) : (xs.attachWith P H).head h = ⟨xs.head (by simpa using h), H _ (head_mem _)⟩
Full source
@[simp] theorem head_attachWith {P : α → Prop} {xs : List α}
    {H : ∀ (a : α), a ∈ xs → P a} (h : xs.attachWith P H ≠ []) :
    (xs.attachWith P H).head h = ⟨xs.head (by simpa using h), H _ (head_mem _)⟩ := by
  cases xs with
  | nil => simp at h
  | cons x xs => simp [head_attachWith, h]
Head of Attached List Equals Head with Proof
Informal description
Let $P$ be a predicate on elements of type $\alpha$, and let $xs$ be a list of elements of type $\alpha$ such that every element $a \in xs$ satisfies $P(a)$. If the attached list $xs.\text{attachWith}\ P\ H$ is non-empty, then its head element is equal to the pair consisting of: 1. The head element of $xs$ (which exists since $xs.\text{attachWith}\ P\ H$ is non-empty) 2. The proof that $P$ holds for this head element (using the fact that the head element is in $xs$) In other words, for a non-empty attached list, the head is the head of the original list paired with its proof.
List.head?_attach theorem
{xs : List α} : xs.attach.head? = xs.head?.pbind (fun a h => some ⟨a, mem_of_mem_head? h⟩)
Full source
@[simp] theorem head?_attach {xs : List α} :
    xs.attach.head? = xs.head?.pbind (fun a h => some ⟨a, mem_of_mem_head? h⟩) := by
  cases xs <;> simp_all
Optional Head of Attached List Equals Partial Bind of Optional Head with Membership Proof
Informal description
For any list `xs` of elements of type `α`, the optional head of the attached list `xs.attach` is equal to the partial bind of the optional head of `xs` with a function that pairs each element `a` with a proof that `a` is in `xs`. More precisely, we have: \[ \text{head?}(xs.\text{attach}) = xs.\text{head?}.\text{pbind}\ (\lambda a\ h, \text{some}\ \langle a, \text{mem\_of\_mem\_head?}\ h \rangle) \]
List.head_attach theorem
{xs : List α} (h) : xs.attach.head h = ⟨xs.head (by simpa using h), head_mem (by simpa using h)⟩
Full source
@[simp] theorem head_attach {xs : List α} (h) :
    xs.attach.head h = ⟨xs.head (by simpa using h), head_mem (by simpa using h)⟩ := by
  cases xs with
  | nil => simp at h
  | cons x xs => simp [head_attach, h]
Head of Attached List Equals Head with Membership Proof
Informal description
For any non-empty list $xs$ of elements of type $\alpha$, the head of the attached list $xs.\text{attach}$ is equal to the pair consisting of the head of $xs$ and a proof that this head element is a member of $xs$.
List.tail_pmap theorem
{P : α → Prop} {f : (a : α) → P a → β} {xs : List α} (H : ∀ (a : α), a ∈ xs → P a) : (xs.pmap f H).tail = xs.tail.pmap f (fun a h => H a (mem_of_mem_tail h))
Full source
@[simp] theorem tail_pmap {P : α → Prop} {f : (a : α) → P a → β} {xs : List α}
    (H : ∀ (a : α), a ∈ xs → P a) :
    (xs.pmap f H).tail = xs.tail.pmap f (fun a h => H a (mem_of_mem_tail h)) := by
  cases xs <;> simp
Tail of Partial Map Equals Partial Map of Tail
Informal description
For any predicate $P$ on elements of type $\alpha$, any function $f$ mapping elements $a \in \alpha$ satisfying $P(a)$ to elements of type $\beta$, and any list $xs$ of elements of type $\alpha$ where every element satisfies $P$ (as witnessed by $H$), the tail of the list obtained by mapping $f$ over $xs$ is equal to the list obtained by mapping $f$ over the tail of $xs$ with the corresponding restriction of $H$.
List.tail_attachWith theorem
{P : α → Prop} {xs : List α} {H : ∀ (a : α), a ∈ xs → P a} : (xs.attachWith P H).tail = xs.tail.attachWith P (fun a h => H a (mem_of_mem_tail h))
Full source
@[simp] theorem tail_attachWith {P : α → Prop} {xs : List α}
    {H : ∀ (a : α), a ∈ xs → P a} :
    (xs.attachWith P H).tail = xs.tail.attachWith P (fun a h => H a (mem_of_mem_tail h)) := by
  cases xs <;> simp
Tail of Attached List Equals Attached Tail
Informal description
For any predicate $P$ on elements of type $\alpha$, any list $xs$ of elements of type $\alpha$, and any proof $H$ that every element in $xs$ satisfies $P$, the tail of the list obtained by attaching $P$ and $H$ to $xs$ is equal to the list obtained by attaching $P$ and the restricted proof to the tail of $xs$. Specifically, the restricted proof ensures that for any element $a$ in the tail of $xs$, $H a$ holds because $a$ is also in $xs$. In mathematical notation: \[ \text{tail}(\text{attachWith}(xs, P, H)) = \text{attachWith}(\text{tail}(xs), P, \lambda a h. H a (\text{mem\_of\_mem\_tail} h)) \]
List.tail_attach theorem
{xs : List α} : xs.attach.tail = xs.tail.attach.map (fun ⟨x, h⟩ => ⟨x, mem_of_mem_tail h⟩)
Full source
@[simp] theorem tail_attach {xs : List α} :
    xs.attach.tail = xs.tail.attach.map (fun ⟨x, h⟩ => ⟨x, mem_of_mem_tail h⟩) := by
  cases xs <;> simp
Tail of Attached List Equals Mapped Attached Tail
Informal description
For any list $xs$ of elements of type $\alpha$, the tail of the attached list $xs.\text{attach}$ is equal to the list obtained by mapping each element $\langle x, h \rangle$ in the attached list of the tail of $xs$ to $\langle x, \text{mem\_of\_mem\_tail}\ h \rangle$. In mathematical notation: \[ \text{tail}(xs.\text{attach}) = \text{map}\ (\lambda \langle x, h \rangle, \langle x, \text{mem\_of\_mem\_tail}\ h \rangle)\ (\text{tail}(xs).\text{attach}) \]
List.foldl_pmap theorem
{l : List α} {P : α → Prop} {f : (a : α) → P a → β} (H : ∀ (a : α), a ∈ l → P a) (g : γ → β → γ) (x : γ) : (l.pmap f H).foldl g x = l.attach.foldl (fun acc a => g acc (f a.1 (H _ a.2))) x
Full source
theorem foldl_pmap {l : List α} {P : α → Prop} {f : (a : α) → P a → β}
    (H : ∀ (a : α), a ∈ l → P a) (g : γ → β → γ) (x : γ) :
    (l.pmap f H).foldl g x = l.attach.foldl (fun acc a => g acc (f a.1 (H _ a.2))) x := by
  rw [pmap_eq_map_attach, foldl_map]
Left Fold of Partial Map Equals Fold over Attached List with Predicate Proofs
Informal description
For any list $l$ of elements of type $\alpha$, any predicate $P$ on $\alpha$, any function $f$ defined on elements $a$ of $\alpha$ that satisfy $P(a)$, and any proof $H$ that all elements in $l$ satisfy $P$, the left fold of the partial map of $f$ over $l$ with operation $g$ and initial value $x$ is equal to the left fold over the attached list of $l$ with operation $\lambda \text{acc}\ a, g\ \text{acc}\ (f\ a.1\ (H\ a.1\ a.2))$ and initial value $x$. In symbols: $$(\text{pmap}\ f\ l\ H).\text{foldl}\ g\ x = l.\text{attach}.\text{foldl}\ (\lambda \text{acc}\ a, g\ \text{acc}\ (f\ a.1\ (H\ a.1\ a.2)))\ x$$
List.foldr_pmap theorem
{l : List α} {P : α → Prop} {f : (a : α) → P a → β} (H : ∀ (a : α), a ∈ l → P a) (g : β → γ → γ) (x : γ) : (l.pmap f H).foldr g x = l.attach.foldr (fun a acc => g (f a.1 (H _ a.2)) acc) x
Full source
theorem foldr_pmap {l : List α} {P : α → Prop} {f : (a : α) → P a → β}
    (H : ∀ (a : α), a ∈ l → P a) (g : β → γ → γ) (x : γ) :
    (l.pmap f H).foldr g x = l.attach.foldr (fun a acc => g (f a.1 (H _ a.2)) acc) x := by
  rw [pmap_eq_map_attach, foldr_map]
Right Fold of Partial Map Equals Fold of Attached List with Composed Operation
Informal description
For any list $l$ of elements of type $\alpha$, any predicate $P$ on $\alpha$, any function $f$ defined on elements $a \in \alpha$ that satisfy $P(a)$, and any binary operation $g : \beta \to \gamma \to \gamma$ with initial value $x \in \gamma$, the right fold of the partial map of $f$ over $l$ (with respect to $P$) using $g$ and $x$ is equal to the right fold of the attached list of $l$ using the operation $\lambda a\ \text{acc}, g (f a.1 (H a.1 a.2)) \text{acc}$ and initial value $x$. In symbols: $$(\text{pmap}\ f\ l\ H).\text{foldr}\ g\ x = l.\text{attach}.\text{foldr}\ (\lambda a\ \text{acc}, g (f a.1 (H a.1 a.2)) \text{acc})\ x$$ where $a.1$ is the value and $a.2$ is the proof that $a.1 \in l$.
List.foldl_attachWith theorem
{l : List α} {q : α → Prop} (H : ∀ a, a ∈ l → q a) {f : β → { x // q x } → β} {b} : (l.attachWith q H).foldl f b = l.attach.foldl (fun b ⟨a, h⟩ => f b ⟨a, H _ h⟩) b
Full source
@[simp] theorem foldl_attachWith
    {l : List α} {q : α → Prop} (H : ∀ a, a ∈ l → q a) {f : β → { x // q x } → β} {b} :
    (l.attachWith q H).foldl f b = l.attach.foldl (fun b ⟨a, h⟩ => f b ⟨a, H _ h⟩) b := by
  induction l generalizing b with
  | nil => simp
  | cons a l ih => simp [ih, foldl_map]
Left Fold Equivalence for Attached Lists with Predicate
Informal description
For any list $l$ of elements of type $\alpha$, any predicate $q$ on $\alpha$ such that $H : \forall a \in l, q(a)$ holds, any function $f : \beta \to \{x \mid q(x)\} \to \beta$, and any initial value $b : \beta$, the left fold of the list $l.\text{attachWith}\ q\ H$ with operation $f$ and initial value $b$ is equal to the left fold of the attached list $l.\text{attach}$ with operation $\lambda b \langle a, h\rangle, f b \langle a, H a h\rangle$ and initial value $b$. In symbols: $$(l.\text{attachWith}\ q\ H).\text{foldl}\ f\ b = l.\text{attach}.\text{foldl}\ (\lambda b \langle a, h\rangle, f b \langle a, H a h\rangle)\ b$$
List.foldr_attachWith theorem
{l : List α} {q : α → Prop} (H : ∀ a, a ∈ l → q a) {f : { x // q x } → β → β} {b} : (l.attachWith q H).foldr f b = l.attach.foldr (fun a acc => f ⟨a.1, H _ a.2⟩ acc) b
Full source
@[simp] theorem foldr_attachWith
    {l : List α} {q : α → Prop} (H : ∀ a, a ∈ l → q a) {f : { x // q x } → β → β} {b} :
    (l.attachWith q H).foldr f b = l.attach.foldr (fun a acc => f ⟨a.1, H _ a.2⟩ acc) b := by
  induction l generalizing b with
  | nil => simp
  | cons a l ih => simp [ih, foldr_map]
Right Fold Equivalence for Attached Lists with Predicate
Informal description
For any list $l$ of elements of type $\alpha$, any predicate $q$ on $\alpha$ such that every element $a$ in $l$ satisfies $q(a)$, any function $f$ from the subtype $\{x \mid q(x)\}$ to $\beta \to \beta$, and any initial value $b$ of type $\beta$, the right fold of the list $l.\text{attachWith}\, q\, H$ with operation $f$ and initial value $b$ is equal to the right fold of the attached list $l.\text{attach}$ with operation $\lambda a\, \text{acc}, f \langle a.1, H\, a.1\, a.2 \rangle\, \text{acc}$ and initial value $b$. In symbols: $$(l.\text{attachWith}\, q\, H).\text{foldr}\, f\, b = l.\text{attach}.\text{foldr}\, (\lambda a\, \text{acc}, f \langle a.1, H\, a.1\, a.2 \rangle\, \text{acc})\, b$$
List.foldl_attach theorem
{l : List α} {f : β → α → β} {b : β} : l.attach.foldl (fun acc t => f acc t.1) b = l.foldl f b
Full source
/--
If we fold over `l.attach` with a function that ignores the membership predicate,
we get the same results as folding over `l` directly.

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
  induction l generalizing b with
  | nil => simp
  | cons a l ih => rw [foldl_cons, attach_cons, foldl_cons, foldl_map, ih]
Left Fold Equality Between Attached and Original List
Informal description
For any list $l$ of elements of type $\alpha$, any binary operation $f : \beta \to \alpha \to \beta$, and any initial value $b : \beta$, the left fold of $f$ over the attached list $l.\text{attach}$ (where each element is paired with its membership proof) is equal to the left fold of $f$ over the original list $l$. More precisely, if we define the operation on attached elements as $\lambda \text{acc}\ t, f\ \text{acc}\ t.1$ (where $t.1$ extracts the underlying value from the attached pair), then: $$ l.\text{attach}.\text{foldl}\ (\lambda \text{acc}\ t, f\ \text{acc}\ t.1)\ b = l.\text{foldl}\ f\ b $$
List.foldr_attach theorem
{l : List α} {f : α → β → β} {b : β} : l.attach.foldr (fun t acc => f t.1 acc) b = l.foldr f b
Full source
/--
If we fold over `l.attach` with a function that ignores the membership predicate,
we get the same results as folding over `l` directly.

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
  induction l generalizing b with
  | nil => simp
  | cons a l ih => rw [foldr_cons, attach_cons, foldr_cons, foldr_map, ih]
Right Fold Invariance under List Attachment
Informal description
For any list $l$ of elements of type $\alpha$, any binary function $f : \alpha \to \beta \to \beta$, and any initial value $b : \beta$, the right fold of $f$ over the attached list $l.\text{attach}$ (where each element is paired with a membership proof) is equal to the right fold of $f$ over the original list $l$. In symbols: $$\text{foldr } (\lambda t\ acc, f\ t.1\ acc)\ b\ (l.\text{attach}) = \text{foldr } f\ b\ l$$
List.attach_map theorem
{l : List α} {f : α → β} : (l.map f).attach = l.attach.map (fun ⟨x, h⟩ => ⟨f x, mem_map_of_mem h⟩)
Full source
theorem attach_map {l : List α} {f : α → β} :
    (l.map f).attach = l.attach.map (fun ⟨x, h⟩ => ⟨f x, mem_map_of_mem h⟩) := by
  induction l <;> simp [*]
Commutativity of List Mapping and Attachment: $(f(l)).\text{attach} = \text{map}\ (\lambda \langle x, h\rangle, \langle f(x), \text{mem\_map\_of\_mem}\ h\rangle)\ l.\text{attach}$
Informal description
For any list $l$ of elements of type $\alpha$ and any function $f : \alpha \to \beta$, the attached list of the mapped list $f(l)$ is equal to the result of mapping each element $\langle x, h\rangle$ in the attached list of $l$ to $\langle f(x), \text{mem\_map\_of\_mem}\ h\rangle$. In other words, the following diagram commutes: 1. First map $f$ over $l$, then attach the resulting list 2. First attach $l$, then map the attached elements by applying $f$ to the values while preserving the membership proofs.
List.attachWith_map theorem
{l : List α} {f : α → β} {P : β → Prop} (H : ∀ (b : β), b ∈ l.map f → P b) : (l.map f).attachWith P H = (l.attachWith (P ∘ f) (fun _ h => H _ (mem_map_of_mem h))).map fun ⟨x, h⟩ => ⟨f x, h⟩
Full source
theorem attachWith_map {l : List α} {f : α → β} {P : β → Prop} (H : ∀ (b : β), b ∈ l.map f → P b) :
    (l.map f).attachWith P H = (l.attachWith (P ∘ f) (fun _ h => H _ (mem_map_of_mem h))).map
      fun ⟨x, h⟩ => ⟨f x, h⟩ := by
  induction l <;> simp [*]
Commutativity of List Mapping and Attachment with Predicate
Informal description
Let $l$ be a list of elements of type $\alpha$, $f : \alpha \to \beta$ a function, and $P : \beta \to \text{Prop}$ a predicate. If for every $b \in \beta$ in the mapped list $f(l)$ we have $P(b)$ holds, then the attached list $(l.map\, f).attachWith\, P\, H$ is equal to first attaching $l$ with the composed predicate $P \circ f$ and then mapping each element $\langle x, h \rangle$ to $\langle f(x), h \rangle$. In other words, the following diagram commutes: 1. First map $f$ over $l$, then attach with $P$ 2. First attach $l$ with $P \circ f$, then map the attached elements by applying $f$ to the values while preserving the proofs
List.map_attachWith theorem
{l : List α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a} {f : { x // P x } → β} : (l.attachWith P H).map f = l.attach.map fun ⟨x, h⟩ => f ⟨x, H _ h⟩
Full source
@[simp] theorem map_attachWith {l : List α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a}
    {f : { x // P x } → β} :
    (l.attachWith P H).map f = l.attach.map fun ⟨x, h⟩ => f ⟨x, H _ h⟩ := by
  induction l <;> simp_all
Equality of Mapping Attached List with Predicate and Mapping Attached List with Membership Proofs
Informal description
For any list $l$ of elements of type $\alpha$, any predicate $P$ on $\alpha$ such that $P(a)$ holds for all $a \in l$, and any function $f$ defined on the subtype $\{x \mid P(x)\}$, the following equality holds: \[ \text{map}\ f\ (l.\text{attachWith}\ P\ H) = \text{map}\ (\lambda \langle x, h \rangle, f \langle x, H\ x\ h \rangle)\ l.\text{attach} \] where $l.\text{attachWith}\ P\ H$ is the list obtained by attaching proofs of $P$ to each element of $l$, and $l.\text{attach}$ is the list obtained by attaching membership proofs to each element of $l$.
List.map_attachWith_eq_pmap theorem
{l : List α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a} {f : { x // P x } → β} : (l.attachWith P H).map f = l.pmap (fun a (h : a ∈ l ∧ P a) => f ⟨a, H _ h.1⟩) (fun a h => ⟨h, H a h⟩)
Full source
theorem map_attachWith_eq_pmap {l : List α} {P : α → Prop} {H : ∀ (a : α), a ∈ l → P a}
    {f : { x // P x } → β} :
    (l.attachWith P H).map f =
      l.pmap (fun a (h : a ∈ la ∈ l ∧ P a) => f ⟨a, H _ h.1⟩) (fun a h => ⟨h, H a h⟩) := by
  induction l with
  | nil => rfl
  | cons x xs ih =>
    simp only [attachWith_cons, map_cons, ih, pmap, cons.injEq, true_and]
    apply pmap_congr_left
    simp
Equality of Mapping Attached List and Partial Map with Predicate
Informal description
For any list $l$ of elements of type $\alpha$, any predicate $P$ on $\alpha$, and any function $f$ defined on the subtype $\{x \mid P x\}$, the following equality holds: \[ \text{map}\ f\ (l.\text{attachWith}\ P\ H) = \text{pmap}\ (\lambda a\ h, f\ \langle a, H\ a\ h.1 \rangle)\ l\ (\lambda a\ h, \langle h, H\ a\ h \rangle) \] where $H$ is a proof that for every $a \in l$, $P(a)$ holds, and $\text{pmap}$ is the partial map function that applies $f$ only to elements satisfying $P$.
List.map_attach_eq_pmap theorem
{l : List α} {f : { x // x ∈ l } → β} : l.attach.map f = l.pmap (fun a h => f ⟨a, h⟩) (fun _ => id)
Full source
/-- See also `pmap_eq_map_attach` for writing `pmap` in terms of `map` and `attach`. -/
theorem map_attach_eq_pmap {l : List α} {f : { x // x ∈ l } → β} :
    l.attach.map f = l.pmap (fun a h => f ⟨a, h⟩) (fun _ => id) := by
  induction l with
  | nil => rfl
  | cons x xs ih =>
    simp only [attach_cons, map_cons, map_map, Function.comp_apply, pmap, cons.injEq, true_and, ih]
    apply pmap_congr_left
    simp
Equality of Map over Attached List and Partial Map with Membership Proof
Informal description
For any list $l$ of elements of type $\alpha$ and any function $f$ defined on the subtype $\{x \mid x \in l\}$, the map of $f$ over the attached list $l.\text{attach}$ is equal to the partial map of $\lambda a\ h, f \langle a, h \rangle$ over $l$ with the trivial proof $\lambda \_, \text{id}$. That is, \[ \text{map}\ f\ (l.\text{attach}) = \text{pmap}\ (\lambda a\ h, f \langle a, h \rangle)\ l\ (\lambda \_, \text{id}). \]
List.map_attach abbrev
Full source
@[deprecated map_attach_eq_pmap (since := "2025-02-09")]
abbrev map_attach := @map_attach_eq_pmap
Equality of Mapping Attached List and Mapping with Membership Proof
Informal description
Given a list $l$ of elements of type $\alpha$ and a function $f$ defined on the subtype $\{x \mid x \in l\}$, the map of $f$ over the attached list $l.\text{attach}$ is equal to the map of $\lambda x, f \langle x, \text{mem}_x \rangle$ over $l$, where $\text{mem}_x$ is the proof that $x \in l$.
List.attach_filterMap theorem
{l : List α} {f : α → Option β} : (l.filterMap f).attach = l.attach.filterMap fun ⟨x, h⟩ => (f x).pbind (fun b m => some ⟨b, mem_filterMap.mpr ⟨x, h, m⟩⟩)
Full source
theorem attach_filterMap {l : List α} {f : α → Option β} :
    (l.filterMap f).attach = l.attach.filterMap
      fun ⟨x, h⟩ => (f x).pbind (fun b m => some ⟨b, mem_filterMap.mpr ⟨x, h, m⟩⟩) := by
  induction l with
  | nil => rfl
  | cons x xs ih =>
    simp only [filterMap_cons, attach_cons, ih, filterMap_map]
    split <;> rename_i h
    · simp only [Option.pbind_eq_none_iff, reduceCtorEq, Option.mem_def, exists_false,
        or_false] at h
      rw [attach_congr]
      rotate_left
      · simp only [h]
        rfl
      rw [ih]
      simp only [map_filterMap, Option.map_pbind, Option.map_some']
      rfl
    · simp only [Option.pbind_eq_some_iff] at h
      obtain ⟨a, h, w⟩ := h
      simp only [Option.some.injEq] at w
      subst w
      simp only [Option.mem_def] at h
      rw [attach_congr]
      rotate_left
      · simp only [h]
        rfl
      rw [attach_cons, map_cons, map_map, ih, map_filterMap]
      congr
      ext
      simp
Attached Filtered Map Equals Filtered Map of Attached List via Partial Bind
Informal description
For any list $l$ of elements of type $\alpha$ and any function $f : \alpha \to \text{Option } \beta$, the attached list of the filtered map $\text{filterMap } f l$ is equal to the filtered map obtained by: 1. First attaching $l$ to get pairs $\langle x, h\rangle$ where $x \in l$ with proof $h$ 2. Then applying a partial bind operation that for each $\langle x, h\rangle$: - Takes $f x$ (an optional $\beta$) - For each $b$ in $f x$ (with proof $m$ that $b$ is in $f x$) - Returns $\langle b, \text{mem\_filterMap.mpr } \langle x, h, m\rangle\rangle$ In other words: $$(l.\text{filterMap } f).\text{attach} = l.\text{attach}.\text{filterMap } (\lambda \langle x, h\rangle, (f x).\text{pbind } (\lambda b m, \text{some } \langle b, \text{mem\_filterMap.mpr } \langle x, h, m\rangle\rangle))$$
List.attach_filter theorem
{l : List α} (p : α → Bool) : (l.filter p).attach = l.attach.filterMap fun x => if w : p x.1 then some ⟨x.1, mem_filter.mpr ⟨x.2, w⟩⟩ else none
Full source
theorem attach_filter {l : List α} (p : α → Bool) :
    (l.filter p).attach = l.attach.filterMap
      fun x => if w : p x.1 then some ⟨x.1, mem_filter.mpr ⟨x.2, w⟩⟩ else none := by
  rw [attach_congr (congrFun filterMap_eq_filter.symm _), attach_filterMap, map_filterMap]
  simp only [Option.guard]
  congr
  ext1
  split <;> simp
Attached Filtered List Equals Filtered Map of Attached List via Predicate Check
Informal description
For any list $l$ of elements of type $\alpha$ and any predicate $p : \alpha \to \text{Bool}$, the attached list of the filtered list $\text{filter } p l$ is equal to the filtered map obtained by: 1. Attaching $l$ to get pairs $\langle x, h\rangle$ where $x \in l$ with proof $h$ 2. For each $\langle x, h\rangle$, checking if $p(x)$ holds (with proof $w$ if it does) 3. If $p(x)$ holds, returning $\langle x, \text{mem\_filter.mpr } \langle h, w\rangle\rangle$ (wrapped in `some`), otherwise returning `none` In other words: $$(l.\text{filter } p).\text{attach} = l.\text{attach}.\text{filterMap } (\lambda \langle x, h\rangle, \text{if } p(x) \text{ then some } \langle x, \text{mem\_filter.mpr } \langle h, p(x)\rangle\rangle \text{ else none})$$
List.filterMap_attachWith theorem
{q : α → Prop} {l : List α} {f : { x // q x } → Option β} (H) : (l.attachWith q H).filterMap f = l.attach.filterMap (fun ⟨x, h⟩ => f ⟨x, H _ h⟩)
Full source
@[simp]
theorem filterMap_attachWith {q : α → Prop} {l : List α} {f : {x // q x}Option β} (H) :
    (l.attachWith q H).filterMap f = l.attach.filterMap (fun ⟨x, h⟩ => f ⟨x, H _ h⟩) := by
  induction l with
  | nil => rfl
  | cons x xs ih =>
    simp only [attachWith_cons, filterMap_cons]
    split <;> simp_all [Function.comp_def]
Filter-Map Equality for Attached Lists with Predicate
Informal description
For any predicate $q$ on a type $\alpha$, any list $l$ of elements of type $\alpha$, and any function $f$ from the subtype $\{x \mid q(x)\}$ to optional values of type $\beta$, given a proof $H$ that every element $x$ in $l$ satisfies $q(x)$, the following equality holds: \[ \text{filterMap } f (l.\text{attachWith } q H) = l.\text{attach}.\text{filterMap } (\lambda \langle x, h\rangle, f \langle x, H x h\rangle) \]
List.filter_attachWith theorem
{q : α → Prop} {l : List α} {p : { x // q x } → Bool} (H) : (l.attachWith q H).filter p = (l.attach.filter (fun ⟨x, h⟩ => p ⟨x, H _ h⟩)).map (fun ⟨x, h⟩ => ⟨x, H _ h⟩)
Full source
@[simp]
theorem filter_attachWith {q : α → Prop} {l : List α} {p : {x // q x}Bool} (H) :
    (l.attachWith q H).filter p =
      (l.attach.filter (fun ⟨x, h⟩ => p ⟨x, H _ h⟩)).map (fun ⟨x, h⟩ => ⟨x, H _ h⟩) := by
  induction l with
  | nil => rfl
  | cons x xs ih =>
    simp only [attachWith_cons, filter_cons]
    split <;> simp_all [Function.comp_def, filter_map]
Filtering Attached List with Predicate Equals Mapping Filtered Attached List
Informal description
For any predicate $q$ on elements of type $\alpha$, any list $l$ of elements of type $\alpha$, any predicate $p$ on the subtype $\{x \mid q(x)\}$, and any proof $H$ that all elements in $l$ satisfy $q$, filtering the list obtained by attaching $q$-proofs to $l$ with $p$ is equal to mapping the corresponding elements with their proofs over the result of filtering the attached list of $l$ with the composed predicate $p \circ (\lambda \langle x, h\rangle, \langle x, H x h\rangle)$. In symbols: \[ \text{filter}\ p\ (l.\text{attachWith}\ q\ H) = \text{map}\ (\lambda \langle x, h\rangle, \langle x, H x h\rangle)\ (\text{filter}\ (\lambda \langle x, h\rangle, p\ \langle x, H x h\rangle)\ l.\text{attach}) \]
List.pmap_pmap theorem
{p : α → Prop} {q : β → Prop} {g : ∀ a, p a → β} {f : ∀ b, q b → γ} {l} (H₁ H₂) : pmap f (pmap g l H₁) H₂ = pmap (α := { x // x ∈ l }) (fun a h => f (g a h) (H₂ (g a h) (mem_pmap_of_mem a.2))) l.attach (fun a _ => H₁ a a.2)
Full source
theorem pmap_pmap {p : α → Prop} {q : β → Prop} {g : ∀ a, p a → β} {f : ∀ b, q b → γ} {l} (H₁ H₂) :
    pmap f (pmap g l H₁) H₂ =
      pmap (α := { x // x ∈ l }) (fun a h => f (g a h) (H₂ (g a h) (mem_pmap_of_mem a.2))) l.attach
        (fun a _ => H₁ a a.2) := by
  simp [pmap_eq_map_attach, attach_map]
Composition of Partial Maps via Attached List
Informal description
For any predicates $p$ on $\alpha$ and $q$ on $\beta$, any functions $g : \forall a \in \alpha, p(a) \to \beta$ and $f : \forall b \in \beta, q(b) \to \gamma$, and any list $l$ of elements of type $\alpha$ with proofs $H_1$ that all elements in $l$ satisfy $p$ and $H_2$ that all elements in the image of $g$ satisfy $q$, the following equality holds: \[ \text{pmap}\ f\ (\text{pmap}\ g\ l\ H_1)\ H_2 = \text{pmap}\ (\lambda a\ h, f\ (g\ a\ h)\ (H_2\ (g\ a\ h)\ (\text{mem\_pmap\_of\_mem}\ a.2)))\ l.\text{attach}\ (\lambda a\ \_, H_1\ a\ a.2) \] where $a.2$ denotes the proof that $a \in l$.
List.pmap_append theorem
{p : ι → Prop} {f : ∀ a : ι, p a → α} {l₁ l₂ : List ι} (h : ∀ a ∈ l₁ ++ l₂, p a) : (l₁ ++ l₂).pmap f h = (l₁.pmap f fun a ha => h a (mem_append_left l₂ ha)) ++ l₂.pmap f fun a ha => h a (mem_append_right l₁ ha)
Full source
@[simp] theorem pmap_append {p : ι → Prop} {f : ∀ a : ι, p a → α} {l₁ l₂ : List ι}
    (h : ∀ a ∈ l₁ ++ l₂, p a) :
    (l₁ ++ l₂).pmap f h =
      (l₁.pmap f fun a ha => h a (mem_append_left l₂ ha)) ++
        l₂.pmap f fun a ha => h a (mem_append_right l₁ ha) := by
  induction l₁ with
  | nil => rfl
  | cons _ _ ih =>
    dsimp only [pmap, cons_append]
    rw [ih]
Partial Map Distributes Over List Concatenation
Informal description
For any predicate $p$ on elements of type $\iota$ and any function $f$ defined on elements $a \in \iota$ satisfying $p(a)$, the partial map operation `pmap` distributes over list concatenation. Specifically, for any lists $l_1$ and $l_2$ of type $\text{List } \iota$ and any proof $h$ that all elements in $l_1 \mathbin{+\kern-0.5em+} l_2$ satisfy $p$, we have: $$(l_1 \mathbin{+\kern-0.5em+} l_2).\text{pmap}\ f\ h = (l_1.\text{pmap}\ f\ h_1) \mathbin{+\kern-0.5em+} (l_2.\text{pmap}\ f\ h_2)$$ where $h_1$ and $h_2$ are the restrictions of $h$ to $l_1$ and $l_2$ respectively.
List.pmap_append' theorem
{p : α → Prop} {f : ∀ a : α, p a → β} {l₁ l₂ : List α} (h₁ : ∀ a ∈ l₁, p a) (h₂ : ∀ a ∈ l₂, p a) : ((l₁ ++ l₂).pmap f fun a ha => (List.mem_append.1 ha).elim (h₁ a) (h₂ a)) = l₁.pmap f h₁ ++ l₂.pmap f h₂
Full source
theorem pmap_append' {p : α → Prop} {f : ∀ a : α, p a → β} {l₁ l₂ : List α}
    (h₁ : ∀ a ∈ l₁, p a) (h₂ : ∀ a ∈ l₂, p a) :
    ((l₁ ++ l₂).pmap f fun a ha => (List.mem_append.1 ha).elim (h₁ a) (h₂ a)) =
      l₁.pmap f h₁ ++ l₂.pmap f h₂ :=
  pmap_append _
Partial Map Distributes Over List Concatenation (Case Analysis Version)
Informal description
For any predicate $p$ on elements of type $\alpha$ and any function $f$ defined on elements $a \in \alpha$ satisfying $p(a)$, given two lists $l_1$ and $l_2$ of type $\text{List } \alpha$ with proofs $h_1$ and $h_2$ that all elements in $l_1$ and $l_2$ respectively satisfy $p$, the partial map operation `pmap` satisfies: $$(l_1 \mathbin{+\kern-0.5em+} l_2).\text{pmap}\ f\ h = (l_1.\text{pmap}\ f\ h_1) \mathbin{+\kern-0.5em+} (l_2.\text{pmap}\ f\ h_2)$$ where $h$ is constructed from $h_1$ and $h_2$ by case analysis on membership in $l_1$ or $l_2$.
List.attach_append theorem
{xs ys : List α} : (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⟩
Full source
@[simp] theorem attach_append {xs ys : List α} :
    (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
Attachment Distributes Over List Concatenation
Informal description
For any two lists $xs$ and $ys$ of elements of type $\alpha$, the attached version of their concatenation $xs \mathbin{+\kern-0.5em+} ys$ is equal to the concatenation of: 1. The attached version of $xs$ with membership proofs adjusted to show membership in $xs \mathbin{+\kern-0.5em+} ys$ via left inclusion, and 2. The attached version of $ys$ with membership proofs adjusted to show membership in $xs \mathbin{+\kern-0.5em+} ys$ via right inclusion. In other words: $$(xs \mathbin{+\kern-0.5em+} ys).\text{attach} = \text{map}\ (\lambda \langle x, h \rangle \mapsto \langle x, \text{mem\_append\_left}\ ys\ h \rangle)\ xs.\text{attach} \mathbin{+\kern-0.5em+} \text{map}\ (\lambda \langle x, h \rangle \mapsto \langle x, \text{mem\_append\_right}\ xs\ h \rangle)\ ys.\text{attach}$$
List.attachWith_append theorem
{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_left ys h)) ++ ys.attachWith P (fun a h => H a (mem_append_right xs h))
Full source
@[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_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]
AttachWith Distributes Over List Concatenation
Informal description
For any predicate $P$ on elements of type $\alpha$ and any lists $xs, ys$ of type $\text{List } \alpha$, if $H$ is a proof that every element in the concatenated list $xs \mathbin{+\kern-0.5em+} ys$ satisfies $P$, then the attached list with predicate $P$ of the concatenation is equal to the concatenation of the attached lists of $xs$ and $ys$ with appropriately restricted proofs. Specifically: $$(xs \mathbin{+\kern-0.5em+} ys).\text{attachWith}\ P\ H = xs.\text{attachWith}\ P\ H_1 \mathbin{+\kern-0.5em+} ys.\text{attachWith}\ P\ H_2$$ where $H_1$ and $H_2$ are the restrictions of $H$ to $xs$ and $ys$ respectively.
List.pmap_reverse theorem
{P : α → Prop} {f : (a : α) → P a → β} {xs : List α} (H : ∀ (a : α), a ∈ xs.reverse → P a) : xs.reverse.pmap f H = (xs.pmap f (fun a h => H a (by simpa using h))).reverse
Full source
@[simp] theorem pmap_reverse {P : α → Prop} {f : (a : α) → P a → β} {xs : List α}
    (H : ∀ (a : α), a ∈ xs.reverse → P a) :
    xs.reverse.pmap f H = (xs.pmap f (fun a h => H a (by simpa using h))).reverse := by
  induction xs <;> simp_all
Partial Map Commutes with List Reversal: $\text{pmap}\ f\ H\ (xs.\text{reverse}) = (\text{pmap}\ f\ H'\ xs).\text{reverse}$
Informal description
For any predicate $P$ on elements of type $\alpha$, any function $f$ defined on elements $a \in \alpha$ satisfying $P(a)$, and any list $xs$ of type $\text{List } \alpha$, if $H$ is a proof that all elements in the reversed list $xs.\text{reverse}$ satisfy $P$, then the partial map of $f$ over the reversed list $xs.\text{reverse}$ with proof $H$ is equal to the reverse of the partial map of $f$ over the original list $xs$ with appropriately adjusted proofs. In symbols: $$ \text{pmap}\ f\ H\ (xs.\text{reverse}) = (\text{pmap}\ f\ H'\ xs).\text{reverse} $$ where $H'$ is the proof that all elements in $xs$ satisfy $P$, derived from $H$ via the reverse membership relation.
List.reverse_pmap theorem
{P : α → Prop} {f : (a : α) → P a → β} {xs : List α} (H : ∀ (a : α), a ∈ xs → P a) : (xs.pmap f H).reverse = xs.reverse.pmap f (fun a h => H a (by simpa using h))
Full source
theorem reverse_pmap {P : α → Prop} {f : (a : α) → P a → β} {xs : List α}
    (H : ∀ (a : α), a ∈ xs → P a) :
    (xs.pmap f H).reverse = xs.reverse.pmap f (fun a h => H a (by simpa using h)) := by
  rw [pmap_reverse]
Reversed Partial Map Equals Partial Map of Reversed List
Informal description
For any predicate $P$ on elements of type $\alpha$, any function $f$ defined on elements $a \in \alpha$ satisfying $P(a)$, and any list $xs$ of type $\text{List } \alpha$, if $H$ is a proof that all elements in $xs$ satisfy $P$, then the reverse of the partial map of $f$ over $xs$ with proof $H$ is equal to the partial map of $f$ over the reversed list $xs.\text{reverse}$ with appropriately adjusted proofs. In symbols: $$ (\text{pmap}\ f\ H\ xs).\text{reverse} = \text{pmap}\ f\ H'\ (xs.\text{reverse}) $$ where $H'$ is the proof that all elements in $xs.\text{reverse}$ satisfy $P$, derived from $H$ via the reverse membership relation.
List.attachWith_reverse theorem
{P : α → Prop} {xs : List α} {H : ∀ (a : α), a ∈ xs.reverse → P a} : xs.reverse.attachWith P H = (xs.attachWith P (fun a h => H a (by simpa using h))).reverse
Full source
@[simp] theorem attachWith_reverse {P : α → Prop} {xs : List α}
    {H : ∀ (a : α), a ∈ xs.reverse → P a} :
    xs.reverse.attachWith P H =
      (xs.attachWith P (fun a h => H a (by simpa using h))).reverse :=
  pmap_reverse ..
Reversed List Attachment Equals Reverse of Attachment: $xs.\text{reverse}.\text{attachWith}\ P\ H = (xs.\text{attachWith}\ P\ H').\text{reverse}$
Informal description
For any type $\alpha$, predicate $P$ on $\alpha$, list $xs$ of elements in $\alpha$, and proof $H$ that all elements in the reversed list $xs.\text{reverse}$ satisfy $P$, we have: $$ xs.\text{reverse}.\text{attachWith}\ P\ H = (xs.\text{attachWith}\ P\ H').\text{reverse} $$ where $H'$ is the proof that all elements in $xs$ satisfy $P$, derived from $H$ via the reverse membership relation.
List.reverse_attachWith theorem
{P : α → Prop} {xs : List α} {H : ∀ (a : α), a ∈ xs → P a} : (xs.attachWith P H).reverse = (xs.reverse.attachWith P (fun a h => H a (by simpa using h)))
Full source
theorem reverse_attachWith {P : α → Prop} {xs : List α}
    {H : ∀ (a : α), a ∈ xs → P a} :
    (xs.attachWith P H).reverse = (xs.reverse.attachWith P (fun a h => H a (by simpa using h))) :=
  reverse_pmap ..
Reversed Attached List Equals Attached Reversed List
Informal description
For any type $\alpha$, predicate $P$ on $\alpha$, list $xs$ of elements in $\alpha$, and proof $H$ that all elements in $xs$ satisfy $P$, the reverse of the attached list with predicate $P$ is equal to the attached list with predicate $P$ of the reversed list with appropriately adjusted proofs. In symbols: $$ (xs.\text{attachWith}\ P\ H).\text{reverse} = xs.\text{reverse}.\text{attachWith}\ P\ H' $$ where $H'$ is the proof that all elements in $xs.\text{reverse}$ satisfy $P$, derived from $H$ via the reverse membership relation.
List.attach_reverse theorem
{xs : List α} : xs.reverse.attach = xs.attach.reverse.map fun ⟨x, h⟩ => ⟨x, by simpa using h⟩
Full source
@[simp] theorem attach_reverse {xs : List α} :
    xs.reverse.attach = xs.attach.reverse.map fun ⟨x, h⟩ => ⟨x, by simpa using h⟩ := by
  simp only [attach, attachWith, reverse_pmap, map_pmap]
  apply pmap_congr_left
  intros
  rfl
Reversed List Attachment Equals Mapped Reverse Attachment
Informal description
For any list $xs$ of elements of type $\alpha$, the attached version of the reversed list $xs.\text{reverse}$ is equal to the reversed attached list $xs.\text{attach}$ with each element's membership proof adjusted to account for the reversal. In symbols: $$ \text{attach}(xs.\text{reverse}) = \text{map}\ (\lambda \langle x, h \rangle, \langle x, \text{by simpa using } h \rangle)\ (\text{attach}(xs).\text{reverse}) $$
List.reverse_attach theorem
{xs : List α} : xs.attach.reverse = xs.reverse.attach.map fun ⟨x, h⟩ => ⟨x, by simpa using h⟩
Full source
theorem reverse_attach {xs : List α} :
    xs.attach.reverse = xs.reverse.attach.map fun ⟨x, h⟩ => ⟨x, by simpa using h⟩ := by
  simp only [attach, attachWith, reverse_pmap, map_pmap]
  apply pmap_congr_left
  intros
  rfl
Reversed Attached List Equals Mapped Attached Reversed List
Informal description
For any list `xs` of elements of type `α`, the reverse of the attached list `xs.attach` (where each element is paired with a proof of its membership in `xs`) is equal to the attached list of the reversed `xs`, with each element paired with a proof of its membership in the reversed list. In symbols: $$ \text{reverse}(\text{attach}(xs)) = \text{map}\ (\lambda \langle x, h \rangle, \langle x, \text{by simpa using } h \rangle)\ (\text{attach}(\text{reverse}(xs))) $$
List.getLast?_pmap theorem
{P : α → Prop} {f : (a : α) → P a → β} {xs : List α} (H : ∀ (a : α), a ∈ xs → P a) : (xs.pmap f H).getLast? = xs.attach.getLast?.map fun ⟨a, m⟩ => f a (H a m)
Full source
@[simp] theorem getLast?_pmap {P : α → Prop} {f : (a : α) → P a → β} {xs : List α}
    (H : ∀ (a : α), a ∈ xs → P a) :
    (xs.pmap f H).getLast? = xs.attach.getLast?.map fun ⟨a, m⟩ => f a (H a m) := by
  simp only [getLast?_eq_head?_reverse]
  rw [reverse_pmap, reverse_attach, head?_map, pmap_eq_map_attach, head?_map]
  simp only [Option.map_map]
  congr
Optional Last Element of Partial Map Equals Mapped Last Element of Attached List
Informal description
For any predicate $P$ on elements of type $\alpha$, any function $f$ defined on elements $a \in \alpha$ satisfying $P(a)$, and any list $xs$ of type $\text{List } \alpha$, if $H$ is a proof that all elements in $xs$ satisfy $P$, then the optional last element of the partial map of $f$ over $xs$ with proof $H$ is equal to mapping $f$ over the optional last element of the attached list $xs.\text{attach}$ (where each element is paired with its membership proof). In symbols: $$ \text{getLast?}(\text{pmap}\ f\ H\ xs) = \text{Option.map}\ (\lambda \langle a, m \rangle, f\ a\ (H\ a\ m))\ (\text{getLast?}(xs.\text{attach})) $$
List.getLast_pmap theorem
{P : α → Prop} {f : (a : α) → P a → β} {xs : List α} (H : ∀ (a : α), a ∈ xs → P a) (h : xs.pmap f H ≠ []) : (xs.pmap f H).getLast h = f (xs.getLast (by simpa using h)) (H _ (getLast_mem _))
Full source
@[simp] theorem getLast_pmap {P : α → Prop} {f : (a : α) → P a → β} {xs : List α}
    (H : ∀ (a : α), a ∈ xs → P a) (h : xs.pmap f H ≠ []) :
    (xs.pmap f H).getLast h = f (xs.getLast (by simpa using h)) (H _ (getLast_mem _)) := by
  simp only [getLast_eq_head_reverse]
  simp only [reverse_pmap, head_pmap, head_reverse]
Last Element of Partial Map Equals Function Applied to Last Element
Informal description
For any predicate $P$ on elements of type $\alpha$, any function $f : (a : \alpha) \to P(a) \to \beta$, and any non-empty list $xs$ of elements of type $\alpha$ where every element satisfies $P$, the last element of the partially mapped list $\text{pmap}\ f\ xs\ H$ is equal to $f$ applied to the last element of $xs$ and its corresponding proof that $P$ holds for it. In symbols: Given $H : \forall a \in xs, P(a)$ and $h : \text{pmap}\ f\ xs\ H \neq []$, we have: $$ \text{getLast}(\text{pmap}\ f\ xs\ H, h) = f(\text{getLast}(xs, h')) (H(\text{getLast}(xs, h')) (\text{getLast\_mem}(xs, h'))) $$ where $h'$ is derived from $h$ via simplification.
List.getLast?_attachWith theorem
{P : α → Prop} {xs : List α} {H : ∀ (a : α), a ∈ xs → P a} : (xs.attachWith P H).getLast? = xs.getLast?.pbind (fun a h => some ⟨a, H _ (mem_of_getLast? h)⟩)
Full source
@[simp] theorem getLast?_attachWith {P : α → Prop} {xs : List α}
    {H : ∀ (a : α), a ∈ xs → P a} :
    (xs.attachWith P H).getLast? = xs.getLast?.pbind (fun a h => some ⟨a, H _ (mem_of_getLast? h)⟩) := by
  rw [getLast?_eq_head?_reverse, reverse_attachWith, head?_attachWith]
  simp
Optional Last Element of Attached List Equals Partial Bind of Optional Last Element with Proof Pairing
Informal description
For any type $\alpha$, predicate $P$ on $\alpha$, list $xs$ of elements in $\alpha$, and proof $H$ that all elements in $xs$ satisfy $P$, the optional last element of the attached list $xs.\text{attachWith}\ P\ H$ is equal to the partial bind of the optional last element of $xs$ with a function that pairs each element $a$ with its proof that $P(a)$ holds (using the fact that $a$ is in $xs$). In symbols: $$ \text{getLast?}(xs.\text{attachWith}\ P\ H) = xs.\text{getLast?}.\text{pbind}\ (\lambda a\ h, \text{some}\ \langle a, H\ a\ (\text{mem\_of\_getLast?}\ h) \rangle) $$
List.getLast_attachWith theorem
{P : α → Prop} {xs : List α} {H : ∀ (a : α), a ∈ xs → P a} (h : xs.attachWith P H ≠ []) : (xs.attachWith P H).getLast h = ⟨xs.getLast (by simpa using h), H _ (getLast_mem _)⟩
Full source
@[simp] theorem getLast_attachWith {P : α → Prop} {xs : List α}
    {H : ∀ (a : α), a ∈ xs → P a} (h : xs.attachWith P H ≠ []) :
    (xs.attachWith P H).getLast h = ⟨xs.getLast (by simpa using h), H _ (getLast_mem _)⟩ := by
  simp only [getLast_eq_head_reverse, reverse_attachWith, head_attachWith, head_map]
Last Element of Attached List Equals Last with Proof
Informal description
For any type $\alpha$, predicate $P$ on $\alpha$, list $xs$ of elements in $\alpha$, and proof $H$ that all elements in $xs$ satisfy $P$, if the attached list $xs.\text{attachWith}\ P\ H$ is non-empty, then its last element is equal to the pair consisting of: 1. The last element of $xs$ (which exists since $xs.\text{attachWith}\ P\ H$ is non-empty) 2. The proof that $P$ holds for this last element (using the fact that the last element is in $xs$) In other words, for a non-empty attached list, the last element is the last element of the original list paired with its proof.
List.getLast?_attach theorem
{xs : List α} : xs.attach.getLast? = xs.getLast?.pbind fun a h => some ⟨a, mem_of_getLast? h⟩
Full source
@[simp]
theorem getLast?_attach {xs : List α} :
    xs.attach.getLast? = xs.getLast?.pbind fun a h => some ⟨a, mem_of_getLast? h⟩ := by
  rw [getLast?_eq_head?_reverse, reverse_attach, head?_map, head?_attach]
  simp
Last Element Preservation under List Attachment: $\text{getLast?}(\text{attach}(xs)) = \text{pbind}(\text{getLast?}(xs), \lambda a\ h, \text{some}\ \langle a, h \rangle)$
Informal description
For any list $xs$ of elements of type $\alpha$, the optional last element of the attached list $xs.\text{attach}$ (where each element is paired with a proof of membership in $xs$) is equal to the partial bind of the optional last element of $xs$ with a function that pairs each element $a$ with a proof that $a$ is in $xs$. In symbols: $$ \text{getLast?}(xs.\text{attach}) = xs.\text{getLast?}.\text{pbind}\ (\lambda a\ h, \text{some}\ \langle a, \text{mem\_of\_getLast?}\ h \rangle) $$
List.getLast_attach theorem
{xs : List α} (h : xs.attach ≠ []) : xs.attach.getLast h = ⟨xs.getLast (by simpa using h), getLast_mem (by simpa using h)⟩
Full source
@[simp]
theorem getLast_attach {xs : List α} (h : xs.attach ≠ []) :
    xs.attach.getLast h = ⟨xs.getLast (by simpa using h), getLast_mem (by simpa using h)⟩ := by
  simp only [getLast_eq_head_reverse, reverse_attach, head_map, head_attach]
Last Element of Attached List Equals Last Element with Membership Proof
Informal description
For any non-empty attached list `xs.attach` (where each element is paired with a proof of membership in `xs`), the last element of the attached list is equal to the pair consisting of the last element of `xs` and a proof that this element is a member of `xs`.
List.countP_attach theorem
{l : List α} {p : α → Bool} : l.attach.countP (fun a : { x // x ∈ l } => p a) = l.countP p
Full source
@[simp]
theorem countP_attach {l : List α} {p : α → Bool} :
    l.attach.countP (fun a : {x // x ∈ l} => p a) = l.countP p := by
  simp only [← Function.comp_apply (g := Subtype.val), ← countP_map, attach_map_subtype_val]
Count Preservation under List Attachment: $\text{countP } p\ (l.\text{attach}) = \text{countP } p\ l$
Informal description
For any list $l$ of elements of type $\alpha$ and any Boolean predicate $p$ on $\alpha$, the count of elements in the attached list (where each element is paired with a proof of membership in $l$) that satisfy $p$ is equal to the count of elements in the original list $l$ that satisfy $p$. In mathematical notation: \[ \text{countP } p\ (l.\text{attach}) = \text{countP } p\ l \]
List.countP_attachWith theorem
{p : α → Prop} {q : α → Bool} {l : List α} (H : ∀ a ∈ l, p a) : (l.attachWith p H).countP (fun a : { x // p x } => q a) = l.countP q
Full source
@[simp]
theorem countP_attachWith {p : α → Prop} {q : α → Bool} {l : List α} (H : ∀ a ∈ l, p a) :
    (l.attachWith p H).countP (fun a : {x // p x} => q a) = l.countP q := by
  simp only [← Function.comp_apply (g := Subtype.val), ← countP_map, attachWith_map_subtype_val]
Count Preservation under List Attachment with Predicate: $\text{countP } q \circ \text{attachWith } p = \text{countP } q$
Informal description
For any predicate $p$ on elements of type $\alpha$, any Boolean-valued predicate $q$ on $\alpha$, and any list $l$ of elements of type $\alpha$ where every element satisfies $p$, the count of elements in the attached list (with proofs of $p$) that satisfy $q$ is equal to the count of elements in the original list $l$ that satisfy $q$. More formally: \[ \text{countP } q\ (\text{attachWith } p\ l\ H) = \text{countP } q\ l \] where $\text{attachWith } p\ l\ H$ pairs each element of $l$ with a proof that it satisfies $p$.
List.count_attach theorem
[DecidableEq α] {l : List α} {a : { x // x ∈ l }} : l.attach.count a = l.count ↑a
Full source
@[simp]
theorem count_attach [DecidableEq α] {l : List α} {a : {x // x ∈ l}} :
    l.attach.count a = l.count ↑a :=
  Eq.trans (countP_congr fun _ _ => by simp [Subtype.ext_iff]) <| countP_attach
Count Preservation under List Attachment: $\text{count}(a, l.\text{attach}) = \text{count}(x, l)$
Informal description
For any list $l$ of elements of type $\alpha$ with decidable equality, and any element $a$ in the attached list (i.e., $a$ is a pair consisting of an element $x \in l$ and a proof that $x \in l$), the count of $a$ in $l.\text{attach}$ is equal to the count of the underlying element $x$ in the original list $l$. In mathematical notation: \[ \text{count}(a, l.\text{attach}) = \text{count}(x, l) \] where $x = a.1$ is the underlying element of $a$.
List.count_attachWith theorem
[DecidableEq α] {p : α → Prop} {l : List α} (H : ∀ a ∈ l, p a) {a : { x // p x }} : (l.attachWith p H).count a = l.count ↑a
Full source
@[simp]
theorem count_attachWith [DecidableEq α] {p : α → Prop} {l : List α} (H : ∀ a ∈ l, p a) {a : {x // p x}} :
    (l.attachWith p H).count a = l.count ↑a :=
  Eq.trans (countP_congr fun _ _ => by simp [Subtype.ext_iff]) <| countP_attachWith _
Count Preservation under List Attachment with Predicate: $\text{count}(a, \text{attachWith}\ p\ l\ H) = \text{count}(a.1, l)$
Informal description
For any type $\alpha$ with decidable equality, any predicate $p$ on $\alpha$, and any list $l$ of elements of type $\alpha$ where every element satisfies $p$, the count of an element $a$ in the list obtained by attaching proofs (via `attachWith`) equals the count of the underlying value of $a$ in the original list $l$. More formally, for $a \in \{x \mid p x\}$: \[ \text{count}(a, \text{attachWith}\ p\ l\ H) = \text{count}(a.1, l) \] where $a.1$ denotes the underlying value of $a$.
List.countP_pmap theorem
{p : α → Prop} {g : ∀ a, p a → β} {f : β → Bool} {l : List α} (H₁) : (l.pmap g H₁).countP f = l.attach.countP (fun ⟨a, m⟩ => f (g a (H₁ a m)))
Full source
@[simp] theorem countP_pmap {p : α → Prop} {g : ∀ a, p a → β} {f : β → Bool} {l : List α} (H₁) :
    (l.pmap g H₁).countP f =
      l.attach.countP (fun ⟨a, m⟩ => f (g a (H₁ a m))) := by
  simp [pmap_eq_map_attach, countP_map, Function.comp_def]
Count Preservation under Partial Mapping and Attachment: $\text{countP}_f(\text{pmap}\ g\ l\ H_1) = \text{countP}_{f \circ (g \circ \cdot)}\ l.\text{attach}$
Informal description
For any predicate $p$ on elements of type $\alpha$, any function $g$ defined on elements $a$ of $\alpha$ that satisfy $p(a)$, any Boolean-valued function $f$ on $\beta$, and any list $l$ of elements of type $\alpha$ with a proof $H_1$ that all elements in $l$ satisfy $p$, the count of elements in the partial map $\text{pmap}\ g\ l\ H_1$ that satisfy $f$ is equal to the count of elements in the attached list $l.\text{attach}$ that satisfy the composed predicate $\lambda \langle a, m \rangle, f(g\ a\ (H_1\ a\ m))$. In other words: \[ \text{countP}\ f\ (\text{pmap}\ g\ l\ H_1) = \text{countP}\ (\lambda \langle a, m \rangle, f(g\ a\ (H_1\ a\ m)))\ l.\text{attach} \]
List.unattach definition
{α : Type _} {p : α → Prop} (l : List { x // p x }) : List α
Full source
/--
Maps a list of terms in a subtype to the corresponding terms in the type by forgetting that they
satisfy the predicate.

This is the inverse of `List.attachWith` and a synonym for `l.map (·.val)`.

Mostly this should not be needed by users. It is introduced as an intermediate step by lemmas such
as `map_subtype`, and is ideally subsequently simplified away by `unattach_attach`.

This function is usually inserted automatically by Lean as an intermediate step while proving
termination. It is rarely used explicitly in code. It is introduced as an intermediate step during
the elaboration of definitions by [well-founded
recursion](lean-manual://section/well-founded-recursion). If this function is encountered in a proof
state, the right approach is usually the tactic `simp [List.unattach, -List.map_subtype]`.
-/
def unattach {α : Type _} {p : α → Prop} (l : List { x // p x }) : List α := l.map (·.val)
List of values from a subtype list
Informal description
Given a list $l$ of elements of type $\{x : \alpha \mid p x\}$ (i.e., elements of $\alpha$ paired with proofs that they satisfy the predicate $p$), the function `List.unattach` returns the list obtained by extracting the underlying values from each element, forgetting their proofs. This is equivalent to applying the projection function $\cdot.\text{val}$ to each element of $l$.
List.unattach_nil theorem
{p : α → Prop} : ([] : List { x // p x }).unattach = []
Full source
@[simp] theorem unattach_nil {p : α → Prop} : ([] : List { x // p x }).unattach = [] := rfl
Unattach of Empty List is Empty
Informal description
For any predicate $p$ on a type $\alpha$, the unattach operation applied to the empty list of subtype elements $\{x \mid p x\}$ yields the empty list, i.e., $[].\text{unattach} = []$.
List.unattach_cons theorem
{p : α → Prop} {a : { x // p x }} {l : List { x // p x }} : (a :: l).unattach = a.val :: l.unattach
Full source
@[simp] theorem unattach_cons {p : α → Prop} {a : { x // p x }} {l : List { x // p x }} :
  (a :: l).unattach = a.val :: l.unattach := rfl
Cons Preservation under List Unattachment: $(a :: l).\text{unattach} = a.\text{val} :: l.\text{unattach}$
Informal description
For any predicate $p : \alpha \to \text{Prop}$, any element $a$ of the subtype $\{x : \alpha \mid p x\}$, and any list $l$ of such elements, the unattached version of the cons list $(a :: l)$ is equal to $a.\text{val} :: l.\text{unattach}$.
List.mem_unattach theorem
{p : α → Prop} {l : List { x // p x }} {a} : a ∈ l.unattach ↔ ∃ h : p a, ⟨a, h⟩ ∈ l
Full source
@[simp] theorem mem_unattach {p : α → Prop} {l : List { x // p x }} {a} :
    a ∈ l.unattacha ∈ l.unattach ↔ ∃ h : p a, ⟨a, h⟩ ∈ l := by
  simp only [unattach, mem_map, Subtype.exists, exists_and_right, exists_eq_right]
Membership in Unattached List of Subtype Elements
Informal description
For any list $l$ of elements of type $\{x : \alpha \mid p x\}$ (subtype of $\alpha$ satisfying predicate $p$) and any element $a : \alpha$, the element $a$ appears in the list obtained by extracting the values from $l$ (i.e., $a \in l.\text{unattach}$) if and only if there exists a proof $h$ that $p(a)$ holds and the element $\langle a, h \rangle$ is in the original list $l$.
List.length_unattach theorem
{p : α → Prop} {l : List { x // p x }} : l.unattach.length = l.length
Full source
@[simp] theorem length_unattach {p : α → Prop} {l : List { x // p x }} :
    l.unattach.length = l.length := by
  unfold unattach
  simp
Length Preservation under List Unattachment: $\text{length}(l.\text{unattach}) = \text{length}(l)$
Informal description
For any list $l$ of elements of type $\{x : \alpha \mid p x\}$ (where $p$ is a predicate on $\alpha$), the length of the list obtained by extracting the underlying values (via `unattach`) is equal to the length of the original list $l$. That is, $\text{length}(l.\text{unattach}) = \text{length}(l)$.
List.unattach_attach theorem
{l : List α} : l.attach.unattach = l
Full source
@[simp] theorem unattach_attach {l : List α} : l.attach.unattach = l := by
  unfold unattach
  induction l with
  | nil => simp
  | cons a l ih => simp [ih, Function.comp_def]
Unattachment of Attached List Preserves Original List
Informal description
For any list $l$ of elements of type $\alpha$, the unattachment of the attached list $l.\text{attach}$ is equal to the original list $l$. That is, $l.\text{attach}.\text{unattach} = l$.
List.unattach_attachWith theorem
{p : α → Prop} {l : List α} {H : ∀ a ∈ l, p a} : (l.attachWith p H).unattach = l
Full source
@[simp] theorem unattach_attachWith {p : α → Prop} {l : List α}
    {H : ∀ a ∈ l, p a} :
    (l.attachWith p H).unattach = l := by
  unfold unattach
  induction l with
  | nil => simp
  | cons a l ih => simp [ih, Function.comp_def]
Unattachment of Attached List with Predicate Proofs Preserves Original List
Informal description
Given a list $l$ of elements of type $\alpha$ and a predicate $p$ on $\alpha$ such that every element $a$ in $l$ satisfies $p(a)$, the unattachment of the list obtained by attaching proofs to each element of $l$ via `attachWith` is equal to the original list $l$. In other words, $(l.\text{attachWith}\ p\ H).\text{unattach} = l$.
List.getElem?_unattach theorem
{p : α → Prop} {l : List { x // p x }} (i : Nat) : l.unattach[i]? = l[i]?.map Subtype.val
Full source
@[simp] theorem getElem?_unattach {p : α → Prop} {l : List { x // p x }} (i : Nat) :
    l.unattach[i]? = l[i]?.map Subtype.val := by
  simp [unattach]
Optional Indexing Commutes with Unattachment: $(l.\text{unattach})[i]? = \text{Option.map}\ \text{val}\ (l[i]?)$
Informal description
For any list $l$ of elements of type $\{x : \alpha \mid p x\}$ (where $p$ is a predicate on $\alpha$) and any natural number index $i$, the optional indexing operation on the unattached list $l.\text{unattach}$ at position $i$ is equal to the optional indexing operation on $l$ at position $i$ followed by mapping the underlying value projection $\text{Subtype.val}$. In other words, $(l.\text{unattach})[i]? = \text{Option.map}\ \text{Subtype.val}\ (l[i]?)$.
List.getElem_unattach theorem
{p : α → Prop} {l : List { x // p x }} (i : Nat) (h : i < l.unattach.length) : l.unattach[i] = (l[i]'(by simpa using h)).1
Full source
@[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]
Indexing Commutes with Unattachment: $l.\text{unattach}[i] = l[i].\text{val}$
Informal description
For any list $l$ of elements of type $\{x : \alpha \mid p x\}$ (where $p$ is a predicate on $\alpha$) and any natural number index $i$ such that $i < \text{length}(l.\text{unattach})$, the element at position $i$ in the unattached list $l.\text{unattach}[i]$ is equal to the underlying value of the element at position $i$ in the original list $l[i]$.
List.foldl_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : β → { x // p x } → β} {g : β → α → β} {x : β} (hf : ∀ b x h, f b ⟨x, h⟩ = g b x) : l.foldl f x = l.unattach.foldl g x
Full source
/--
This lemma identifies folds over lists of subtypes, where the function only depends on the value, not the proposition,
and simplifies these to the function directly taking the value.
-/
@[simp] theorem foldl_subtype {p : α → Prop} {l : List { x // p x }}
    {f : β → { x // p x } → β} {g : β → α → β} {x : β}
    (hf : ∀ b x h, f b ⟨x, h⟩ = g b x) :
    l.foldl f x = l.unattach.foldl g x := by
  unfold unattach
  induction l generalizing x with
  | nil => simp
  | cons a l ih => simp [ih, hf]
Left Fold Equality for Subtype Lists via Unattach
Informal description
Let $p : \alpha \to \text{Prop}$ be a predicate on a type $\alpha$, and let $l$ be a list of elements of the subtype $\{x \mid p x\}$. Given functions $f : \beta \to \{x \mid p x\} \to \beta$ and $g : \beta \to \alpha \to \beta$ such that for all $b \in \beta$, $x \in \alpha$, and $h : p x$, we have $f(b, \langle x, h \rangle) = g(b, x)$, then the left fold of $f$ over $l$ with initial value $x \in \beta$ is equal to the left fold of $g$ over $l.\text{unattach}$ with the same initial value $x$.
List.foldr_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → β → β} {g : α → β → β} {x : β} (hf : ∀ x h b, f ⟨x, h⟩ b = g x b) : l.foldr f x = l.unattach.foldr g x
Full source
/--
This lemma identifies folds over lists of subtypes, where the function only depends on the value, not the proposition,
and simplifies these to the function directly taking the value.
-/
@[simp] theorem foldr_subtype {p : α → Prop} {l : List { x // p x }}
    {f : { x // p x } → β → β} {g : α → β → β} {x : β}
    (hf : ∀ x h b, f ⟨x, h⟩ b = g x b) :
    l.foldr f x = l.unattach.foldr g x := by
  unfold unattach
  induction l generalizing x with
  | nil => simp
  | cons a l ih => simp [ih, hf]
Right Fold Equivalence for Subtype Lists
Informal description
For any predicate $p : \alpha \to \text{Prop}$ and any list $l$ of elements of type $\{x \mid p x\}$, if $f : \{x \mid p x\} \to \beta \to \beta$ and $g : \alpha \to \beta \to \beta$ are functions such that for all $x$ with proof $h$ of $p(x)$ and all $b \in \beta$, we have $f(\langle x, h \rangle, b) = g(x, b)$, then the right fold of $f$ over $l$ with initial value $x$ is equal to the right fold of $g$ over $l.\text{unattach}$ with the same initial value $x$.
List.map_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → β} {g : α → β} (hf : ∀ x h, f ⟨x, h⟩ = g x) : l.map f = l.unattach.map g
Full source
/--
This lemma identifies maps over lists of subtypes, where the function only depends on the value, not the proposition,
and simplifies these to the function directly taking the value.
-/
@[simp] theorem map_subtype {p : α → Prop} {l : List { x // p x }}
    {f : { x // p x } → β} {g : α → β} (hf : ∀ x h, f ⟨x, h⟩ = g x) :
    l.map f = l.unattach.map g := by
  unfold unattach
  induction l with
  | nil => simp
  | cons a l ih => simp [ih, hf]
Map over Subtype List Equals Map over Unattached List
Informal description
For any predicate $p : \alpha \to \text{Prop}$, any list $l$ of elements of type $\{x : \alpha \mid p x\}$, and any functions $f : \{x : \alpha \mid p x\} \to \beta$ and $g : \alpha \to \beta$ such that $f(\langle x, h\rangle) = g(x)$ for all $x \in \alpha$ and $h : p x$, the map of $f$ over $l$ is equal to the map of $g$ over the list obtained by extracting the values from $l$ (i.e., $l.\text{unattach}$). That is, $\text{map}\ f\ l = \text{map}\ g\ (l.\text{unattach})$.
List.filterMap_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → Option β} {g : α → Option β} (hf : ∀ x h, f ⟨x, h⟩ = g x) : l.filterMap f = l.unattach.filterMap g
Full source
@[simp] theorem filterMap_subtype {p : α → Prop} {l : List { x // p x }}
    {f : { x // p x }Option β} {g : α → Option β} (hf : ∀ x h, f ⟨x, h⟩ = g x) :
    l.filterMap f = l.unattach.filterMap g := by
  unfold unattach
  induction l with
  | nil => simp
  | cons a l ih => simp [ih, hf, filterMap_cons]
FilterMap Equivalence for Subtype Lists: $\text{filterMap}\ f\ l = \text{filterMap}\ g\ l.\text{unattach}$
Informal description
For any predicate $p : \alpha \to \text{Prop}$ and any list $l$ of elements of type $\{x : \alpha \mid p x\}$, if $f : \{x : \alpha \mid p x\} \to \text{Option } \beta$ and $g : \alpha \to \text{Option } \beta$ are functions such that $f(\langle x, h\rangle) = g(x)$ for all $x \in \alpha$ and $h : p x$, then the filterMap of $f$ over $l$ is equal to the filterMap of $g$ over the list obtained by extracting the values from $l$ (i.e., $l.\text{unattach}$). That is, $\text{filterMap}\ f\ l = \text{filterMap}\ g\ (l.\text{unattach})$.
List.flatMap_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → List β} {g : α → List β} (hf : ∀ x h, f ⟨x, h⟩ = g x) : (l.flatMap f) = l.unattach.flatMap g
Full source
@[simp] theorem flatMap_subtype {p : α → Prop} {l : List { x // p x }}
    {f : { x // p x }List β} {g : α → List β} (hf : ∀ x h, f ⟨x, h⟩ = g x) :
    (l.flatMap f) = l.unattach.flatMap g := by
  unfold unattach
  induction l with
  | nil => simp
  | cons a l ih => simp [ih, hf]
FlatMap Equivalence for Subtype Lists: $\text{flatMap}\ f\ l = \text{flatMap}\ g\ l.\text{unattach}$
Informal description
For any predicate $p : \alpha \to \text{Prop}$, any list $l$ of elements of type $\{x : \alpha \mid p x\}$, and any functions $f : \{x : \alpha \mid p x\} \to \text{List}\ \beta$ and $g : \alpha \to \text{List}\ \beta$ such that $f(\langle x, h\rangle) = g(x)$ for all $x \in \alpha$ and $h : p x$, the flatMap of $f$ over $l$ is equal to the flatMap of $g$ over the list obtained by extracting the values from $l$ (i.e., $l.\text{unattach}$). That is, \[ \text{flatMap}\ f\ l = \text{flatMap}\ g\ (l.\text{unattach}). \]
List.bind_subtype abbrev
Full source
@[deprecated flatMap_subtype (since := "2024-10-16")] abbrev bind_subtype := @flatMap_subtype
Bind Operation Equivalence for Subtype Lists: $\text{bind}\ f\ l = \text{bind}\ g\ l.\text{unattach}$
Informal description
For any predicate $p : \alpha \to \text{Prop}$, any list $l$ of elements of type $\{x : \alpha \mid p x\}$, and any functions $f : \{x : \alpha \mid p x\} \to \text{List}\ \beta$ and $g : \alpha \to \text{List}\ \beta$ such that $f(\langle x, h\rangle) = g(x)$ for all $x \in \alpha$ and $h : p x$, the bind operation of $f$ over $l$ is equal to the bind operation of $g$ over the list obtained by extracting the values from $l$ (i.e., $l.\text{unattach}$). That is, \[ \text{bind}\ f\ l = \text{bind}\ g\ (l.\text{unattach}). \]
List.findSome?_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → Option β} {g : α → Option β} (hf : ∀ x h, f ⟨x, h⟩ = g x) : l.findSome? f = l.unattach.findSome? g
Full source
@[simp] theorem findSome?_subtype {p : α → Prop} {l : List { x // p x }}
    {f : { x // p x }Option β} {g : α → Option β} (hf : ∀ x h, f ⟨x, h⟩ = g x) :
    l.findSome? f = l.unattach.findSome? g := by
  unfold unattach
  induction l with
  | nil => simp
  | cons a l ih => simp [ih, hf, findSome?_cons]
First Non-None Result Preservation under List Unattachment
Informal description
For any predicate $p : \alpha \to \text{Prop}$, any list $l$ of elements of type $\{x : \alpha \mid p x\}$, and any functions $f : \{x : \alpha \mid p x\} \to \text{Option}\ \beta$ and $g : \alpha \to \text{Option}\ \beta$ such that $f(\langle x, h\rangle) = g(x)$ for all $x \in \alpha$ and $h : p x$, the first non-`none` result of applying $f$ to elements of $l$ is equal to the first non-`none` result of applying $g$ to the list obtained by extracting the values from $l$ (i.e., $l.\text{unattach}$). That is, $\text{findSome?}\ f\ l = \text{findSome?}\ g\ (l.\text{unattach})$.
List.find?_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → Bool} {g : α → Bool} (hf : ∀ x h, f ⟨x, h⟩ = g x) : (l.find? f).map Subtype.val = l.unattach.find? g
Full source
@[simp] theorem find?_subtype {p : α → Prop} {l : List { x // p x }}
    {f : { x // p x }Bool} {g : α → Bool} (hf : ∀ x h, f ⟨x, h⟩ = g x) :
    (l.find? f).map Subtype.val = l.unattach.find? g := by
  unfold unattach
  induction l with
  | nil => simp
  | cons a l ih =>
    simp [hf, find?_cons]
    split <;> simp [ih]
Equality of First Satisfying Element in Subtype List and Unattached List
Informal description
For any predicate $p : \alpha \to \text{Prop}$, any list $l$ of elements of type $\{x : \alpha \mid p x\}$, and any functions $f : \{x : \alpha \mid p x\} \to \text{Bool}$ and $g : \alpha \to \text{Bool}$ such that $f(\langle x, h\rangle) = g(x)$ for all $x \in \alpha$ and $h : p x$, the first element of $l$ satisfying $f$ (when mapped to its underlying value) is equal to the first element of $l.\text{unattach}$ satisfying $g$. That is, $\text{map}\ \text{Subtype.val}\ (l.\text{find?}\ f) = l.\text{unattach}.\text{find?}\ g$.
List.all_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → Bool} {g : α → Bool} (hf : ∀ x h, f ⟨x, h⟩ = g x) : l.all f = l.unattach.all g
Full source
@[simp] theorem all_subtype {p : α → Prop} {l : List { x // p x }} {f : { x // p x }Bool} {g : α → Bool}
    (hf : ∀ x h, f ⟨x, h⟩ = g x) :
    l.all f = l.unattach.all g := by
  simp [all_eq, hf]
Equivalence of Universal Quantification on Subtype List and Unattached List
Informal description
For any predicate $p : \alpha \to \text{Prop}$, any list $l$ of elements of type $\{x : \alpha \mid p x\}$, and any functions $f : \{x : \alpha \mid p x\} \to \text{Bool}$ and $g : \alpha \to \text{Bool}$ such that $f(\langle x, h\rangle) = g(x)$ for all $x \in \alpha$ and $h : p x$, the boolean evaluation of `l.all f` is equal to the boolean evaluation of `l.unattach.all g`. That is, the universal quantification over the list $l$ with predicate $f$ is equivalent to the universal quantification over the unattached list with predicate $g$.
List.any_subtype theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → Bool} {g : α → Bool} (hf : ∀ x h, f ⟨x, h⟩ = g x) : l.any f = l.unattach.any g
Full source
@[simp] theorem any_subtype {p : α → Prop} {l : List { x // p x }} {f : { x // p x }Bool} {g : α → Bool}
    (hf : ∀ x h, f ⟨x, h⟩ = g x) :
    l.any f = l.unattach.any g := by
  simp [any_eq, hf]
Equivalence of `any` on Subtype List and Unattached List: $\text{any}(l, f) = \text{any}(\text{unattach}(l), g)$
Informal description
For any predicate $p : \alpha \to \text{Prop}$, any list $l$ of elements of type $\{x : \alpha \mid p x\}$, and any functions $f : \{x : \alpha \mid p x\} \to \text{Bool}$ and $g : \alpha \to \text{Bool}$ such that $f(\langle x, h\rangle) = g(x)$ for all $x \in \alpha$ and $h : p x$, the boolean evaluation of `l.any f` is equal to the boolean evaluation of `l.unattach.any g`. That is, the existence of an element in $l$ satisfying $f$ is equivalent to the existence of an element in the unattached list satisfying $g$.
List.unattach_filter theorem
{p : α → Prop} {l : List { x // p x }} {f : { x // p x } → Bool} {g : α → Bool} {hf : ∀ x h, f ⟨x, h⟩ = g x} : (l.filter f).unattach = l.unattach.filter g
Full source
@[simp] theorem unattach_filter {p : α → Prop} {l : List { x // p x }}
    {f : { x // p x }Bool} {g : α → Bool} {hf : ∀ x h, f ⟨x, h⟩ = g x} :
    (l.filter f).unattach = l.unattach.filter g := by
  induction l with
  | nil => simp
  | cons a l ih =>
    simp only [filter_cons, hf, unattach_cons]
    split <;> simp [ih]
Filtering Commutes with Unattach: $\text{unattach}(\text{filter}\, f\, l) = \text{filter}\, g\, (\text{unattach}\, l)$
Informal description
For any predicate $p : \alpha \to \text{Prop}$, list $l$ of elements of type $\{x : \alpha \mid p x\}$, and boolean-valued functions $f$ (on the subtype) and $g$ (on $\alpha$) such that $f(\langle x, h \rangle) = g(x)$ for all $x$ and proof $h$ of $p(x)$, the following equality holds: $$ \text{unattach}(l.\text{filter}\, f) = (\text{unattach}\, l).\text{filter}\, g. $$
List.unattach_reverse theorem
{p : α → Prop} {l : List { x // p x }} : l.reverse.unattach = l.unattach.reverse
Full source
@[simp] theorem unattach_reverse {p : α → Prop} {l : List { x // p x }} :
    l.reverse.unattach = l.unattach.reverse := by
  simp [unattach, -map_subtype]
Unattach Commutes with List Reversal: $\text{unattach}(l^{\text{rev}}) = (\text{unattach}(l))^{\text{rev}}$
Informal description
For any list $l$ of elements of type $\{x : \alpha \mid p x\}$ (i.e., elements of $\alpha$ paired with proofs that they satisfy the predicate $p$), the unattached version of the reversed list is equal to the reverse of the unattached list. That is, $$ \text{unattach}(l^{\text{rev}}) = (\text{unattach}(l))^{\text{rev}}. $$
List.unattach_append theorem
{p : α → Prop} {l₁ l₂ : List { x // p x }} : (l₁ ++ l₂).unattach = l₁.unattach ++ l₂.unattach
Full source
@[simp] theorem unattach_append {p : α → Prop} {l₁ l₂ : List { x // p x }} :
    (l₁ ++ l₂).unattach = l₁.unattach ++ l₂.unattach := by
  simp [unattach, -map_subtype]
Unattach Preserves List Concatenation: $\text{unattach}(l_1 ++ l_2) = \text{unattach}(l_1) ++ \text{unattach}(l_2)$
Informal description
For any predicate $p : \alpha \to \text{Prop}$ and any two lists $l_1, l_2$ of elements of type $\{x : \alpha \mid p x\}$, the unattach operation commutes with list concatenation. That is, $$ \text{unattach}(l_1 ++ l_2) = \text{unattach}(l_1) ++ \text{unattach}(l_2). $$
List.unattach_flatten theorem
{p : α → Prop} {l : List (List { x // p x })} : l.flatten.unattach = (l.map unattach).flatten
Full source
@[simp] theorem unattach_flatten {p : α → Prop} {l : List (List { x // p x })} :
    l.flatten.unattach = (l.map unattach).flatten := by
  unfold unattach
  induction l <;> simp_all
Unattach-Flatten Commutativity: $\text{unattach} \circ \text{flatten} = \text{flatten} \circ \text{map}(\text{unattach})$
Informal description
For any predicate $p : \alpha \to \text{Prop}$ and any list of lists $l$ where each sublist contains elements of the subtype $\{x : \alpha \mid p x\}$, the unattach operation commutes with flattening. That is, flattening $l$ and then applying unattach is equivalent to first applying unattach to each sublist in $l$ and then flattening the result. Formally: $$ \text{unattach}(\text{flatten}(l)) = \text{flatten}(\text{map}(\text{unattach})(l)). $$
List.unattach_join abbrev
Full source
@[deprecated unattach_flatten (since := "2024-10-14")] abbrev unattach_join := @unattach_flatten
Unattach-Join Commutativity: $\text{unattach} \circ \text{join} = \text{join} \circ \text{map}(\text{unattach})$
Informal description
For any predicate $p : \alpha \to \text{Prop}$ and any list of lists $l$ where each sublist contains elements of the subtype $\{x : \alpha \mid p x\}$, the unattach operation commutes with list joining. That is, joining $l$ and then applying unattach is equivalent to first applying unattach to each sublist in $l$ and then joining the result. Formally: $$ \text{unattach}(\text{join}(l)) = \text{join}(\text{map}(\text{unattach})(l)). $$
List.unattach_replicate theorem
{p : α → Prop} {n : Nat} {x : { x // p x }} : (List.replicate n x).unattach = List.replicate n x.1
Full source
@[simp] theorem unattach_replicate {p : α → Prop} {n : Nat} {x : { x // p x }} :
    (List.replicate n x).unattach = List.replicate n x.1 := by
  simp [unattach, -map_subtype]
Unattach of Replicated Subtype List: $(\operatorname{replicate}(n, x)).\text{unattach} = \operatorname{replicate}(n, x.1)$
Informal description
For any predicate $p : \alpha \to \text{Prop}$, natural number $n$, and element $x$ of the subtype $\{x : \alpha \mid p x\}$, the unattached version of a list containing $n$ copies of $x$ is equal to a list containing $n$ copies of the underlying value $x.1$. That is, $$(\operatorname{replicate}(n, x)).\text{unattach} = \operatorname{replicate}(n, x.1)$$
List.map_wfParam theorem
{xs : List α} {f : α → β} : (wfParam xs).map f = xs.attach.unattach.map f
Full source
@[wf_preprocess] theorem map_wfParam {xs : List α} {f : α → β} :
    (wfParam xs).map f = xs.attach.unattach.map f := by
  simp [wfParam]
Well-Founded Parameter Mapping Equals Unattached Mapping: `(wfParam xs).map f = xs.attach.unattach.map f`
Informal description
For any list `xs` of elements of type `α` and any function `f : α → β`, mapping `f` over the well-founded recursion parameter `wfParam xs` is equal to mapping `f` over the unattached list obtained from attaching `xs`. That is, `(wfParam xs).map f = xs.attach.unattach.map f`.
List.map_unattach theorem
{P : α → Prop} {xs : List (Subtype P)} {f : α → β} : xs.unattach.map f = xs.map fun ⟨x, h⟩ => binderNameHint x f <| binderNameHint h () <| f (wfParam x)
Full source
@[wf_preprocess] theorem map_unattach {P : α → Prop} {xs : List (Subtype P)} {f : α → β} :
    xs.unattach.map f = xs.map fun ⟨x, h⟩ =>
      binderNameHint x f <| binderNameHint h () <| f (wfParam x) := by
  simp [wfParam]
Map over Unattached List Equals Map over Subtype List with Binder Preservation
Informal description
For any predicate $P : \alpha \to \text{Prop}$, any list $xs$ of elements of type $\{x : \alpha \mid P x\}$, and any function $f : \alpha \to \beta$, the map of $f$ over the list obtained by extracting the values from $xs$ (i.e., $xs.\text{unattach}$) is equal to the map over $xs$ of the function that applies $f$ to each value while preserving the binder names during simplification.
List.foldl_wfParam theorem
{xs : List α} {f : β → α → β} {x : β} : (wfParam xs).foldl f x = xs.attach.unattach.foldl f x
Full source
@[wf_preprocess] theorem foldl_wfParam {xs : List α} {f : β → α → β} {x : β} :
    (wfParam xs).foldl f x = xs.attach.unattach.foldl f x := by
  simp [wfParam]
Equality of Left Folds via Well-Founded Parameter and Unattachment
Informal description
For any list $xs$ of elements of type $\alpha$, any function $f : \beta \to \alpha \to \beta$, and any initial value $x : \beta$, the left fold of $f$ over the well-founded recursion parameter $xs$ with initial value $x$ is equal to the left fold of $f$ over the unattached version of the attached list $xs.\text{attach}$ with the same initial value $x$. That is, $\text{foldl } f \ x \ (\text{wfParam } xs) = \text{foldl } f \ x \ (xs.\text{attach}.\text{unattach})$.
List.foldl_unattach theorem
{P : α → Prop} {xs : List (Subtype P)} {f : β → α → β} {x : β} : xs.unattach.foldl f x = xs.foldl (fun s ⟨x, h⟩ => binderNameHint s f <| binderNameHint x (f s) <| binderNameHint h () <| f s (wfParam x)) x
Full source
@[wf_preprocess] theorem foldl_unattach {P : α → Prop} {xs : List (Subtype P)} {f : β → α → β} {x : β} :
    xs.unattach.foldl f x = xs.foldl (fun s ⟨x, h⟩ =>
      binderNameHint s f <| binderNameHint x (f s) <| binderNameHint h () <| f s (wfParam x)) x := by
  simp [wfParam]
Left Fold Equality via Unattach for Subtype Lists
Informal description
Let $P : \alpha \to \text{Prop}$ be a predicate on a type $\alpha$, and let $xs$ be a list of elements of the subtype $\{x \mid P x\}$. For any function $f : \beta \to \alpha \to \beta$ and initial value $x : \beta$, the left fold of $f$ over $xs.\text{unattach}$ with initial value $x$ is equal to the left fold of the function $\lambda s \ \langle y, h \rangle, f s y$ over $xs$ with the same initial value $x$.
List.foldr_wfParam theorem
{xs : List α} {f : α → β → β} {x : β} : (wfParam xs).foldr f x = xs.attach.unattach.foldr f x
Full source
@[wf_preprocess] theorem foldr_wfParam {xs : List α} {f : α → β → β} {x : β} :
    (wfParam xs).foldr f x = xs.attach.unattach.foldr f x := by
  simp [wfParam]
Right Fold Equality via Well-Founded Parameter and Unattachment
Informal description
For any list `xs` of elements of type `α`, any binary function `f : α → β → β`, and any initial value `x : β`, the right fold of `f` over the well-founded recursion parameter `wfParam xs` with initial value `x` is equal to the right fold of `f` over the unattached version of the attached list `xs.attach` with the same initial value `x`. That is, `foldr f x (wfParam xs) = foldr f x (xs.attach.unattach)`.
List.foldr_unattach theorem
{P : α → Prop} {xs : List (Subtype P)} {f : α → β → β} {x : β} : xs.unattach.foldr f x = xs.foldr (fun ⟨x, h⟩ s => binderNameHint x f <| binderNameHint s (f x) <| binderNameHint h () <| f (wfParam x) s) x
Full source
@[wf_preprocess] theorem foldr_unattach {P : α → Prop} {xs : List (Subtype P)} {f : α → β → β} {x : β} :
    xs.unattach.foldr f x = xs.foldr (fun ⟨x, h⟩ s =>
      binderNameHint x f <| binderNameHint s (f x) <| binderNameHint h () <| f (wfParam x) s) x := by
  simp [wfParam]
Right Fold Equivalence for Unattached Subtype Lists
Informal description
For any predicate $P : \alpha \to \text{Prop}$ and any list $xs$ of elements of type $\{x : \alpha \mid P x\}$, given a binary function $f : \alpha \to \beta \to \beta$ and an initial value $x : \beta$, the right fold of $f$ over $xs.\text{unattach}$ with initial value $x$ is equal to the right fold of the function $\lambda \langle x, h \rangle \, s \mapsto f\,x\,s$ over $xs$ with the same initial value $x$.
List.filter_wfParam theorem
{xs : List α} {f : α → Bool} : (wfParam xs).filter f = xs.attach.unattach.filter f
Full source
@[wf_preprocess] theorem filter_wfParam {xs : List α} {f : α → Bool} :
    (wfParam xs).filter f = xs.attach.unattach.filter f:= by
  simp [wfParam]
Filtering Preserved Under Well-Founded Parameter and Unattachment
Informal description
For any list `xs` of elements of type `α` and any predicate `f : α → Bool`, filtering the list `wfParam xs` with `f` yields the same result as filtering the unattached version of the attached list `xs.attach` with `f`. That is, `(wfParam xs).filter f = xs.attach.unattach.filter f`.
List.filter_unattach theorem
{P : α → Prop} {xs : List (Subtype P)} {f : α → Bool} : xs.unattach.filter f = (xs.filter (fun ⟨x, h⟩ => binderNameHint x f <| binderNameHint h () <| f (wfParam x))).unattach
Full source
@[wf_preprocess] theorem filter_unattach {P : α → Prop} {xs : List (Subtype P)} {f : α → Bool} :
    xs.unattach.filter f = (xs.filter (fun ⟨x, h⟩ =>
      binderNameHint x f <| binderNameHint h () <| f (wfParam x))).unattach := by
  simp [wfParam]
Filtering Commutes with Unattachment: $\text{filter}\, f \circ \text{unattach} = \text{unattach} \circ \text{filter}\, (\lambda \langle x, h \rangle \Rightarrow f\, x)$
Informal description
For any predicate $P : \alpha \to \text{Prop}$ and any list $xs$ of elements of type $\{x : \alpha \mid P x\}$, given a boolean-valued function $f : \alpha \to \text{Bool}$, the following equality holds: $$ \text{filter}\, f\, (\text{unattach}\, xs) = \text{unattach}\, (\text{filter}\, (\lambda \langle x, h \rangle \Rightarrow f\, x)\, xs). $$
List.reverse_wfParam theorem
{xs : List α} : (wfParam xs).reverse = xs.attach.unattach.reverse
Full source
@[wf_preprocess] theorem reverse_wfParam {xs : List α} :
    (wfParam xs).reverse = xs.attach.unattach.reverse := by simp [wfParam]
Reversing Well-Founded Parameter Equals Reversing Unattached Attached List
Informal description
For any list `xs` of elements of type `α`, reversing the well-founded parameter `wfParam xs` yields the same result as reversing the unattached version of the attached list `xs.attach.unattach`. That is, $$ \text{reverse}(\text{wfParam}(xs)) = \text{reverse}(\text{unattach}(\text{attach}(xs))). $$
List.reverse_unattach theorem
{P : α → Prop} {xs : List (Subtype P)} : xs.unattach.reverse = xs.reverse.unattach
Full source
@[wf_preprocess] theorem reverse_unattach {P : α → Prop} {xs : List (Subtype P)} :
    xs.unattach.reverse = xs.reverse.unattach := by simp
Commutation of Unattach and List Reversal: $\text{unattach}(xs)^{\text{rev}} = \text{unattach}(xs^{\text{rev}})$
Informal description
For any list $xs$ of elements of type $\{x : \alpha \mid P x\}$ (i.e., elements of $\alpha$ paired with proofs that they satisfy the predicate $P$), the reverse of the unattached list is equal to the unattached version of the reversed list. That is, $$ \text{unattach}(xs)^{\text{rev}} = \text{unattach}(xs^{\text{rev}}). $$
List.filterMap_wfParam theorem
{xs : List α} {f : α → Option β} : (wfParam xs).filterMap f = xs.attach.unattach.filterMap f
Full source
@[wf_preprocess] theorem filterMap_wfParam {xs : List α} {f : α → Option β} :
    (wfParam xs).filterMap f = xs.attach.unattach.filterMap f := by
  simp [wfParam]
FilterMap Commutes with Well-founded Parameter and Unattachment: $\text{filterMap}\ f\ (\text{wfParam}\ xs) = \text{filterMap}\ f\ (xs.\text{attach}.\text{unattach})$
Informal description
For any list $xs$ of elements of type $\alpha$ and any function $f : \alpha \to \text{Option } \beta$, the filterMap of $f$ applied to the well-founded recursion parameter of $xs$ is equal to the filterMap of $f$ applied to the unattached version of the attached list $xs.\text{attach}$. That is, $$ \text{filterMap}\ f\ (\text{wfParam}\ xs) = \text{filterMap}\ f\ (xs.\text{attach}.\text{unattach}). $$
List.filterMap_unattach theorem
{P : α → Prop} {xs : List (Subtype P)} {f : α → Option β} : xs.unattach.filterMap f = xs.filterMap fun ⟨x, h⟩ => binderNameHint x f <| binderNameHint h () <| f (wfParam x)
Full source
@[wf_preprocess] theorem filterMap_unattach {P : α → Prop} {xs : List (Subtype P)} {f : α → Option β} :
    xs.unattach.filterMap f = xs.filterMap fun ⟨x, h⟩ =>
      binderNameHint x f <| binderNameHint h () <| f (wfParam x) := by
  simp [wfParam]
FilterMap Commutes with Unattach: $\text{filterMap}\ f\ (xs.\text{unattach}) = \text{filterMap}\ (\lambda \langle x, h\rangle .\ f(x))\ xs$
Informal description
For any predicate $P : \alpha \to \text{Prop}$, any list $xs$ of elements of type $\{x : \alpha \mid P x\}$, and any function $f : \alpha \to \text{Option } \beta$, the filterMap of $f$ over the unattached list $xs.\text{unattach}$ is equal to the filterMap over $xs$ of the function that applies $f$ to the underlying value of each element (ignoring the proof component). That is, $$ \text{filterMap}\ f\ (xs.\text{unattach}) = \text{filterMap}\ (\lambda \langle x, h\rangle .\ f(x))\ xs. $$
List.flatMap_wfParam theorem
{xs : List α} {f : α → List β} : (wfParam xs).flatMap f = xs.attach.unattach.flatMap f
Full source
@[wf_preprocess] theorem flatMap_wfParam {xs : List α} {f : α → List β} :
    (wfParam xs).flatMap f = xs.attach.unattach.flatMap f := by
  simp [wfParam]
FlatMap Identity via Unattachment: $\text{flatMap}\ f\ (wfParam\ xs) = \text{flatMap}\ f\ (xs.\text{attach}.\text{unattach})$
Informal description
For any list $xs$ of elements of type $\alpha$ and any function $f : \alpha \to \text{List}\ \beta$, the flatMap operation satisfies: $$ \text{flatMap}\ f\ (wfParam\ xs) = \text{flatMap}\ f\ (xs.\text{attach}.\text{unattach}) $$ where $wfParam$ is the identity function and $\text{attach}$ pairs each element with a proof of its membership in the list.
List.flatMap_unattach theorem
{P : α → Prop} {xs : List (Subtype P)} {f : α → List β} : xs.unattach.flatMap f = xs.flatMap fun ⟨x, h⟩ => binderNameHint x f <| binderNameHint h () <| f (wfParam x)
Full source
@[wf_preprocess] theorem flatMap_unattach {P : α → Prop} {xs : List (Subtype P)} {f : α → List β} :
    xs.unattach.flatMap f = xs.flatMap fun ⟨x, h⟩ =>
      binderNameHint x f <| binderNameHint h () <| f (wfParam x) := by
  simp [wfParam]
FlatMap Commutes with Unattachment: $\text{flatMap}\ f\ (xs.\text{unattach}) = \text{flatMap}\ (\lambda \langle x, h\rangle .\ f(x))\ xs$
Informal description
For any predicate $P : \alpha \to \text{Prop}$, any list $xs$ of elements of type $\{x : \alpha \mid P x\}$, and any function $f : \alpha \to \text{List}\ \beta$, the flatMap operation satisfies: \[ \text{flatMap}\ f\ (xs.\text{unattach}) = \text{flatMap}\ (\lambda \langle x, h\rangle .\ f(x))\ xs. \] Here, $xs.\text{unattach}$ denotes the list obtained by extracting the underlying values from each element of $xs$ (forgetting the proofs $h : P x$).