Tune Data.String.Iterator.

This commit is contained in:
Matus Tejiscak 2020-09-20 09:56:58 +02:00
parent f73fa55075
commit d26a9c55bf
6 changed files with 60 additions and 36 deletions

View File

@ -8,7 +8,7 @@ module Data.List.Lazy
public export
data LazyList : Type -> Type where
Nil : LazyList a
(::) : a -> Lazy (LazyList a) -> LazyList a
(::) : (1 x : a) -> (1 xs : Lazy (LazyList a)) -> LazyList a
public export
Semigroup (LazyList a) where
@ -53,16 +53,3 @@ public export
traverse : Applicative f => (a -> f b) -> LazyList a -> f (List b)
traverse g [] = pure []
traverse g (x :: xs) = [| g x :: traverse g xs |]
-- This function is unsafe because it asserts
-- that the seed gets smaller in every iteration
-- but the type system does not check that.
--
-- Use cautiously.
public export
unsafeUnfoldr : (b -> Maybe (a, b)) -> b -> LazyList a
unsafeUnfoldr next seed =
case next seed of
Nothing => []
Just (x, seed') =>
x :: Delay (unsafeUnfoldr next $ assert_smaller seed seed')

View File

@ -4,19 +4,59 @@ import public Data.List.Lazy
%default total
-- Backend-dependent string iteration type.
export
data StringIterator : Type where [external]
-- This function is private
-- to avoid subverting the linearity guarantees of withString.
%foreign
"scheme:blodwen-string-iterator-new"
export
private
fromString : String -> StringIterator
-- This function uses a linear string iterator
-- so that backends can use mutating iterators.
export
withString : String -> ((1 it : StringIterator) -> a) -> a
withString str f = f (fromString str)
-- We use a custom data type instead of Maybe (Char, StringIterator)
-- to remove one level of pointer indirection
-- in every iteration of something that's likely to be a hot loop,
-- and to avoid one allocation per character.
--
-- The Char field of Character is unrestricted for flexibility.
public export
data UnconsResult : Type where
EOF : UnconsResult
Character : (c : Char) -> (1 it : StringIterator) -> UnconsResult
-- We pass the whole string to the uncons function
-- to avoid yet another allocation per character
-- because for many backends, StringIterator can be simply an integer
-- (e.g. byte offset into an UTF-8 string).
%foreign
"scheme:blodwen-string-iterator-next"
export
uncons : StringIterator -> Maybe (Char, StringIterator)
uncons : String -> (1 it : StringIterator) -> UnconsResult
export
foldl : (accTy -> Char -> accTy) -> accTy -> String -> accTy
foldl op acc str = withString str (loop acc)
where
loop : accTy -> (1 it : StringIterator) -> accTy
loop acc it =
case uncons str it of
EOF => acc
Character c it' => loop (acc `op` c) (assert_smaller it it')
export
unpack : String -> LazyList Char
unpack = unsafeUnfoldr uncons . fromString
unpack str = withString str unpack'
where
unpack' : (1 it : StringIterator) -> LazyList Char
unpack' it =
case uncons str it of
EOF => []
Character c it' => c :: Delay (unpack' $ assert_smaller it it')

View File

@ -58,7 +58,7 @@ Hashable a => Hashable (Maybe a) where
export
Hashable String where
hashWithSalt h = foldl hashWithSalt h . String.Iterator.unpack
hashWithSalt h = String.Iterator.foldl hashWithSalt h
export
Hashable Namespace where

View File

@ -87,13 +87,12 @@
(substring s b end))))
(define (blodwen-string-iterator-new s)
(cons s 0))
0)
(define (blodwen-string-iterator-next s-ofs)
(let ((s (car s-ofs)) (ofs (cdr s-ofs)))
(if (>= ofs (string-length s))
(vector 0) ; Nothing
(vector 1 (vector 0 (string-ref s ofs) (cons s (+ ofs 1)))))))
(define (blodwen-string-iterator-next s ofs)
(if (>= ofs (string-length s))
(vector 0) ; EOF
(vector 1 (string-ref s ofs) (+ ofs 1))))
(define either-left
(lambda (x)

View File

@ -99,13 +99,12 @@
(define-macro (get-tag x) `(vector-ref ,x 0))
(define (blodwen-string-iterator-new s)
(cons s 0))
0)
(define (blodwen-string-iterator-next s-ofs)
(let ((s (car s-ofs)) (ofs (cdr s-ofs)))
(if (>= ofs (string-length s))
(vector 0) ; Nothing
(vector 1 (vector 0 (string-ref s ofs) (cons s (+ ofs 1)))))))
(define (blodwen-string-iterator-next s ofs)
(if (>= ofs (string-length s))
(vector 0) ; EOF
(vector 1 (string-ref s ofs) (+ ofs 1))))
;; These two are only used in this file
(define-macro (either-left x) `(vector 0 ,x))

View File

@ -78,13 +78,12 @@
(substring s b end)))
(define (blodwen-string-iterator-new s)
(cons s 0))
0)
(define (blodwen-string-iterator-next s-ofs)
(let ((s (car s-ofs)) (ofs (cdr s-ofs)))
(if (>= ofs (string-length s))
(vector 0) ; Nothing
(vector 1 (vector 0 (string-ref s ofs) (cons s (+ ofs 1)))))))
(define (blodwen-string-iterator-next s ofs)
(if (>= ofs (string-length s))
(vector 0) ; EOF
(vector 1 (string-ref s ofs) (+ ofs 1))))
(define either-left
(lambda (x)