diff --git a/CHANGELOG.md b/CHANGELOG.md index 81f599b13..520781854 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -154,6 +154,7 @@ assumption in `setByte` that the value is between 0 and 255. * Adds RefC support for 16- and 32-bit access in `Data.Buffer`. + * Add `Show` instance to `Data.Vect.Quantifiers.All` and add a few helpers for listy computations on the `All` type. * Add an alias for `HVect` to `All id` in `Data.Vect.Quantifiers.All`. This is the @@ -248,6 +249,8 @@ * A more generalised way of applicative mapping of `TTImp` expression was added, called `mapATTImp`; the original `mapMTTimp` was implemented through the new one. +* Adds `Data.Vect.foldrImplGoLemma`. + #### System * Changes `getNProcessors` to return the number of online processors rather than diff --git a/CONTRIBUTORS b/CONTRIBUTORS index 6170c090c..0d8237250 100644 --- a/CONTRIBUTORS +++ b/CONTRIBUTORS @@ -7,6 +7,7 @@ Aleksei Volkov Alex Gryzlov Alex Silva Alissa Tung +Andor Penzes Andre Kuhlenschmidt André Videla Andy Lok diff --git a/libs/base/Data/Vect.idr b/libs/base/Data/Vect.idr index 413959e42..8ddf862e5 100644 --- a/libs/base/Data/Vect.idr +++ b/libs/base/Data/Vect.idr @@ -6,8 +6,9 @@ import Data.Nat import public Data.Fin import public Data.Zippable -import Decidable.Equality import Control.Function +import Decidable.Equality +import Syntax.PreorderReasoning %default total @@ -413,6 +414,17 @@ foldrImpl : (t -> acc -> acc) -> acc -> (acc -> acc) -> Vect n t -> acc foldrImpl f e go [] = go e foldrImpl f e go (x::xs) = foldrImpl f e (go . (f x)) xs +export +foldrImplGoLemma + : (x : a) -> (xs : Vect n a) -> (f : a -> b -> b) -> (e : b) -> (go : b -> b) + -> go (foldrImpl f e (f x) xs) === foldrImpl f e (go . (f x)) xs +foldrImplGoLemma z [] f e go = Refl +foldrImplGoLemma z (y :: ys) f e go = Calc $ + |~ go (foldrImpl f e ((f z) . (f y)) ys) + ~~ go ((f z) (foldrImpl f e (f y) ys)) ... (cong go (sym (foldrImplGoLemma y ys f e (f z)))) + ~~ (go . (f z)) (foldrImpl f e (f y) ys) ... (cong go Refl) + ~~ foldrImpl f e (go . (f z) . (f y)) ys ... (foldrImplGoLemma y ys f e (go . (f z))) + public export implementation Foldable (Vect n) where foldr f e xs = foldrImpl f e id xs