mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-24 12:14:26 +03:00
Add foldrImplGoLemma to Data.Vect (#2835)
Co-authored-by: CodingCellist <teh6@st-andrews.ac.uk>
This commit is contained in:
parent
6dd25cd9ec
commit
82e4fd9da7
@ -154,6 +154,7 @@
|
|||||||
assumption in `setByte` that the value is between 0 and 255.
|
assumption in `setByte` that the value is between 0 and 255.
|
||||||
|
|
||||||
* Adds RefC support for 16- and 32-bit access in `Data.Buffer`.
|
* 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
|
* Add `Show` instance to `Data.Vect.Quantifiers.All` and add a few helpers for listy
|
||||||
computations on the `All` type.
|
computations on the `All` type.
|
||||||
* Add an alias for `HVect` to `All id` in `Data.Vect.Quantifiers.All`. This is the
|
* 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,
|
* A more generalised way of applicative mapping of `TTImp` expression was added,
|
||||||
called `mapATTImp`; the original `mapMTTimp` was implemented through the new one.
|
called `mapATTImp`; the original `mapMTTimp` was implemented through the new one.
|
||||||
|
|
||||||
|
* Adds `Data.Vect.foldrImplGoLemma`.
|
||||||
|
|
||||||
#### System
|
#### System
|
||||||
|
|
||||||
* Changes `getNProcessors` to return the number of online processors rather than
|
* Changes `getNProcessors` to return the number of online processors rather than
|
||||||
|
@ -7,6 +7,7 @@ Aleksei Volkov
|
|||||||
Alex Gryzlov
|
Alex Gryzlov
|
||||||
Alex Silva
|
Alex Silva
|
||||||
Alissa Tung
|
Alissa Tung
|
||||||
|
Andor Penzes
|
||||||
Andre Kuhlenschmidt
|
Andre Kuhlenschmidt
|
||||||
André Videla
|
André Videla
|
||||||
Andy Lok
|
Andy Lok
|
||||||
|
@ -6,8 +6,9 @@ import Data.Nat
|
|||||||
import public Data.Fin
|
import public Data.Fin
|
||||||
import public Data.Zippable
|
import public Data.Zippable
|
||||||
|
|
||||||
import Decidable.Equality
|
|
||||||
import Control.Function
|
import Control.Function
|
||||||
|
import Decidable.Equality
|
||||||
|
import Syntax.PreorderReasoning
|
||||||
|
|
||||||
%default total
|
%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 [] = go e
|
||||||
foldrImpl f e go (x::xs) = foldrImpl f e (go . (f x)) xs
|
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
|
public export
|
||||||
implementation Foldable (Vect n) where
|
implementation Foldable (Vect n) where
|
||||||
foldr f e xs = foldrImpl f e id xs
|
foldr f e xs = foldrImpl f e id xs
|
||||||
|
Loading…
Reference in New Issue
Block a user