[ fix ] fastConcat for JS backends

This commit is contained in:
stefan-hoeck 2021-10-14 12:34:41 +00:00 committed by G. Allais
parent 1877e66309
commit 62237f74ea
6 changed files with 22 additions and 2 deletions

View File

@ -464,7 +464,7 @@ Traversable List where
%foreign
"scheme:string-concat"
"RefC:fastConcat"
"javascript:lambda:(xs)=>''.concat(...__prim_idris2js_array(xs))"
"javascript:lambda:(xs)=>__prim_idris2js_array(xs).join('')"
export
fastConcat : List String -> String
@ -592,7 +592,7 @@ pack (x :: xs) = strCons x (pack xs)
%foreign
"scheme:string-pack"
"RefC:fastPack"
"javascript:lambda:(xs)=>''.concat(...__prim_idris2js_array(xs))"
"javascript:lambda:(xs)=>__prim_idris2js_array(xs).join('')"
export
fastPack : List Char -> String

View File

@ -271,6 +271,7 @@ nodeTests = MkTestPool "Node backend" [] (Just Node)
, "bitops"
, "casts"
, "memo"
, "fastConcat"
, "newints"
, "reg001", "reg002"
, "stringcast"

View File

@ -0,0 +1,9 @@
replicateTR : Nat -> a -> List a
replicateTR n v = go n Nil
where go : Nat -> List a -> List a
go 0 as = as
go (S k) as = go k (v :: as)
main : IO ()
main = printLn (length $ fastPack $ replicateTR 1000000 '0')
>> putStrLn (fastConcat $ replicateTR 1000000 "")

View File

@ -0,0 +1,4 @@
1/1: Building FastConcat (FastConcat.idr)
Main> 1000000
Main> Bye for now!

View File

@ -0,0 +1,2 @@
:exec main
:q

View File

@ -0,0 +1,4 @@
rm -rf build
$1 --cg node --no-banner --no-color --console-width 0 FastConcat.idr < input