add postulate, #[compile-scheme], #[main]
This commit is contained in:
parent
cc0bade747
commit
050346e344
14 changed files with 579 additions and 321 deletions
50
exe/Main.idr
50
exe/Main.idr
|
@ -58,16 +58,19 @@ newState = pure $ MkState {
|
||||||
}
|
}
|
||||||
|
|
||||||
private
|
private
|
||||||
data Error
|
data Error =
|
||||||
= ParseError String Parser.Error
|
ParseError String Parser.Error
|
||||||
| FromParserError FromParser.Error
|
| FromParserError FromParser.Error
|
||||||
| EraseError Erase.Error
|
| EraseError Erase.Error
|
||||||
| WriteError FilePath FileError
|
| WriteError FilePath FileError
|
||||||
|
| NoMain
|
||||||
|
| MultipleMains (List Id)
|
||||||
%hide FromParser.Error
|
%hide FromParser.Error
|
||||||
%hide Erase.Error
|
%hide Erase.Error
|
||||||
%hide Lexer.Error
|
%hide Lexer.Error
|
||||||
%hide Parser.Error
|
%hide Parser.Error
|
||||||
|
|
||||||
|
|
||||||
private
|
private
|
||||||
loadError : Loc -> FilePath -> FileError -> Error
|
loadError : Loc -> FilePath -> FileError -> Error
|
||||||
loadError loc file err = FromParserError $ LoadError loc file err
|
loadError loc file err = FromParserError $ LoadError loc file err
|
||||||
|
@ -77,6 +80,10 @@ prettyError : {opts : LayoutOpts} -> Error -> Eff Pretty (Doc opts)
|
||||||
prettyError (ParseError file e) = prettyParseError file e
|
prettyError (ParseError file e) = prettyParseError file e
|
||||||
prettyError (FromParserError e) = FromParser.prettyError True e
|
prettyError (FromParserError e) = FromParser.prettyError True e
|
||||||
prettyError (EraseError e) = Erase.prettyError True e
|
prettyError (EraseError e) = Erase.prettyError True e
|
||||||
|
prettyError NoMain = pure "no #[main] function given"
|
||||||
|
prettyError (MultipleMains xs) =
|
||||||
|
pure $ sep ["multiple #[main] functions given:",
|
||||||
|
separateLoose "," !(traverse prettyId xs)]
|
||||||
prettyError (WriteError file e) = pure $
|
prettyError (WriteError file e) = pure $
|
||||||
hangSingle 2 (text "couldn't write file \{file}:") (pshow e)
|
hangSingle 2 (text "couldn't write file \{file}:") (pshow e)
|
||||||
|
|
||||||
|
@ -133,9 +140,10 @@ outputStr str =
|
||||||
rethrow $ mapFst (WriteError f) res
|
rethrow $ mapFst (WriteError f) res
|
||||||
|
|
||||||
private
|
private
|
||||||
outputDocs : {opts : LayoutOpts} -> List (Doc opts) -> Eff Compile ()
|
outputDocs : (opts : Options) ->
|
||||||
outputDocs {opts = Opts _} doc =
|
({opts : LayoutOpts} -> List (Doc opts)) -> Eff Compile ()
|
||||||
outputStr $ concat $ map (render _) doc
|
outputDocs opts doc =
|
||||||
|
outputStr $ concat $ map (render (Opts opts.width)) doc
|
||||||
|
|
||||||
private
|
private
|
||||||
outputDocStopIf : Phase ->
|
outputDocStopIf : Phase ->
|
||||||
|
@ -144,7 +152,7 @@ outputDocStopIf : Phase ->
|
||||||
outputDocStopIf p docs = do
|
outputDocStopIf p docs = do
|
||||||
opts <- askAt OPTS
|
opts <- askAt OPTS
|
||||||
when (opts.until == Just p) $ Prelude.do
|
when (opts.until == Just p) $ Prelude.do
|
||||||
lift $ outputDocs (runPretty opts docs) {opts = Opts opts.width}
|
lift $ outputDocs !(askAt OPTS) (runPretty opts docs)
|
||||||
stopHere
|
stopHere
|
||||||
|
|
||||||
private
|
private
|
||||||
|
@ -166,10 +174,19 @@ liftErase defs act =
|
||||||
handleStateIORef !(asksAt STATE suf)]
|
handleStateIORef !(asksAt STATE suf)]
|
||||||
|
|
||||||
private
|
private
|
||||||
liftScheme : Eff Scheme a -> Eff CompileStop a
|
liftScheme : Eff Scheme a -> Eff CompileStop (a, List Id)
|
||||||
liftScheme act = runEff act [handleStateIORef !(newIORef empty)]
|
liftScheme act = do
|
||||||
|
runEff [|MkPair act (getAt MAIN)|]
|
||||||
|
[handleStateIORef !(newIORef empty),
|
||||||
|
handleStateIORef !(newIORef [])]
|
||||||
|
|
||||||
|
|
||||||
|
private
|
||||||
|
oneMain : Has (Except Error) fs => List Id -> Eff fs Id
|
||||||
|
oneMain [] = throw NoMain
|
||||||
|
oneMain [x] = pure x
|
||||||
|
oneMain mains = throw $ MultipleMains mains
|
||||||
|
|
||||||
private
|
private
|
||||||
processFile : String -> Eff Compile ()
|
processFile : String -> Eff Compile ()
|
||||||
processFile file = withEarlyStop $ do
|
processFile file = withEarlyStop $ do
|
||||||
|
@ -187,11 +204,16 @@ processFile file = withEarlyStop $ do
|
||||||
traverse (\(x, d) => (x,) <$> eraseDef x d) defList
|
traverse (\(x, d) => (x,) <$> eraseDef x d) defList
|
||||||
outputDocStopIf Erase $
|
outputDocStopIf Erase $
|
||||||
traverse (uncurry U.prettyDef) erased
|
traverse (uncurry U.prettyDef) erased
|
||||||
scheme <- liftScheme $ map catMaybes $
|
(scheme, mains) <- liftScheme $ map catMaybes $
|
||||||
traverse (uncurry defToScheme) erased
|
traverse (uncurry defToScheme) erased
|
||||||
outputDocStopIf Scheme $
|
outputDocStopIf Scheme $
|
||||||
(text Scheme.prelude ::) <$> traverse prettySexp scheme
|
intersperse empty <$> traverse prettySexp scheme
|
||||||
die "that's all folks"
|
opts <- askAt OPTS
|
||||||
|
main <- oneMain mains
|
||||||
|
lift $ outputDocs opts $ intersperse empty $ runPretty opts $ do
|
||||||
|
res <- traverse prettySexp scheme
|
||||||
|
runner <- makeRunMain main
|
||||||
|
pure $ text Scheme.prelude :: res ++ [runner]
|
||||||
|
|
||||||
export
|
export
|
||||||
main : IO ()
|
main : IO ()
|
||||||
|
|
|
@ -26,18 +26,24 @@ namespace DefBody
|
||||||
public export
|
public export
|
||||||
record Definition where
|
record Definition where
|
||||||
constructor MkDef
|
constructor MkDef
|
||||||
qty : GQty
|
qty : GQty
|
||||||
type0 : Term 0 0
|
type0 : Term 0 0
|
||||||
body0 : DefBody
|
body0 : DefBody
|
||||||
loc_ : Loc
|
scheme : Maybe String
|
||||||
|
isMain : Bool
|
||||||
|
loc_ : Loc
|
||||||
|
|
||||||
public export %inline
|
public export %inline
|
||||||
mkPostulate : GQty -> (type0 : Term 0 0) -> Loc -> Definition
|
mkPostulate : GQty -> (type0 : Term 0 0) -> Maybe String -> Bool -> Loc ->
|
||||||
mkPostulate qty type0 loc_ = MkDef {qty, type0, body0 = Postulate, loc_}
|
Definition
|
||||||
|
mkPostulate qty type0 scheme isMain loc_ =
|
||||||
|
MkDef {qty, type0, body0 = Postulate, scheme, isMain, loc_}
|
||||||
|
|
||||||
public export %inline
|
public export %inline
|
||||||
mkDef : GQty -> (type0, term0 : Term 0 0) -> Loc -> Definition
|
mkDef : GQty -> (type0, term0 : Term 0 0) -> Maybe String -> Bool -> Loc ->
|
||||||
mkDef qty type0 term0 loc_ = MkDef {qty, type0, body0 = Concrete term0, loc_}
|
Definition
|
||||||
|
mkDef qty type0 term0 scheme isMain loc_ =
|
||||||
|
MkDef {qty, type0, body0 = Concrete term0, scheme, isMain, loc_}
|
||||||
|
|
||||||
export Located Definition where def.loc = def.loc_
|
export Located Definition where def.loc = def.loc_
|
||||||
export Relocatable Definition where setLoc loc = {loc_ := loc}
|
export Relocatable Definition where setLoc loc = {loc_ := loc}
|
||||||
|
@ -108,10 +114,10 @@ lookupElim0 = lookupElim
|
||||||
|
|
||||||
export
|
export
|
||||||
prettyDef : {opts : LayoutOpts} -> Name -> Definition -> Eff Pretty (Doc opts)
|
prettyDef : {opts : LayoutOpts} -> Name -> Definition -> Eff Pretty (Doc opts)
|
||||||
prettyDef name (MkDef qty type _ _) = withPrec Outer $ do
|
prettyDef name def = withPrec Outer $ do
|
||||||
qty <- prettyQty qty.qty
|
qty <- prettyQty def.qty.qty
|
||||||
dot <- dotD
|
dot <- dotD
|
||||||
name <- prettyFree name
|
name <- prettyFree name
|
||||||
colon <- colonD
|
colon <- colonD
|
||||||
type <- prettyTerm [<] [<] type
|
type <- prettyTerm [<] [<] def.type
|
||||||
hangDSingle (hsep [hcat [qty, dot, name], colon]) type
|
hangDSingle (hsep [hcat [qty, dot, name], colon]) type
|
||||||
|
|
|
@ -323,35 +323,39 @@ liftTC tc = runEff tc $ with Union.(::)
|
||||||
\g => send g]
|
\g => send g]
|
||||||
|
|
||||||
private
|
private
|
||||||
addDef : Has DefsState fs => Name -> GQty -> Term 0 0 -> Term 0 0 -> Loc ->
|
addDef : Has DefsState fs => Name -> Definition -> Eff fs NDefinition
|
||||||
Eff fs NDefinition
|
addDef name def = do
|
||||||
addDef name gqty type term loc = do
|
|
||||||
let def = mkDef gqty type term loc
|
|
||||||
modifyAt DEFS $ insert name def
|
modifyAt DEFS $ insert name def
|
||||||
pure (name, def)
|
pure (name, def)
|
||||||
|
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
fromPDef : PDefinition -> Eff FromParserPure NDefinition
|
fromPDef : PDefinition -> Maybe String -> Bool ->
|
||||||
fromPDef (MkPDef qty pname ptype pterm defLoc) = do
|
Eff FromParserPure NDefinition
|
||||||
|
fromPDef (MkPDef qty pname pbody defLoc) scheme isMain = do
|
||||||
name <- fromPBaseNameNS pname
|
name <- fromPBaseNameNS pname
|
||||||
when !(getsAt DEFS $ isJust . lookup name) $ do
|
when !(getsAt DEFS $ isJust . lookup name) $ do
|
||||||
throw $ AlreadyExists defLoc name
|
throw $ AlreadyExists defLoc name
|
||||||
gqty <- globalPQty qty.val qty.loc
|
gqty <- globalPQty qty.val qty.loc
|
||||||
let sqty = globalToSubj gqty
|
let sqty = globalToSubj gqty
|
||||||
type <- traverse fromPTerm ptype
|
case pbody of
|
||||||
term <- fromPTerm pterm
|
PConcrete ptype pterm => do
|
||||||
case type of
|
type <- traverse fromPTerm ptype
|
||||||
Just type => do
|
term <- fromPTerm pterm
|
||||||
ignore $ liftTC $ do
|
case type of
|
||||||
checkTypeC empty type Nothing
|
Just type => do
|
||||||
checkC empty sqty term type
|
ignore $ liftTC $ do
|
||||||
addDef name gqty type term defLoc
|
checkTypeC empty type Nothing
|
||||||
Nothing => do
|
checkC empty sqty term type
|
||||||
let E elim = term
|
addDef name $ mkDef gqty type term scheme isMain defLoc
|
||||||
| _ => throw $ AnnotationNeeded term.loc empty term
|
Nothing => do
|
||||||
res <- liftTC $ inferC empty sqty elim
|
let E elim = term
|
||||||
addDef name gqty res.type term defLoc
|
| _ => throw $ AnnotationNeeded term.loc empty term
|
||||||
|
res <- liftTC $ inferC empty sqty elim
|
||||||
|
addDef name $ mkDef gqty res.type term scheme isMain defLoc
|
||||||
|
PPostulate ptype => do
|
||||||
|
type <- fromPTerm ptype
|
||||||
|
addDef name $ mkPostulate gqty type scheme isMain defLoc
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
|
@ -359,31 +363,50 @@ data HasFail = NoFail | AnyFail | FailWith String
|
||||||
|
|
||||||
export
|
export
|
||||||
hasFail : List PDeclMod -> HasFail
|
hasFail : List PDeclMod -> HasFail
|
||||||
hasFail [] = NoFail
|
hasFail [] = NoFail
|
||||||
hasFail (PFail str _ :: _) = maybe AnyFail FailWith str
|
hasFail (PFail str :: _) = maybe AnyFail FailWith str
|
||||||
|
hasFail (_ :: rest) = hasFail rest
|
||||||
|
|
||||||
|
export
|
||||||
|
getScheme : List PDeclMod -> Maybe String
|
||||||
|
getScheme [] = Nothing
|
||||||
|
getScheme (PCompileScheme str :: _) = Just str
|
||||||
|
getScheme (_ :: rest) = getScheme rest
|
||||||
|
|
||||||
|
export
|
||||||
|
isMain : List PDeclMod -> Bool
|
||||||
|
isMain [] = False
|
||||||
|
isMain (PMain :: _) = True
|
||||||
|
isMain (_ :: rest) = isMain rest
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
fromPDecl : PDecl -> Eff FromParserPure (List NDefinition)
|
fromPDecl : PDecl -> Eff FromParserPure (List NDefinition)
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
fromPDeclBody : PDeclBody -> Eff FromParserPure (List NDefinition)
|
fromPDeclBody : PDeclBody -> Maybe String -> Bool -> Loc ->
|
||||||
fromPDeclBody (PDef def) = singleton <$> fromPDef def
|
Eff FromParserPure (List NDefinition)
|
||||||
fromPDeclBody (PNs ns) =
|
fromPDeclBody (PDef def) scheme isMain loc =
|
||||||
|
singleton <$> fromPDef def scheme isMain
|
||||||
|
fromPDeclBody (PNs ns) scheme isMain loc = do
|
||||||
|
when (isJust scheme) $ throw $ SchemeOnNamespace loc ns.name
|
||||||
|
when isMain $ throw $ MainOnNamespace loc ns.name
|
||||||
localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls
|
localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
expectFail : PDeclBody -> Eff FromParserPure Error
|
expectFail : PDeclBody -> Loc -> Eff FromParserPure Error
|
||||||
expectFail body =
|
expectFail body loc =
|
||||||
case fromParserPure !(getAt GEN) !(getAt DEFS) $ fromPDeclBody body of
|
let res = fromParserPure !(getAt GEN) !(getAt DEFS) $
|
||||||
|
fromPDeclBody body Nothing False loc in
|
||||||
|
case res of
|
||||||
Left err => pure err
|
Left err => pure err
|
||||||
Right _ => throw $ ExpectedFail body.loc
|
Right _ => throw $ ExpectedFail body.loc
|
||||||
|
|
||||||
|
|
||||||
fromPDecl (MkPDecl mods decl loc) = case hasFail mods of
|
fromPDecl (MkPDecl mods decl loc) = case hasFail mods of
|
||||||
NoFail => fromPDeclBody decl
|
NoFail => fromPDeclBody decl (getScheme mods) (isMain mods) loc
|
||||||
AnyFail => expectFail decl $> []
|
AnyFail => expectFail decl loc $> []
|
||||||
FailWith str => do
|
FailWith str => do
|
||||||
err <- expectFail decl
|
err <- expectFail decl loc
|
||||||
let msg = runPretty $ prettyError False err {opts = Opts 10_000} -- w/e
|
let msg = runPretty $ prettyError False err {opts = Opts 10_000} -- w/e
|
||||||
if str `isInfixOf` renderInfinite msg
|
if str `isInfixOf` renderInfinite msg
|
||||||
then pure []
|
then pure []
|
||||||
|
|
|
@ -35,6 +35,8 @@ data Error =
|
||||||
| AlreadyExists Loc Name
|
| AlreadyExists Loc Name
|
||||||
| LoadError Loc FilePath FileError
|
| LoadError Loc FilePath FileError
|
||||||
| ExpectedFail Loc
|
| ExpectedFail Loc
|
||||||
|
| SchemeOnNamespace Loc Mods
|
||||||
|
| MainOnNamespace Loc Mods
|
||||||
| WrongFail String Error Loc
|
| WrongFail String Error Loc
|
||||||
| WrapParseError String ParseError
|
| WrapParseError String ParseError
|
||||||
|
|
||||||
|
@ -123,12 +125,22 @@ parameters {opts : LayoutOpts} (showContext : Bool)
|
||||||
text $ show err]
|
text $ show err]
|
||||||
|
|
||||||
prettyError (ExpectedFail loc) = pure $
|
prettyError (ExpectedFail loc) = pure $
|
||||||
sep [!(prettyLoc loc), "expected error"]
|
vsep [!(prettyLoc loc), "expected error"]
|
||||||
|
|
||||||
|
prettyError (SchemeOnNamespace loc ns) = pure $
|
||||||
|
vsep [!(prettyLoc loc),
|
||||||
|
hsep ["namespace", !(hl Free $ text $ joinBy "." $ toList ns),
|
||||||
|
"cannot have #[compile-scheme] attached"]]
|
||||||
|
|
||||||
|
prettyError (MainOnNamespace loc ns) = pure $
|
||||||
|
vsep [!(prettyLoc loc),
|
||||||
|
hsep ["namespace", !(hl Free $ text $ joinBy "." $ toList ns),
|
||||||
|
"cannot have #[main] attached"]]
|
||||||
|
|
||||||
prettyError (WrongFail str err loc) = pure $
|
prettyError (WrongFail str err loc) = pure $
|
||||||
sep [!(prettyLoc loc),
|
vsep [!(prettyLoc loc),
|
||||||
"wrong error, expected to match", !(hl Tag $ text "\"\{str}\""),
|
"wrong error, expected to match", !(hl Tag $ text "\"\{str}\""),
|
||||||
"but got", !(prettyError err)]
|
"but got", !(prettyError err)]
|
||||||
|
|
||||||
prettyError (WrapParseError file err) =
|
prettyError (WrapParseError file err) =
|
||||||
prettyParseError file err
|
prettyParseError file err
|
||||||
|
|
|
@ -227,6 +227,9 @@ reserved =
|
||||||
Word1 "def",
|
Word1 "def",
|
||||||
Word1 "def0",
|
Word1 "def0",
|
||||||
Word "defω" `Or` Word "def#",
|
Word "defω" `Or` Word "def#",
|
||||||
|
Word1 "postulate",
|
||||||
|
Word1 "postulate0",
|
||||||
|
Word "postulateω" `Or` Word "postulate#",
|
||||||
Sym1 "=",
|
Sym1 "=",
|
||||||
Word1 "load",
|
Word1 "load",
|
||||||
Word1 "namespace"]
|
Word1 "namespace"]
|
||||||
|
|
|
@ -587,32 +587,60 @@ pragma : Grammar True a -> Grammar True a
|
||||||
pragma p = resC "#[" *> p <* mustWork (resC "]")
|
pragma p = resC "#[" *> p <* mustWork (resC "]")
|
||||||
|
|
||||||
export
|
export
|
||||||
declMod : FileName -> Grammar True PDeclMod
|
declMod : Grammar True PDeclMod
|
||||||
declMod fname = withLoc fname $ pragma $
|
declMod = pragma $
|
||||||
exactName "fail" *> [|PFail $ optional strLit|]
|
exactName "fail" *> [|PFail $ optional strLit|]
|
||||||
|
<|> exactName "compile-scheme" *> [|PCompileScheme strLit|]
|
||||||
|
<|> exactName "main" $> PMain
|
||||||
|
<|> do other <- qname
|
||||||
|
fatalError "unknown declaration flag \{show other}" {c = False}
|
||||||
|
|
||||||
export
|
export
|
||||||
decl : FileName -> Grammar True PDecl
|
decl : FileName -> Grammar True PDecl
|
||||||
|
|
||||||
||| `def` alone means `defω`
|
||| `def` alone means `defω`; same for `postulate`
|
||||||
export
|
export
|
||||||
defIntro : FileName -> Grammar True PQty
|
defIntro' : (bare, zero, omega : String) ->
|
||||||
defIntro fname =
|
(0 _ : IsReserved bare) =>
|
||||||
withLoc fname (PQ Zero <$ resC "def0")
|
(0 _ : IsReserved zero) =>
|
||||||
<|> withLoc fname (PQ Any <$ resC "defω")
|
(0 _ : IsReserved omega) =>
|
||||||
<|> do pos <- bounds $ resC "def"
|
FileName -> Grammar True PQty
|
||||||
|
defIntro' bare zero omega fname =
|
||||||
|
withLoc fname (PQ Zero <$ resC zero)
|
||||||
|
<|> withLoc fname (PQ Any <$ resC omega)
|
||||||
|
<|> do pos <- bounds $ resC bare
|
||||||
let any = PQ Any $ makeLoc fname pos.bounds
|
let any = PQ Any $ makeLoc fname pos.bounds
|
||||||
option any $ qty fname <* needRes "."
|
option any $ qty fname <* needRes "."
|
||||||
|
|
||||||
export
|
export
|
||||||
definition : FileName -> Grammar True PDefinition
|
defIntro : FileName -> Grammar True PQty
|
||||||
definition fname = withLoc fname $ do
|
defIntro = defIntro' "def" "def0" "defω"
|
||||||
|
|
||||||
|
export
|
||||||
|
postulateIntro : FileName -> Grammar True PQty
|
||||||
|
postulateIntro = defIntro' "postulate" "postulate0" "postulateω"
|
||||||
|
|
||||||
|
export
|
||||||
|
postulate : FileName -> Grammar True PDefinition
|
||||||
|
postulate fname = withLoc fname $ Core.do
|
||||||
|
qty <- postulateIntro fname
|
||||||
|
name <- baseName
|
||||||
|
type <- resC ":" *> mustWork (term fname)
|
||||||
|
pure $ MkPDef qty name $ PPostulate type
|
||||||
|
|
||||||
|
export
|
||||||
|
concrete : FileName -> Grammar True PDefinition
|
||||||
|
concrete fname = withLoc fname $ do
|
||||||
qty <- defIntro fname
|
qty <- defIntro fname
|
||||||
name <- baseName
|
name <- baseName
|
||||||
type <- optional $ resC ":" *> mustWork (term fname)
|
type <- optional $ resC ":" *> mustWork (term fname)
|
||||||
term <- needRes "=" *> mustWork (term fname)
|
term <- needRes "=" *> mustWork (term fname)
|
||||||
optRes ";"
|
optRes ";"
|
||||||
pure $ MkPDef qty name type term
|
pure $ MkPDef qty name $ PConcrete type term
|
||||||
|
|
||||||
|
export
|
||||||
|
definition : FileName -> Grammar True PDefinition
|
||||||
|
definition fname = try (postulate fname) <|> concrete fname
|
||||||
|
|
||||||
export
|
export
|
||||||
namespace_ : FileName -> Grammar True PNamespace
|
namespace_ : FileName -> Grammar True PNamespace
|
||||||
|
@ -629,7 +657,7 @@ export
|
||||||
declBody : FileName -> Grammar True PDeclBody
|
declBody : FileName -> Grammar True PDeclBody
|
||||||
declBody fname = [|PDef $ definition fname|] <|> [|PNs $ namespace_ fname|]
|
declBody fname = [|PDef $ definition fname|] <|> [|PNs $ namespace_ fname|]
|
||||||
|
|
||||||
decl fname = withLoc fname [|MkPDecl (many $ declMod fname) (declBody fname)|]
|
decl fname = withLoc fname [|MkPDecl (many declMod) (declBody fname)|]
|
||||||
|
|
||||||
export
|
export
|
||||||
load : FileName -> Grammar True PTopLevel
|
load : FileName -> Grammar True PTopLevel
|
||||||
|
|
|
@ -141,13 +141,18 @@ Located PCaseBody where
|
||||||
(CaseBox _ _ loc).loc = loc
|
(CaseBox _ _ loc).loc = loc
|
||||||
|
|
||||||
|
|
||||||
|
public export
|
||||||
|
data PBody = PConcrete (Maybe PTerm) PTerm | PPostulate PTerm
|
||||||
|
%name PBody body
|
||||||
|
%runElab derive "PBody" [Eq, Ord, Show, PrettyVal]
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
record PDefinition where
|
record PDefinition where
|
||||||
constructor MkPDef
|
constructor MkPDef
|
||||||
qty : PQty
|
qty : PQty
|
||||||
name : PBaseName
|
name : PBaseName
|
||||||
type : Maybe PTerm
|
body : PBody
|
||||||
term : PTerm
|
|
||||||
loc_ : Loc
|
loc_ : Loc
|
||||||
%name PDefinition def
|
%name PDefinition def
|
||||||
%runElab derive "PDefinition" [Eq, Ord, Show, PrettyVal]
|
%runElab derive "PDefinition" [Eq, Ord, Show, PrettyVal]
|
||||||
|
@ -156,7 +161,9 @@ export Located PDefinition where def.loc = def.loc_
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data PDeclMod =
|
data PDeclMod =
|
||||||
PFail (Maybe String) Loc
|
PFail (Maybe String)
|
||||||
|
| PCompileScheme String
|
||||||
|
| PMain
|
||||||
%name PDeclMod mod
|
%name PDeclMod mod
|
||||||
%runElab derive "PDeclMod" [Eq, Ord, Show, PrettyVal]
|
%runElab derive "PDeclMod" [Eq, Ord, Show, PrettyVal]
|
||||||
|
|
||||||
|
|
|
@ -46,6 +46,7 @@ data Error =
|
||||||
| WrapTypeError TypeError
|
| WrapTypeError TypeError
|
||||||
| Postulate Loc Name
|
| Postulate Loc Name
|
||||||
| WhileErasing Name Q.Definition Error
|
| WhileErasing Name Q.Definition Error
|
||||||
|
| MainIsErased Loc Name
|
||||||
%name Error err
|
%name Error err
|
||||||
|
|
||||||
private %inline
|
private %inline
|
||||||
|
@ -58,6 +59,7 @@ Located Error where
|
||||||
(WrapTypeError err).loc = err.loc
|
(WrapTypeError err).loc = err.loc
|
||||||
(Postulate loc _).loc = loc
|
(Postulate loc _).loc = loc
|
||||||
(WhileErasing _ def e).loc = e.loc `or` def.loc
|
(WhileErasing _ def e).loc = e.loc `or` def.loc
|
||||||
|
(MainIsErased loc _).loc = loc
|
||||||
|
|
||||||
|
|
||||||
parameters {opts : LayoutOpts} (showContext : Bool)
|
parameters {opts : LayoutOpts} (showContext : Bool)
|
||||||
|
@ -71,9 +73,11 @@ parameters {opts : LayoutOpts} (showContext : Bool)
|
||||||
prettyErrorNoLoc showContext err
|
prettyErrorNoLoc showContext err
|
||||||
prettyErrorNoLoc (Postulate _ x) =
|
prettyErrorNoLoc (Postulate _ x) =
|
||||||
pure $ sep [!(prettyFree x), "is a postulate with no definition"]
|
pure $ sep [!(prettyFree x), "is a postulate with no definition"]
|
||||||
prettyErrorNoLoc (WhileErasing name def err) = pure $
|
prettyErrorNoLoc (WhileErasing x def err) = pure $
|
||||||
vsep [hsep ["while erasing the definition", !(prettyFree name)],
|
vsep [hsep ["while erasing the definition", !(prettyFree x)],
|
||||||
!(prettyErrorNoLoc err)]
|
!(prettyErrorNoLoc err)]
|
||||||
|
prettyErrorNoLoc (MainIsErased _ x) =
|
||||||
|
pure $ hsep [!(prettyFree x), "is marked #[main] but is erased"]
|
||||||
|
|
||||||
export
|
export
|
||||||
prettyError : Error -> Eff Pretty (Doc opts)
|
prettyError : Error -> Eff Pretty (Doc opts)
|
||||||
|
@ -485,10 +489,16 @@ trimLets (Erased loc) = Erased loc
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
eraseDef : Name -> Q.Definition -> Eff Erase U.Definition
|
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) $
|
wrapErr (WhileErasing name def) $
|
||||||
case isErased qty.qty of
|
case isErased qty.qty of
|
||||||
Erased => pure ErasedDef
|
Erased => do
|
||||||
Kept => case body of
|
when isMain $ throw $ MainIsErased loc name
|
||||||
Postulate => throw $ Postulate loc name
|
pure ErasedDef
|
||||||
Concrete body => KeptDef . trimLets <$> eraseTerm empty type body
|
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.EffExtra
|
||||||
import Quox.CharExtra
|
import Quox.CharExtra
|
||||||
|
import Data.DPair
|
||||||
|
import Data.List1
|
||||||
import Data.String
|
import Data.String
|
||||||
import Data.SortedSet
|
import Data.SortedSet
|
||||||
import Data.Vect
|
import Data.Vect
|
||||||
import Data.List1
|
|
||||||
import Derive.Prelude
|
import Derive.Prelude
|
||||||
|
|
||||||
%default total
|
%default total
|
||||||
|
@ -19,165 +20,6 @@ import Derive.Prelude
|
||||||
%hide TT.Name
|
%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
|
export
|
||||||
isSchemeInitial : Char -> Bool
|
isSchemeInitial : Char -> Bool
|
||||||
|
@ -222,55 +64,306 @@ where
|
||||||
doEsc '\'' = "^"
|
doEsc '\'' = "^"
|
||||||
doEsc c = singleton c
|
doEsc c = singleton c
|
||||||
|
|
||||||
|
|
||||||
|
public export
|
||||||
|
data Id = I String Nat
|
||||||
|
%runElab derive "Id" [Eq, Ord]
|
||||||
|
|
||||||
|
private
|
||||||
prettyId' : {opts : LayoutOpts} -> Id -> Doc opts
|
prettyId' : {opts : LayoutOpts} -> Id -> Doc opts
|
||||||
prettyId' (I str 0) = text $ escId str
|
prettyId' (I str 0) = text $ escId str
|
||||||
prettyId' (I str k) = text $ escId "\{str}:\{show k}"
|
prettyId' (I str k) = text $ escId "\{str}:\{show k}"
|
||||||
|
|
||||||
|
export
|
||||||
prettyId : {opts : LayoutOpts} -> Id -> Eff Pretty (Doc opts)
|
prettyId : {opts : LayoutOpts} -> Id -> Eff Pretty (Doc opts)
|
||||||
prettyId x = hl TVar $ prettyId' x
|
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 : {opts : LayoutOpts} -> Doc opts -> Doc opts -> Doc opts
|
||||||
orIndent a b =
|
orIndent a b =
|
||||||
parens $ ifMultiline (a <++> b) (a `vappend` indent 2 b)
|
parens $ ifMultiline (a <++> b) (a `vappend` indent 2 b)
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
prettySexp : {opts : LayoutOpts} -> Sexp -> Eff Pretty (Doc opts)
|
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 (V x) = prettyId x
|
||||||
prettySexp (L []) = hl Delim "()"
|
prettySexp (L []) = hl Delim "()"
|
||||||
prettySexp (L (x :: xs)) = do
|
prettySexp (L (x :: xs)) = do
|
||||||
d <- prettySexp x
|
d <- prettySexp x
|
||||||
ds <- Prelude.traverse prettySexp xs
|
ds <- traverse prettySexp xs
|
||||||
parens $ ifMultiline (hsep $ d :: ds) (vsep $ d :: map (indent 2) ds)
|
parens $ (hsep $ d :: ds) <|> (hsep [d, vsep ds]) <|>
|
||||||
|
(vsep $ d :: map (indent 2) ds)
|
||||||
prettySexp (Q (V x)) = hl Tag $ "'" <+> prettyId' x
|
prettySexp (Q (V x)) = hl Tag $ "'" <+> prettyId' x
|
||||||
prettySexp (Q x) = pure $ hcat [!(hl Tag "'"), !(prettySexp x)]
|
prettySexp (Q x) = pure $ hcat [!(hl Tag "'"), !(prettySexp x)]
|
||||||
prettySexp (N n) = hl Tag $ pshow n
|
prettySexp (N n) = hl Tag $ pshow n
|
||||||
prettySexp (Lambda xs e) =
|
prettySexp (Lambda xs e) = prettyLambda "lambda" xs e
|
||||||
pure $ orIndent
|
prettySexp (LambdaC xs e) = prettyLambda "lambda%" xs e
|
||||||
(hsep [!(hl Syntax "lambda"), !(prettySexp $ L $ map V xs)])
|
prettySexp (Let x rhs e) = prettyLet [< (x, rhs)] e
|
||||||
!(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
|
prettySexp (Case h as) = do
|
||||||
header' <- prettySexp h
|
header' <- prettySexp h
|
||||||
case_ <- caseD
|
case_ <- caseD
|
||||||
let header = ifMultiline (case_ <++> header')
|
let header = ifMultiline (case_ <++> header')
|
||||||
(case_ `vappend` indent 2 header')
|
(case_ `vappend` indent 2 header')
|
||||||
bodys <- traverse prettyCase $ toList as
|
arms <- traverse prettyCase $ toList as
|
||||||
pure $ ifMultiline
|
pure $ ifMultiline
|
||||||
(parens $ header <++> hsep bodys)
|
(parens $ header <++> hsep arms)
|
||||||
(parens $ header `vappend` indent 2 (vsep bodys))
|
(parens $ vsep $ header :: map (indent 2) arms)
|
||||||
where
|
where
|
||||||
prettyCase : (List Sexp, Sexp) -> Eff Pretty (Doc opts)
|
prettyCase : (List Sexp, Sexp) -> Eff Pretty (Doc opts)
|
||||||
prettyCase (ps, e) = bracks $
|
prettyCase (ps, e) = bracks $
|
||||||
ifMultiline
|
ifMultiline
|
||||||
(hsep [!(parens . hsep =<< traverse prettySexp ps), !(prettySexp e)])
|
(hsep [!(parens . hsep =<< traverse prettySexp ps), !(prettySexp e)])
|
||||||
(vsep [!(parens . sep =<< traverse prettySexp ps), !(prettySexp e)])
|
(vsep [!(parens . sep =<< traverse prettySexp ps), !(prettySexp e)])
|
||||||
prettySexp (Define x e) =
|
prettySexp (Define x e) = case e of
|
||||||
pure $ orIndent
|
LambdaC xs e => prettyDefine "define%" (Right $ x :: xs) e
|
||||||
(hsep [!(hl Syntax "define"), !(prettyId x)])
|
Lambda xs e => prettyDefine "define" (Right $ x :: xs) e
|
||||||
!(prettySexp 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
|
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
|
public export
|
||||||
0 Definitions : Type
|
0 Definitions : Type
|
||||||
|
@ -94,27 +98,33 @@ letD, inD : {opts : LayoutOpts} -> Eff Pretty (Doc opts)
|
||||||
letD = hl Syntax "let"
|
letD = hl Syntax "let"
|
||||||
inD = hl Syntax "in"
|
inD = hl Syntax "in"
|
||||||
|
|
||||||
export
|
export covering
|
||||||
prettyTerm : {opts : LayoutOpts} -> BContext n ->
|
prettyTerm : {opts : LayoutOpts} -> BContext n ->
|
||||||
Term n -> Eff Pretty (Doc opts)
|
Term n -> Eff Pretty (Doc opts)
|
||||||
|
|
||||||
export
|
export covering
|
||||||
prettyArg : {opts : LayoutOpts} -> BContext n -> Term n -> Eff Pretty (Doc opts)
|
prettyArg : {opts : LayoutOpts} -> BContext n -> Term n -> Eff Pretty (Doc opts)
|
||||||
prettyArg xs arg = withPrec Arg $ prettyTerm xs arg
|
prettyArg xs arg = withPrec Arg $ prettyTerm xs arg
|
||||||
|
|
||||||
export
|
export covering
|
||||||
prettyApp' : {opts : LayoutOpts} -> BContext n -> Doc opts ->
|
prettyAppHead : {opts : LayoutOpts} -> BContext n ->
|
||||||
Term n -> Eff Pretty (Doc opts)
|
Term n -> Eff Pretty (Doc opts)
|
||||||
prettyApp' xs fun arg =
|
prettyAppHead xs fun = parensIfM App =<< prettyTerm xs fun
|
||||||
parensIfM App =<< do
|
|
||||||
arg <- prettyArg xs arg
|
|
||||||
pure $ sep [fun, arg]
|
|
||||||
|
|
||||||
export
|
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 ->
|
prettyApp : {opts : LayoutOpts} -> BContext n ->
|
||||||
Term n -> Term n -> Eff Pretty (Doc opts)
|
Doc opts -> SnocList (Term n) -> Eff Pretty (Doc opts)
|
||||||
prettyApp xs fun arg =
|
prettyApp xs fun args = prettyApp' fun =<< traverse (prettyTerm xs) args
|
||||||
prettyApp' xs !(withPrec App $ prettyTerm xs fun) arg
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
record PrettyCaseArm a n where
|
record PrettyCaseArm a n where
|
||||||
|
@ -124,7 +134,7 @@ record PrettyCaseArm a n where
|
||||||
vars : Vect len BindName
|
vars : Vect len BindName
|
||||||
rhs : Term (len + n)
|
rhs : Term (len + n)
|
||||||
|
|
||||||
export
|
export covering
|
||||||
prettyCase : {opts : LayoutOpts} -> BContext n ->
|
prettyCase : {opts : LayoutOpts} -> BContext n ->
|
||||||
(a -> Eff Pretty (Doc opts)) ->
|
(a -> Eff Pretty (Doc opts)) ->
|
||||||
Term n -> List (PrettyCaseArm a n) ->
|
Term n -> List (PrettyCaseArm a n) ->
|
||||||
|
@ -145,19 +155,24 @@ private
|
||||||
sucPat : {opts : LayoutOpts} -> BindName -> Eff Pretty (Doc opts)
|
sucPat : {opts : LayoutOpts} -> BindName -> Eff Pretty (Doc opts)
|
||||||
sucPat x = pure $ !succD <++> !(prettyTBind x)
|
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 ->
|
splitLam : Telescope' BindName a b -> Term b ->
|
||||||
Exists $ \c => (Telescope' BindName a c, Term c)
|
Exists $ \c => (Telescope' BindName a c, Term c)
|
||||||
splitLam ys (Lam x body _) = splitLam (ys :< x) body
|
splitLam ys (Lam x body _) = splitLam (ys :< x) body
|
||||||
splitLam ys t = Evidence _ (ys, t)
|
splitLam ys t = Evidence _ (ys, t)
|
||||||
|
|
||||||
private
|
export
|
||||||
splitLet : Telescope (\i => (BindName, Term i)) a b -> Term b ->
|
splitLet : Telescope (\i => (BindName, Term i)) a b -> Term b ->
|
||||||
Exists $ \c => (Telescope (\i => (BindName, Term i)) a c, Term c)
|
Exists $ \c => (Telescope (\i => (BindName, Term i)) a c, Term c)
|
||||||
splitLet ys (Let x rhs body _) = splitLet (ys :< (x, rhs)) body
|
splitLet ys (Let x rhs body _) = splitLet (ys :< (x, rhs)) body
|
||||||
splitLet ys t = Evidence _ (ys, t)
|
splitLet ys t = Evidence _ (ys, t)
|
||||||
|
|
||||||
private
|
private covering
|
||||||
prettyLets : {opts : LayoutOpts} ->
|
prettyLets : {opts : LayoutOpts} ->
|
||||||
BContext a -> Telescope (\i => (BindName, Term i)) a b ->
|
BContext a -> Telescope (\i => (BindName, Term i)) a b ->
|
||||||
Eff Pretty (SnocList (Doc opts))
|
Eff Pretty (SnocList (Doc opts))
|
||||||
|
@ -168,7 +183,7 @@ prettyLets xs lets = sequence $ snd $ go lets where
|
||||||
go (lets :< (x, rhs)) =
|
go (lets :< (x, rhs)) =
|
||||||
let (ys, docs) = go lets
|
let (ys, docs) = go lets
|
||||||
doc = hsep <$> sequence
|
doc = hsep <$> sequence
|
||||||
[letD, prettyTBind x, cstD, assert_total prettyTerm ys rhs, inD]
|
[letD, prettyTBind x, cstD, prettyTerm ys rhs, inD]
|
||||||
in
|
in
|
||||||
(ys :< x, docs :< doc)
|
(ys :< x, docs :< doc)
|
||||||
|
|
||||||
|
@ -180,51 +195,71 @@ sucCaseArm (NSRec x ih s) = pure $
|
||||||
sucCaseArm (NSNonrec x s) = pure $
|
sucCaseArm (NSNonrec x s) = pure $
|
||||||
MkPrettyCaseArm !(sucPat x) [x] s
|
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 _ (F x _) = prettyFree x
|
||||||
prettyTerm xs (B i _) = prettyTBind $ xs !!! i
|
prettyTerm xs (B i _) = prettyTBind $ xs !!! i
|
||||||
prettyTerm xs (Lam x body _) =
|
prettyTerm xs (Lam x body _) =
|
||||||
parensIfM Outer =<< do
|
parensIfM Outer =<< do
|
||||||
let Evidence n' (ys, body) = splitLam [< x] body
|
let Evidence n' (ys, body) = splitLam [< x] body
|
||||||
vars <- hsep . toList' <$> traverse prettyTBind ys
|
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
|
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 _) =
|
prettyTerm xs (Pair fst snd _) =
|
||||||
parens =<< separateTight !commaD <$>
|
parens =<< separateTight !commaD <$>
|
||||||
sequence {t = List} [prettyTerm xs fst, prettyTerm xs snd]
|
sequence {t = List} [prettyTerm xs fst, prettyTerm xs snd]
|
||||||
prettyTerm xs (Fst pair _) = prettyApp' xs !fstD pair
|
prettyTerm xs (Fst pair _) = prettyApp xs !fstD [< pair]
|
||||||
prettyTerm xs (Snd pair _) = prettyApp' xs !sndD pair
|
prettyTerm xs (Snd pair _) = prettyApp xs !sndD [< pair]
|
||||||
prettyTerm xs (Tag tag _) = prettyTag tag
|
prettyTerm xs (Tag tag _) = prettyTag tag
|
||||||
prettyTerm xs (CaseEnum tag cases _) =
|
prettyTerm xs (CaseEnum tag cases _) =
|
||||||
assert_total
|
|
||||||
prettyCase xs prettyTag tag $
|
prettyCase xs prettyTag tag $
|
||||||
map (\(t, rhs) => MkPrettyCaseArm t [] rhs) $ toList cases
|
map (\(t, rhs) => MkPrettyCaseArm t [] rhs) $ toList cases
|
||||||
prettyTerm xs (Absurd _) = hl Syntax "absurd"
|
prettyTerm xs (Absurd _) = hl Syntax "absurd"
|
||||||
prettyTerm xs (Zero _) = zeroD
|
prettyTerm xs (Zero _) = hl Tag "0"
|
||||||
prettyTerm xs (Succ nat _) = prettyApp' xs !succD nat
|
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 _) =
|
prettyTerm xs (CaseNat nat zer suc _) =
|
||||||
assert_total
|
|
||||||
prettyCase xs pure nat [MkPrettyCaseArm !zeroD [] zer, !(sucCaseArm suc)]
|
prettyCase xs pure nat [MkPrettyCaseArm !zeroD [] zer, !(sucCaseArm suc)]
|
||||||
prettyTerm xs (Let x rhs body _) =
|
prettyTerm xs (Let x rhs body _) =
|
||||||
parensIfM Outer =<< do
|
parensIfM Outer =<< do
|
||||||
let Evidence n' (lets, body) = splitLet [< (x, rhs)] body
|
let Evidence n' (lets, body) = splitLet [< (x, rhs)] body
|
||||||
heads <- prettyLets xs lets
|
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
|
let lines = toList $ heads :< body
|
||||||
pure $ ifMultiline (hsep lines) (vsep lines)
|
pure $ ifMultiline (hsep lines) (vsep lines)
|
||||||
prettyTerm _ (Erased _) =
|
prettyTerm _ (Erased _) =
|
||||||
hl Syntax =<< ifUnicode "⌷" "[]"
|
hl Syntax =<< ifUnicode "⌷" "[]"
|
||||||
|
|
||||||
export
|
export covering
|
||||||
prettyDef : {opts : LayoutOpts} -> Name ->
|
prettyDef : {opts : LayoutOpts} -> Name ->
|
||||||
Definition -> Eff Pretty (Doc opts)
|
Definition -> Eff Pretty (Doc opts)
|
||||||
prettyDef name ErasedDef =
|
prettyDef name ErasedDef =
|
||||||
pure $ hsep [!(prettyFree name), !cstD, !(prettyTerm [<] $ Erased noLoc)]
|
pure $ hsep [!(prettyFree name), !cstD, !(prettyTerm [<] $ Erased noLoc)]
|
||||||
prettyDef name (KeptDef rhs) = do
|
prettyDef name (KeptDef isMain rhs) = do
|
||||||
name <- prettyFree name
|
name <- prettyFree name {opts}
|
||||||
eq <- cstD
|
eq <- cstD
|
||||||
rhs <- prettyTerm [<] rhs
|
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
|
public export
|
||||||
|
|
|
@ -2,7 +2,7 @@ module Tests.Equal
|
||||||
|
|
||||||
import Quox.Equal
|
import Quox.Equal
|
||||||
import Quox.Typechecker
|
import Quox.Typechecker
|
||||||
import Quox.ST
|
import Control.Monad.ST
|
||||||
import public TypingImpls
|
import public TypingImpls
|
||||||
import TAP
|
import TAP
|
||||||
import Quox.EffExtra
|
import Quox.EffExtra
|
||||||
|
|
|
@ -403,36 +403,53 @@ tests = "parser" :- [
|
||||||
"definitions" :- [
|
"definitions" :- [
|
||||||
parseMatch definition "defω x : {a} × {b} = ('a, 'b);"
|
parseMatch definition "defω x : {a} × {b} = ('a, 'b);"
|
||||||
`(MkPDef (PQ Any _) "x"
|
`(MkPDef (PQ Any _) "x"
|
||||||
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
(PConcrete
|
||||||
(Pair (Tag "a" _) (Tag "b" _) _) _),
|
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
||||||
|
(Pair (Tag "a" _) (Tag "b" _) _)) _),
|
||||||
parseMatch definition "def# x : {a} ** {b} = ('a, 'b)"
|
parseMatch definition "def# x : {a} ** {b} = ('a, 'b)"
|
||||||
`(MkPDef (PQ Any _) "x"
|
`(MkPDef (PQ Any _) "x"
|
||||||
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
(PConcrete
|
||||||
(Pair (Tag "a" _) (Tag "b" _) _) _),
|
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
||||||
|
(Pair (Tag "a" _) (Tag "b" _) _)) _),
|
||||||
parseMatch definition "def ω.x : {a} × {b} = ('a, 'b)"
|
parseMatch definition "def ω.x : {a} × {b} = ('a, 'b)"
|
||||||
`(MkPDef (PQ Any _) "x"
|
`(MkPDef (PQ Any _) "x"
|
||||||
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
(PConcrete
|
||||||
(Pair (Tag "a" _) (Tag "b" _) _) _),
|
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
||||||
|
(Pair (Tag "a" _) (Tag "b" _) _)) _),
|
||||||
parseMatch definition "def x : {a} × {b} = ('a, 'b)"
|
parseMatch definition "def x : {a} × {b} = ('a, 'b)"
|
||||||
`(MkPDef (PQ Any _) "x"
|
`(MkPDef (PQ Any _) "x"
|
||||||
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
(PConcrete
|
||||||
(Pair (Tag "a" _) (Tag "b" _) _) _),
|
(Just (Sig (Unused _) (Enum ["a"] _) (Enum ["b"] _) _))
|
||||||
|
(Pair (Tag "a" _) (Tag "b" _) _)) _),
|
||||||
parseMatch definition "def0 A : ★⁰ = {a, b, c}"
|
parseMatch definition "def0 A : ★⁰ = {a, b, c}"
|
||||||
`(MkPDef (PQ Zero _) "A" (Just $ TYPE 0 _)
|
`(MkPDef (PQ Zero _) "A"
|
||||||
(Enum ["a", "b", "c"] _) _)
|
(PConcrete (Just $ TYPE 0 _) (Enum ["a", "b", "c"] _)) _),
|
||||||
|
parseMatch definition "postulate yeah : ℕ"
|
||||||
|
`(MkPDef (PQ Any _) "yeah" (PPostulate (Nat _)) _),
|
||||||
|
parseMatch definition "postulateω yeah : ℕ"
|
||||||
|
`(MkPDef (PQ Any _) "yeah" (PPostulate (Nat _)) _),
|
||||||
|
parseMatch definition "postulate0 FileHandle : ★"
|
||||||
|
`(MkPDef (PQ Zero _) "FileHandle" (PPostulate (TYPE 0 _)) _),
|
||||||
|
parseFails definition "postulate not-a-postulate : ℕ = 69",
|
||||||
|
parseFails definition "postulate not-a-postulate = 69",
|
||||||
|
parseFails definition "def not-a-def : ℕ"
|
||||||
],
|
],
|
||||||
|
|
||||||
"top level" :- [
|
"top level" :- [
|
||||||
parseMatch input "def0 A : ★⁰ = {}; def0 B : ★¹ = A;"
|
parseMatch input "def0 A : ★⁰ = {}; def0 B : ★¹ = A;"
|
||||||
`([PD $ MkPDecl []
|
`([PD $ MkPDecl []
|
||||||
(PDef $ MkPDef (PQ Zero _) "A" (Just $ TYPE 0 _) (Enum [] _) _) _,
|
(PDef $ MkPDef (PQ Zero _) "A"
|
||||||
|
(PConcrete (Just $ TYPE 0 _) (Enum [] _)) _) _,
|
||||||
PD $ MkPDecl []
|
PD $ MkPDecl []
|
||||||
(PDef $ MkPDef (PQ Zero _) "B" (Just $ TYPE 1 _) (V "A" {}) _) _]),
|
(PDef $ MkPDef (PQ Zero _) "B"
|
||||||
|
(PConcrete (Just $ TYPE 1 _) (V "A" {})) _) _]),
|
||||||
parseMatch input "def0 A : ★⁰ = {} def0 B : ★¹ = A" $
|
parseMatch input "def0 A : ★⁰ = {} def0 B : ★¹ = A" $
|
||||||
`([PD $ MkPDecl []
|
`([PD $ MkPDecl []
|
||||||
(PDef $ MkPDef (PQ Zero _) "A" (Just $ TYPE 0 _) (Enum [] _) _) _,
|
(PDef $ MkPDef (PQ Zero _) "A"
|
||||||
|
(PConcrete (Just $ TYPE 0 _) (Enum [] _)) _) _,
|
||||||
PD $ MkPDecl []
|
PD $ MkPDecl []
|
||||||
(PDef $ MkPDef (PQ Zero _) "B" (Just $ TYPE 1 _) (V "A" {}) _) _]),
|
(PDef $ MkPDef (PQ Zero _) "B"
|
||||||
|
(PConcrete (Just $ TYPE 1 _) (V "A" {})) _) _]),
|
||||||
note "empty input",
|
note "empty input",
|
||||||
parsesAs input "" [],
|
parsesAs input "" [],
|
||||||
parseFails input ";;;;;;;;;;;;;;;;;;;;;;;;;;",
|
parseFails input ";;;;;;;;;;;;;;;;;;;;;;;;;;",
|
||||||
|
@ -449,21 +466,23 @@ tests = "parser" :- [
|
||||||
`([PD (MkPDecl []
|
`([PD (MkPDecl []
|
||||||
(PNs $ MkPNamespace [< "a"]
|
(PNs $ MkPNamespace [< "a"]
|
||||||
[MkPDecl []
|
[MkPDecl []
|
||||||
(PDef $ MkPDef (PQ Any _) "x" Nothing
|
(PDef $ MkPDef (PQ Any _) "x"
|
||||||
(Ann (Tag "t" _) (Enum ["t"] _) _) _) _] _) _)]),
|
(PConcrete Nothing
|
||||||
parseMatch input "namespace a {def x = 't ∷ {t}} def y = a.x"
|
(Ann (Tag "t" _) (Enum ["t"] _) _)) _) _] _) _)]),
|
||||||
|
parseMatch input "namespace a {def x : {t} = 't} def y = a.x"
|
||||||
`([PD (MkPDecl []
|
`([PD (MkPDecl []
|
||||||
(PNs $ MkPNamespace [< "a"]
|
(PNs $ MkPNamespace [< "a"]
|
||||||
[MkPDecl []
|
[MkPDecl []
|
||||||
(PDef $ MkPDef (PQ Any _) "x" Nothing
|
(PDef $ MkPDef (PQ Any _) "x"
|
||||||
(Ann (Tag "t" _) (Enum ["t"] _) _) _) _] _) _),
|
(PConcrete (Just (Enum ["t"] _))
|
||||||
|
(Tag "t" _)) _) _] _) _),
|
||||||
PD (MkPDecl []
|
PD (MkPDecl []
|
||||||
(PDef $ MkPDef (PQ Any _) "y" Nothing
|
(PDef $ MkPDef (PQ Any _) "y"
|
||||||
(V (MakePName [< "a"] "x") Nothing _) _) _)]),
|
(PConcrete Nothing (V (MakePName [< "a"] "x") Nothing _)) _) _)]),
|
||||||
parseMatch input #" load "a.quox"; def b = a.b "#
|
parseMatch input #" load "a.quox"; def b = a.b "#
|
||||||
`([PLoad "a.quox" _,
|
`([PLoad "a.quox" _,
|
||||||
PD (MkPDecl []
|
PD (MkPDecl []
|
||||||
(PDef $ MkPDef (PQ Any _) "b" Nothing
|
(PDef $ MkPDef (PQ Any _) "b"
|
||||||
(V (MakePName [< "a"] "b") Nothing _) _) _)])
|
(PConcrete Nothing (V (MakePName [< "a"] "b") Nothing _)) _) _)])
|
||||||
]
|
]
|
||||||
]
|
]
|
||||||
|
|
|
@ -14,8 +14,8 @@ import Control.Eff
|
||||||
|
|
||||||
runWhnf : Eff Whnf a -> Either Error a
|
runWhnf : Eff Whnf a -> Either Error a
|
||||||
runWhnf act = runSTErr $ do
|
runWhnf act = runSTErr $ do
|
||||||
runEff act [handleStateSTRef !(liftST $ newSTRef 0),
|
runEff act [handleExcept (\e => stLeft e),
|
||||||
handleExcept (\e => stLeft e)]
|
handleStateSTRef !(liftST $ newSTRef 0)]
|
||||||
|
|
||||||
parameters {0 isRedex : RedexTest tm} {auto _ : CanWhnf tm isRedex} {d, n : Nat}
|
parameters {0 isRedex : RedexTest tm} {auto _ : CanWhnf tm isRedex} {d, n : Nat}
|
||||||
{auto _ : (Eq (tm d n), Show (tm d n))}
|
{auto _ : (Eq (tm d n), Show (tm d n))}
|
||||||
|
@ -32,7 +32,7 @@ parameters {0 isRedex : RedexTest tm} {auto _ : CanWhnf tm isRedex} {d, n : Nat}
|
||||||
testNoStep label ctx e = testWhnf label ctx e e
|
testNoStep label ctx e = testWhnf label ctx e e
|
||||||
|
|
||||||
private
|
private
|
||||||
ctx : Context (\n => (BindName, Term 0 n)) n -> WhnfContext 0 n
|
ctx : {n : Nat} -> Context (\n => (BindName, Term 0 n)) n -> WhnfContext 0 n
|
||||||
ctx xs = let (ns, ts) = unzip xs in MkWhnfContext [<] ns ts
|
ctx xs = let (ns, ts) = unzip xs in MkWhnfContext [<] ns ts
|
||||||
|
|
||||||
|
|
||||||
|
|
|
@ -2,7 +2,7 @@ module Tests.Typechecker
|
||||||
|
|
||||||
import Quox.Syntax
|
import Quox.Syntax
|
||||||
import Quox.Typechecker as Lib
|
import Quox.Typechecker as Lib
|
||||||
import Quox.ST
|
import Control.Monad.ST
|
||||||
import public TypingImpls
|
import public TypingImpls
|
||||||
import TAP
|
import TAP
|
||||||
import Quox.EffExtra
|
import Quox.EffExtra
|
||||||
|
|
Loading…
Reference in a new issue