mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-18 08:42:11 +03:00
23 lines
417 B
Idris
23 lines
417 B
Idris
module Data.String.Iterator
|
|
|
|
import public Data.List.Lazy
|
|
|
|
%default total
|
|
|
|
export
|
|
data StringIterator : Type where [external]
|
|
|
|
%foreign
|
|
"scheme:blodwen-string-iterator-new"
|
|
export
|
|
fromString : String -> StringIterator
|
|
|
|
%foreign
|
|
"scheme:blodwen-string-iterator-next"
|
|
export
|
|
uncons : StringIterator -> Maybe (Char, StringIterator)
|
|
|
|
export
|
|
unpack : String -> LazyList Char
|
|
unpack = unsafeUnfoldr uncons . fromString
|