Idris2/libs/base/System/REPL.idr
Edwin Brady 28855088c2 Split HasIO into HasIO and MonadIO
For things which don't require (>>=), HasIO is fine, otherwise MonadIO
gives access to the monad interface.
2020-06-21 14:46:14 +01:00

38 lines
1.1 KiB
Idris

module System.REPL
import System.File
||| A basic read-eval-print loop, maintaining a state
||| @ state the input state
||| @ prompt the prompt to show
||| @ onInput the function to run on reading input, returning a String to
||| output and a new state. Returns Nothing if the repl should exit
export
replWith : MonadIO io =>
(state : a) -> (prompt : String) ->
(onInput : a -> String -> Maybe (String, a)) -> io ()
replWith acc prompt fn
= do eof <- fEOF stdin
if eof
then pure ()
else do putStr prompt
fflush stdout
x <- getLine
case fn acc x of
Just (out, acc') =>
do putStr out
replWith acc' prompt fn
Nothing => pure ()
||| A basic read-eval-print loop
||| @ prompt the prompt to show
||| @ onInput the function to run on reading input, returning a String to
||| output
export
repl : MonadIO io =>
(prompt : String) -> (onInput : String -> String) -> io ()
repl prompt fn
= replWith () prompt (\x, s => Just (fn s, ()))