2020-05-18 15:59:07 +03:00
|
|
|
module Control.App.Console
|
|
|
|
|
|
|
|
import public Control.App
|
|
|
|
|
|
|
|
public export
|
|
|
|
interface Console e where
|
2020-05-30 13:41:13 +03:00
|
|
|
putChar : Char -> App {l} e ()
|
2020-05-18 15:59:07 +03:00
|
|
|
putStr : String -> App {l} e ()
|
2020-05-30 13:41:13 +03:00
|
|
|
getChar : App {l} e Char
|
|
|
|
getLine : App {l} e String
|
2020-05-18 15:59:07 +03:00
|
|
|
|
2021-01-16 10:03:45 +03:00
|
|
|
export
|
2020-05-18 15:59:07 +03:00
|
|
|
PrimIO e => Console e where
|
2020-05-30 13:41:13 +03:00
|
|
|
putChar c = primIO $ putChar c
|
2020-05-18 15:59:07 +03:00
|
|
|
putStr str = primIO $ putStr str
|
2020-05-30 13:41:13 +03:00
|
|
|
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")
|
2020-05-30 13:41:13 +03:00
|
|
|
|
|
|
|
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 ()
|
2021-01-16 10:03:45 +03:00
|
|
|
printLn x = putStrLn $ show x
|