Make fastAppend a foreign call.

This commit is contained in:
Matus Tejiscak 2020-06-11 21:06:56 +02:00
parent ef5299733a
commit 362d2204ab
4 changed files with 21 additions and 6 deletions

View File

@ -16,15 +16,12 @@ partial
foldl1 : (a -> a -> a) -> List a -> a
foldl1 f (x::xs) = foldl f x xs
-- This works quickly because when string-append builds the result, it allocates
-- This works quickly because when string-concat builds the result, it allocates
-- enough room in advance so there's only one allocation, rather than lots!
%foreign
"scheme:string-concat"
export
fastAppend : List String -> String
fastAppend xs = unsafePerformIO (schemeCall String "string-append" (toFArgs xs))
where
toFArgs : List String -> FArgList
toFArgs [] = []
toFArgs (x :: xs) = x :: toFArgs xs
||| Splits a character list into a list of whitespace separated character lists.
|||

View File

@ -61,6 +61,12 @@
(define cast-string-double
(lambda (x)
(cast-num (string->number (destroy-prefix x)))))
(define (from-idris-list xs)
(if (= (vector-ref xs 0) 0)
'()
(cons (vector-ref xs 1) (from-idris-list (vector-ref xs 2)))))
(define (string-concat xs) (apply string-append (from-idris-list xs)))
(define string-cons (lambda (x y) (string-append (string x) y)))
(define get-tag (lambda (x) (vector-ref x 0)))
(define string-reverse (lambda (x)

View File

@ -68,6 +68,13 @@
(define-macro (cast-string-int x)
`(floor (cast-string-double ,x)))
(define (from-idris-list xs)
(if (= (vector-ref xs 0) 0)
'()
(cons (vector-ref xs 1) (from-idris-list (vector-ref xs 2)))))
(define-macro (string-concat xs)
`(apply string-append (from-idris-list ,xs)))
(define-macro (string-cons x y)
`(string-append (string ,x) ,y))

View File

@ -55,6 +55,11 @@
(define cast-string-double
(lambda (x)
(cast-num (string->number (destroy-prefix x)))))
(define (from-idris-list xs)
(if (= (vector-ref xs 0) 0)
'()
(cons (vector-ref xs 1) (from-idris-list (vector-ref xs 2)))))
(define (string-concat xs) (apply string-append (from-idris-list xs)))
(define string-cons (lambda (x y) (string-append (string x) y)))
(define get-tag (lambda (x) (vector-ref x 0)))
(define string-reverse (lambda (x)