Documentation

Init.Data.Vector.Attach

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

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.Vector.Attach.0.Vector.attachWithImpl]
    def Vector.attachWith {α : Type u_1} {n : Nat} (xs : Vector α n) (P : αProp) (H : ∀ (x : α), x xsP x) :
    Vector { x : α // P x } n

    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
    Instances For
      @[inline]
      def Vector.attach {α : Type u_1} {n : Nat} (xs : Vector α n) :
      Vector { x : α // x xs } n

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

      Equations
      Instances For
        @[simp]
        theorem Vector.attachWith_mk {α : Type u_1} {n : Nat} {xs : Array α} {h : xs.size = n} {P : αProp} {H : ∀ (x : α), x { toArray := xs, size_toArray := h }P x} :
        { toArray := xs, size_toArray := h }.attachWith P H = { toArray := xs.attachWith P , size_toArray := }
        @[simp]
        theorem Vector.attach_mk {α : Type u_1} {n : Nat} {xs : Array α} {h : xs.size = n} :
        { toArray := xs, size_toArray := h }.attach = { toArray := xs.attachWith (fun (x : α) => x { toArray := xs, size_toArray := h }) , size_toArray := }
        @[simp]
        theorem Vector.pmap_mk {α : Type u_1} {n : Nat} {β : Type u_2} {xs : Array α} {h : xs.size = n} {P : αProp} {f : (a : α) → P aβ} {H : ∀ (a : α), a { toArray := xs, size_toArray := h }P a} :
        pmap f { toArray := xs, size_toArray := h } H = { toArray := Array.pmap f xs , size_toArray := }
        @[simp]
        theorem Vector.toArray_attachWith {α : Type u_1} {n : Nat} {l : Vector α n} {P : αProp} {H : ∀ (x : α), x lP x} :
        (l.attachWith P H).toArray = l.attachWith P
        @[simp]
        theorem Vector.toArray_attach {n : Nat} {α : Type u_1} {l : Vector α n} :
        l.attach.toArray = l.attachWith (fun (x : α) => x l)
        @[simp]
        theorem Vector.toArray_pmap {α : Type u_1} {n : Nat} {β : Type u_2} {l : Vector α n} {P : αProp} {f : (a : α) → P aβ} {H : ∀ (a : α), a lP a} :
        (pmap f l H).toArray = Array.pmap f l.toArray
        @[simp]
        theorem Vector.toList_attachWith {α : Type u_1} {n : Nat} {l : Vector α n} {P : αProp} {H : ∀ (x : α), x lP x} :
        @[simp]
        theorem Vector.toList_attach {n : Nat} {α : Type u_1} {l : Vector α n} :
        l.attach.toList = l.toList.attachWith (fun (x : α) => x l)
        @[simp]
        theorem Vector.toList_pmap {α : Type u_1} {n : Nat} {β : Type u_2} {l : Vector α n} {P : αProp} {f : (a : α) → P aβ} {H : ∀ (a : α), a lP a} :
        (pmap f l H).toList = List.pmap f l.toList
        @[simp]
        theorem Vector.pmap_empty {α : Type u_1} {β : Type u_2} {P : αProp} (f : (a : α) → P aβ) :
        pmap f { toArray := #[], size_toArray := } = { toArray := #[], size_toArray := }
        @[simp]
        theorem Vector.pmap_push {α : Type u_1} {β : Type u_2} {n : Nat} {P : αProp} (f : (a : α) → P aβ) (a : α) (l : Vector α n) (h : ∀ (b : α), b l.push aP b) :
        pmap f (l.push a) h = (pmap f l ).push (f a )
        @[simp]
        theorem Vector.attach_empty {α : Type u_1} :
        { toArray := #[], size_toArray := }.attach = { toArray := #[], size_toArray := }
        @[simp]
        theorem Vector.attachWith_empty {α : Type u_1} {P : αProp} (H : ∀ (x : α), x { toArray := #[], size_toArray := }P x) :
        { toArray := #[], size_toArray := }.attachWith P H = { toArray := #[], size_toArray := }
        @[simp]
        theorem Vector.pmap_eq_map {α : Type u_1} {β : Type u_2} {n : Nat} (p : αProp) (f : αβ) (l : Vector α n) (H : ∀ (a : α), a lp a) :
        pmap (fun (a : α) (x : p a) => f a) l H = map f l
        theorem Vector.pmap_congr_left {α : Type u_1} {β : Type u_2} {n : Nat} {p q : αProp} {f : (a : α) → p aβ} {g : (a : α) → q aβ} (l : Vector α n) {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₂) :
        pmap f l H₁ = pmap g l H₂
        theorem Vector.map_pmap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {n : Nat} {p : αProp} (g : βγ) (f : (a : α) → p aβ) (l : Vector α n) (H : ∀ (a : α), a lp a) :
        map g (pmap f l H) = pmap (fun (a : α) (h : p a) => g (f a h)) l H
        theorem Vector.pmap_map {β : Type u_1} {γ : Type u_2} {α : Type u_3} {n : Nat} {p : βProp} (g : (b : β) → p bγ) (f : αβ) (l : Vector α n) (H : ∀ (a : β), a map f lp a) :
        pmap g (map f l) H = pmap (fun (a : α) (h : p (f a)) => g (f a) h) l
        theorem Vector.attach_congr {α : Type u_1} {n : Nat} {l₁ l₂ : Vector α n} (h : l₁ = l₂) :
        l₁.attach = map (fun (x : { x : α // x l₂ }) => x.val, ) l₂.attach
        theorem Vector.attachWith_congr {α : Type u_1} {n : Nat} {l₁ l₂ : Vector α n} (w : l₁ = l₂) {P : αProp} {H : ∀ (x : α), x l₁P x} :
        l₁.attachWith P H = l₂.attachWith P
        @[simp]
        theorem Vector.attach_push {α : Type u_1} {n : Nat} {a : α} {l : Vector α n} :
        (l.push a).attach = (map (fun (x : { x : α // x l }) => match x with | x, h => x, ) l.attach).push a,
        @[simp]
        theorem Vector.attachWith_push {α : Type u_1} {n : Nat} {a : α} {l : Vector α n} {P : αProp} {H : ∀ (x : α), x l.push aP x} :
        (l.push a).attachWith P H = (l.attachWith P ).push a,
        theorem Vector.pmap_eq_map_attach {α : Type u_1} {β : Type u_2} {n : Nat} {p : αProp} (f : (a : α) → p aβ) (l : Vector α n) (H : ∀ (a : α), a lp a) :
        pmap f l H = map (fun (x : { x : α // x l }) => f x.val ) l.attach
        @[simp]
        theorem Vector.pmap_eq_attachWith {α : Type u_1} {n : Nat} {p q : αProp} (f : ∀ (a : α), p aq a) (l : Vector α n) (H : ∀ (a : α), a lp a) :
        pmap (fun (a : α) (h : p a) => a, ) l H = l.attachWith q
        theorem Vector.attach_map_coe {α : Type u_1} {n : Nat} {β : Type u_2} (l : Vector α n) (f : αβ) :
        map (fun (i : { i : α // i l }) => f i.val) l.attach = map f l
        theorem Vector.attach_map_val {α : Type u_1} {n : Nat} {β : Type u_2} (l : Vector α n) (f : αβ) :
        map (fun (i : { x : α // x l }) => f i.val) l.attach = map f l
        theorem Vector.attach_map_subtype_val {α : Type u_1} {n : Nat} (l : Vector α n) :
        theorem Vector.attachWith_map_coe {α : Type u_1} {β : Type u_2} {n : Nat} {p : αProp} (f : αβ) (l : Vector α n) (H : ∀ (a : α), a lp a) :
        map (fun (i : { i : α // p i }) => f i.val) (l.attachWith p H) = map f l
        theorem Vector.attachWith_map_val {α : Type u_1} {β : Type u_2} {n : Nat} {p : αProp} (f : αβ) (l : Vector α n) (H : ∀ (a : α), a lp a) :
        map (fun (i : { x : α // p x }) => f i.val) (l.attachWith p H) = map f l
        theorem Vector.attachWith_map_subtype_val {α : Type u_1} {n : Nat} {p : αProp} (l : Vector α n) (H : ∀ (a : α), a lp a) :
        @[simp]
        theorem Vector.mem_attach {α : Type u_1} {n : Nat} (l : Vector α n) (x : { x : α // x l }) :
        @[simp]
        theorem Vector.mem_attachWith {α : Type u_1} {n : Nat} (l : Vector α n) {q : αProp} (H : ∀ (x : α), x lq x) (x : { x : α // q x }) :
        x l.attachWith q H x.val l
        @[simp]
        theorem Vector.mem_pmap {α : Type u_1} {β : Type u_2} {n : Nat} {p : αProp} {f : (a : α) → p aβ} {l : Vector α n} {H : ∀ (a : α), a lp a} {b : β} :
        b pmap f l H (a : α), (h : a l), f a = b
        theorem Vector.mem_pmap_of_mem {α : Type u_1} {β : Type u_2} {n : Nat} {p : αProp} {f : (a : α) → p aβ} {l : Vector α n} {H : ∀ (a : α), a lp a} {a : α} (h : a l) :
        f a pmap f l H
        theorem Vector.pmap_eq_self {α : Type u_1} {n : Nat} {l : Vector α n} {p : αProp} {hp : ∀ (a : α), a lp a} {f : (a : α) → p aα} :
        pmap f l hp = l ∀ (a : α) (h : a l), f a = a
        @[simp]
        theorem Vector.getElem?_pmap {α : Type u_1} {β : Type u_2} {n : Nat} {p : αProp} (f : (a : α) → p aβ) {l : Vector α n} (h : ∀ (a : α), a lp a) (i : Nat) :
        (pmap f l h)[i]? = Option.pmap f l[i]?
        @[simp]
        theorem Vector.getElem_pmap {α : Type u_1} {β : Type u_2} {n : Nat} {p : αProp} (f : (a : α) → p aβ) {l : Vector α n} (h : ∀ (a : α), a lp a) {i : Nat} (hn : i < n) :
        (pmap f l h)[i] = f l[i]
        @[simp]
        theorem Vector.getElem?_attachWith {α : Type u_1} {n : Nat} {xs : Vector α n} {i : Nat} {P : αProp} {H : ∀ (a : α), a xsP a} :
        @[simp]
        theorem Vector.getElem?_attach {α : Type u_1} {n : Nat} {xs : Vector α n} {i : Nat} :
        @[simp]
        theorem Vector.getElem_attachWith {α : Type u_1} {n : Nat} {xs : Vector α n} {P : αProp} {H : ∀ (a : α), a xsP a} {i : Nat} (h : i < n) :
        (xs.attachWith P H)[i] = xs[i],
        @[simp]
        theorem Vector.getElem_attach {α : Type u_1} {n : Nat} {xs : Vector α n} {i : Nat} (h : i < n) :
        xs.attach[i] = xs[i],
        @[simp]
        theorem Vector.pmap_attach {α : Type u_1} {n : Nat} {β : Type u_2} (l : Vector α n) {p : { x : α // x l }Prop} (f : (a : { x : α // x l }) → p aβ) (H : ∀ (a : { x : α // x l }), a l.attachp a) :
        pmap f l.attach H = pmap (fun (a : α) (h : (fun (a : α) => (h : a l), p a, h) a) => f a, ) l
        @[simp]
        theorem Vector.pmap_attachWith {α : Type u_1} {n : Nat} {q : αProp} {β : Type u_2} (l : Vector α n) {p : { x : α // q x }Prop} (f : (a : { x : α // q x }) → p aβ) (H₁ : ∀ (x : α), x lq x) (H₂ : ∀ (a : { x : α // q x }), a l.attachWith q H₁p a) :
        pmap f (l.attachWith q H₁) H₂ = pmap (fun (a : α) (h : (fun (a : α) => (h : q a), p a, h) a) => f a, ) l
        theorem Vector.foldl_pmap {α : Type u_1} {n : Nat} {β : Type u_2} {γ : Type u_3} (l : Vector α n) {P : αProp} (f : (a : α) → P aβ) (H : ∀ (a : α), a lP a) (g : γβγ) (x : γ) :
        foldl g x (pmap f l H) = foldl (fun (acc : γ) (a : { x : α // x l }) => g acc (f a.val )) x l.attach
        theorem Vector.foldr_pmap {α : Type u_1} {n : Nat} {β : Type u_2} {γ : Type u_3} (l : Vector α n) {P : αProp} (f : (a : α) → P aβ) (H : ∀ (a : α), a lP a) (g : βγγ) (x : γ) :
        foldr g x (pmap f l H) = foldr (fun (a : { x : α // x l }) (acc : γ) => g (f a.val ) acc) x l.attach
        theorem Vector.foldl_attach {α : Type u_1} {n : Nat} {β : Type u_2} (l : Vector α n) (f : βαβ) (b : β) :
        foldl (fun (acc : β) (t : { x : α // x l }) => f acc t.val) b l.attach = 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 Vector.foldr_attach {α : Type u_1} {n : Nat} {β : Type u_2} (l : Vector α n) (f : αββ) (b : β) :
        foldr (fun (t : { x : α // x l }) (acc : β) => f t.val acc) b l.attach = 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 Vector.attach_map {α : Type u_1} {n : Nat} {β : Type u_2} {l : Vector α n} (f : αβ) :
        (map f l).attach = map (fun (x : { x : α // x l }) => match x with | x, h => f x, ) l.attach
        theorem Vector.attachWith_map {α : Type u_1} {n : Nat} {β : Type u_2} {l : Vector α n} (f : αβ) {P : βProp} {H : ∀ (b : β), b map f lP b} :
        (map f l).attachWith P H = map (fun (x : { x : α // (P f) x }) => match x with | x, h => f x, h) (l.attachWith (P f) )
        theorem Vector.map_attachWith {α : Type u_1} {n : Nat} {β : Type u_2} {l : Vector α n} {P : αProp} {H : ∀ (a : α), a lP a} (f : { x : α // P x }β) :
        map f (l.attachWith P H) = pmap (fun (a : α) (h : a l P a) => f a, ) l
        theorem Vector.map_attach {α : Type u_1} {n : Nat} {β : Type u_2} {l : Vector α n} (f : { x : α // x l }β) :
        map f l.attach = 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 Vector.pmap_pmap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {n : Nat} {p : αProp} {q : βProp} (g : (a : α) → p aβ) (f : (b : β) → q bγ) (l : Vector α n) (H₁ : ∀ (a : α), a lp a) (H₂ : ∀ (a : β), a pmap g l H₁q a) :
        pmap f (pmap g l H₁) H₂ = pmap (fun (a : { x : α // x l }) (h : p a.val) => f (g a.val h) ) l.attach
        @[simp]
        theorem Vector.pmap_append {ι : Type u_1} {α : Type u_2} {n m : Nat} {p : ιProp} (f : (a : ι) → p aα) (l₁ : Vector ι n) (l₂ : Vector ι m) (h : ∀ (a : ι), a l₁ ++ l₂p a) :
        pmap f (l₁ ++ l₂) h = pmap f l₁ ++ pmap f l₂
        theorem Vector.pmap_append' {α : Type u_1} {β : Type u_2} {n m : Nat} {p : αProp} (f : (a : α) → p aβ) (l₁ : Vector α n) (l₂ : Vector α m) (h₁ : ∀ (a : α), a l₁p a) (h₂ : ∀ (a : α), a l₂p a) :
        pmap f (l₁ ++ l₂) = pmap f l₁ h₁ ++ pmap f l₂ h₂
        @[simp]
        theorem Vector.attach_append {α : Type u_1} {n m : Nat} (xs : Vector α n) (ys : Vector α m) :
        (xs ++ ys).attach = map (fun (x : { x : α // x xs }) => match x with | x, h => x, ) xs.attach ++ map (fun (x : { x : α // x ys }) => match x with | y, h => y, ) ys.attach
        @[simp]
        theorem Vector.attachWith_append {α : Type u_1} {n m : Nat} {P : αProp} {xs : Vector α n} {ys : Vector α m} {H : ∀ (a : α), a xs ++ ysP a} :
        (xs ++ ys).attachWith P H = xs.attachWith P ++ ys.attachWith P
        @[simp]
        theorem Vector.pmap_reverse {α : Type u_1} {β : Type u_2} {n : Nat} {P : αProp} (f : (a : α) → P aβ) (xs : Vector α n) (H : ∀ (a : α), a xs.reverseP a) :
        pmap f xs.reverse H = (pmap f xs ).reverse
        theorem Vector.reverse_pmap {α : Type u_1} {β : Type u_2} {n : Nat} {P : αProp} (f : (a : α) → P aβ) (xs : Vector α n) (H : ∀ (a : α), a xsP a) :
        (pmap f xs H).reverse = pmap f xs.reverse
        @[simp]
        theorem Vector.attachWith_reverse {α : Type u_1} {n : Nat} {P : αProp} {xs : Vector α n} {H : ∀ (a : α), a xs.reverseP a} :
        theorem Vector.reverse_attachWith {α : Type u_1} {n : Nat} {P : αProp} {xs : Vector α n} {H : ∀ (a : α), a xsP a} :
        @[simp]
        theorem Vector.attach_reverse {α : Type u_1} {n : Nat} (xs : Vector α n) :
        xs.reverse.attach = map (fun (x : { x : α // x xs }) => match x with | x, h => x, ) xs.attach.reverse
        theorem Vector.reverse_attach {α : Type u_1} {n : Nat} (xs : Vector α n) :
        xs.attach.reverse = map (fun (x : { x : α // x xs.reverse }) => match x with | x, h => x, ) xs.reverse.attach
        @[simp]
        theorem Vector.back?_pmap {α : Type u_1} {β : Type u_2} {n : Nat} {P : αProp} (f : (a : α) → P aβ) (xs : Vector α n) (H : ∀ (a : α), a xsP a) :
        (pmap f xs H).back? = Option.map (fun (x : { x : α // x xs }) => match x with | a, m => f a ) xs.attach.back?
        @[simp]
        theorem Vector.back?_attachWith {α : Type u_1} {n : Nat} {P : αProp} {xs : Vector α n} {H : ∀ (a : α), a xsP a} :
        (xs.attachWith P H).back? = xs.back?.pbind fun (a : α) (h : a xs.back?) => some a,
        @[simp]
        theorem Vector.back?_attach {α : Type u_1} {n : Nat} {xs : Vector α n} :
        xs.attach.back? = xs.back?.pbind fun (a : α) (h : a xs.back?) => some a,
        @[simp]
        theorem Vector.countP_attach {α : Type u_1} {n : Nat} (l : Vector α n) (p : αBool) :
        countP (fun (a : { x : α // x l }) => p a.val) l.attach = countP p l
        @[simp]
        theorem Vector.countP_attachWith {α : Type u_1} {n : Nat} {p : αProp} (l : Vector α n) (H : ∀ (a : α), a lp a) (q : αBool) :
        countP (fun (a : { x : α // p x }) => q a.val) (l.attachWith p H) = countP q l
        @[simp]
        theorem Vector.count_attach {α : Type u_1} {n : Nat} [DecidableEq α] (l : Vector α n) (a : { x : α // x l }) :
        @[simp]
        theorem Vector.count_attachWith {α : Type u_1} {n : Nat} [DecidableEq α] {p : αProp} (l : Vector α n) (H : ∀ (a : α), a lp a) (a : { x : α // p x }) :
        count a (l.attachWith p H) = count a.val l
        @[simp]
        theorem Vector.countP_pmap {α : Type u_1} {β : Type u_2} {n : Nat} {p : αProp} (g : (a : α) → p aβ) (f : βBool) (l : Vector α n) (H₁ : ∀ (a : α), a lp a) :
        countP f (pmap g l H₁) = countP (fun (x : { x : α // x l }) => match x with | a, m => f (g a )) l.attach

        unattach #

        Vector.unattach is the (one-sided) inverse of Vector.attach. It is a synonym for Vector.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 : Vector { 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 Vector.unattach {n : Nat} {α : Type u_1} {p : αProp} (l : Vector { x : α // p x } n) :
        Vector α n

        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 [Vector.unattach, -Vector.map_subtype] to unfold.

        Equations
        Instances For
          @[simp]
          theorem Vector.unattach_nil {α : Type u_1} {p : αProp} :
          { toArray := #[], size_toArray := }.unattach = { toArray := #[], size_toArray := }
          @[simp]
          theorem Vector.unattach_push {α : Type u_1} {n : Nat} {p : αProp} {a : { x : α // p x }} {l : Vector { x : α // p x } n} :
          @[simp]
          theorem Vector.unattach_mk {α : Type u_1} {n : Nat} {p : αProp} {l : Array { x : α // p x }} {h : l.size = n} :
          { toArray := l, size_toArray := h }.unattach = { toArray := l.unattach, size_toArray := }
          @[simp]
          theorem Vector.toArray_unattach {α : Type u_1} {n : Nat} {p : αProp} {l : Vector { x : α // p x } n} :
          @[simp]
          theorem Vector.toList_unattach {α : Type u_1} {p : αProp} {l : Array { x : α // p x }} :
          @[simp]
          theorem Vector.unattach_attach {α : Type u_1} {n : Nat} {l : Vector α n} :
          @[simp]
          theorem Vector.unattach_attachWith {α : Type u_1} {n : Nat} {p : αProp} {l : Vector α n} {H : ∀ (a : α), a lp a} :
          @[simp]
          theorem Vector.getElem?_unattach {α : Type u_1} {n : Nat} {p : αProp} {l : Vector { x : α // p x } n} (i : Nat) :
          @[simp]
          theorem Vector.getElem_unattach {α : Type u_1} {n : Nat} {p : αProp} {l : Vector { x : α // p x } n} (i : Nat) (h : i < n) :

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

          @[simp]
          theorem Vector.foldl_subtype {α : Type u_1} {n : Nat} {β : Type u_2} {p : αProp} {l : Vector { x : α // p x } n} {f : β{ x : α // p x }β} {g : βαβ} {x : β} (hf : ∀ (b : β) (x : α) (h : p x), f b x, h = g b x) :
          foldl f x l = 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 Vector.foldr_subtype {α : Type u_1} {n : Nat} {β : Type u_2} {p : αProp} {l : Vector { x : α // p x } n} {f : { x : α // p x }ββ} {g : αββ} {x : β} (hf : ∀ (x : α) (h : p x) (b : β), f x, h b = g x b) :
          foldr f x l = 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 Vector.map_subtype {α : Type u_1} {n : Nat} {β : Type u_2} {p : αProp} {l : Vector { x : α // p x } n} {f : { x : α // p x }β} {g : αβ} (hf : ∀ (x : α) (h : p x), f x, h = g x) :
          map f l = 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 Vector.findSome?_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 Vector.find?_subtype {α : 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) :

          Simp lemmas pushing unattach inwards. #

          @[simp]
          theorem Vector.unattach_reverse {α : Type u_1} {n : Nat} {p : αProp} {l : Vector { x : α // p x } n} :
          @[simp]
          theorem Vector.unattach_append {α : Type u_1} {n : Nat} {p : αProp} {l₁ l₂ : Vector { x : α // p x } n} :
          (l₁ ++ l₂).unattach = l₁.unattach ++ l₂.unattach
          @[simp]
          theorem Vector.unattach_flatten {α : Type u_1} {n : Nat} {p : αProp} {l : Vector (Vector { x : α // p x } n) n} :
          @[simp]
          theorem Vector.unattach_mkVector {α : Type u_1} {p : αProp} {n : Nat} {x : { x : α // p x }} :