[ fix #423 ] --repl should load the main file

This commit is contained in:
Guillaume ALLAIS 2020-07-07 20:24:37 +01:00 committed by G. Allais
parent e277054fda
commit 301666b91d
7 changed files with 34 additions and 6 deletions

View File

@ -450,17 +450,21 @@ getParseErrorLoc fname (LexFail (l, c, _)) = MkFC fname (l, c) (l, c)
getParseErrorLoc fname (LitFail _) = MkFC fname (0, 0) (0, 0) -- TODO: Remove this unused case
getParseErrorLoc fname _ = replFC
-- Just load the 'Main' module, if it exists, which will involve building
-- Just load the given module, if it exists, which will involve building
-- it if necessary
runRepl : {auto c : Ref Ctxt Defs} ->
{auto s : Ref Syn SyntaxInfo} ->
{auto o : Ref ROpts REPLOpts} ->
PkgDesc ->
List CLOpt ->
Maybe String ->
Core ()
runRepl pkg opts = do
runRepl fname = do
u <- newRef UST initUState
m <- newRef MD initMetadata
case fname of
Nothing => pure ()
Just fn => do
errs <- loadMainFile fn
displayErrors errs
repl {u} {s}
@ -495,7 +499,7 @@ processPackage cmd file opts
REPL => do
[] <- build pkg opts
| errs => coreLift (exitWith (ExitFailure 1))
runRepl pkg opts
runRepl (map snd $ mainmod pkg)
record POptsFilterResult where
constructor MkPFR

View File

@ -85,7 +85,7 @@ idrisTests
"perror001", "perror002", "perror003", "perror004", "perror005",
"perror006",
-- Packages and ipkg files
"pkg001", "pkg002", "pkg003", "pkg004",
"pkg001", "pkg002", "pkg003", "pkg004", "pkg005",
-- Larger programs arising from real usage. Typically things with
-- interesting interactions between features
"real001", "real002",

View File

@ -0,0 +1,4 @@
module Foo
hi : IO String
hi = pure "Hi!"

View File

@ -0,0 +1,7 @@
1/1: Building Foo (Foo.idr)
Foo> Foo.hi : IO String
Foo> MkIO (\1 w => (MkIORes "Hi!" w))
Foo> Bye for now!
Foo> Foo.hi : IO String
Foo> MkIO (\1 w => (MkIORes "Hi!" w))
Foo> Bye for now!

View File

@ -0,0 +1,5 @@
package foo
modules = Foo
main = Foo

View File

@ -0,0 +1,3 @@
:t hi
hi
:q

5
tests/idris2/pkg005/run Executable file
View File

@ -0,0 +1,5 @@
$1 --clean foo.ipkg
$1 --repl foo.ipkg < input
$1 --repl foo.ipkg < input
rm -rf build