scheme output
This commit is contained in:
parent
cd08a0fd98
commit
cc0bade747
4 changed files with 326 additions and 44 deletions
89
exe/Main.idr
89
exe/Main.idr
|
@ -6,6 +6,7 @@ import Quox.Definition as Q
|
|||
import Quox.Pretty
|
||||
import Quox.Untyped.Syntax as U
|
||||
import Quox.Untyped.Erase
|
||||
import Quox.Untyped.Scheme
|
||||
import Options
|
||||
|
||||
import Data.IORef
|
||||
|
@ -20,22 +21,11 @@ import Control.Eff
|
|||
%hide Core.(>>=)
|
||||
|
||||
|
||||
parameters {auto _ : HasIO io} (width : Nat)
|
||||
private
|
||||
putDoc : Doc (Opts width) -> io ()
|
||||
putDoc = putStr . render _
|
||||
|
||||
private
|
||||
fPutDoc : File -> Doc (Opts width) -> io (Either FileError ())
|
||||
fPutDoc h = fPutStr h . render _
|
||||
|
||||
private
|
||||
putDocErr : Doc (Opts width) -> io ()
|
||||
putDocErr = ignore . fPutDoc stderr
|
||||
|
||||
private
|
||||
die : Doc (Opts width) -> io a
|
||||
die err = do putDocErr err; exitFailure
|
||||
private
|
||||
die : HasIO io => (opts : LayoutOpts) -> Doc opts -> io a
|
||||
die opts err = do
|
||||
ignore $ fPutStr stderr $ render opts err
|
||||
exitFailure
|
||||
|
||||
private
|
||||
runPretty : Options -> Eff Pretty a -> a
|
||||
|
@ -82,6 +72,14 @@ private
|
|||
loadError : Loc -> FilePath -> FileError -> Error
|
||||
loadError loc file err = FromParserError $ LoadError loc file err
|
||||
|
||||
private
|
||||
prettyError : {opts : LayoutOpts} -> Error -> Eff Pretty (Doc opts)
|
||||
prettyError (ParseError file e) = prettyParseError file e
|
||||
prettyError (FromParserError e) = FromParser.prettyError True e
|
||||
prettyError (EraseError e) = Erase.prettyError True e
|
||||
prettyError (WriteError file e) = pure $
|
||||
hangSingle 2 (text "couldn't write file \{file}:") (pshow e)
|
||||
|
||||
private
|
||||
data CompileTag = OPTS | STATE
|
||||
|
||||
|
@ -125,21 +123,28 @@ FlexDoc = {opts : LayoutOpts} -> Doc opts
|
|||
|
||||
|
||||
private
|
||||
outputDoc : FlexDoc -> Eff Compile ()
|
||||
outputDoc doc =
|
||||
outputStr : Lazy String -> Eff Compile ()
|
||||
outputStr str =
|
||||
case !(asksAt OPTS outFile) of
|
||||
None => pure ()
|
||||
Stdout => putDoc !(asksAt OPTS width) doc
|
||||
Stdout => putStr str
|
||||
File f => do
|
||||
res <- withFile f WriteTruncate pure $ \h =>
|
||||
fPutDoc !(asksAt OPTS width) h doc
|
||||
res <- withFile f WriteTruncate pure $ \h => fPutStr h str
|
||||
rethrow $ mapFst (WriteError f) res
|
||||
|
||||
private
|
||||
outputDocStopIf : Phase -> FlexDoc -> Eff CompileStop ()
|
||||
outputDocStopIf p doc =
|
||||
when (!(asksAt OPTS until) == Just p) $ do
|
||||
lift (outputDoc doc)
|
||||
outputDocs : {opts : LayoutOpts} -> List (Doc opts) -> Eff Compile ()
|
||||
outputDocs {opts = Opts _} doc =
|
||||
outputStr $ concat $ map (render _) doc
|
||||
|
||||
private
|
||||
outputDocStopIf : Phase ->
|
||||
({opts : LayoutOpts} -> Eff Pretty (List (Doc opts))) ->
|
||||
Eff CompileStop ()
|
||||
outputDocStopIf p docs = do
|
||||
opts <- askAt OPTS
|
||||
when (opts.until == Just p) $ Prelude.do
|
||||
lift $ outputDocs (runPretty opts docs) {opts = Opts opts.width}
|
||||
stopHere
|
||||
|
||||
private
|
||||
|
@ -160,41 +165,41 @@ liftErase defs act =
|
|||
\case Ask => pure defs,
|
||||
handleStateIORef !(asksAt STATE suf)]
|
||||
|
||||
private
|
||||
liftScheme : Eff Scheme a -> Eff CompileStop a
|
||||
liftScheme act = runEff act [handleStateIORef !(newIORef empty)]
|
||||
|
||||
|
||||
private
|
||||
processFile : String -> Eff Compile ()
|
||||
processFile file = withEarlyStop $ do
|
||||
Just ast <- loadFile noLoc file
|
||||
| Nothing => pure ()
|
||||
putErrLn "checking \{file}"
|
||||
outputDocStopIf Parse $ dumpDoc ast
|
||||
-- putErrLn "checking \{file}"
|
||||
when (!(asksAt OPTS until) == Just Parse) $ do
|
||||
lift $ outputStr $ show ast
|
||||
stopHere
|
||||
defList <- liftFromParser $ concat <$> traverse fromPTopLevel ast
|
||||
outputDocStopIf Check $ runPretty !(askAt OPTS) $
|
||||
vsep <$> traverse (uncurry Q.prettyDef) defList
|
||||
outputDocStopIf Check $
|
||||
traverse (uncurry Q.prettyDef) defList
|
||||
let defs = SortedMap.fromList defList
|
||||
erased <- liftErase defs $
|
||||
traverse (\(x, d) => (x,) <$> eraseDef x d) defList
|
||||
outputDocStopIf Erase $ runPretty !(askAt OPTS) $
|
||||
vsep . catMaybes <$> traverse (uncurry U.prettyDef) erased
|
||||
outputDocStopIf Erase $
|
||||
traverse (uncurry U.prettyDef) erased
|
||||
scheme <- liftScheme $ map catMaybes $
|
||||
traverse (uncurry defToScheme) erased
|
||||
outputDocStopIf Scheme $
|
||||
(text Scheme.prelude ::) <$> traverse prettySexp scheme
|
||||
die "that's all folks"
|
||||
|
||||
private
|
||||
dieError : HasIO io => Options -> Error -> io a
|
||||
dieError opts e = do
|
||||
die opts.width $ runPretty opts $ case e of
|
||||
ParseError file e => prettyParseError file e
|
||||
FromParserError e => FromParser.prettyError True e
|
||||
EraseError e => Erase.prettyError True e
|
||||
WriteError file e => pure $
|
||||
hangSingle 2 (text "couldn't write file \{file}:") (pshow e)
|
||||
|
||||
export
|
||||
main : IO ()
|
||||
main = do
|
||||
(_, opts, files) <- options
|
||||
case !(runCompile opts !newState $ traverse_ processFile files) of
|
||||
Right () => pure ()
|
||||
Left e => dieError opts e
|
||||
Left e => die (Opts opts.width) $ runPretty opts $ prettyError e
|
||||
|
||||
|
||||
-----------------------------------
|
||||
|
|
|
@ -16,7 +16,7 @@ data OutFile = File String | Stdout | None
|
|||
%runElab derive "OutFile" [Eq, Ord, Show]
|
||||
|
||||
public export
|
||||
data Phase = Parse | Check | Erase
|
||||
data Phase = Parse | Check | Erase | Scheme
|
||||
%name Phase p
|
||||
%runElab derive "Phase" [Eq, Ord, Show]
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue