mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 20:51:43 +03:00
[ base ] Move the worker outside of the Vect's reverse
definition
This commit is contained in:
parent
7d9d39c045
commit
c6a8c9e7a7
@ -257,6 +257,17 @@ replaceAtDiffIndexPreserves (_::_) (FS z) (FS w) co y = replaceAtDiffIndexPreser
|
|||||||
-- Transformations
|
-- Transformations
|
||||||
--------------------------------------------------------------------------------
|
--------------------------------------------------------------------------------
|
||||||
|
|
||||||
|
||| Reverse the second vector, prepending the result to the first.
|
||||||
|
|||
|
||||||
|
||| ```idris example
|
||||||
|
||| reverseOnto [0, 1] [10, 11, 12]
|
||||||
|
||| ```
|
||||||
|
public export
|
||||||
|
reverseOnto : Vect n elem -> Vect m elem -> Vect (n+m) elem
|
||||||
|
reverseOnto {n} acc [] = rewrite plusZeroRightNeutral n in acc
|
||||||
|
reverseOnto {n} {m=S m} acc (x :: xs) = rewrite sym $ plusSuccRightSucc n m
|
||||||
|
in reverseOnto (x::acc) xs
|
||||||
|
|
||||||
||| Reverse the order of the elements of a vector
|
||| Reverse the order of the elements of a vector
|
||||||
|||
|
|||
|
||||||
||| ```idris example
|
||| ```idris example
|
||||||
@ -264,11 +275,7 @@ replaceAtDiffIndexPreserves (_::_) (FS z) (FS w) co y = replaceAtDiffIndexPreser
|
|||||||
||| ```
|
||| ```
|
||||||
public export
|
public export
|
||||||
reverse : (xs : Vect len elem) -> Vect len elem
|
reverse : (xs : Vect len elem) -> Vect len elem
|
||||||
reverse xs = go [] xs
|
reverse = reverseOnto []
|
||||||
where go : Vect n elem -> Vect m elem -> Vect (n+m) elem
|
|
||||||
go {n} acc [] = rewrite plusZeroRightNeutral n in acc
|
|
||||||
go {n} {m=S m} acc (x :: xs) = rewrite sym $ plusSuccRightSucc n m
|
|
||||||
in go (x::acc) xs
|
|
||||||
|
|
||||||
||| Alternate an element between the other elements of a vector
|
||| Alternate an element between the other elements of a vector
|
||||||
||| @ sep the element to intersperse
|
||| @ sep the element to intersperse
|
||||||
|
Loading…
Reference in New Issue
Block a user