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
- Vector.pmap f xs H = Vector.mk (Array.pmap f xs.toArray ⋯) ⋯
Instances For
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 = Vector.mk (xs.toArray.attachWith P ⋯) ⋯
Instances For
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
- xs.attach = xs.attachWith (Membership.mem xs) ⋯
Instances For
Implementation of pmap
using the zero-copy version of attach
.
Equations
- Vector.pmapImpl f xs H = Vector.map (fun (x : { x : α // P x }) => match x with | ⟨x, h'⟩ => f x h') (xs.attachWith P H)
Instances For
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.
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.
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 xs.attach.unattach = xs
, and simp lemmas which recognize higher order
functions applied to xs : Vector { x // p x }
which only depend on the value, not the predicate, and rewrite these
in terms of a simpler function applied to xs.unattach
.
Further, we provide simp lemmas that push unattach
inwards.
A synonym for xs.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.
Instances For
Equations
Instances For
Recognizing higher order functions using a function that only depends on the value. #
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.
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.
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.