Skip to content

Commit

Permalink
":let" command for REPL
Browse files Browse the repository at this point in the history
  • Loading branch information
Nikita Vilunov committed Jun 7, 2020
1 parent 8af183d commit cc92338
Show file tree
Hide file tree
Showing 3 changed files with 34 additions and 0 deletions.
17 changes: 17 additions & 0 deletions src/Idris/Parser.idr
Original file line number Diff line number Diff line change
Expand Up @@ -1707,6 +1707,9 @@ data CmdArg : Type where
||| The command takes an expression.
ExprArg : CmdArg

||| The command takes a list of declarations
DeclsArg : CmdArg

||| The command takes a number.
NumberArg : CmdArg

Expand All @@ -1721,6 +1724,7 @@ Show CmdArg where
show NoArg = ""
show NameArg = "<name>"
show ExprArg = "<expr>"
show DeclsArg = "<decls>"
show NumberArg = "<number>"
show OptionArg = "<option>"
show FileArg = "<filename>"
Expand Down Expand Up @@ -1785,6 +1789,18 @@ exprArgCmd parseCmd command doc = (names, ExprArg, doc, parse)
tm <- expr pdef "(interactive)" init
pure (command tm)

declsArgCmd : ParseCmd -> (List PDecl -> REPLCmd) -> String -> CommandDefinition
declsArgCmd parseCmd command doc = (names, DeclsArg, doc, parse)
where
names : List String
names = extractNames parseCmd
parse : Rule REPLCmd
parse = do
symbol ":"
runParseCmd parseCmd
tm <- topDecl "(interactive)" init
pure (command tm)

optArgCmd : ParseCmd -> (REPLOpt -> REPLCmd) -> Bool -> String -> CommandDefinition
optArgCmd parseCmd command set doc = (names, OptionArg, doc, parse)
where
Expand Down Expand Up @@ -1845,6 +1861,7 @@ parserCommandsForHelp =
, noArgCmd (ParseREPLCmd ["m", "metavars"]) Metavars "Show remaining proof obligations (metavariables or holes)"
, noArgCmd (ParseREPLCmd ["version"]) ShowVersion "Display the Idris version"
, noArgCmd (ParseREPLCmd ["?", "h", "help"]) Help "Display this help text"
, declsArgCmd (ParseKeywordCmd "let") NewDefn "Define a new value"
]

export
Expand Down
16 changes: 16 additions & 0 deletions src/Idris/REPL.idr
Original file line number Diff line number Diff line change
Expand Up @@ -424,6 +424,7 @@ data REPLResult : Type where
OptionsSet : List REPLOpt -> REPLResult
LogLevelSet : Nat -> REPLResult
VersionIs : Version -> REPLResult
DefDeclared : REPLResult
Exited : REPLResult
Edited : EditResult -> REPLResult

Expand All @@ -443,6 +444,20 @@ execExp ctm
pure $ Executed ctm


execDecls : {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} ->
{auto s : Ref Syn SyntaxInfo} ->
{auto m : Ref MD Metadata} ->
List PDecl -> Core REPLResult
execDecls decls = do
traverse_ execDecl decls
pure DefDeclared
where
execDecl : PDecl -> Core ()
execDecl decl = do
i <- desugarDecl [] decl
traverse_ (processDecl [] (MkNested []) []) i

export
compileExp : {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} ->
Expand Down Expand Up @@ -492,6 +507,7 @@ process : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto o : Ref ROpts REPLOpts} ->
REPLCmd -> Core REPLResult
process (NewDefn decls) = execDecls decls
process (Eval itm)
= do opts <- get ROpts
case evalMode opts of
Expand Down
1 change: 1 addition & 0 deletions src/Idris/Syntax.idr
Original file line number Diff line number Diff line change
Expand Up @@ -304,6 +304,7 @@ data EditCmd : Type where

public export
data REPLCmd : Type where
NewDefn : List PDecl -> REPLCmd
Eval : PTerm -> REPLCmd
Check : PTerm -> REPLCmd
PrintDef : Name -> REPLCmd
Expand Down

0 comments on commit cc92338

Please sign in to comment.