mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-24 06:52:19 +03:00
[ fix #454 ] compiling nonexisting file
This commit is contained in:
parent
acda3b44a9
commit
62a5406533
@ -173,14 +173,17 @@ stMain cgs opts
|
||||
then findIpkg fname
|
||||
else pure fname
|
||||
setMainFile fname
|
||||
the (Core ()) $ case fname of
|
||||
Nothing => logTime "Loading prelude" $
|
||||
result <- case fname of
|
||||
Nothing => logTime "Loading prelude" $ do
|
||||
when (not $ noprelude session) $
|
||||
readPrelude True
|
||||
Just f => logTime "Loading main file" $
|
||||
(loadMainFile f >>= displayErrors)
|
||||
pure Done
|
||||
Just f => logTime "Loading main file" $ do
|
||||
res <- loadMainFile f
|
||||
displayErrors res
|
||||
pure res
|
||||
|
||||
doRepl <- postOptions opts
|
||||
doRepl <- postOptions result opts
|
||||
if doRepl then
|
||||
if ide || ideSocket then
|
||||
if not ideSocket
|
||||
|
@ -125,23 +125,26 @@ postOptions : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto o : Ref ROpts REPLOpts} ->
|
||||
List CLOpt -> Core Bool
|
||||
postOptions [] = pure True
|
||||
postOptions (OutputFile outfile :: rest)
|
||||
REPLResult -> List CLOpt -> Core Bool
|
||||
postOptions _ [] = pure True
|
||||
postOptions res@(ErrorLoadingFile _ _) (OutputFile _ :: rest)
|
||||
= do postOptions res rest
|
||||
pure False
|
||||
postOptions res (OutputFile outfile :: rest)
|
||||
= do compileExp (PRef (MkFC "(script)" (0, 0) (0, 0)) (UN "main")) outfile
|
||||
postOptions rest
|
||||
postOptions res rest
|
||||
pure False
|
||||
postOptions (ExecFn str :: rest)
|
||||
postOptions res (ExecFn str :: rest)
|
||||
= do execExp (PRef (MkFC "(script)" (0, 0) (0, 0)) (UN str))
|
||||
postOptions rest
|
||||
postOptions res rest
|
||||
pure False
|
||||
postOptions (CheckOnly :: rest)
|
||||
= do postOptions rest
|
||||
postOptions res (CheckOnly :: rest)
|
||||
= do postOptions res rest
|
||||
pure False
|
||||
postOptions (RunREPL str :: rest)
|
||||
postOptions res (RunREPL str :: rest)
|
||||
= do replCmd str
|
||||
pure False
|
||||
postOptions (_ :: rest) = postOptions rest
|
||||
postOptions res (_ :: rest) = postOptions res rest
|
||||
|
||||
export
|
||||
ideMode : List CLOpt -> Bool
|
||||
|
@ -50,7 +50,7 @@ idrisTests
|
||||
-- Error messages
|
||||
"error001", "error002", "error003", "error004", "error005",
|
||||
"error006", "error007", "error008", "error009", "error010",
|
||||
"error011",
|
||||
"error011", "error012",
|
||||
-- Modules and imports
|
||||
"import001", "import002", "import003", "import004", "import005",
|
||||
-- Interactive editing support
|
||||
|
1
tests/idris2/error012/expected
Normal file
1
tests/idris2/error012/expected
Normal file
@ -0,0 +1 @@
|
||||
File error in nothere.idr: File Not Found
|
1
tests/idris2/error012/run
Executable file
1
tests/idris2/error012/run
Executable file
@ -0,0 +1 @@
|
||||
$1 nothere.idr -o argh
|
Loading…
Reference in New Issue
Block a user