mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-19 17:21:59 +03:00
[ base ] Prove anyToFin
preserves the property witnessed by Any
This commit is contained in:
parent
8d5caaa137
commit
32b639ca3c
@ -67,6 +67,14 @@ namespace Any
|
||||
anyToFin (Here _) = FZ
|
||||
anyToFin (There later) = FS (anyToFin later)
|
||||
|
||||
||| `anyToFin`'s return type satisfies the predicate
|
||||
export
|
||||
anyToFinCorrect : {0 xs : Vect n a} ->
|
||||
(witness : Any p xs) ->
|
||||
p (anyToFin witness `index` xs)
|
||||
anyToFinCorrect (Here prf) = prf
|
||||
anyToFinCorrect (There later) = anyToFinCorrect later
|
||||
|
||||
namespace All
|
||||
||| A proof that all elements of a vector satisfy a property. It is a list of
|
||||
||| proofs, corresponding element-wise to the `Vect`.
|
||||
|
Loading…
Reference in New Issue
Block a user