mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-24 06:52:19 +03:00
Update stringOps.c (#848)
Co-authored-by: Guillaume ALLAIS <guillaume.allais@ens-lyon.org>
This commit is contained in:
parent
d076bcd649
commit
d04b28b62d
@ -14,14 +14,25 @@ Value *head(Value *str)
|
||||
return (Value *)c;
|
||||
}
|
||||
|
||||
Value *tail(Value *str)
|
||||
Value *tail(Value *input)
|
||||
{
|
||||
Value_Char *c = (Value_Char *)newValue();
|
||||
c->header.tag = CHAR_TAG;
|
||||
Value_String *s = (Value_String *)str;
|
||||
Value_String *tailStr = (Value_String *)newValue();
|
||||
tailStr->header.tag = STRING_TAG;
|
||||
Value_String *s = (Value_String *)input;
|
||||
int l = strlen(s->str);
|
||||
c->c = s->str[l - 1];
|
||||
return (Value *)c;
|
||||
if(l != 0)
|
||||
{
|
||||
tailStr->str = malloc(l);
|
||||
memset(tailStr->str, 0, l);
|
||||
memcpy(tailStr->str, s->str + 1, l - 1);
|
||||
return (Value *)tailStr;
|
||||
}
|
||||
else
|
||||
{
|
||||
tailStr->str = malloc(1);
|
||||
tailStr->str[0] = '\0';
|
||||
return (Value *)tailStr;
|
||||
}
|
||||
}
|
||||
|
||||
Value *reverse(Value *str)
|
||||
|
@ -225,10 +225,11 @@ pathLookup names = do
|
||||
||| Some test may involve Idris' backends and have requirements.
|
||||
||| We define here the ones supported by Idris
|
||||
public export
|
||||
data Requirement = Chez | Node | Racket
|
||||
data Requirement = C | Chez | Node | Racket
|
||||
|
||||
export
|
||||
Show Requirement where
|
||||
show C = "C"
|
||||
show Chez = "Chez"
|
||||
show Node = "node"
|
||||
show Racket = "racket"
|
||||
@ -242,6 +243,7 @@ checkRequirement req
|
||||
|
||||
where
|
||||
requirement : Requirement -> (String, List String)
|
||||
requirement C = ("CC", ["cc"])
|
||||
requirement Chez = ("CHEZ", ["chez", "chezscheme9.5", "scheme", "scheme.exe"])
|
||||
requirement Node = ("NODE", ["node"])
|
||||
requirement Racket = ("RACKET", ["racket"])
|
||||
|
@ -171,6 +171,10 @@ chezTests = MkTestPool [Chez]
|
||||
"perf001",
|
||||
"reg001"]
|
||||
|
||||
refcTests : TestPool
|
||||
refcTests = MkTestPool [C]
|
||||
[ "refc001" ]
|
||||
|
||||
racketTests : TestPool
|
||||
racketTests = MkTestPool [Racket]
|
||||
["concurrency001"]
|
||||
@ -216,6 +220,7 @@ main = runner
|
||||
, testPaths "ideMode" ideModeTests
|
||||
, testPaths "prelude" preludeTests
|
||||
, testPaths "chez" chezTests
|
||||
, testPaths "refc" refcTests
|
||||
, testPaths "racket" racketTests
|
||||
, testPaths "node" nodeTests
|
||||
, testPaths "templates" templateTests
|
||||
|
8
tests/refc/refc001/Tail.idr
Normal file
8
tests/refc/refc001/Tail.idr
Normal file
@ -0,0 +1,8 @@
|
||||
module Tail
|
||||
|
||||
import Data.Strings
|
||||
|
||||
main : IO ()
|
||||
main = assert_total $ do
|
||||
putStrLn $ strTail ""
|
||||
putStrLn $ strTail " Hello World!"
|
2
tests/refc/refc001/expected
Normal file
2
tests/refc/refc001/expected
Normal file
@ -0,0 +1,2 @@
|
||||
|
||||
Hello World!
|
4
tests/refc/refc001/run
Normal file
4
tests/refc/refc001/run
Normal file
@ -0,0 +1,4 @@
|
||||
$1 --no-banner --no-color --console-width 0 --cg refc -o refc001 Tail.idr > /dev/null
|
||||
./build/exec/refc001
|
||||
|
||||
rm -rf build
|
Loading…
Reference in New Issue
Block a user