mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-24 20:23:11 +03:00
[ fix ] fastConcat for JS backends
This commit is contained in:
parent
1877e66309
commit
62237f74ea
@ -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
|
||||
|
||||
|
@ -271,6 +271,7 @@ nodeTests = MkTestPool "Node backend" [] (Just Node)
|
||||
, "bitops"
|
||||
, "casts"
|
||||
, "memo"
|
||||
, "fastConcat"
|
||||
, "newints"
|
||||
, "reg001", "reg002"
|
||||
, "stringcast"
|
||||
|
9
tests/node/fastConcat/FastConcat.idr
Normal file
9
tests/node/fastConcat/FastConcat.idr
Normal 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 "")
|
4
tests/node/fastConcat/expected
Normal file
4
tests/node/fastConcat/expected
Normal file
@ -0,0 +1,4 @@
|
||||
1/1: Building FastConcat (FastConcat.idr)
|
||||
Main> 1000000
|
||||
|
||||
Main> Bye for now!
|
2
tests/node/fastConcat/input
Normal file
2
tests/node/fastConcat/input
Normal file
@ -0,0 +1,2 @@
|
||||
:exec main
|
||||
:q
|
4
tests/node/fastConcat/run
Normal file
4
tests/node/fastConcat/run
Normal file
@ -0,0 +1,4 @@
|
||||
rm -rf build
|
||||
|
||||
$1 --cg node --no-banner --no-color --console-width 0 FastConcat.idr < input
|
||||
|
Loading…
Reference in New Issue
Block a user