Documentation

Init.Data.Array.Attach

def Array.pmap {α : Type u_1} {β : Type u_2} {P : αProp} (f : (a : α) → P aβ) (l : Array α) (H : ∀ (a : α), a lP a) :

O(n). Partial map. If f : Π a, P a → β is a partial function defined on a : α satisfying P, then pmap f l h is essentially the same as map f l but is defined only when all members of l satisfy P, using the proof to apply f.

We replace this at runtime with a more efficient version via the csimp lemma pmap_eq_pmapImpl.

Equations
Instances For
    @[implemented_by _private.Init.Data.Array.Attach.0.Array.attachWithImpl]
    def Array.attachWith {α : Type u_1} (xs : Array α) (P : αProp) (H : ∀ (x : α), x xsP x) :
    Array { x : α // P x }

    O(1). "Attach" a proof P x that holds for all the elements of xs to produce a new array with the same elements but in the type {x // P x}.

    Equations
    • xs.attachWith P H = { toList := xs.toList.attachWith P }
    Instances For
      @[inline]
      def Array.attach {α : Type u_1} (xs : Array α) :
      Array { x : α // x xs }

      O(1). "Attach" the proof that the elements of xs are in xs to produce a new array with the same elements but in the type {x // x ∈ xs}.

      Equations
      Instances For
        @[simp]
        theorem List.attachWith_toArray {α : Type u_1} {l : List α} {P : αProp} {H : ∀ (x : α), x l.toArrayP x} :
        l.toArray.attachWith P H = (l.attachWith P ).toArray
        @[simp]
        theorem List.attach_toArray {α : Type u_1} {l : List α} :
        l.toArray.attach = (l.attachWith (fun (x : α) => x l.toArray) ).toArray
        @[simp]
        theorem List.pmap_toArray {α : Type u_1} {β : Type u_2} {l : List α} {P : αProp} {f : (a : α) → P aβ} {H : ∀ (a : α), a l.toArrayP a} :
        Array.pmap f l.toArray H = (List.pmap f l ).toArray
        @[simp]
        theorem Array.toList_attachWith {α : Type u_1} {l : Array α} {P : αProp} {H : ∀ (x : α), x lP x} :
        (l.attachWith P H).toList = l.toList.attachWith P
        @[simp]
        theorem Array.toList_attach {α : Type u_1} {l : Array α} :
        l.attach.toList = l.toList.attachWith (fun (x : α) => x l)
        @[simp]
        theorem Array.toList_pmap {α : Type u_1} {β : Type u_2} {l : Array α} {P : αProp} {f : (a : α) → P aβ} {H : ∀ (a : α), a lP a} :
        (Array.pmap f l H).toList = List.pmap f l.toList
        @[simp]
        theorem Array.pmap_empty {α : Type u_1} {β : Type u_2} {P : αProp} (f : (a : α) → P aβ) :
        Array.pmap f #[] = #[]
        @[simp]
        theorem Array.pmap_push {α : Type u_1} {β : Type u_2} {P : αProp} (f : (a : α) → P aβ) (a : α) (l : Array α) (h : ∀ (b : α), b l.push aP b) :
        Array.pmap f (l.push a) h = (Array.pmap f l ).push (f a )
        @[simp]
        theorem Array.attach_empty {α : Type u_1} :
        #[].attach = #[]
        @[simp]
        theorem Array.attachWith_empty {α : Type u_1} {P : αProp} (H : ∀ (x : α), x #[]P x) :
        #[].attachWith P H = #[]
        @[simp]
        theorem List.attachWith_mem_toArray {α : Type u_1} {l : List α} :
        l.attachWith (fun (x : α) => x l.toArray) = List.map (fun (x : { x : α // x l }) => match x with | x, h => x, ) l.attach
        @[simp]
        theorem Array.pmap_eq_map {α : Type u_1} {β : Type u_2} (p : αProp) (f : αβ) (l : Array α) (H : ∀ (a : α), a lp a) :
        Array.pmap (fun (a : α) (x : p a) => f a) l H = Array.map f l
        theorem Array.pmap_congr_left {α : Type u_1} {β : Type u_2} {p q : αProp} {f : (a : α) → p aβ} {g : (a : α) → q aβ} (l : Array α) {H₁ : ∀ (a : α), a lp a} {H₂ : ∀ (a : α), a lq a} (h : ∀ (a : α), a l∀ (h₁ : p a) (h₂ : q a), f a h₁ = g a h₂) :
        Array.pmap f l H₁ = Array.pmap g l H₂
        theorem Array.map_pmap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {p : αProp} (g : βγ) (f : (a : α) → p aβ) (l : Array α) (H : ∀ (a : α), a lp a) :
        Array.map g (Array.pmap f l H) = Array.pmap (fun (a : α) (h : p a) => g (f a h)) l H
        theorem Array.pmap_map {β : Type u_1} {γ : Type u_2} {α : Type u_3} {p : βProp} (g : (b : β) → p bγ) (f : αβ) (l : Array α) (H : ∀ (a : β), a Array.map f lp a) :
        Array.pmap g (Array.map f l) H = Array.pmap (fun (a : α) (h : p (f a)) => g (f a) h) l
        theorem Array.attach_congr {α : Type u_1} {l₁ l₂ : Array α} (h : l₁ = l₂) :
        l₁.attach = Array.map (fun (x : { x : α // x l₂ }) => x.val, ) l₂.attach
        theorem Array.attachWith_congr {α : Type u_1} {l₁ l₂ : Array α} (w : l₁ = l₂) {P : αProp} {H : ∀ (x : α), x l₁P x} :
        l₁.attachWith P H = l₂.attachWith P
        @[simp]
        theorem Array.attach_push {α : Type u_1} {a : α} {l : Array α} :
        (l.push a).attach = (Array.map (fun (x : { x : α // x l }) => match x with | x, h => x, ) l.attach).push a,
        @[simp]
        theorem Array.attachWith_push {α : Type u_1} {a : α} {l : Array α} {P : αProp} {H : ∀ (x : α), x l.push aP x} :
        (l.push a).attachWith P H = (l.attachWith P ).push a,
        theorem Array.pmap_eq_map_attach {α : Type u_1} {β : Type u_2} {p : αProp} (f : (a : α) → p aβ) (l : Array α) (H : ∀ (a : α), a lp a) :
        Array.pmap f l H = Array.map (fun (x : { x : α // x l }) => f x.val ) l.attach
        theorem Array.attach_map_coe {α : Type u_1} {β : Type u_2} (l : Array α) (f : αβ) :
        Array.map (fun (i : { i : α // i l }) => f i.val) l.attach = Array.map f l
        theorem Array.attach_map_val {α : Type u_1} {β : Type u_2} (l : Array α) (f : αβ) :
        Array.map (fun (i : { x : α // x l }) => f i.val) l.attach = Array.map f l
        @[simp]
        theorem Array.attach_map_subtype_val {α : Type u_1} (l : Array α) :
        Array.map Subtype.val l.attach = l
        theorem Array.attachWith_map_coe {α : Type u_1} {β : Type u_2} {p : αProp} (f : αβ) (l : Array α) (H : ∀ (a : α), a lp a) :
        Array.map (fun (i : { i : α // p i }) => f i.val) (l.attachWith p H) = Array.map f l
        theorem Array.attachWith_map_val {α : Type u_1} {β : Type u_2} {p : αProp} (f : αβ) (l : Array α) (H : ∀ (a : α), a lp a) :
        Array.map (fun (i : { x : α // p x }) => f i.val) (l.attachWith p H) = Array.map f l
        @[simp]
        theorem Array.attachWith_map_subtype_val {α : Type u_1} {p : αProp} (l : Array α) (H : ∀ (a : α), a lp a) :
        Array.map Subtype.val (l.attachWith p H) = l
        @[simp]
        theorem Array.mem_attach {α : Type u_1} (l : Array α) (x : { x : α // x l }) :
        x l.attach
        @[simp]
        theorem Array.mem_pmap {α : Type u_1} {β : Type u_2} {p : αProp} {f : (a : α) → p aβ} {l : Array α} {H : ∀ (a : α), a lp a} {b : β} :
        b Array.pmap f l H ∃ (a : α), ∃ (h : a l), f a = b
        theorem Array.mem_pmap_of_mem {α : Type u_1} {β : Type u_2} {p : αProp} {f : (a : α) → p aβ} {l : Array α} {H : ∀ (a : α), a lp a} {a : α} (h : a l) :
        f a Array.pmap f l H
        @[simp]
        theorem Array.size_pmap {α : Type u_1} {β : Type u_2} {p : αProp} {f : (a : α) → p aβ} {l : Array α} {H : ∀ (a : α), a lp a} :
        (Array.pmap f l H).size = l.size
        @[simp]
        theorem Array.size_attach {α : Type u_1} {L : Array α} :
        L.attach.size = L.size
        @[simp]
        theorem Array.size_attachWith {α : Type u_1} {p : αProp} {l : Array α} {H : ∀ (x : α), x lp x} :
        (l.attachWith p H).size = l.size
        @[simp]
        theorem Array.pmap_eq_empty_iff {α : Type u_1} {β : Type u_2} {p : αProp} {f : (a : α) → p aβ} {l : Array α} {H : ∀ (a : α), a lp a} :
        Array.pmap f l H = #[] l = #[]
        theorem Array.pmap_ne_empty_iff {α : Type u_1} {β : Type u_2} {P : αProp} (f : (a : α) → P aβ) {xs : Array α} (H : ∀ (a : α), a xsP a) :
        Array.pmap f xs H #[] xs #[]
        theorem Array.pmap_eq_self {α : Type u_1} {l : Array α} {p : αProp} (hp : ∀ (a : α), a lp a) (f : (a : α) → p aα) :
        Array.pmap f l hp = l ∀ (a : α) (h : a l), f a = a
        @[simp]
        theorem Array.attach_eq_empty_iff {α : Type u_1} {l : Array α} :
        l.attach = #[] l = #[]
        theorem Array.attach_ne_empty_iff {α : Type u_1} {l : Array α} :
        l.attach #[] l #[]
        @[simp]
        theorem Array.attachWith_eq_empty_iff {α : Type u_1} {l : Array α} {P : αProp} {H : ∀ (a : α), a lP a} :
        l.attachWith P H = #[] l = #[]
        theorem Array.attachWith_ne_empty_iff {α : Type u_1} {l : Array α} {P : αProp} {H : ∀ (a : α), a lP a} :
        l.attachWith P H #[] l #[]
        @[simp]
        theorem Array.getElem?_pmap {α : Type u_1} {β : Type u_2} {p : αProp} (f : (a : α) → p aβ) {l : Array α} (h : ∀ (a : α), a lp a) (n : Nat) :
        (Array.pmap f l h)[n]? = Option.pmap f l[n]?
        @[simp]
        theorem Array.getElem_pmap {α : Type u_1} {β : Type u_2} {p : αProp} (f : (a : α) → p aβ) {l : Array α} (h : ∀ (a : α), a lp a) {n : Nat} (hn : n < (Array.pmap f l h).size) :
        (Array.pmap f l h)[n] = f l[n]
        @[simp]
        theorem Array.getElem?_attachWith {α : Type u_1} {xs : Array α} {i : Nat} {P : αProp} {H : ∀ (a : α), a xsP a} :
        (xs.attachWith P H)[i]? = Option.pmap Subtype.mk xs[i]?
        @[simp]
        theorem Array.getElem?_attach {α : Type u_1} {xs : Array α} {i : Nat} :
        xs.attach[i]? = Option.pmap Subtype.mk xs[i]?
        @[simp]
        theorem Array.getElem_attachWith {α : Type u_1} {xs : Array α} {P : αProp} {H : ∀ (a : α), a xsP a} {i : Nat} (h : i < (xs.attachWith P H).size) :
        (xs.attachWith P H)[i] = xs[i],
        @[simp]
        theorem Array.getElem_attach {α : Type u_1} {xs : Array α} {i : Nat} (h : i < xs.attach.size) :
        xs.attach[i] = xs[i],
        theorem Array.foldl_pmap {α : Type u_1} {β : Type u_2} {γ : Type u_3} (l : Array α) {P : αProp} (f : (a : α) → P aβ) (H : ∀ (a : α), a lP a) (g : γβγ) (x : γ) :
        Array.foldl g x (Array.pmap f l H) = Array.foldl (fun (acc : γ) (a : { x : α // x l }) => g acc (f a.val )) x l.attach
        theorem Array.foldr_pmap {α : Type u_1} {β : Type u_2} {γ : Type u_3} (l : Array α) {P : αProp} (f : (a : α) → P aβ) (H : ∀ (a : α), a lP a) (g : βγγ) (x : γ) :
        Array.foldr g x (Array.pmap f l H) = Array.foldr (fun (a : { x : α // x l }) (acc : γ) => g (f a.val ) acc) x l.attach
        theorem Array.foldl_attach {α : Type u_1} {β : Type u_2} (l : Array α) (f : βαβ) (b : β) :
        Array.foldl (fun (acc : β) (t : { x : α // x l }) => f acc t.val) b l.attach = Array.foldl f b l

        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 Array.foldr_attach {α : Type u_1} {β : Type u_2} (l : Array α) (f : αββ) (b : β) :
        Array.foldr (fun (t : { x : α // x l }) (acc : β) => f t.val acc) b l.attach = Array.foldr f b l

        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 Array.attach_map {α : Type u_1} {β : Type u_2} {l : Array α} (f : αβ) :
        (Array.map f l).attach = Array.map (fun (x : { x : α // x l }) => match x with | x, h => f x, ) l.attach
        theorem Array.attachWith_map {α : Type u_1} {β : Type u_2} {l : Array α} (f : αβ) {P : βProp} {H : ∀ (b : β), b Array.map f lP b} :
        (Array.map f l).attachWith P H = Array.map (fun (x : { x : α // (P f) x }) => match x with | x, h => f x, h) (l.attachWith (P f) )
        theorem Array.map_attachWith {α : Type u_1} {β : Type u_2} {l : Array α} {P : αProp} {H : ∀ (a : α), a lP a} (f : { x : α // P x }β) :
        Array.map f (l.attachWith P H) = Array.pmap (fun (a : α) (h : a l P a) => f a, ) l
        theorem Array.map_attach {α : Type u_1} {β : Type u_2} {l : Array α} (f : { x : α // x l }β) :
        Array.map f l.attach = Array.pmap (fun (a : α) (h : a l) => f a, h) l

        See also pmap_eq_map_attach for writing pmap in terms of map and attach.

        theorem Array.attach_filterMap {α : Type u_1} {β : Type u_2} {l : Array α} {f : αOption β} :
        (Array.filterMap f l).attach = Array.filterMap (fun (x : { x : α // x l }) => match x with | x, h => (f x).pbind fun (b : β) (m : b f x) => some b, ) l.attach
        theorem Array.attach_filter {α : Type u_1} {l : Array α} (p : αBool) :
        (Array.filter p l).attach = Array.filterMap (fun (x : { x : α // x l }) => if w : p x.val = true then some x.val, else none) l.attach
        theorem Array.pmap_pmap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {p : αProp} {q : βProp} (g : (a : α) → p aβ) (f : (b : β) → q bγ) (l : Array α) (H₁ : ∀ (a : α), a lp a) (H₂ : ∀ (a : β), a Array.pmap g l H₁q a) :
        Array.pmap f (Array.pmap g l H₁) H₂ = Array.pmap (fun (a : { x : α // x l }) (h : p a.val) => f (g a.val h) ) l.attach
        @[simp]
        theorem Array.pmap_append {ι : Type u_1} {α : Type u_2} {p : ιProp} (f : (a : ι) → p aα) (l₁ l₂ : Array ι) (h : ∀ (a : ι), a l₁ ++ l₂p a) :
        Array.pmap f (l₁ ++ l₂) h = Array.pmap f l₁ ++ Array.pmap f l₂
        theorem Array.pmap_append' {α : Type u_1} {β : Type u_2} {p : αProp} (f : (a : α) → p aβ) (l₁ l₂ : Array α) (h₁ : ∀ (a : α), a l₁p a) (h₂ : ∀ (a : α), a l₂p a) :
        Array.pmap f (l₁ ++ l₂) = Array.pmap f l₁ h₁ ++ Array.pmap f l₂ h₂
        @[simp]
        theorem Array.attach_append {α : Type u_1} (xs ys : Array α) :
        (xs ++ ys).attach = Array.map (fun (x : { x : α // x xs }) => match x with | x, h => x, ) xs.attach ++ Array.map (fun (x : { x : α // x ys }) => match x with | x, h => x, ) ys.attach
        @[simp]
        theorem Array.attachWith_append {α : Type u_1} {P : αProp} {xs ys : Array α} {H : ∀ (a : α), a xs ++ ysP a} :
        (xs ++ ys).attachWith P H = xs.attachWith P ++ ys.attachWith P
        @[simp]
        theorem Array.pmap_reverse {α : Type u_1} {β : Type u_2} {P : αProp} (f : (a : α) → P aβ) (xs : Array α) (H : ∀ (a : α), a xs.reverseP a) :
        Array.pmap f xs.reverse H = (Array.pmap f xs ).reverse
        theorem Array.reverse_pmap {α : Type u_1} {β : Type u_2} {P : αProp} (f : (a : α) → P aβ) (xs : Array α) (H : ∀ (a : α), a xsP a) :
        (Array.pmap f xs H).reverse = Array.pmap f xs.reverse
        @[simp]
        theorem Array.attachWith_reverse {α : Type u_1} {P : αProp} {xs : Array α} {H : ∀ (a : α), a xs.reverseP a} :
        xs.reverse.attachWith P H = (xs.attachWith P ).reverse
        theorem Array.reverse_attachWith {α : Type u_1} {P : αProp} {xs : Array α} {H : ∀ (a : α), a xsP a} :
        (xs.attachWith P H).reverse = xs.reverse.attachWith P
        @[simp]
        theorem Array.attach_reverse {α : Type u_1} (xs : Array α) :
        xs.reverse.attach = Array.map (fun (x : { x : α // x xs }) => match x with | x, h => x, ) xs.attach.reverse
        theorem Array.reverse_attach {α : Type u_1} (xs : Array α) :
        xs.attach.reverse = Array.map (fun (x : { x : α // x xs.reverse }) => match x with | x, h => x, ) xs.reverse.attach
        @[simp]
        theorem Array.back?_pmap {α : Type u_1} {β : Type u_2} {P : αProp} (f : (a : α) → P aβ) (xs : Array α) (H : ∀ (a : α), a xsP a) :
        (Array.pmap f xs H).back? = Option.map (fun (x : { x : α // x xs }) => match x with | a, m => f a ) xs.attach.back?
        @[simp]
        theorem Array.back?_attachWith {α : Type u_1} {P : αProp} {xs : Array α} {H : ∀ (a : α), a xsP a} :
        (xs.attachWith P H).back? = xs.back?.pbind fun (a : α) (h : a xs.back?) => some a,
        @[simp]
        theorem Array.back?_attach {α : Type u_1} {xs : Array α} :
        xs.attach.back? = xs.back?.pbind fun (a : α) (h : a xs.back?) => some a,

        unattach #

        Array.unattach is the (one-sided) inverse of Array.attach. It is a synonym for Array.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 : Array { 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.

        def Array.unattach {α : Type u_1} {p : αProp} (l : Array { x : α // p x }) :

        A synonym for l.map (·.val). Mostly this should not be needed by users. It is introduced as in intermediate step by lemmas such as map_subtype, and is ideally subsequently simplified away by unattach_attach.

        If not, usually the right approach is simp [Array.unattach, -Array.map_subtype] to unfold.

        Equations
        • l.unattach = Array.map (fun (x : { x : α // p x }) => x.val) l
        Instances For
          @[simp]
          theorem Array.unattach_nil {α : Type u_1} {p : αProp} :
          #[].unattach = #[]
          @[simp]
          theorem Array.unattach_push {α : Type u_1} {p : αProp} {a : { x : α // p x }} {l : Array { x : α // p x }} :
          (l.push a).unattach = l.unattach.push a.val
          @[simp]
          theorem Array.size_unattach {α : Type u_1} {p : αProp} {l : Array { x : α // p x }} :
          l.unattach.size = l.size
          @[simp]
          theorem List.unattach_toArray {α : Type u_1} {p : αProp} {l : List { x : α // p x }} :
          l.toArray.unattach = l.unattach.toArray
          @[simp]
          theorem Array.toList_unattach {α : Type u_1} {p : αProp} {l : Array { x : α // p x }} :
          l.unattach.toList = l.toList.unattach
          @[simp]
          theorem Array.unattach_attach {α : Type u_1} {l : Array α} :
          l.attach.unattach = l
          @[simp]
          theorem Array.unattach_attachWith {α : Type u_1} {p : αProp} {l : Array α} {H : ∀ (a : α), a lp a} :
          (l.attachWith p H).unattach = l
          @[simp]
          theorem Array.getElem?_unattach {α : Type u_1} {p : αProp} {l : Array { x : α // p x }} (i : Nat) :
          l.unattach[i]? = Option.map Subtype.val l[i]?
          @[simp]
          theorem Array.getElem_unattach {α : Type u_1} {p : αProp} {l : Array { x : α // p x }} (i : Nat) (h : i < l.unattach.size) :
          l.unattach[i] = l[i].val

          Recognizing higher order functions using a function that only depends on the value. #

          theorem Array.foldl_subtype {α : Type u_1} {β : Type u_2} {p : αProp} {l : Array { x : α // p x }} {f : β{ x : α // p x }β} {g : βαβ} {x : β} {hf : ∀ (b : β) (x : α) (h : p x), f b x, h = g b x} :
          Array.foldl f x l = Array.foldl g x l.unattach

          This lemma identifies folds over arrays 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 Array.foldl_subtype' {α : Type u_1} {β : Type u_2} {stop : Nat} {p : αProp} {l : Array { x : α // p x }} {f : β{ x : α // p x }β} {g : βαβ} {x : β} {hf : ∀ (b : β) (x : α) (h : p x), f b x, h = g b x} (h : stop = l.size) :
          Array.foldl f x l 0 stop = Array.foldl g x l.unattach

          Variant of foldl_subtype with side condition to check stop = l.size.

          theorem Array.foldr_subtype {α : Type u_1} {β : Type u_2} {p : αProp} {l : Array { x : α // p x }} {f : { x : α // p x }ββ} {g : αββ} {x : β} {hf : ∀ (x : α) (h : p x) (b : β), f x, h b = g x b} :
          Array.foldr f x l = Array.foldr g x l.unattach

          This lemma identifies folds over arrays 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 Array.foldr_subtype' {α : Type u_1} {β : Type u_2} {start : Nat} {p : αProp} {l : Array { x : α // p x }} {f : { x : α // p x }ββ} {g : αββ} {x : β} {hf : ∀ (x : α) (h : p x) (b : β), f x, h b = g x b} (h : start = l.size) :
          Array.foldr f x l start = Array.foldr g x l.unattach

          Variant of foldr_subtype with side condition to check stop = l.size.

          @[simp]
          theorem Array.map_subtype {α : Type u_1} {β : Type u_2} {p : αProp} {l : Array { x : α // p x }} {f : { x : α // p x }β} {g : αβ} {hf : ∀ (x : α) (h : p x), f x, h = g x} :
          Array.map f l = Array.map g l.unattach

          This lemma identifies maps over arrays 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 Array.filterMap_subtype {α : Type u_1} {β : Type u_2} {p : αProp} {l : Array { x : α // p x }} {f : { x : α // p x }Option β} {g : αOption β} {hf : ∀ (x : α) (h : p x), f x, h = g x} :
          @[simp]
          theorem Array.unattach_filter {α : Type u_1} {p : αProp} {l : Array { x : α // p x }} {f : { x : α // p x }Bool} {g : αBool} {hf : ∀ (x : α) (h : p x), f x, h = g x} :
          (Array.filter f l).unattach = Array.filter g l.unattach

          Simp lemmas pushing unattach inwards. #

          @[simp]
          theorem Array.unattach_reverse {α : Type u_1} {p : αProp} {l : Array { x : α // p x }} :
          l.reverse.unattach = l.unattach.reverse
          @[simp]
          theorem Array.unattach_append {α : Type u_1} {p : αProp} {l₁ l₂ : Array { x : α // p x }} :
          (l₁ ++ l₂).unattach = l₁.unattach ++ l₂.unattach