diff --git a/exe/Main.idr b/exe/Main.idr index be03398..f67aaf2 100644 --- a/exe/Main.idr +++ b/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 ----------------------------------- diff --git a/exe/Options.idr b/exe/Options.idr index 7698730..8707fe3 100644 --- a/exe/Options.idr +++ b/exe/Options.idr @@ -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] diff --git a/lib/Quox/Untyped/Scheme.idr b/lib/Quox/Untyped/Scheme.idr new file mode 100644 index 0000000..b68edc3 --- /dev/null +++ b/lib/Quox/Untyped/Scheme.idr @@ -0,0 +1,276 @@ +module Quox.Untyped.Scheme + +import Quox.Name +import Quox.Context +import Quox.Untyped.Syntax +import Quox.Pretty + +import Quox.EffExtra +import Quox.CharExtra +import Data.String +import Data.SortedSet +import Data.Vect +import Data.List1 +import Derive.Prelude + +%default total +%language ElabReflection + +%hide TT.Name + + +public export +data Id = I String Nat +%runElab derive "Id" [Eq, Ord] + +public export +Scheme : List (Type -> Type) +Scheme = [State (SortedSet Id)] + + +public export +data Sexp = + V Id +| L (List Sexp) +| Q Sexp +| N Nat +| Lambda (List Id) Sexp +| Let (List (Id, Sexp)) Sexp +| Case Sexp (List1 (List Sexp, Sexp)) +| Define Id Sexp + +export +FromString Sexp where fromString s = V $ I s 0 + + +private +makeIdBase : Mods -> String -> String +makeIdBase mods str = joinBy "." $ toList $ mods :< str + +export +makeId : Name -> Id +makeId (MakeName mods (UN str)) = I (makeIdBase mods str) 0 +makeId (MakeName mods (MN str k)) = I (makeIdBase mods str) (S k) +makeId (MakeName mods Unused) = I (makeIdBase mods "_") 0 + +export +makeIdB : BindName -> Id +makeIdB (BN name _) = makeId $ MakeName [<] name + +private +bump : Id -> Id +bump (I x i) = I x (S i) + +export covering +getFresh : SortedSet Id -> Id -> Id +getFresh used x = + if contains x used then getFresh used (bump x) else x + +export covering +freshIn : Id -> (Id -> Eff Scheme a) -> Eff Scheme a +freshIn x k = + let x = getFresh !get x in + local (insert x) $ k x + +export covering +freshInB : BindName -> (Id -> Eff Scheme a) -> Eff Scheme a +freshInB x = freshIn (makeIdB x) + +export covering +freshInBN : Vect n BindName -> (Vect n Id -> Eff Scheme a) -> + Eff Scheme a +freshInBN xs act = do + let (xs', used') = go (map makeIdB xs) !get + local_ used' $ act xs' +where + go : forall n. Vect n Id -> SortedSet Id -> (Vect n Id, SortedSet Id) + go [] used = ([], used) + go (x :: xs) used = + let x = getFresh used x + (xs, used) = go xs (insert x used) + in + (x :: xs, used) + +export covering +toScheme : Context' Id n -> Term n -> Eff Scheme Sexp +toScheme xs (F x loc) = pure $ V $ makeId x + +toScheme xs (B i loc) = pure $ V $ xs !!! i + +toScheme xs (Lam x body loc) = + freshInB x $ \x => + pure $ Lambda [x] !(toScheme (xs :< x) body) + +toScheme xs (App fun arg loc) = + pure $ L [!(toScheme xs fun), !(toScheme xs arg)] + +toScheme xs (Pair fst snd loc) = + pure $ L ["cons", !(toScheme xs fst), !(toScheme xs snd)] + +toScheme xs (Fst pair loc) = + pure $ L ["car", !(toScheme xs pair)] + +toScheme xs (Snd pair loc) = + pure $ L ["cdr", !(toScheme xs pair)] + +toScheme xs (Tag tag loc) = + pure $ Q $ fromString tag + +toScheme xs (CaseEnum tag cases loc) = + Case <$> toScheme xs tag + <*> for cases (\(t, rhs) => ([fromString t],) <$> toScheme xs rhs) + +toScheme xs (Absurd loc) = + pure $ Q "absurd" + +toScheme xs (Zero loc) = + pure $ N 0 + +toScheme xs (Succ nat loc) = + case !(toScheme xs nat) of + N n => pure $ N n + s => pure $ L ["1+", s] + +toScheme xs (CaseNat nat zer (NSRec p ih suc) loc) = + freshInBN [p, ih] $ \[p, ih] => + pure $ + L ["case-nat-rec", + Lambda [] !(toScheme xs zer), + Lambda [p, ih] !(toScheme (xs :< p :< ih) suc), + !(toScheme xs nat)] + +toScheme xs (CaseNat nat zer (NSNonrec p suc) loc) = + freshInB p $ \p => + pure $ + L ["case-nat-nonrec", + Lambda [] !(toScheme xs zer), + Lambda [p] !(toScheme (xs :< p) suc), + !(toScheme xs nat)] + +toScheme xs (Let x rhs body loc) = + freshInB x $ \x => + pure $ Let [(x, !(toScheme xs rhs))] !(toScheme (xs :< x) body) + +toScheme xs (Erased loc) = + pure $ Q "erased" + + +export +prelude : String +prelude = """ +(define (case-nat-rec z s n) + (if (= n 0) + (z) + (let* [(p (1- n)) + (ih (case-nat-rec z s p))] + (s p ih)))) + +(define (case-nat-nonrec z s n) + (if (= n 0) (z) (s (1- n)))) + +""" + +export covering +defToScheme : Name -> Definition -> Eff Scheme (Maybe Sexp) +defToScheme x ErasedDef = pure Nothing +defToScheme x (KeptDef def) = do + let x = makeId x + modify $ insert x + pure $ Just $ Define x !(toScheme [<] def) + + +export +isSchemeInitial : Char -> Bool +isSchemeInitial c = + let gc = genCat c in + isLetter gc || isSymbol gc && c /= '|' || + gc == Number Letter || + gc == Number Other || + gc == Mark NonSpacing || + gc == Punctuation Dash || + gc == Punctuation Connector || + gc == Punctuation Other && c /= '\'' && c /= '\\' || + gc == Other PrivateUse || + (c `elem` unpack "!$%&*/:<=>?~_^") + +export +isSchemeSubsequent : Char -> Bool +isSchemeSubsequent c = + let gc = genCat c in + isSchemeInitial c || + isNumber gc || + isMark gc || + (c `elem` unpack ".+-@") + +export +isSchemeId : String -> Bool +isSchemeId str = + str == "1+" || str == "1-" || + case unpack str of + [] => False + c :: cs => isSchemeInitial c && all isSchemeSubsequent cs + +export +escId : String -> String +escId str = + let str' = concatMap doEsc $ unpack str in + if isSchemeId str' then str' else "|\{str}|" +where + doEsc : Char -> String + doEsc '\\' = "\\\\" + doEsc '|' = "\\|" + doEsc '\'' = "^" + doEsc c = singleton c + +prettyId' : {opts : LayoutOpts} -> Id -> Doc opts +prettyId' (I str 0) = text $ escId str +prettyId' (I str k) = text $ escId "\{str}:\{show k}" + +prettyId : {opts : LayoutOpts} -> Id -> Eff Pretty (Doc opts) +prettyId x = hl TVar $ prettyId' x + +orIndent : {opts : LayoutOpts} -> Doc opts -> Doc opts -> Doc opts +orIndent a b = + parens $ ifMultiline (a <++> b) (a `vappend` indent 2 b) + +export covering +prettySexp : {opts : LayoutOpts} -> Sexp -> Eff Pretty (Doc opts) +prettySexp (V x) = prettyId x +prettySexp (L []) = hl Delim "()" +prettySexp (L (x :: xs)) = do + d <- prettySexp x + ds <- Prelude.traverse prettySexp xs + parens $ ifMultiline (hsep $ d :: ds) (vsep $ d :: map (indent 2) ds) +prettySexp (Q (V x)) = hl Tag $ "'" <+> prettyId' x +prettySexp (Q x) = pure $ hcat [!(hl Tag "'"), !(prettySexp x)] +prettySexp (N n) = hl Tag $ pshow n +prettySexp (Lambda xs e) = + pure $ orIndent + (hsep [!(hl Syntax "lambda"), !(prettySexp $ L $ map V xs)]) + !(prettySexp e) +prettySexp (Let ps e) = + pure $ orIndent + (hsep [!(hl Syntax "let"), !(bracks . sep =<< traverse prettyBind ps)]) + !(prettySexp e) +where + prettyBind : (Id, Sexp) -> Eff Pretty (Doc opts) + prettyBind (x, e) = pure $ sep [!(prettyId x), !(prettySexp e)] +prettySexp (Case h as) = do + header' <- prettySexp h + case_ <- caseD + let header = ifMultiline (case_ <++> header') + (case_ `vappend` indent 2 header') + bodys <- traverse prettyCase $ toList as + pure $ ifMultiline + (parens $ header <++> hsep bodys) + (parens $ header `vappend` indent 2 (vsep bodys)) +where + prettyCase : (List Sexp, Sexp) -> Eff Pretty (Doc opts) + prettyCase (ps, e) = bracks $ + ifMultiline + (hsep [!(parens . hsep =<< traverse prettySexp ps), !(prettySexp e)]) + (vsep [!(parens . sep =<< traverse prettySexp ps), !(prettySexp e)]) +prettySexp (Define x e) = + pure $ orIndent + (hsep [!(hl Syntax "define"), !(prettyId x)]) + !(prettySexp e) diff --git a/lib/quox-lib.ipkg b/lib/quox-lib.ipkg index 27063b1..f5d188a 100644 --- a/lib/quox-lib.ipkg +++ b/lib/quox-lib.ipkg @@ -61,4 +61,5 @@ modules = Quox.Parser.FromParser.Error, Quox.Parser, Quox.Untyped.Syntax, - Quox.Untyped.Erase + Quox.Untyped.Erase, + Quox.Untyped.Scheme