{-# OPTIONS --cubical-compatible --guardedness #-}
module IO.Finite where
open import Data.Unit.Polymorphic.Base using (⊤)
open import Agda.Builtin.String using (String)
import Data.Unit.Base as Unit0 using (⊤)
open import IO.Base as Base using (IO; lift; lift′)
import IO.Primitive.Core as Prim using (IO; _>>=_; _>>_)
import IO.Primitive.Finite as Prim
using (getLine; readFile; writeFile; appendFile; putStr; putStrLn)
open import Level using (Level; Lift; 0ℓ; suc)
private
variable
a : Level
getLine : IO String
getLine = lift Prim.getLine
readFile : String → IO String
readFile f = lift (Prim.readFile f)
writeFile : String → String → IO {a} ⊤
writeFile f s = lift′ (Prim.writeFile f s)
appendFile : String → String → IO {a} ⊤
appendFile f s = lift′ (Prim.appendFile f s)
putStr : String → IO {a} ⊤
putStr s = lift′ (Prim.putStr s)
putStrLn : String → IO {a} ⊤
putStrLn s = lift′ (Prim.putStrLn s)