mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-12 18:50:17 +03:00
Merge pull request #1736 from stepancheg/test-discovery
Implement test discovery
This commit is contained in:
commit
5576d30c27
@ -1,5 +1,6 @@
|
||||
module System.Directory
|
||||
|
||||
import System.Errno
|
||||
import public System.File
|
||||
|
||||
%default total
|
||||
@ -97,9 +98,40 @@ removeDir : HasIO io => String -> io ()
|
||||
removeDir dirName = primIO (prim__removeDir dirName)
|
||||
|
||||
export
|
||||
dirEntry : HasIO io => Directory -> io (Either FileError String)
|
||||
dirEntry (MkDir d)
|
||||
nextDirEntry : HasIO io => Directory -> io (Either FileError (Maybe String))
|
||||
nextDirEntry (MkDir d)
|
||||
= do res <- primIO (prim__dirEntry d)
|
||||
if prim__nullPtr res /= 0
|
||||
then returnError
|
||||
else ok (prim__getString res)
|
||||
then if !(getErrno) /= 0
|
||||
then returnError
|
||||
else pure $ Right Nothing
|
||||
else do let n = prim__getString res
|
||||
if n == "." || n == ".."
|
||||
then assert_total $ nextDirEntry (MkDir d)
|
||||
else pure $ Right (Just n)
|
||||
|
||||
-- This function is deprecated; to be removed after the next version bump
|
||||
export
|
||||
dirEntry : HasIO io => Directory -> io (Either FileError String)
|
||||
dirEntry d = do r <- nextDirEntry d
|
||||
pure $ case r of
|
||||
Left e => Left e
|
||||
Right (Just n) => Right n
|
||||
Right Nothing => Left FileNotFound
|
||||
|
||||
collectDir : HasIO io => Directory -> io (Either FileError (List String))
|
||||
collectDir d
|
||||
= liftIO $ do let (>>=) : (IO . Either e) a -> (a -> (IO . Either e) b) -> (IO . Either e) b
|
||||
(>>=) = Prelude.(>>=) @{Monad.Compose {m = IO} {t = Either e}}
|
||||
Just n <- nextDirEntry d
|
||||
| Nothing => pure $ Right []
|
||||
ns <- assert_total $ collectDir d
|
||||
pure $ Right (n :: ns)
|
||||
|
||||
export
|
||||
listDir : HasIO io => String -> io (Either FileError (List String))
|
||||
listDir name = do Right d <- openDir name
|
||||
| Left e => pure $ Left e
|
||||
ns <- collectDir d
|
||||
ignore <- closeDir d
|
||||
pure $ ns
|
||||
|
@ -380,6 +380,34 @@ record TestPool where
|
||||
codegen : Codegen
|
||||
testCases : List String
|
||||
|
||||
||| Find all the test in the given directory.
|
||||
export
|
||||
testsInDir : (dirName : String) -> (testNameFilter : String -> Bool) -> (poolName : String) -> List Requirement -> Codegen -> IO TestPool
|
||||
testsInDir dirName testNameFilter poolName reqs cg = do
|
||||
Right names <- listDir dirName
|
||||
| Left e => do putStrLn ("failed to list " ++ dirName ++ ": " ++ show e)
|
||||
exitFailure
|
||||
let names = [n | n <- names, testNameFilter n]
|
||||
let testNames = [dirName ++ "/" ++ n | n <- names]
|
||||
testNames <- filter testNames
|
||||
when (length testNames == 0) $ do
|
||||
putStrLn ("no tests found in " ++ dirName)
|
||||
exitFailure
|
||||
pure $ MkTestPool poolName reqs cg testNames
|
||||
where
|
||||
-- Directory without `run` file is not a test
|
||||
isTest : (path : String) -> IO Bool
|
||||
isTest path = exists (path ++ "/run")
|
||||
|
||||
filter : (testPaths : List String) -> IO (List String)
|
||||
filter [] = pure []
|
||||
filter (p :: ps) =
|
||||
do rem <- filter ps
|
||||
case !(isTest p) of
|
||||
True => pure $ p :: rem
|
||||
False => pure rem
|
||||
|
||||
|
||||
||| Only keep the tests that have been asked for
|
||||
export
|
||||
filterTests : Options -> List String -> List String
|
||||
|
@ -60,6 +60,10 @@ int idris2_removeDir(char* path) {
|
||||
|
||||
char* idris2_nextDirEntry(void* d) {
|
||||
DirInfo* di = (DirInfo*)d;
|
||||
// `readdir` keeps `errno` unchanged on end of stream
|
||||
// so we need to reset `errno` to distinguish between
|
||||
// end of stream and failure.
|
||||
errno = 0;
|
||||
struct dirent* de = readdir(di->dirptr);
|
||||
|
||||
if (de == NULL) {
|
||||
|
@ -1,5 +1,9 @@
|
||||
module Main
|
||||
|
||||
import System
|
||||
import System.Directory
|
||||
import System.File
|
||||
|
||||
import Test.Golden
|
||||
|
||||
%default covering
|
||||
@ -199,12 +203,8 @@ idrisTests = MkTestPool "Misc" [] Nothing
|
||||
-- golden file testing
|
||||
"golden001"]
|
||||
|
||||
typeddTests : TestPool
|
||||
typeddTests = MkTestPool "Type Driven Development" [] Nothing
|
||||
[ "chapter01", "chapter02", "chapter03", "chapter04", "chapter05"
|
||||
, "chapter06", "chapter07", "chapter08", "chapter09", "chapter10"
|
||||
, "chapter11", "chapter12", "chapter13", "chapter14"
|
||||
]
|
||||
typeddTests : IO TestPool
|
||||
typeddTests = testsInDir "typedd-book" (const True) "Type Driven Development" [] Nothing
|
||||
|
||||
chezTests : TestPool
|
||||
chezTests = MkTestPool "Chez backend" [] (Just Chez)
|
||||
@ -226,12 +226,8 @@ chezTests = MkTestPool "Chez backend" [] (Just Chez)
|
||||
, "channels001", "channels002", "channels003", "channels004", "channels005"
|
||||
]
|
||||
|
||||
refcTests : TestPool
|
||||
refcTests = MkTestPool "Reference counting C backend" [] (Just C)
|
||||
[ "refc001" , "refc002"
|
||||
, "strings", "integers", "doubles"
|
||||
, "buffer", "clock", "args"
|
||||
]
|
||||
refcTests : IO TestPool
|
||||
refcTests = testsInDir "refc" (const True) "Reference counting C backend" [] (Just C)
|
||||
|
||||
racketTests : TestPool
|
||||
racketTests = MkTestPool "Racket backend" [] (Just Racket)
|
||||
@ -264,56 +260,32 @@ nodeTests = MkTestPool "Node backend" [] (Just Node)
|
||||
, "integers"
|
||||
]
|
||||
|
||||
vmcodeInterpTests : TestPool
|
||||
vmcodeInterpTests = MkTestPool "VMCode interpreter" [] Nothing
|
||||
[ "basic001"
|
||||
]
|
||||
vmcodeInterpTests : IO TestPool
|
||||
vmcodeInterpTests = testsInDir "vmcode" (const True) "VMCode interpreter" [] Nothing
|
||||
|
||||
ideModeTests : TestPool
|
||||
ideModeTests = MkTestPool "IDE mode" [] Nothing
|
||||
[ "ideMode001", "ideMode002", "ideMode003", "ideMode004", "ideMode005"
|
||||
]
|
||||
ideModeTests : IO TestPool
|
||||
ideModeTests = testsInDir "ideMode" (const True) "IDE mode" [] Nothing
|
||||
|
||||
preludeTests : TestPool
|
||||
preludeTests = MkTestPool "Prelude library" [] Nothing
|
||||
[ "reg001"
|
||||
]
|
||||
preludeTests : IO TestPool
|
||||
preludeTests = testsInDir "prelude" (const True) "Prelude library" [] Nothing
|
||||
|
||||
templateTests : TestPool
|
||||
templateTests = MkTestPool "Test templates" [] Nothing
|
||||
[ "simple-test", "ttimp", "with-ipkg"
|
||||
]
|
||||
templateTests : IO TestPool
|
||||
templateTests = testsInDir "templates" (const True) "Test templates" [] Nothing
|
||||
|
||||
-- base library tests are run against
|
||||
-- each codegen supported and to keep
|
||||
-- things simple it's all one test group
|
||||
-- that only runs if all backends are
|
||||
-- available.
|
||||
baseLibraryTests : TestPool
|
||||
baseLibraryTests = MkTestPool "Base library" [Chez, Node] Nothing
|
||||
[ "control_app001"
|
||||
, "system_file001"
|
||||
, "system_info_os001"
|
||||
, "system_system"
|
||||
, "data_bits001"
|
||||
, "data_string_lines001"
|
||||
, "data_string_unlines001"
|
||||
, "system_errno"
|
||||
, "system_info001"
|
||||
, "system_signal001", "system_signal002", "system_signal003", "system_signal004"
|
||||
]
|
||||
baseLibraryTests : IO TestPool
|
||||
baseLibraryTests = testsInDir "base" (const True) "Base library" [Chez, Node] Nothing
|
||||
|
||||
-- same behavior as `baseLibraryTests`
|
||||
contribLibraryTests : TestPool
|
||||
contribLibraryTests = MkTestPool "Contrib library" [Chez, Node] Nothing
|
||||
[ "json_001"
|
||||
]
|
||||
contribLibraryTests : IO TestPool
|
||||
contribLibraryTests = testsInDir "contrib" (const True) "Contrib library" [Chez, Node] Nothing
|
||||
|
||||
codegenTests : TestPool
|
||||
codegenTests = MkTestPool "Code generation" [] Nothing
|
||||
[ "con001"
|
||||
, "builtin001"
|
||||
]
|
||||
codegenTests : IO TestPool
|
||||
codegenTests = testsInDir "codegen" (const True) "Code generation" [] Nothing
|
||||
|
||||
main : IO ()
|
||||
main = runner $
|
||||
@ -333,18 +305,18 @@ main = runner $
|
||||
, testPaths "idris2" idrisTestsBuiltin
|
||||
, testPaths "idris2" idrisTestsEvaluator
|
||||
, testPaths "idris2" idrisTests
|
||||
, testPaths "typedd-book" typeddTests
|
||||
, testPaths "ideMode" ideModeTests
|
||||
, testPaths "prelude" preludeTests
|
||||
, testPaths "base" baseLibraryTests
|
||||
, testPaths "contrib" contribLibraryTests
|
||||
, !typeddTests
|
||||
, !ideModeTests
|
||||
, !preludeTests
|
||||
, !baseLibraryTests
|
||||
, !contribLibraryTests
|
||||
, testPaths "chez" chezTests
|
||||
, testPaths "refc" refcTests
|
||||
, !refcTests
|
||||
, testPaths "racket" racketTests
|
||||
, testPaths "node" nodeTests
|
||||
, testPaths "vmcode" vmcodeInterpTests
|
||||
, testPaths "templates" templateTests
|
||||
, testPaths "codegen" codegenTests
|
||||
, !vmcodeInterpTests
|
||||
, !templateTests
|
||||
, !codegenTests
|
||||
]
|
||||
++ map (testPaths "allbackends" . idrisTestsAllBackends) [Chez, Node, Racket]
|
||||
|
||||
|
12
tests/base/system_directory/ReadDir.idr
Normal file
12
tests/base/system_directory/ReadDir.idr
Normal file
@ -0,0 +1,12 @@
|
||||
import System
|
||||
import System.Directory
|
||||
|
||||
panic : String -> IO a
|
||||
panic s = do putStrLn s
|
||||
exitFailure
|
||||
|
||||
main : IO ()
|
||||
main = do Right ["a"] <- listDir "dir"
|
||||
| Left e => panic (show e)
|
||||
| Right x => panic ("wrong entries: " ++ (show x))
|
||||
pure ()
|
0
tests/base/system_directory/dir/a
Normal file
0
tests/base/system_directory/dir/a
Normal file
2
tests/base/system_directory/expected
Normal file
2
tests/base/system_directory/expected
Normal file
@ -0,0 +1,2 @@
|
||||
1/1: Building ReadDir (ReadDir.idr)
|
||||
Main> Main> Bye for now!
|
2
tests/base/system_directory/input
Normal file
2
tests/base/system_directory/input
Normal file
@ -0,0 +1,2 @@
|
||||
:exec main
|
||||
:q
|
3
tests/base/system_directory/run
Executable file
3
tests/base/system_directory/run
Executable file
@ -0,0 +1,3 @@
|
||||
rm -rf build
|
||||
|
||||
$1 --no-color --console-width 0 --no-banner ReadDir.idr < input
|
Loading…
Reference in New Issue
Block a user