Idris2/libs/base/Control/App/Console.idr

34 lines
731 B
Idris
Raw Normal View History

2020-05-18 15:59:07 +03:00
module Control.App.Console
import public Control.App
public export
interface Console e where
putChar : Char -> App {l} e ()
2020-05-18 15:59:07 +03:00
putStr : String -> App {l} e ()
getChar : App {l} e Char
getLine : App {l} e String
2020-05-18 15:59:07 +03:00
export
2020-05-18 15:59:07 +03:00
PrimIO e => Console e where
putChar c = primIO $ putChar c
2020-05-18 15:59:07 +03:00
putStr str = primIO $ putStr str
getChar = primIO getChar
getLine = primIO getLine
2020-05-18 15:59:07 +03:00
export
putStrLn : Console e => String -> App {l} e ()
putStrLn str = putStr (str ++ "\n")
export
putCharLn : Console e => Char -> App {l} e ()
putCharLn c = putStrLn $ strCons c ""
export
print : (Console e, Show a) => a -> App {l} e ()
print x = putStr $ show x
export
printLn : (Console e, Show a) => a -> App {l} e ()
printLn x = putStrLn $ show x