Skip to content

Commit

Permalink
[ test ] Add tests for file errors
Browse files Browse the repository at this point in the history
  • Loading branch information
dunhamsteve committed Mar 1, 2025
1 parent e2d1e45 commit 51ff83e
Show file tree
Hide file tree
Showing 12 changed files with 35 additions and 21 deletions.
1 change: 1 addition & 0 deletions libs/base/System/Errno.idr
Original file line number Diff line number Diff line change
Expand Up @@ -12,6 +12,7 @@ prim__getErrno : PrimIO Int
prim__strerror : Int -> PrimIO String

||| Fetch libc `errno` global variable.
||| This sometimes returns 0 on windows.
export
getErrno : HasIO io => io Int
getErrno = primIO prim__getErrno
Expand Down
4 changes: 2 additions & 2 deletions support/js/support_system_file.js
Original file line number Diff line number Diff line change
Expand Up @@ -8,7 +8,7 @@ function support_system_file_fileErrno(){
switch(n) {
// these are documented in include/uv/errno.h of libuv
case -4058: return 2
case -4092: return 3
case -4048: return 3
case -4075: return 4
default: return -n + 5
}
Expand Down Expand Up @@ -136,7 +136,7 @@ function support_system_file_popen (cmd, m) {
case "r":
io_setting = ['ignore', write_fd, 2]
break
case "w":
case "w":
case "a":
io_setting = [write_fd, 'ignore', 2]
break
Expand Down
5 changes: 4 additions & 1 deletion tests/base/system_directory/ReadDir.idr
Original file line number Diff line number Diff line change
Expand Up @@ -6,7 +6,10 @@ panic s = do putStrLn s
exitFailure

main : IO ()
main = do Right ["a"] <- listDir "dir"
main = do -- Verify that errorno does not cause subsequent listDir to fail
Left _ <- listDir "doesNotExist"
| Right _ => panic "Read of \"doesNotExist\" succeeded"
Right ["a"] <- listDir "dir"
| Left e => panic (show e)
| Right x => panic ("wrong entries: " ++ (show x))
pure ()
3 changes: 2 additions & 1 deletion tests/base/system_directory/run
Original file line number Diff line number Diff line change
@@ -1,3 +1,4 @@
. ../../testutils.sh

run ReadDir.idr
run --cg chez ReadDir.idr
run --cg node ReadDir.idr
18 changes: 18 additions & 0 deletions tests/base/system_fileError/FileErrors.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,18 @@
import System.File
import System.Directory

main : IO ()
main = do
Left FileNotFound <- openFile "doesNotExist" Read
| Right _ => putStrLn "expected FileNotFound error"
| Left err => putStrLn "expected FileNotFound, got \{show err}"

Left PermissionDenied <- openFile "read_only.txt" ReadWrite
| Right _ => putStrLn "expected PermissionDenied error"
| Left err => putStrLn "expected PermissionDenied, got \{show err}"

Left FileExists <- createDir "build"
| Right _ => putStrLn "expected FileExists error"
| Left err => putStrLn "expected FileExists, got \{show err}"

putStrLn "done"
2 changes: 2 additions & 0 deletions tests/base/system_fileError/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,2 @@
done
done
1 change: 1 addition & 0 deletions tests/base/system_fileError/read_only.txt
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
This file should be read only.
5 changes: 5 additions & 0 deletions tests/base/system_fileError/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,5 @@
. ../../testutils.sh

chmod 444 read_only.txt
run --cg chez FileErrors.idr
run --cg node FileErrors.idr
14 changes: 0 additions & 14 deletions tests/node/node029/ReadDir.idr

This file was deleted.

Empty file removed tests/node/node029/dir/a
Empty file.
Empty file removed tests/node/node029/expected
Empty file.
3 changes: 0 additions & 3 deletions tests/node/node029/run

This file was deleted.

0 comments on commit 51ff83e

Please sign in to comment.