add postulate, #[compile-scheme], #[main]
This commit is contained in:
parent
cc0bade747
commit
050346e344
14 changed files with 579 additions and 321 deletions
|
@ -46,6 +46,7 @@ data Error =
|
|||
| WrapTypeError TypeError
|
||||
| Postulate Loc Name
|
||||
| WhileErasing Name Q.Definition Error
|
||||
| MainIsErased Loc Name
|
||||
%name Error err
|
||||
|
||||
private %inline
|
||||
|
@ -58,6 +59,7 @@ Located Error where
|
|||
(WrapTypeError err).loc = err.loc
|
||||
(Postulate loc _).loc = loc
|
||||
(WhileErasing _ def e).loc = e.loc `or` def.loc
|
||||
(MainIsErased loc _).loc = loc
|
||||
|
||||
|
||||
parameters {opts : LayoutOpts} (showContext : Bool)
|
||||
|
@ -71,9 +73,11 @@ parameters {opts : LayoutOpts} (showContext : Bool)
|
|||
prettyErrorNoLoc showContext err
|
||||
prettyErrorNoLoc (Postulate _ x) =
|
||||
pure $ sep [!(prettyFree x), "is a postulate with no definition"]
|
||||
prettyErrorNoLoc (WhileErasing name def err) = pure $
|
||||
vsep [hsep ["while erasing the definition", !(prettyFree name)],
|
||||
prettyErrorNoLoc (WhileErasing x def err) = pure $
|
||||
vsep [hsep ["while erasing the definition", !(prettyFree x)],
|
||||
!(prettyErrorNoLoc err)]
|
||||
prettyErrorNoLoc (MainIsErased _ x) =
|
||||
pure $ hsep [!(prettyFree x), "is marked #[main] but is erased"]
|
||||
|
||||
export
|
||||
prettyError : Error -> Eff Pretty (Doc opts)
|
||||
|
@ -485,10 +489,16 @@ trimLets (Erased loc) = Erased loc
|
|||
|
||||
export covering
|
||||
eraseDef : Name -> Q.Definition -> Eff Erase U.Definition
|
||||
eraseDef name def@(MkDef qty type body loc) =
|
||||
eraseDef name def@(MkDef qty type body scheme isMain loc) =
|
||||
wrapErr (WhileErasing name def) $
|
||||
case isErased qty.qty of
|
||||
Erased => pure ErasedDef
|
||||
Kept => case body of
|
||||
Postulate => throw $ Postulate loc name
|
||||
Concrete body => KeptDef . trimLets <$> eraseTerm empty type body
|
||||
Erased => do
|
||||
when isMain $ throw $ MainIsErased loc name
|
||||
pure ErasedDef
|
||||
Kept =>
|
||||
case scheme of
|
||||
Just str => pure $ SchemeDef isMain str
|
||||
Nothing => case body of
|
||||
Postulate => throw $ Postulate loc name
|
||||
Concrete body => KeptDef isMain . trimLets <$>
|
||||
eraseTerm empty type body
|
||||
|
|
|
@ -7,10 +7,11 @@ import Quox.Pretty
|
|||
|
||||
import Quox.EffExtra
|
||||
import Quox.CharExtra
|
||||
import Data.DPair
|
||||
import Data.List1
|
||||
import Data.String
|
||||
import Data.SortedSet
|
||||
import Data.Vect
|
||||
import Data.List1
|
||||
import Derive.Prelude
|
||||
|
||||
%default total
|
||||
|
@ -19,165 +20,6 @@ import Derive.Prelude
|
|||
%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
|
||||
|
@ -222,55 +64,306 @@ where
|
|||
doEsc '\'' = "^"
|
||||
doEsc c = singleton c
|
||||
|
||||
|
||||
public export
|
||||
data Id = I String Nat
|
||||
%runElab derive "Id" [Eq, Ord]
|
||||
|
||||
private
|
||||
prettyId' : {opts : LayoutOpts} -> Id -> Doc opts
|
||||
prettyId' (I str 0) = text $ escId str
|
||||
prettyId' (I str k) = text $ escId "\{str}:\{show k}"
|
||||
|
||||
export
|
||||
prettyId : {opts : LayoutOpts} -> Id -> Eff Pretty (Doc opts)
|
||||
prettyId x = hl TVar $ prettyId' x
|
||||
|
||||
|
||||
public export
|
||||
data StateTag = AVOID | MAIN
|
||||
|
||||
public export
|
||||
Scheme : List (Type -> Type)
|
||||
Scheme = [StateL AVOID (SortedSet Id), StateL MAIN (List Id)]
|
||||
-- names to avoid, and functions with #[main] (should only be one)
|
||||
|
||||
|
||||
public export
|
||||
data Sexp =
|
||||
V Id
|
||||
| L (List Sexp)
|
||||
| Q Sexp
|
||||
| N Nat
|
||||
| Lambda (List Id) Sexp
|
||||
| LambdaC (List Id) Sexp -- curried lambda
|
||||
| Let Id Sexp Sexp
|
||||
| Case Sexp (List1 (List Sexp, Sexp))
|
||||
| Define Id Sexp
|
||||
| Literal String
|
||||
|
||||
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) 0
|
||||
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 !(getAt AVOID) x in
|
||||
localAt AVOID (insert x) $ k x
|
||||
|
||||
export covering
|
||||
freshInB : BindName -> (Id -> Eff Scheme a) -> Eff Scheme a
|
||||
freshInB x = freshIn (makeIdB x)
|
||||
|
||||
export covering
|
||||
freshInBT : Telescope' BindName m n ->
|
||||
(Telescope' Id m n -> Eff Scheme a) ->
|
||||
Eff Scheme a
|
||||
freshInBT xs act = do
|
||||
let (xs', used') = go (map makeIdB xs) !(getAt AVOID)
|
||||
localAt_ AVOID used' $ act xs'
|
||||
where
|
||||
go : forall n. Telescope' Id m n ->
|
||||
SortedSet Id -> (Telescope' Id m n, SortedSet Id)
|
||||
go [<] used = ([<], used)
|
||||
go (xs :< x) used =
|
||||
let x = getFresh used x
|
||||
(xs, used) = go xs (insert x used)
|
||||
in
|
||||
(xs :< x, used)
|
||||
|
||||
export covering
|
||||
freshInBC : Context' BindName n -> (Context' Id n -> Eff Scheme a) ->
|
||||
Eff Scheme a
|
||||
freshInBC = freshInBT
|
||||
|
||||
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) =
|
||||
let Evidence n' (ys, body) = splitLam [< x] body in
|
||||
freshInBT ys $ \ys => do
|
||||
pure $ LambdaC (toList' ys) !(toScheme (xs . ys) body)
|
||||
|
||||
toScheme xs (App fun arg loc) = do
|
||||
let (fun, args) = splitApp fun
|
||||
fun <- toScheme xs fun
|
||||
args <- traverse (toScheme xs) args
|
||||
arg <- toScheme xs arg
|
||||
pure $ if null args
|
||||
then L [fun, arg]
|
||||
else L $ "%" :: fun :: toList (args :< 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 $ S n
|
||||
s => pure $ L ["+", s, N 1]
|
||||
|
||||
toScheme xs (CaseNat nat zer (NSRec p ih suc) loc) =
|
||||
freshInBC [< 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 = """
|
||||
#!r6rs
|
||||
|
||||
; curried lambda
|
||||
(define-syntax lambda%
|
||||
(syntax-rules ()
|
||||
[(_ (x0 x1 ...) body ...)
|
||||
(lambda (x0) (lambda% (x1 ...) body ...))]
|
||||
[(_ () body ...)
|
||||
(begin body ...)]))
|
||||
|
||||
; curried application
|
||||
(define-syntax %
|
||||
(syntax-rules ()
|
||||
[(_ e0 e1 e2 ...)
|
||||
(% (e0 e1) e2 ...)]
|
||||
[(_ e) e]))
|
||||
|
||||
; curried function definition
|
||||
(define-syntax define%
|
||||
(syntax-rules ()
|
||||
[(_ (f x ...) body ...)
|
||||
(define f (lambda% (x ...) body ...))]
|
||||
[(_ x body ...)
|
||||
(define x body ...)]))
|
||||
|
||||
(define-syntax builtin-io
|
||||
(syntax-rules ()
|
||||
[(_ body ...)
|
||||
(lambda (s)
|
||||
(let [(res (begin body ...))]
|
||||
(cons res s)))]))
|
||||
|
||||
(define (case-nat-rec z s n)
|
||||
(let go [(acc (z)) (i 0)]
|
||||
(if (= i n) acc (go (s i acc) (+ i 1)))))
|
||||
|
||||
(define (case-nat-nonrec z s n)
|
||||
(if (= n 0) (z) (s (- n 1))))
|
||||
|
||||
(define (run-main f) (f 'io-state) (void))
|
||||
;;;;;;
|
||||
"""
|
||||
|
||||
export covering
|
||||
defToScheme : Name -> Definition -> Eff Scheme (Maybe Sexp)
|
||||
defToScheme x ErasedDef = pure Nothing
|
||||
defToScheme x (KeptDef isMain def) = do
|
||||
let x = makeId x
|
||||
when isMain $ modifyAt MAIN (x ::)
|
||||
modifyAt AVOID $ insert x
|
||||
pure $ Just $ Define x !(toScheme [<] def)
|
||||
defToScheme x (SchemeDef isMain str) = do
|
||||
let x = makeId x
|
||||
when isMain $ modifyAt MAIN (x ::)
|
||||
modifyAt AVOID $ insert x
|
||||
pure $ Just $ Define x $ Literal str
|
||||
|
||||
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)
|
||||
|
||||
private covering
|
||||
prettyLambda : {opts : LayoutOpts} ->
|
||||
String -> List Id -> Sexp -> Eff Pretty (Doc opts)
|
||||
prettyLambda lam xs e =
|
||||
pure $ orIndent
|
||||
(hsep [!(hl Syntax $ text lam), !(prettySexp $ L $ map V xs)])
|
||||
!(prettySexp e)
|
||||
|
||||
private covering
|
||||
prettyBind : {opts : LayoutOpts} -> (Id, Sexp) -> Eff Pretty (Doc opts)
|
||||
prettyBind (x, e) = parens $ sep [!(prettyId x), !(prettySexp e)]
|
||||
|
||||
private covering
|
||||
prettyLet : {opts : LayoutOpts} ->
|
||||
SnocList (Id, Sexp) -> Sexp -> Eff Pretty (Doc opts)
|
||||
prettyLet ps (Let x rhs body) = prettyLet (ps :< (x, rhs)) body
|
||||
prettyLet ps e =
|
||||
pure $ orIndent
|
||||
(hsep [!(hl Syntax "let*"),
|
||||
!(bracks . sep . toList =<< traverse prettyBind ps)])
|
||||
!(prettySexp e)
|
||||
|
||||
private covering
|
||||
prettyDefine : {opts : LayoutOpts} ->
|
||||
String -> Either Id (List Id) -> Sexp -> Eff Pretty (Doc opts)
|
||||
prettyDefine def xs body =
|
||||
parens $ vappend
|
||||
(hsep [!(hl Syntax $ text def),
|
||||
!(either prettyId (prettySexp . L . map V) xs)])
|
||||
(indent 2 !(prettySexp body))
|
||||
|
||||
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)
|
||||
ds <- traverse prettySexp xs
|
||||
parens $ (hsep $ d :: ds) <|> (hsep [d, vsep 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 (Lambda xs e) = prettyLambda "lambda" xs e
|
||||
prettySexp (LambdaC xs e) = prettyLambda "lambda%" xs e
|
||||
prettySexp (Let x rhs e) = prettyLet [< (x, rhs)] 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
|
||||
arms <- traverse prettyCase $ toList as
|
||||
pure $ ifMultiline
|
||||
(parens $ header <++> hsep bodys)
|
||||
(parens $ header `vappend` indent 2 (vsep bodys))
|
||||
(parens $ header <++> hsep arms)
|
||||
(parens $ vsep $ header :: map (indent 2) arms)
|
||||
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)
|
||||
prettySexp (Define x e) = case e of
|
||||
LambdaC xs e => prettyDefine "define%" (Right $ x :: xs) e
|
||||
Lambda xs e => prettyDefine "define" (Right $ x :: xs) e
|
||||
_ => prettyDefine "define" (Left x) e
|
||||
prettySexp (Literal sexp) =
|
||||
pure $ text sexp
|
||||
|
||||
export covering
|
||||
makeRunMain : {opts : LayoutOpts} -> Id -> Eff Pretty (Doc opts)
|
||||
makeRunMain x = prettySexp $ L ["run-main", V x]
|
||||
|
|
|
@ -82,7 +82,11 @@ Located (Term n) where
|
|||
|
||||
|
||||
public export
|
||||
data Definition = ErasedDef | KeptDef (Term 0)
|
||||
data Definition =
|
||||
ErasedDef
|
||||
| KeptDef Bool (Term 0)
|
||||
| SchemeDef Bool String
|
||||
-- bools are presence of #[main] flag
|
||||
|
||||
public export
|
||||
0 Definitions : Type
|
||||
|
@ -94,27 +98,33 @@ letD, inD : {opts : LayoutOpts} -> Eff Pretty (Doc opts)
|
|||
letD = hl Syntax "let"
|
||||
inD = hl Syntax "in"
|
||||
|
||||
export
|
||||
export covering
|
||||
prettyTerm : {opts : LayoutOpts} -> BContext n ->
|
||||
Term n -> Eff Pretty (Doc opts)
|
||||
|
||||
export
|
||||
export covering
|
||||
prettyArg : {opts : LayoutOpts} -> BContext n -> Term n -> Eff Pretty (Doc opts)
|
||||
prettyArg xs arg = withPrec Arg $ prettyTerm xs arg
|
||||
|
||||
export
|
||||
prettyApp' : {opts : LayoutOpts} -> BContext n -> Doc opts ->
|
||||
Term n -> Eff Pretty (Doc opts)
|
||||
prettyApp' xs fun arg =
|
||||
parensIfM App =<< do
|
||||
arg <- prettyArg xs arg
|
||||
pure $ sep [fun, arg]
|
||||
export covering
|
||||
prettyAppHead : {opts : LayoutOpts} -> BContext n ->
|
||||
Term n -> Eff Pretty (Doc opts)
|
||||
prettyAppHead xs fun = parensIfM App =<< prettyTerm xs fun
|
||||
|
||||
export
|
||||
prettyApp' : {opts : LayoutOpts} ->
|
||||
Doc opts -> SnocList (Doc opts) -> Eff Pretty (Doc opts)
|
||||
prettyApp' fun args = do
|
||||
d <- askAt INDENT
|
||||
let args = toList args
|
||||
pure $ hsep (fun :: args)
|
||||
<|> hsep [fun, vsep args]
|
||||
<|> vsep (fun :: map (indent d) args)
|
||||
|
||||
export covering
|
||||
prettyApp : {opts : LayoutOpts} -> BContext n ->
|
||||
Term n -> Term n -> Eff Pretty (Doc opts)
|
||||
prettyApp xs fun arg =
|
||||
prettyApp' xs !(withPrec App $ prettyTerm xs fun) arg
|
||||
Doc opts -> SnocList (Term n) -> Eff Pretty (Doc opts)
|
||||
prettyApp xs fun args = prettyApp' fun =<< traverse (prettyTerm xs) args
|
||||
|
||||
public export
|
||||
record PrettyCaseArm a n where
|
||||
|
@ -124,7 +134,7 @@ record PrettyCaseArm a n where
|
|||
vars : Vect len BindName
|
||||
rhs : Term (len + n)
|
||||
|
||||
export
|
||||
export covering
|
||||
prettyCase : {opts : LayoutOpts} -> BContext n ->
|
||||
(a -> Eff Pretty (Doc opts)) ->
|
||||
Term n -> List (PrettyCaseArm a n) ->
|
||||
|
@ -145,19 +155,24 @@ private
|
|||
sucPat : {opts : LayoutOpts} -> BindName -> Eff Pretty (Doc opts)
|
||||
sucPat x = pure $ !succD <++> !(prettyTBind x)
|
||||
|
||||
private
|
||||
export
|
||||
splitApp : Term b -> (Term b, SnocList (Term b))
|
||||
splitApp (App f x _) = mapSnd (:< x) $ splitApp f
|
||||
splitApp f = (f, [<])
|
||||
|
||||
export
|
||||
splitLam : Telescope' BindName a b -> Term b ->
|
||||
Exists $ \c => (Telescope' BindName a c, Term c)
|
||||
splitLam ys (Lam x body _) = splitLam (ys :< x) body
|
||||
splitLam ys t = Evidence _ (ys, t)
|
||||
|
||||
private
|
||||
export
|
||||
splitLet : Telescope (\i => (BindName, Term i)) a b -> Term b ->
|
||||
Exists $ \c => (Telescope (\i => (BindName, Term i)) a c, Term c)
|
||||
splitLet ys (Let x rhs body _) = splitLet (ys :< (x, rhs)) body
|
||||
splitLet ys t = Evidence _ (ys, t)
|
||||
|
||||
private
|
||||
private covering
|
||||
prettyLets : {opts : LayoutOpts} ->
|
||||
BContext a -> Telescope (\i => (BindName, Term i)) a b ->
|
||||
Eff Pretty (SnocList (Doc opts))
|
||||
|
@ -168,7 +183,7 @@ prettyLets xs lets = sequence $ snd $ go lets where
|
|||
go (lets :< (x, rhs)) =
|
||||
let (ys, docs) = go lets
|
||||
doc = hsep <$> sequence
|
||||
[letD, prettyTBind x, cstD, assert_total prettyTerm ys rhs, inD]
|
||||
[letD, prettyTBind x, cstD, prettyTerm ys rhs, inD]
|
||||
in
|
||||
(ys :< x, docs :< doc)
|
||||
|
||||
|
@ -180,51 +195,71 @@ sucCaseArm (NSRec x ih s) = pure $
|
|||
sucCaseArm (NSNonrec x s) = pure $
|
||||
MkPrettyCaseArm !(sucPat x) [x] s
|
||||
|
||||
private covering
|
||||
prettyNat : {opts : LayoutOpts} ->
|
||||
BContext n -> Term n -> Eff Pretty (Either Nat (Doc opts))
|
||||
prettyNat xs (Zero _) = pure $ Left 0
|
||||
prettyNat xs (Succ n _) =
|
||||
case !(withPrec Arg $ prettyNat xs n) of
|
||||
Left n => pure $ Left $ S n
|
||||
Right doc => map Right $ parensIfM App $ sep [!succD, doc]
|
||||
prettyNat xs s = map Right $ prettyTerm xs s
|
||||
|
||||
prettyTerm _ (F x _) = prettyFree x
|
||||
prettyTerm xs (B i _) = prettyTBind $ xs !!! i
|
||||
prettyTerm xs (Lam x body _) =
|
||||
parensIfM Outer =<< do
|
||||
let Evidence n' (ys, body) = splitLam [< x] body
|
||||
vars <- hsep . toList' <$> traverse prettyTBind ys
|
||||
body <- withPrec Outer $ assert_total prettyTerm (xs . ys) body
|
||||
body <- withPrec Outer $ prettyTerm (xs . ys) body
|
||||
hangDSingle (hsep [!lamD, vars, !darrowD]) body
|
||||
prettyTerm xs (App fun arg _) = prettyApp xs fun arg
|
||||
prettyTerm xs (App fun arg _) =
|
||||
let (fun, args) = splitApp fun in
|
||||
prettyApp xs !(prettyAppHead xs fun) (args :< arg)
|
||||
prettyTerm xs (Pair fst snd _) =
|
||||
parens =<< separateTight !commaD <$>
|
||||
sequence {t = List} [prettyTerm xs fst, prettyTerm xs snd]
|
||||
prettyTerm xs (Fst pair _) = prettyApp' xs !fstD pair
|
||||
prettyTerm xs (Snd pair _) = prettyApp' xs !sndD pair
|
||||
prettyTerm xs (Fst pair _) = prettyApp xs !fstD [< pair]
|
||||
prettyTerm xs (Snd pair _) = prettyApp xs !sndD [< pair]
|
||||
prettyTerm xs (Tag tag _) = prettyTag tag
|
||||
prettyTerm xs (CaseEnum tag cases _) =
|
||||
assert_total
|
||||
prettyCase xs prettyTag tag $
|
||||
map (\(t, rhs) => MkPrettyCaseArm t [] rhs) $ toList cases
|
||||
prettyTerm xs (Absurd _) = hl Syntax "absurd"
|
||||
prettyTerm xs (Zero _) = zeroD
|
||||
prettyTerm xs (Succ nat _) = prettyApp' xs !succD nat
|
||||
prettyTerm xs (Zero _) = hl Tag "0"
|
||||
prettyTerm xs (Succ nat _) =
|
||||
case !(prettyNat xs nat) of
|
||||
Left n => hl Tag $ pshow $ S n
|
||||
Right doc => prettyApp' !succD [< doc]
|
||||
prettyTerm xs (CaseNat nat zer suc _) =
|
||||
assert_total
|
||||
prettyCase xs pure nat [MkPrettyCaseArm !zeroD [] zer, !(sucCaseArm suc)]
|
||||
prettyTerm xs (Let x rhs body _) =
|
||||
parensIfM Outer =<< do
|
||||
let Evidence n' (lets, body) = splitLet [< (x, rhs)] body
|
||||
heads <- prettyLets xs lets
|
||||
body <- withPrec Outer $ assert_total prettyTerm (xs . map fst lets) body
|
||||
body <- withPrec Outer $ prettyTerm (xs . map fst lets) body
|
||||
let lines = toList $ heads :< body
|
||||
pure $ ifMultiline (hsep lines) (vsep lines)
|
||||
prettyTerm _ (Erased _) =
|
||||
hl Syntax =<< ifUnicode "⌷" "[]"
|
||||
|
||||
export
|
||||
export covering
|
||||
prettyDef : {opts : LayoutOpts} -> Name ->
|
||||
Definition -> Eff Pretty (Doc opts)
|
||||
prettyDef name ErasedDef =
|
||||
pure $ hsep [!(prettyFree name), !cstD, !(prettyTerm [<] $ Erased noLoc)]
|
||||
prettyDef name (KeptDef rhs) = do
|
||||
name <- prettyFree name
|
||||
prettyDef name (KeptDef isMain rhs) = do
|
||||
name <- prettyFree name {opts}
|
||||
eq <- cstD
|
||||
rhs <- prettyTerm [<] rhs
|
||||
hangDSingle (name <++> eq) rhs
|
||||
let header = if isMain then text "#[main]" <++> name else name
|
||||
hangDSingle (header <++> eq) rhs
|
||||
prettyDef name (SchemeDef isMain str) = do
|
||||
name <- prettyFree name {opts}
|
||||
eq <- cstD
|
||||
let rhs = text $ "scheme:" ++ str
|
||||
let header = if isMain then text "#[main]" <++> name else name
|
||||
hangDSingle (header <++> eq) rhs
|
||||
|
||||
|
||||
public export
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue