Compare commits

...

7 Commits

24 changed files with 1031 additions and 406 deletions

157
exe/CompileMonad.idr Normal file
View File

@ -0,0 +1,157 @@
module CompileMonad
import Quox.Syntax as Q
import Quox.Definition as Q
import Quox.Untyped.Syntax as U
import Quox.Parser
import Quox.Untyped.Erase
import Quox.Untyped.Scheme
import Quox.Pretty
import Quox.Log
import Options
import Output
import Error
import System.File
import Data.IORef
import Control.Eff
%default total
%hide Doc.(>>=)
%hide Core.(>>=)
%hide FromParser.Error
%hide Erase.Error
%hide Lexer.Error
%hide Parser.Error
public export
record State where
constructor MkState
seen : IORef SeenSet
defs : IORef Q.Definitions
ns : IORef Mods
suf : IORef NameSuf
%name CompileMonad.State state
export %inline
newState : HasIO io => io State
newState = pure $ MkState {
seen = !(newIORef empty),
defs = !(newIORef empty),
ns = !(newIORef [<]),
suf = !(newIORef 0)
}
public export
data CompileTag = OPTS | STATE
public export
Compile : List (Type -> Type)
Compile =
[Except Error,
ReaderL STATE State, ReaderL OPTS Options, Log,
LoadFile, IO]
export %inline
handleLog : IORef LevelStack -> OpenFile -> LogL x a -> IOErr Error a
handleLog lvls f l = case f of
OConsole ch => handleLogIO (const $ pure ()) lvls (consoleHandle ch) l
OFile _ h => handleLogIO (const $ pure ()) lvls h l
ONone => handleLogDiscard l
private %inline
withLogFile : Options ->
(IORef LevelStack -> OpenFile -> IO (Either Error a)) ->
IO (Either Error a)
withLogFile opts act = do
lvlStack <- newIORef $ singleton opts.logLevels
withOutFile CErr opts.logFile fromError $ act lvlStack
where
fromError : String -> FileError -> IO (Either Error a)
fromError file err = pure $ Left $ WriteError file err
export covering %inline
runCompile : Options -> State -> Eff Compile a -> IO (Either Error a)
runCompile opts state act = do
withLogFile opts $ \lvls, logFile =>
fromIOErr $ runEff act $ with Union.(::)
[handleExcept (\e => ioLeft e),
handleReaderConst state,
handleReaderConst opts,
handleLog lvls logFile,
handleLoadFileIOE loadError ParseError state.seen opts.include,
liftIO]
private %inline
rethrowFileC : String -> Either FileError a -> Eff Compile a
rethrowFileC f = rethrow . mapFst (WriteError f)
export %inline
outputStr : OpenFile -> Lazy String -> Eff Compile ()
outputStr ONone _ = pure ()
outputStr (OConsole COut) str = putStr str
outputStr (OConsole CErr) str = fPutStr stderr str >>= rethrowFileC "<stderr>"
outputStr (OFile f h) str = fPutStr h str >>= rethrowFileC f
export %inline
outputDocs : OpenFile ->
({opts : LayoutOpts} -> Eff Pretty (List (Doc opts))) ->
Eff Compile ()
outputDocs file docs = do
opts <- askAt OPTS
for_ (runPretty opts (toOutFile file) docs) $ \x =>
outputStr file $ render (Opts opts.width) x
export %inline
outputDoc : OpenFile ->
({opts : LayoutOpts} -> Eff Pretty (Doc opts)) -> Eff Compile ()
outputDoc file doc = outputDocs file $ singleton <$> doc
public export
data StopTag = STOP
public export
CompileStop : List (Type -> Type)
CompileStop = FailL STOP :: Compile
export %inline
withEarlyStop : Eff CompileStop () -> Eff Compile ()
withEarlyStop = ignore . runFailAt STOP
export %inline
stopHere : Has (FailL STOP) fs => Eff fs ()
stopHere = failAt STOP
export %inline
liftFromParser : Eff FromParserIO a -> Eff Compile a
liftFromParser act =
runEff act $ with Union.(::)
[handleExcept $ \err => throw $ FromParserError err,
handleStateIORef !(asksAt STATE defs),
handleStateIORef !(asksAt STATE ns),
handleStateIORef !(asksAt STATE suf),
\g => send g,
\g => send g]
export %inline
liftErase : Q.Definitions -> Eff Erase a -> Eff Compile a
liftErase defs act =
runEff act
[handleExcept $ \err => throw $ EraseError err,
handleStateIORef !(asksAt STATE suf),
\g => send g]
export %inline
liftScheme : Eff Scheme a -> Eff Compile (a, List Id)
liftScheme act = do
runEff [|MkPair act (getAt MAIN)|]
[handleStateIORef !(newIORef empty),
handleStateIORef !(newIORef [])]

49
exe/Error.idr Normal file
View File

@ -0,0 +1,49 @@
module Error
import Quox.Pretty
import Quox.Parser
import Quox.Untyped.Erase
import Quox.Untyped.Scheme
import Options
import Output
import System.File
public export
data Error =
ParseError String Parser.Error
| FromParserError FromParser.Error
| EraseError Erase.Error
| WriteError FilePath FileError
| NoMain
| MultipleMains (List Scheme.Id)
%hide FromParser.Error
%hide Erase.Error
%hide Lexer.Error
%hide Parser.Error
export
loadError : Loc -> FilePath -> FileError -> Error
loadError loc file err = FromParserError $ LoadError loc file err
export
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 NoMain = pure "no #[main] function given"
prettyError (MultipleMains xs) =
pure $ sep ["multiple #[main] functions given:",
separateLoose "," !(traverse prettyId xs)]
prettyError (WriteError file e) = pure $
hangSingle 2 (text "couldn't write file \{file}:") (pshow e)
export
dieError : Options -> Error -> IO a
dieError opts e =
die (Opts opts.width) $
runPretty ({outFile := Console} opts) Console $
prettyError e

View File

@ -7,7 +7,11 @@ import Quox.Parser
import Quox.Untyped.Erase
import Quox.Untyped.Scheme
import Quox.Pretty
import Quox.Log
import Options
import Output
import Error
import CompileMonad
import System
import System.File
@ -19,227 +23,27 @@ import Control.Eff
%hide Doc.(>>=)
%hide Core.(>>=)
private
die : HasIO io => (opts : LayoutOpts) -> Doc opts -> io a
die opts err = do
ignore $ fPutStr stderr $ render opts err
exitFailure
private
hlFor : HLType -> OutFile -> HL -> Highlight
hlFor Guess Console = highlightSGR
hlFor Guess _ = noHighlight
hlFor NoHL _ = noHighlight
hlFor Term _ = highlightSGR
hlFor Html _ = highlightHtml
private
runPretty : Options -> OutFile -> Eff Pretty a -> a
runPretty opts file act =
runPrettyWith Outer opts.flavor (hlFor opts.hlType file) 2 act
private
record State where
constructor MkState
seen : IORef SeenSet
defs : IORef Q.Definitions
ns : IORef Mods
suf : IORef NameSuf
%name Main.State state
private
newState : HasIO io => io State
newState = pure $ MkState {
seen = !(newIORef empty),
defs = !(newIORef empty),
ns = !(newIORef [<]),
suf = !(newIORef 0)
}
private
data Error =
ParseError String Parser.Error
| FromParserError FromParser.Error
| EraseError Erase.Error
| WriteError FilePath FileError
| NoMain
| MultipleMains (List Id)
%hide FromParser.Error
%hide Erase.Error
%hide Lexer.Error
%hide Parser.Error
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 NoMain = pure "no #[main] function given"
prettyError (MultipleMains xs) =
pure $ sep ["multiple #[main] functions given:",
separateLoose "," !(traverse prettyId xs)]
prettyError (WriteError file e) = pure $
hangSingle 2 (text "couldn't write file \{file}:") (pshow e)
private
dieError : Options -> Error -> IO a
dieError opts e =
die (Opts opts.width) $
runPretty ({outFile := Console} opts) Console $
prettyError e
private
data CompileTag = OPTS | STATE
private
Compile : List (Type -> Type)
Compile =
[Except Error,
ReaderL STATE State, ReaderL OPTS Options,
LoadFile, IO]
private covering
runCompile : Options -> State -> Eff Compile a -> IO (Either Error a)
runCompile opts state act =
fromIOErr $ runEff act $ with Union.(::)
[handleExcept (\e => ioLeft e),
handleReaderConst state,
handleReaderConst opts,
handleLoadFileIOE loadError ParseError state.seen opts.include,
liftIO]
private
data StopTag = STOP
private
CompileStop : List (Type -> Type)
CompileStop = FailL STOP :: Compile
private
withEarlyStop : Has (FailL STOP) fs => Eff fs () -> Eff (fs - FailL STOP) ()
withEarlyStop = ignore . runFailAt STOP
private
stopHere : Has (FailL STOP) fs => Eff fs ()
stopHere = failAt STOP
private
data ConsoleChannel = COut | CErr
private
data OpenFile = OConsole ConsoleChannel | OFile String File | ONone
private
rethrowFile : String -> Either FileError a -> Eff Compile a
rethrowFile f = rethrow . mapFst (WriteError f)
private
toOutFile : OpenFile -> OutFile
toOutFile (OConsole _) = Console
toOutFile (OFile f _) = File f
toOutFile ONone = NoOut
private
withFileC : String -> (OpenFile -> Eff Compile a) -> Eff Compile a
withFileC f act =
withFile f WriteTruncate pure (Prelude.map Right . act . OFile f) >>=
rethrowFile f
private
withOutFile : ConsoleChannel -> OutFile ->
(OpenFile -> Eff Compile a) -> Eff Compile a
withOutFile _ (File f) act = withFileC f act
withOutFile ch Console act = act $ OConsole ch
withOutFile _ NoOut act = act ONone
private
outputStr : OpenFile -> Lazy String -> Eff Compile ()
outputStr ONone _ = pure ()
outputStr (OConsole COut) str = putStr str
outputStr (OConsole CErr) str = fPutStr stderr str >>= rethrowFile "<stderr>"
outputStr (OFile f h) str = fPutStr h str >>= rethrowFile f
private
outputDocs : OpenFile ->
({opts : LayoutOpts} -> Eff Pretty (List (Doc opts))) ->
Eff Compile ()
outputDocs file docs = do
opts <- askAt OPTS
for_ (runPretty opts (toOutFile file) docs) $ \x =>
outputStr file $ render (Opts opts.width) x
private
outputDoc : OpenFile ->
({opts : LayoutOpts} -> Eff Pretty (Doc opts)) -> Eff Compile ()
outputDoc file doc = outputDocs file $ singleton <$> doc
private
liftFromParser : Eff FromParserIO a -> Eff Compile a
liftFromParser act =
runEff act $ with Union.(::)
[handleExcept $ \err => throw $ FromParserError err,
handleStateIORef !(asksAt STATE defs),
handleStateIORef !(asksAt STATE ns),
handleStateIORef !(asksAt STATE suf),
\g => send g]
private
liftErase : Q.Definitions -> Eff Erase a -> Eff Compile a
liftErase defs act =
runEff act
[handleExcept $ \err => throw $ EraseError err,
handleStateIORef !(asksAt STATE suf)]
private
liftScheme : Eff Scheme a -> Eff Compile (a, List Id)
liftScheme act = do
runEff [|MkPair act (getAt MAIN)|]
[handleStateIORef !(newIORef empty),
handleStateIORef !(newIORef [])]
private
Step : Type -> Type -> Type
Step i o = OpenFile -> i -> Eff Compile o
-- private
-- processFile : String -> Eff Compile ()
-- processFile file = withEarlyStop $ do
-- Just ast <- loadFile noLoc file
-- | Nothing => pure ()
-- -- putErrLn "checking \{file}"
-- when (!(asksAt OPTS until) == Just Parse) $ do
-- lift $ outputStr $ show ast
-- stopHere
-- defList <- liftFromParser $ concat <$> traverse fromPTopLevel ast
-- outputDocStopIf Check $
-- traverse (uncurry Q.prettyDef) defList
-- let defs = SortedMap.fromList defList
-- erased <- liftErase defs $
-- traverse (\(x, d) => (x,) <$> eraseDef defs x d) defList
-- outputDocStopIf Erase $
-- traverse (uncurry U.prettyDef) erased
-- (scheme, mains) <- liftScheme $ map catMaybes $
-- traverse (uncurry defToScheme) erased
-- outputDocStopIf Scheme $
-- intersperse empty <$> traverse prettySexp scheme
Step a b = OpenFile -> a -> Eff Compile b
private
step : {default CErr console : ConsoleChannel} ->
Phase -> OutFile -> Step i o -> i -> Eff CompileStop o
Phase -> OutFile -> Step a b -> a -> Eff CompileStop b
step phase file act x = do
opts <- askAt OPTS
res <- lift $ withOutFile console file $ \h => act h x
res <- withOutFile console file fromError $ \h => lift $ act h x
when (opts.until == Just phase) stopHere
pure res
where
fromError : String -> FileError -> Eff CompileStop c
fromError file err = throw $ WriteError file err
private covering
@ -268,25 +72,23 @@ erase h defList =
where defs = SortedMap.fromList defList
private covering
scheme : Step (List U.NDefinition) (List Sexp, Id)
scheme : Step (List U.NDefinition) (List Sexp, List Id)
scheme h defs = do
sexps' <- for defs $ \(x, d) => do
(msexp, mains) <- liftScheme $ defToScheme x d
outputDoc h $ maybe (sayErased x) prettySexp msexp
outputDoc h $ case msexp of
Just s => prettySexp s
Nothing => pure $ hsep [";;", prettyName x, "erased"]
pure (msexp, mains)
bitraverse (pure . catMaybes) (oneMain . concat) $ unzip sexps'
where
sayErased : {opts : LayoutOpts} -> Name -> Eff Pretty (Doc opts)
sayErased x = pure $ hsep [";;", prettyName x, "erased"]
oneMain : List Id -> Eff Compile Id
oneMain [m] = pure m
oneMain [] = throw NoMain
oneMain ms = throw $ MultipleMains ms
pure $ bimap catMaybes concat $ unzip sexps'
private covering
output : Step (List Sexp, Id) ()
output h (sexps, main) =
output : Step (List Sexp, List Id) ()
output h (sexps, mains) = do
main <- case mains of
[m] => pure m
[] => throw NoMain
_ => throw $ MultipleMains mains
lift $ outputDocs h $ do
res <- traverse prettySexp sexps
runner <- makeRunMain main

View File

@ -1,6 +1,9 @@
module Options
import Quox.Pretty
import Quox.Log
import Data.DPair
import Data.SortedMap
import System
import System.Console.GetOpt
import System.File
@ -42,13 +45,15 @@ record Dump where
public export
record Options where
constructor MkOpts
hlType : HLType
dump : Dump
outFile : OutFile
until : Maybe Phase
flavor : Pretty.Flavor
width : Nat
include : List String
include : List String
dump : Dump
outFile : OutFile
until : Maybe Phase
hlType : HLType
flavor : Pretty.Flavor
width : Nat
logLevels : LogLevels
logFile : OutFile
%name Options opts
%runElab derive "Options" [Show]
@ -61,13 +66,15 @@ defaultWidth = do
export
defaultOpts : IO Options
defaultOpts = pure $ MkOpts {
hlType = Guess,
dump = MkDump NoOut NoOut NoOut NoOut,
outFile = Console,
until = Nothing,
flavor = Unicode,
width = !defaultWidth,
include = ["."]
include = ["."],
dump = MkDump NoOut NoOut NoOut NoOut,
outFile = Console,
until = Nothing,
hlType = Guess,
flavor = Unicode,
width = !defaultWidth,
logLevels = defaultLogLevels,
logFile = Console
}
private
@ -118,11 +125,52 @@ toHLType str = case toLower str of
"html" => Ok {hlType := Html}
_ => Err "unknown highlighting type \{show str}\ntypes: term, html, none"
||| like ghc, -i '' clears the search path; -i a:b:c adds a,b,c to the end
||| like ghc, `-i ""` clears the search path;
||| `-i a:b:c` adds `a`, `b`, `c` to the end
private
dirListFlag : String -> List String -> List String
dirListFlag arg val =
if null arg then [] else val ++ toList (split (== ':') arg)
dirListFlag "" val = []
dirListFlag dirs val = val ++ toList (split (== ':') dirs)
private
splitLogFlag : String -> Either String (List (Maybe LogCategory, LogLevel))
splitLogFlag = traverse flag1 . toList . split (== ':') where
parseLogCategory : String -> Either String LogCategory
parseLogCategory cat = do
let Just cat = toLogCategory cat
| _ => let catList = joinBy ", " logCategories in
Left "unknown log category. categories are:\n\{catList}"
pure cat
parseLogLevel : String -> Either String LogLevel
parseLogLevel lvl = do
let Just lvl = parsePositive lvl
| _ => Left "log level \{lvl} not a number"
let Just lvl = toLogLevel lvl
| _ => Left "log level \{show lvl} out of range 0\{show maxLogLevel}"
pure lvl
flag1 : String -> Either String (Maybe LogCategory, LogLevel)
flag1 str = do
let (first, second) = break (== '=') str
case strM second of
StrCons '=' lvl => do
cat <- parseLogCategory first
lvl <- parseLogLevel lvl
pure (Just cat, lvl)
StrNil => (Nothing,) <$> parseLogLevel first
_ => Left "invalid log flag \{str}"
private
setLogFlag : LogLevels -> (Maybe LogCategory, LogLevel) -> LogLevels
setLogFlag lvls (Nothing, lvl) = {defLevel := lvl} lvls
setLogFlag lvls (Just name, lvl) = {levels $= ((name, lvl) ::)} lvls
private
logFlag : String -> OptAction
logFlag str = case splitLogFlag str of
Left err => Err err
Right flags => Ok $ \o => {logLevels := foldl setLogFlag o.logLevels flags} o
private
commonOptDescrs' : List (OptDescr OptAction)
@ -133,7 +181,11 @@ commonOptDescrs' = [
MkOpt ['o'] ["output"] (ReqArg (\s => Ok {outFile := toOutFile s}) "<file>")
"output file (\"-\" for stdout, \"\" for no output)",
MkOpt ['P'] ["phase"] (ReqArg toPhase "<phase>")
"stop after the given phase"
"stop after the given phase",
MkOpt ['l'] ["log"] (ReqArg logFlag "[<cat>=]<n>:...")
"set log level",
MkOpt ['L'] ["log-file"] (ReqArg (\s => Ok {logFile := toOutFile s}) "<file>")
"set log output file"
]
private
@ -148,13 +200,17 @@ extraOptDescrs = [
MkOpt [] ["color", "colour"] (ReqArg toHLType "<type>")
"select highlighting type",
MkOpt [] ["dparse"] (ReqArg (\s => Ok {dump.parse := toOutFile s}) "<file>")
MkOpt [] ["dump-parse"]
(ReqArg (\s => Ok {dump.parse := toOutFile s}) "<file>")
"dump AST",
MkOpt [] ["dcheck"] (ReqArg (\s => Ok {dump.check := toOutFile s}) "<file>")
MkOpt [] ["dump-check"]
(ReqArg (\s => Ok {dump.check := toOutFile s}) "<file>")
"dump typechecker output",
MkOpt [] ["derase"] (ReqArg (\s => Ok {dump.erase := toOutFile s}) "<file>")
MkOpt [] ["dump-erase"]
(ReqArg (\s => Ok {dump.erase := toOutFile s}) "<file>")
"dump erasure output",
MkOpt [] ["dscheme"] (ReqArg (\s => Ok {dump.scheme := toOutFile s}) "<file>")
MkOpt [] ["dump-scheme"]
(ReqArg (\s => Ok {dump.scheme := toOutFile s}) "<file>")
"dump scheme output (without prelude)"
]

59
exe/Output.idr Normal file
View File

@ -0,0 +1,59 @@
module Output
import Quox.Pretty
import Options
import System.File
import System
public export
data ConsoleChannel = COut | CErr
export
consoleHandle : ConsoleChannel -> File
consoleHandle COut = stdout
consoleHandle CErr = stderr
public export
data OpenFile = OConsole ConsoleChannel | OFile String File | ONone
export
toOutFile : OpenFile -> OutFile
toOutFile (OConsole _) = Console
toOutFile (OFile f _) = File f
toOutFile ONone = NoOut
export
withFile : HasIO m => String -> (String -> FileError -> m a) ->
(OpenFile -> m a) -> m a
withFile f catch act = Prelude.do
res <- withFile f WriteTruncate pure (Prelude.map Right . act . OFile f)
either (catch f) pure res
export
withOutFile : HasIO m => ConsoleChannel -> OutFile ->
(String -> FileError -> m a) -> (OpenFile -> m a) -> m a
withOutFile _ (File f) catch act = withFile f catch act
withOutFile ch Console catch act = act $ OConsole ch
withOutFile _ NoOut catch act = act ONone
private
hlFor : HLType -> OutFile -> HL -> Highlight
hlFor Guess Console = highlightSGR
hlFor Guess _ = noHighlight
hlFor NoHL _ = noHighlight
hlFor Term _ = highlightSGR
hlFor Html _ = highlightHtml
export
runPretty : Options -> OutFile -> Eff Pretty a -> a
runPretty opts file act =
runPrettyWith Outer opts.flavor (hlFor opts.hlType file) 2 act
export
die : HasIO io => (opts : LayoutOpts) -> Doc opts -> io a
die opts err = do
ignore $ fPutStr stderr $ render opts err
exitFailure

View File

@ -62,3 +62,21 @@ export %inline HasST (STErr e) where liftST = STE . map Right
export
stLeft : e -> STErr e s a
stLeft e = STE $ pure $ Left e
parameters {auto _ : HasST m}
export %inline
newSTRef' : a -> m s (STRef s a)
newSTRef' x = liftST $ newSTRef x
export %inline
readSTRef' : STRef s a -> m s a
readSTRef' r = liftST $ readSTRef r
export %inline
writeSTRef' : STRef s a -> a -> m s ()
writeSTRef' r x = liftST $ writeSTRef r x
export %inline
modifySTRef' : STRef s a -> (a -> a) -> m s ()
modifySTRef' r f = liftST $ modifySTRef r f

View File

@ -97,6 +97,10 @@ export
handleReaderConst : Applicative m => r -> ReaderL lbl r a -> m a
handleReaderConst x Ask = pure x
export
handleWriterST : HasST m => STRef s (SnocList w) -> WriterL lbl w a -> m s a
handleWriterST ref (Tell w) = liftST $ modifySTRef ref (:< w)
public export
record IOErr e a where

View File

@ -2,9 +2,12 @@ module Quox.Equal
import Quox.BoolExtra
import public Quox.Typing
import Data.Maybe
import Quox.EffExtra
import Quox.FreeVars
import Quox.Pretty
import Quox.EffExtra
import Data.List1
import Data.Maybe
%default total
@ -15,11 +18,11 @@ EqModeState = State EqMode
public export
Equal : List (Type -> Type)
Equal = [ErrorEff, DefsReader, NameGen]
Equal = [ErrorEff, DefsReader, NameGen, Log]
public export
EqualInner : List (Type -> Type)
EqualInner = [ErrorEff, NameGen, EqModeState]
EqualInner = [ErrorEff, NameGen, EqModeState, Log]
export %inline
@ -74,9 +77,25 @@ sameTyCon (E {}) _ = False
||| * `[π.A]` is empty if `A` is.
||| * that's it.
public export covering
isEmpty : Definitions -> EqContext n -> SQty -> Term 0 n ->
Eff EqualInner Bool
isEmpty defs ctx sg ty0 = do
isEmpty :
{default 30 logLevel : Nat} -> (0 _ : So (isLogLevel logLevel)) =>
Definitions -> EqContext n -> SQty -> Term 0 n -> Eff EqualInner Bool
private covering
isEmptyNoLog :
Definitions -> EqContext n -> SQty -> Term 0 n -> Eff EqualInner Bool
isEmpty defs ctx sg ty = do
sayMany "equal" ty.loc
[logLevel :> "isEmpty",
95 :> hsep ["ctx =", runPretty $ prettyEqContext ctx],
95 :> hsep ["sg =", runPretty $ prettyQty sg.qty],
logLevel :> hsep ["ty =", runPretty $ prettyTerm [<] ctx.tnames ty]]
res <- isEmptyNoLog defs ctx sg ty
say "equal" logLevel ty.loc $ hsep ["isEmpty ⇝", pshow res]
pure res
isEmptyNoLog defs ctx sg ty0 = do
Element ty0 nc <- whnf defs ctx sg ty0.loc ty0
let Left y = choose $ isTyConE ty0
| Right n => pure False
@ -85,16 +104,17 @@ isEmpty defs ctx sg ty0 = do
IOState {} => pure False
Pi {arg, res, _} => pure False
Sig {fst, snd, _} =>
isEmpty defs ctx sg fst `orM`
isEmpty defs (extendTy0 snd.name fst ctx) sg snd.term
isEmpty defs ctx sg fst {logLevel = 90} `orM`
isEmpty defs (extendTy0 snd.name fst ctx) sg snd.term {logLevel = 90}
Enum {cases, _} =>
pure $ null cases
Eq {} => pure False
NAT {} => pure False
STRING {} => pure False
BOX {ty, _} => isEmpty defs ctx sg ty
BOX {ty, _} => isEmpty defs ctx sg ty {logLevel = 90}
E _ => pure False
||| true if a type is known to be a subsingleton purely by its form.
||| a subsingleton is a type with only zero or one possible values.
||| equality/subtyping accepts immediately on values of subsingleton types.
@ -106,27 +126,42 @@ isEmpty defs ctx sg ty0 = do
||| * an enum type is a subsingleton if it has zero or one tags.
||| * a box type is a subsingleton if its content is
public export covering
isSubSing : Definitions -> EqContext n -> SQty -> Term 0 n ->
Eff EqualInner Bool
isSubSing defs ctx sg ty0 = do
isSubSing :
{default 30 logLevel : Nat} -> (0 _ : So (isLogLevel logLevel)) =>
Definitions -> EqContext n -> SQty -> Term 0 n -> Eff EqualInner Bool
private covering
isSubSingNoLog :
Definitions -> EqContext n -> SQty -> Term 0 n -> Eff EqualInner Bool
isSubSing defs ctx sg ty = do
sayMany "equal" ty.loc
[logLevel :> "isSubSing",
95 :> hsep ["ctx =", runPretty $ prettyEqContext ctx],
95 :> hsep ["sg =", runPretty $ prettyQty sg.qty],
logLevel :> hsep ["ty =", runPretty $ prettyTerm [<] ctx.tnames ty]]
res <- isSubSingNoLog defs ctx sg ty
say "equal" logLevel ty.loc $ hsep ["isSubsing ⇝", pshow res]
pure res
isSubSingNoLog defs ctx sg ty0 = do
Element ty0 nc <- whnf defs ctx sg ty0.loc ty0
let Left y = choose $ isTyConE ty0
| Right n => pure False
let Left y = choose $ isTyConE ty0 | _ => pure False
case ty0 of
TYPE {} => pure False
IOState {} => pure False
Pi {arg, res, _} =>
isEmpty defs ctx sg arg `orM`
isSubSing defs (extendTy0 res.name arg ctx) sg res.term
isEmpty defs ctx sg arg {logLevel = 90} `orM`
isSubSing defs (extendTy0 res.name arg ctx) sg res.term {logLevel = 90}
Sig {fst, snd, _} =>
isSubSing defs ctx sg fst `andM`
isSubSing defs (extendTy0 snd.name fst ctx) sg snd.term
isSubSing defs ctx sg fst {logLevel = 90} `andM`
isSubSing defs (extendTy0 snd.name fst ctx) sg snd.term {logLevel = 90}
Enum {cases, _} =>
pure $ length (SortedSet.toList cases) <= 1
Eq {} => pure True
NAT {} => pure False
STRING {} => pure False
BOX {ty, _} => isSubSing defs ctx sg ty
BOX {ty, _} => isSubSing defs ctx sg ty {logLevel = 90}
E _ => pure False
@ -137,12 +172,21 @@ bigger l r = gets $ \case Super => l; _ => r
export
ensureTyCon : Has ErrorEff fs =>
(loc : Loc) -> (ctx : EqContext n) -> (t : Term 0 n) ->
Eff fs (So (isTyConE t))
ensureTyCon loc ctx t = case nchoose $ isTyConE t of
Left y => pure y
Right n => throw $ NotType loc (toTyContext ctx) (t // shift0 ctx.dimLen)
ensureTyCon, ensureTyConNoLog :
(Has Log fs, Has ErrorEff fs) =>
(loc : Loc) -> (ctx : EqContext n) -> (t : Term 0 n) ->
Eff fs (So (isTyConE t))
ensureTyConNoLog loc ctx ty = do
case nchoose $ isTyConE ty of
Left y => pure y
Right n => throw $ NotType loc (toTyContext ctx) (ty // shift0 ctx.dimLen)
ensureTyCon loc ctx ty = do
sayMany "equal" ty.loc
[60 :> "ensureTyCon",
95 :> hsep ["ctx =", runPretty $ prettyEqContext ctx],
60 :> hsep ["ty =", runPretty $ prettyTerm [<] ctx.tnames ty]]
ensureTyConNoLog loc ctx ty
namespace Term
@ -750,7 +794,11 @@ namespace Elim
namespace Term
compare0 defs ctx sg ty s t =
export covering %inline
compare0NoLog :
Definitions -> EqContext n -> SQty -> (ty, s, t : Term 0 n) ->
Eff EqualInner ()
compare0NoLog defs ctx sg ty s t =
wrapErr (WhileComparingT ctx !mode sg ty s t) $ do
Element ty' _ <- whnf defs ctx SZero ty.loc ty
Element s' _ <- whnf defs ctx sg s.loc s
@ -758,20 +806,72 @@ namespace Term
tty <- ensureTyCon ty.loc ctx ty'
compare0' defs ctx sg ty' s' t'
compare0 defs ctx sg ty s t = do
sayMany "equal" s.loc
[30 :> "Term.compare0",
30 :> hsep ["mode =", pshow !mode],
95 :> hsep ["ctx =", runPretty $ prettyEqContext ctx],
95 :> hsep ["sg =", runPretty $ prettyQty sg.qty],
31 :> hsep ["ty =", runPretty $ prettyTerm [<] ctx.tnames ty],
30 :> hsep ["s =", runPretty $ prettyTerm [<] ctx.tnames s],
30 :> hsep ["t =", runPretty $ prettyTerm [<] ctx.tnames t]]
compare0NoLog defs ctx sg ty s t
namespace Elim
compare0 defs ctx sg e f = do
export covering %inline
compare0NoLog :
Definitions -> EqContext n -> SQty -> (e, f : Elim 0 n) ->
Eff EqualInner (Term 0 n)
compare0NoLog defs ctx sg e f = do
(ty, err) <- runStateAt InnerErr Nothing $ compare0Inner defs ctx sg e f
maybe (pure ty) throw err
compareType defs ctx s t = do
compare0 defs ctx sg e f = do
sayMany "equal" e.loc
[30 :> "Elim.compare0",
30 :> hsep ["mode =", pshow !mode],
95 :> hsep ["ctx =", runPretty $ prettyEqContext ctx],
95 :> hsep ["sg =", runPretty $ prettyQty sg.qty],
30 :> hsep ["e =", runPretty $ prettyElim [<] ctx.tnames e],
30 :> hsep ["f =", runPretty $ prettyElim [<] ctx.tnames f]]
ty <- compare0NoLog defs ctx sg e f
say "equal" 31 e.loc $
hsep ["Elim.compare0 ⇝", runPretty $ prettyTerm [<] ctx.tnames ty]
pure ty
export covering %inline
compareTypeNoLog :
Definitions -> EqContext n -> (s, t : Term 0 n) -> Eff EqualInner ()
compareTypeNoLog defs ctx s t = do
Element s' _ <- whnf defs ctx SZero s.loc s
Element t' _ <- whnf defs ctx SZero t.loc t
ts <- ensureTyCon s.loc ctx s'
tt <- ensureTyCon t.loc ctx t'
st <- either pure (const $ clashTy s.loc ctx s' t') $
nchoose $ sameTyCon s' t'
let Left _ = choose $ sameTyCon s' t' | _ => clashTy s.loc ctx s' t'
compareType' defs ctx s' t'
compareType defs ctx s t = do
sayMany "equal" s.loc
[30 :> "compareType",
30 :> hsep ["mode =", pshow !mode],
95 :> hsep ["ctx =", runPretty $ prettyEqContext ctx],
30 :> hsep ["s =", runPretty $ prettyTerm [<] ctx.tnames s],
30 :> hsep ["t =", runPretty $ prettyTerm [<] ctx.tnames t]]
compareTypeNoLog defs ctx s t
private
getVars : TyContext d _ -> FreeVars d -> List BindName
getVars ctx (FV fvs) = case ctx.dctx of
ZeroIsOne => []
C eqs => toList $ getVars' ctx.dnames eqs fvs
where
getVars' : BContext d' -> DimEq' d' -> FreeVars' d' -> SnocList BindName
getVars' (names :< name) (eqs :< eq) (fvs :< fv) =
let rest = getVars' names eqs fvs in
case eq of Nothing => rest :< name
Just _ => rest
getVars' [<] [<] [<] = [<]
parameters (loc : Loc) (ctx : TyContext d n)
parameters (mode : EqMode)
@ -780,9 +880,11 @@ parameters (loc : Loc) (ctx : TyContext d n)
fromInner = lift . map fst . runState mode
private
eachFace : Applicative f => FreeVars d ->
(EqContext n -> DSubst d 0 -> f ()) -> f ()
eachFace fvs act =
eachCorner : Has Log fs => Loc -> FreeVars d ->
(EqContext n -> DSubst d 0 -> Eff fs ()) -> Eff fs ()
eachCorner loc fvs act = do
say "equal" 50 loc $
hsep $ "eachCorner: split on" :: map prettyBind' (getVars ctx fvs)
for_ (splits loc ctx.dctx fvs) $ \th =>
act (makeEqContext ctx th) th
@ -792,31 +894,36 @@ parameters (loc : Loc) (ctx : TyContext d n)
Definitions -> EqContext n -> DSubst d 0 -> Eff EqualInner ()
private
runCompare : FreeVars d -> CompareAction d n -> Eff Equal ()
runCompare fvs act = fromInner $ eachFace fvs $ act !(askAt DEFS)
runCompare : Loc -> FreeVars d -> CompareAction d n -> Eff Equal ()
runCompare loc fvs act = fromInner $ eachCorner loc fvs $ act !(askAt DEFS)
private
fdvAll : HasFreeDVars t => List (t d n) -> FreeVars d
fdvAll = let Val d = ctx.dimLen in foldMap (fdvWith [|d|] ctx.termLen)
foldMap1 : Semigroup b => (a -> b) -> List1 a -> b
foldMap1 f = foldl1By (\x, y => x <+> f y) f
private
fdvAll : HasFreeDVars t => (xs : List (t d n)) -> (0 _ : NonEmpty xs) =>
FreeVars d
fdvAll (x :: xs) = foldMap1 (fdvWith ctx.dimLen ctx.termLen) (x ::: xs)
namespace Term
export covering
compare : SQty -> (ty, s, t : Term d n) -> Eff Equal ()
compare sg ty s t = runCompare (fdvAll [ty, s, t]) $ \defs, ectx, th =>
compare0 defs ectx sg (ty // th) (s // th) (t // th)
compare sg ty s t = runCompare s.loc (fdvAll [ty, s, t]) $
\defs, ectx, th => compare0 defs ectx sg (ty // th) (s // th) (t // th)
export covering
compareType : (s, t : Term d n) -> Eff Equal ()
compareType s t = runCompare (fdvAll [s, t]) $ \defs, ectx, th =>
compareType defs ectx (s // th) (t // th)
compareType s t = runCompare s.loc (fdvAll [s, t]) $
\defs, ectx, th => compareType defs ectx (s // th) (t // th)
namespace Elim
||| you don't have to pass the type in but the arguments must still be
||| of the same type!!
export covering
compare : SQty -> (e, f : Elim d n) -> Eff Equal ()
compare sg e f = runCompare (fdvAll [e, f]) $ \defs, ectx, th =>
ignore $ compare0 defs ectx sg (e // th) (f // th)
compare sg e f = runCompare e.loc (fdvAll [e, f]) $
\defs, ectx, th => ignore $ compare0 defs ectx sg (e // th) (f // th)
namespace Term
export covering %inline

240
lib/Quox/Log.idr Normal file
View File

@ -0,0 +1,240 @@
module Quox.Log
import Quox.Loc
import Quox.Pretty
import Data.So
import Data.DPair
import Data.Maybe
import Data.List1
import Control.Eff
import Control.Monad.ST.Extra
import Data.IORef
import System.File
import Derive.Prelude
%default total
%language ElabReflection
public export
maxLogLevel : Nat
maxLogLevel = 100
public export
logCategories : List String
logCategories = ["whnf", "equal", "check"]
public export %inline
isLogLevel : Nat -> Bool
isLogLevel l = l <= maxLogLevel
public export
IsLogLevel : Nat -> Type
IsLogLevel l = So $ isLogLevel l
public export %inline
isLogCategory : String -> Bool
isLogCategory cat = cat `elem` logCategories
public export
IsLogCategory : String -> Type
IsLogCategory cat = So $ isLogCategory cat
public export
LogLevel : Type
LogLevel = Subset Nat IsLogLevel
public export
LogCategory : Type
LogCategory = Subset String IsLogCategory
public export %inline
toLogLevel : Nat -> Maybe LogLevel
toLogLevel l =
case choose $ isLogLevel l of
Left y => Just $ Element l y
Right _ => Nothing
public export %inline
toLogCategory : String -> Maybe LogCategory
toLogCategory c =
case choose $ isLogCategory c of
Left y => Just $ Element c y
Right _ => Nothing
public export
LevelMap : Type
LevelMap = List (LogCategory, LogLevel)
-- i tried SortedMap first, but it is too much overhead for LevelMaps
public export
record LogLevels where
constructor MkLogLevels
defLevel : LogLevel
levels : LevelMap
%name LogLevels lvls
%runElab derive "LogLevels" [Eq, Show]
public export
LevelStack : Type
LevelStack = List1 LogLevels
export %inline
defaultLogLevels : LogLevels
defaultLogLevels = MkLogLevels (Element 0 Oh) []
export %inline
initStack : LevelStack
initStack = singleton defaultLogLevels
||| right biased for the default and for overlapping elements
public export %inline
mergeLevels : LogLevels -> LogLevels -> LogLevels
mergeLevels (MkLogLevels _ map1) (MkLogLevels def map2) =
MkLogLevels def $ map1 ++ map2
export %inline
getLevel : LogCategory -> LogLevels -> LogLevel
getLevel cat lvls = fromMaybe lvls.defLevel $ lookup cat lvls.levels
public export
LogDoc : Type
LogDoc = Doc (Opts {lineLength = 80})
public export
data PushArg = SetDefault LogLevel | SetCats LevelMap
%name PushArg push
export %inline
mergePush : PushArg -> LogLevels -> LogLevels
mergePush (SetDefault def) = {defLevel := def}
mergePush (SetCats map) = {levels $= (map ++)}
public export
record LogMsg where
constructor (:>)
level : Nat
{auto 0 levelOk : IsLogLevel level}
message : Lazy LogDoc
infix 0 :>
%name Log.LogMsg msg
public export
data LogL : (lbl : tag) -> Type -> Type where
SayMany : (cat : LogCategory) -> (loc : Loc) ->
(msgs : List LogMsg) -> LogL lbl ()
Push : (push : PushArg) -> LogL lbl ()
Pop : LogL lbl ()
CurLevels : LogL lbl LogLevels
public export
Log : Type -> Type
Log = LogL ()
parameters (0 lbl : tag) {auto _ : Has (LogL lbl) fs}
public export %inline
sayManyAt : (cat : String) -> (0 catOk : IsLogCategory cat) =>
Loc -> List LogMsg -> Eff fs ()
sayManyAt cat loc msgs {catOk} =
send $ SayMany {lbl} (Element cat catOk) loc msgs
public export %inline
sayAt : (cat : String) -> (0 catOk : IsLogCategory cat) =>
(lvl : Nat) -> (0 lvlOk : IsLogLevel lvl) =>
Loc -> Lazy LogDoc -> Eff fs ()
sayAt cat lvl loc msg = sayManyAt cat loc [lvl :> msg]
public export %inline
pushAt : PushArg -> Eff fs ()
pushAt lvls = send $ Push {lbl} lvls
public export %inline
popAt : Eff fs ()
popAt = send $ Pop {lbl}
public export %inline
curLevelsAt : Eff fs LogLevels
curLevelsAt = send $ CurLevels {lbl}
parameters {auto _ : Has Log fs}
public export %inline
sayMany : (cat : String) -> (0 catOk : IsLogCategory cat) =>
Loc -> List LogMsg -> Eff fs ()
sayMany = sayManyAt ()
public export %inline
say : (cat : String) -> (0 _ : IsLogCategory cat) =>
(lvl : Nat) -> (0 _ : IsLogLevel lvl) =>
Loc -> Lazy LogDoc -> Eff fs ()
say = sayAt ()
public export %inline
push : PushArg -> Eff fs ()
push = pushAt ()
public export %inline
pop : Eff fs ()
pop = popAt ()
public export %inline
curLevels : Eff fs LogLevels
curLevels = curLevelsAt ()
export %inline
doPush : PushArg -> LevelStack -> LevelStack
doPush push list = mergePush push (head list) `cons` list
export %inline
doPop : List1 a -> List1 a
doPop (_ ::: x :: xs) = x ::: xs
doPop (x ::: []) = x ::: []
export %inline
doSayMany : Applicative m =>
LevelStack -> (LogDoc -> m ()) ->
LogCategory -> Loc -> List LogMsg -> m ()
doSayMany (lvls ::: _) act cat loc msgs = do
let Element catLvl _ = getLevel cat lvls
loc = runPretty $ prettyLoc loc
for_ msgs $ \msg => when (msg.level <= catLvl) $
act $ hcat [loc, text cat.fst, "@", pshow msg.level, ":"] <++>
msg.message
export %inline
handleLogIO : HasIO m => (FileError -> m ()) ->
IORef LevelStack -> File -> LogL tag a -> m a
handleLogIO th lvls h = \case
Push push => modifyIORef lvls $ doPush push
Pop => modifyIORef lvls doPop
SayMany cat loc msgs => doSayMany !(readIORef lvls) printMsg cat loc msgs
CurLevels => head <$> readIORef lvls
where printMsg : LogDoc -> m ()
printMsg msg = fPutStr h (render _ msg) >>= either th pure
export %inline
handleLogST : (HasST m, Monad (m s)) =>
STRef s (SnocList LogDoc) -> STRef s LevelStack ->
LogL tag a -> m s a
handleLogST docs lvls = \case
Push push => modifySTRef' lvls $ doPush push
Pop => modifySTRef' lvls doPop
SayMany cat loc msgs => doSayMany !(readSTRef' lvls) printMsg cat loc msgs
CurLevels => head <$> readSTRef' lvls
where printMsg : LogDoc -> m s ()
printMsg msg = modifySTRef' docs (:< msg)
export %inline
handleLogDiscard : Applicative m => LogL tag a -> m a
handleLogDiscard = \case
SayMany {} => pure ()
Push {} => pure ()
Pop => pure ()
CurLevels => pure defaultLogLevels

View File

@ -32,41 +32,44 @@ data StateTag = NS | SEEN
public export
FromParserPure : List (Type -> Type)
FromParserPure = [Except Error, DefsState, StateL NS Mods, NameGen]
FromParserPure = [Except Error, DefsState, StateL NS Mods, NameGen, Log]
public export
FromParserIO : List (Type -> Type)
FromParserIO = FromParserPure ++ [LoadFile]
public export
record PureParserResult a where
constructor MkPureParserResult
val : a
suf : NameSuf
defs : Definitions
log : SnocList LogDoc
logLevels : LevelStack
export
fromParserPure : {default [<] ns : Mods} ->
NameSuf -> Definitions ->
Eff FromParserPure a ->
Either Error (a, NameSuf, Definitions)
fromParserPure suf defs act = runSTErr $ do
suf <- liftST $ newSTRef suf
defs <- liftST $ newSTRef defs
NameSuf -> Definitions -> LevelStack ->
Eff FromParserPure a -> Either Error (PureParserResult a)
fromParserPure suf defs lvls act = runSTErr $ do
suf <- newSTRef' suf
defs <- newSTRef' defs
log <- newSTRef' [<]
lvls <- newSTRef' lvls
res <- runEff act $ with Union.(::)
[handleExcept (\e => stLeft e),
[handleExcept $ \e => stLeft e,
handleStateSTRef defs,
handleStateSTRef !(liftST $ newSTRef ns),
handleStateSTRef suf]
pure (res, !(liftST $ readSTRef suf), !(liftST $ readSTRef defs))
export covering
fromParserIO : (MonadRec io, HasIO io) =>
IncludePath -> IORef SeenSet ->
IORef NameSuf -> IORef Definitions ->
Eff FromParserIO a -> io (Either Error a)
fromParserIO inc seen suf defs act =
liftIO $ fromIOErr $ runEff act $ with Union.(::)
[handleExcept (\e => ioLeft e),
handleStateIORef defs,
handleStateIORef !(newIORef [<]),
handleStateIORef suf,
handleLoadFileIOE LoadError WrapParseError seen inc]
handleStateSTRef !(newSTRef' ns),
handleStateSTRef suf,
handleLogST log lvls]
pure $ MkPureParserResult {
val = res,
suf = !(readSTRef' suf),
defs = !(readSTRef' defs),
log = !(readSTRef' log),
logLevels = !(readSTRef' lvls)
}
parameters {auto _ : Functor m} (b : Var n -> m a) (f : PName -> m a)
@ -328,6 +331,7 @@ liftTC : Eff TC a -> Eff FromParserPure a
liftTC tc = runEff tc $ with Union.(::)
[handleExcept $ \e => throw $ WrapTypeError e,
handleReaderConst !(getAt DEFS),
\g => send g,
\g => send g]
private
@ -370,8 +374,9 @@ data HasFail = NoFail | AnyFail | FailWith String
export covering
expectFail : Loc -> Eff FromParserPure a -> Eff FromParserPure Error
expectFail loc act =
case fromParserPure !(getAt GEN) !(getAt DEFS) {ns = !(getAt NS)} act of
expectFail loc act = do
gen <- getAt GEN; defs <- getAt DEFS; ns <- getAt NS; lvl <- curLevels
case fromParserPure {ns} gen defs (singleton lvl) act of
Left err => pure err
Right _ => throw $ ExpectedFail loc

View File

@ -59,10 +59,15 @@ Traversable (IfConsistent eqs) where
traverse f Nothing = pure Nothing
traverse f (Just x) = Just <$> f x
public export
ifConsistentElse : Applicative f => (eqs : DimEq d) ->
f a -> f () -> f (IfConsistent eqs a)
ifConsistentElse ZeroIsOne yes no = Nothing <$ no
ifConsistentElse (C _) yes no = Just <$> yes
public export
ifConsistent : Applicative f => (eqs : DimEq d) -> f a -> f (IfConsistent eqs a)
ifConsistent ZeroIsOne act = pure Nothing
ifConsistent (C _) act = Just <$> act
ifConsistent eqs act = ifConsistentElse eqs act (pure ())
public export
toMaybe : IfConsistent eqs a -> Maybe a

View File

@ -3,6 +3,7 @@ module Quox.Typechecker
import public Quox.Typing
import public Quox.Equal
import Quox.Displace
import Quox.Pretty
import Data.List
import Data.SnocVect
@ -14,7 +15,7 @@ import Quox.EffExtra
public export
0 TC : List (Type -> Type)
TC = [ErrorEff, DefsReader, NameGen]
TC = [ErrorEff, DefsReader, NameGen, Log]
parameters (loc : Loc)
@ -41,6 +42,24 @@ lubs ctx [] = zeroFor ctx
lubs ctx (x :: xs) = lubs1 $ x ::: xs
private
prettyTermTC : {opts : LayoutOpts} ->
TyContext d n -> Term d n -> Eff Pretty (Doc opts)
prettyTermTC ctx s = prettyTerm ctx.dnames ctx.tnames s
private
checkLogs : String -> TyContext d n -> SQty ->
Term d n -> Maybe (Term d n) -> Eff TC ()
checkLogs fun ctx sg subj ty = do
let tyDoc = delay $ maybe (text "none") (runPretty . prettyTermTC ctx) ty
sayMany "check" subj.loc
[10 :> text fun,
95 :> hsep ["ctx =", runPretty $ prettyTyContext ctx],
95 :> hsep ["sg =", runPretty $ prettyQty sg.qty],
10 :> hsep ["subj =", runPretty $ prettyTermTC ctx subj],
10 :> hsep ["ty =", tyDoc]]
mutual
||| "Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ"
|||
@ -53,7 +72,11 @@ mutual
export covering %inline
check : (ctx : TyContext d n) -> SQty -> Term d n -> Term d n ->
Eff TC (CheckResult ctx.dctx n)
check ctx sg subj ty = ifConsistent ctx.dctx $ checkC ctx sg subj ty
check ctx sg subj ty =
ifConsistentElse ctx.dctx
(do checkLogs "check" ctx sg subj (Just ty)
checkC ctx sg subj ty)
(say "check" 20 subj.loc "check: 0=1")
||| "Ψ | Γ ⊢₀ s ⇐ A"
|||
@ -84,7 +107,12 @@ mutual
||| universe doesn't matter, only that a term is _a_ type, so it is optional.
export covering %inline
checkType : TyContext d n -> Term d n -> Maybe Universe -> Eff TC ()
checkType ctx subj l = ignore $ ifConsistent ctx.dctx $ checkTypeC ctx subj l
checkType ctx subj l = do
let univ = TYPE <$> l <*> pure noLoc
ignore $ ifConsistentElse ctx.dctx
(do checkLogs "checkType" ctx SZero subj univ
checkTypeC ctx subj l)
(say "check" 20 subj.loc "checkType: 0=1")
export covering %inline
checkTypeC : TyContext d n -> Term d n -> Maybe Universe -> Eff TC ()
@ -107,7 +135,11 @@ mutual
export covering %inline
infer : (ctx : TyContext d n) -> SQty -> Elim d n ->
Eff TC (InferResult ctx.dctx d n)
infer ctx sg subj = ifConsistent ctx.dctx $ inferC ctx sg subj
infer ctx sg subj = do
ifConsistentElse ctx.dctx
(do checkLogs "infer" ctx sg (E subj) Nothing
inferC ctx sg subj)
(say "check" 20 subj.loc "infer: 0=1")
||| `infer`, assuming the dimension context is consistent
export covering %inline

View File

@ -7,6 +7,7 @@ import public Quox.Typing.Error as Typing
import public Quox.Syntax
import public Quox.Definition
import public Quox.Whnf
import public Quox.Pretty
import Language.Reflection
import Control.Eff
@ -46,16 +47,15 @@ lookupFree x loc defs = maybe (throw $ NotInScope loc x) pure $ lookup x defs
public export
substCasePairRet : BContext 2 -> Term d n -> ScopeTerm d n -> Term d (2 + n)
substCasePairRet [< x, y] dty retty =
let tm = Pair (BVT 1 x.loc) (BVT 0 y.loc) $ x.loc `extendL` y.loc
arg = Ann tm (dty // fromNat 2) tm.loc
in
let tm = Pair (BVT 1 x.loc) (BVT 0 y.loc) $ x.loc `extendL` y.loc
arg = Ann tm (dty // fromNat 2) tm.loc in
retty.term // (arg ::: shift 2)
public export
substCaseSuccRet : BContext 2 -> ScopeTerm d n -> Term d (2 + n)
substCaseSuccRet [< p, ih] retty =
let arg = Ann (Succ (BVT 1 p.loc) p.loc) (NAT p.loc) $ p.loc `extendL` ih.loc
in
let loc = p.loc `extendL` ih.loc
arg = Ann (Succ (BVT 1 p.loc) p.loc) (NAT p.loc) loc in
retty.term // (arg ::: shift 2)
public export
@ -65,23 +65,31 @@ substCaseBoxRet x dty retty =
retty.term // (arg ::: shift 1)
parameters (defs : Definitions) {auto _ : (Has ErrorEff fs, Has NameGen fs)}
private
0 ExpectErrorConstructor : Type
ExpectErrorConstructor =
forall d, n. Loc -> NameContexts d n -> Term d n -> Error
parameters (defs : Definitions)
{auto _ : (Has ErrorEff fs, Has NameGen fs, Has Log fs)}
namespace TyContext
parameters (ctx : TyContext d n) (sg : SQty) (loc : Loc)
export covering
whnf : {0 isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
tm d n -> Eff fs (NonRedex tm d n defs ? sg)
tm d n -> Eff fs (NonRedex tm d n defs (toWhnfContext ctx) sg)
whnf tm = do
let Val n = ctx.termLen; Val d = ctx.dimLen
res <- lift $ runExcept $ whnf defs (toWhnfContext ctx) sg tm
rethrow res
private covering %macro
expect : (forall d, n. Loc -> NameContexts d n -> Term d n -> Error) ->
TTImp -> TTImp -> Elab (Term d n -> Eff fs a)
expect k l r = do
f <- check `(\case ~(l) => Just ~(r); _ => Nothing)
pure $ \t => maybe (throw $ k loc ctx.names t) pure . f . fst =<< whnf t
expect : ExpectErrorConstructor -> TTImp -> TTImp ->
Elab (Term d n -> Eff fs a)
expect err pat rhs = Prelude.do
match <- check `(\case ~(pat) => Just ~(rhs); _ => Nothing)
pure $ \term => do
res <- whnf term
maybe (throw $ err loc ctx.names term) pure $ match $ fst res
export covering %inline
expectTYPE : Term d n -> Eff fs Universe
@ -120,19 +128,20 @@ parameters (defs : Definitions) {auto _ : (Has ErrorEff fs, Has NameGen fs)}
parameters (ctx : EqContext n) (sg : SQty) (loc : Loc)
export covering
whnf : {0 isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
tm 0 n -> Eff fs (NonRedex tm 0 n defs ? sg)
tm 0 n -> Eff fs (NonRedex tm 0 n defs (toWhnfContext ctx) sg)
whnf tm = do
res <- lift $ runExcept $ whnf defs (toWhnfContext ctx) sg tm
rethrow res
private covering %macro
expect : (forall d, n. Loc -> NameContexts d n -> Term d n -> Error) ->
TTImp -> TTImp -> Elab (Term 0 n -> Eff fs a)
expect k l r = do
f <- check `(\case ~(l) => Just ~(r); _ => Nothing)
pure $ \t =>
let err = throw $ k loc ctx.names (t // shift0 ctx.dimLen) in
maybe err pure . f . fst =<< whnf t
expect : ExpectErrorConstructor -> TTImp -> TTImp ->
Elab (Term 0 n -> Eff fs a)
expect err pat rhs = do
match <- check `(\case ~(pat) => Just ~(rhs); _ => Nothing)
pure $ \term => do
res <- whnf term
let t0 = delay $ term // shift0 ctx.dimLen
maybe (throw $ err loc ctx.names t0) pure $ match $ fst res
export covering %inline
expectTYPE : Term 0 n -> Eff fs Universe

View File

@ -339,9 +339,10 @@ namespace WhnfContext
private
prettyTContextElt : {opts : _} ->
BContext d -> BContext n ->
Qty -> BindName -> LocalVar d n -> Eff Pretty (Doc opts)
Doc opts -> BindName -> LocalVar d n ->
Eff Pretty (Doc opts)
prettyTContextElt dnames tnames q x s = do
q <- prettyQty q; dot <- dotD
dot <- dotD
x <- prettyTBind x; colon <- colonD
ty <- withPrec Outer $ prettyTerm dnames tnames s.type; eq <- cstD
tm <- traverse (withPrec Outer . prettyTerm dnames tnames) s.term
@ -356,7 +357,7 @@ prettyTContextElt dnames tnames q x s = do
private
prettyTContext' : {opts : _} ->
BContext d -> QContext n -> BContext n ->
BContext d -> Context' (Doc opts) n -> BContext n ->
TContext d n -> Eff Pretty (SnocList (Doc opts))
prettyTContext' _ [<] [<] [<] = pure [<]
prettyTContext' dnames (qtys :< q) (tnames :< x) (tys :< t) =
@ -369,6 +370,7 @@ prettyTContext : {opts : _} ->
TContext d n -> Eff Pretty (Doc opts)
prettyTContext dnames qtys tnames tys = do
comma <- commaD
qtys <- traverse prettyQty qtys
sepSingle . exceptLast (<+> comma) . toList <$>
prettyTContext' dnames qtys tnames tys
@ -384,3 +386,10 @@ prettyTyContext (MkTyContext dctx dnames tctx tnames qtys) =
export
prettyEqContext : {opts : _} -> EqContext n -> Eff Pretty (Doc opts)
prettyEqContext ctx = prettyTyContext $ toTyContext ctx
export
prettyWhnfContext : {opts : _} -> WhnfContext d n -> Eff Pretty (Doc opts)
prettyWhnfContext ctx =
let Val n = ctx.termLen in
separateTight !commaD <$>
prettyTContext' ctx.dnames (replicate n "_") ctx.tnames ctx.tctx

View File

@ -88,7 +88,7 @@ parameters {opts : LayoutOpts} (showContext : Bool)
public export
Erase : List (Type -> Type)
Erase = [Except Error, NameGen]
Erase = [Except Error, NameGen, Log]
export
liftWhnf : Eff Whnf a -> Eff Erase a

View File

@ -2,6 +2,7 @@ module Quox.Whnf.ComputeElimType
import Quox.Whnf.Interface
import Quox.Displace
import Quox.Pretty
%default total
@ -18,7 +19,6 @@ computeElimType :
(e : Elim d n) -> (0 ne : No (isRedexE defs ctx sg e)) =>
Eff Whnf (Term d n)
||| computes a type and then reduces it to whnf
export covering
computeWhnfElimType0 :
@ -28,7 +28,16 @@ computeWhnfElimType0 :
(e : Elim d n) -> (0 ne : No (isRedexE defs ctx sg e)) =>
Eff Whnf (Term d n)
computeElimType defs ctx sg e =
private covering
computeElimTypeNoLog, computeWhnfElimType0NoLog :
CanWhnf Term Interface.isRedexT =>
CanWhnf Elim Interface.isRedexE =>
(defs : Definitions) -> WhnfContext d n -> (0 sg : SQty) ->
(e : Elim d n) -> (0 ne : No (isRedexE defs ctx sg e)) =>
Eff Whnf (Term d n)
computeElimTypeNoLog defs ctx sg e =
case e of
F x u loc => do
let Just def = lookup x defs
@ -39,7 +48,7 @@ computeElimType defs ctx sg e =
pure (ctx.tctx !! i).type
App f s loc =>
case !(computeWhnfElimType0 defs ctx sg f {ne = noOr1 ne}) of
case !(computeWhnfElimType0NoLog defs ctx sg f {ne = noOr1 ne}) of
Pi {arg, res, _} => pure $ sub1 res $ Ann s arg loc
ty => throw $ ExpectedPi loc ctx.names ty
@ -47,12 +56,12 @@ computeElimType defs ctx sg e =
pure $ sub1 ret pair
Fst pair loc =>
case !(computeWhnfElimType0 defs ctx sg pair {ne = noOr1 ne}) of
case !(computeWhnfElimType0NoLog defs ctx sg pair {ne = noOr1 ne}) of
Sig {fst, _} => pure fst
ty => throw $ ExpectedSig loc ctx.names ty
Snd pair loc =>
case !(computeWhnfElimType0 defs ctx sg pair {ne = noOr1 ne}) of
case !(computeWhnfElimType0NoLog defs ctx sg pair {ne = noOr1 ne}) of
Sig {snd, _} => pure $ sub1 snd $ Fst pair loc
ty => throw $ ExpectedSig loc ctx.names ty
@ -66,7 +75,7 @@ computeElimType defs ctx sg e =
pure $ sub1 ret box
DApp {fun = f, arg = p, loc} =>
case !(computeWhnfElimType0 defs ctx sg f {ne = noOr1 ne}) of
case !(computeWhnfElimType0NoLog defs ctx sg f {ne = noOr1 ne}) of
Eq {ty, _} => pure $ dsub1 ty p
t => throw $ ExpectedEq loc ctx.names t
@ -82,5 +91,20 @@ computeElimType defs ctx sg e =
TypeCase {ret, _} =>
pure ret
computeElimType defs ctx sg e {ne} = do
let Val n = ctx.termLen
sayMany "whnf" e.loc
[90 :> "computeElimType",
95 :> hsep ["ctx =", runPretty $ prettyWhnfContext ctx],
90 :> hsep ["e =", runPretty $ prettyElim ctx.dnames ctx.tnames e]]
res <- computeElimTypeNoLog defs ctx sg e {ne}
say "whnf" 91 e.loc $
hsep ["computeElimType ⇝",
runPretty $ prettyTerm ctx.dnames ctx.tnames res]
pure res
computeWhnfElimType0 defs ctx sg e =
computeElimType defs ctx sg e >>= whnf0 defs ctx SZero
computeWhnfElimType0NoLog defs ctx sg e {ne} =
computeElimTypeNoLog defs ctx sg e {ne} >>= whnf0 defs ctx SZero

View File

@ -1,6 +1,7 @@
module Quox.Whnf.Interface
import public Quox.No
import public Quox.Log
import public Quox.Syntax
import public Quox.Definition
import public Quox.Typing.Context
@ -13,7 +14,7 @@ import public Control.Eff
public export
Whnf : List (Type -> Type)
Whnf = [Except Error, NameGen]
Whnf = [Except Error, NameGen, Log]
public export
@ -24,17 +25,20 @@ RedexTest tm =
public export
interface CanWhnf (0 tm : TermLike) (0 isRedex : RedexTest tm) | tm
where
whnf : (defs : Definitions) -> (ctx : WhnfContext d n) -> (q : SQty) ->
tm d n -> Eff Whnf (Subset (tm d n) (No . isRedex defs ctx q))
whnf, whnfNoLog :
(defs : Definitions) -> (ctx : WhnfContext d n) -> (q : SQty) ->
tm d n -> Eff Whnf (Subset (tm d n) (No . isRedex defs ctx q))
-- having isRedex be part of the class header, and needing to be explicitly
-- quantified on every use since idris can't infer its type, is a little ugly.
-- but none of the alternatives i've thought of so far work. e.g. in some
-- cases idris can't tell that `isRedex` and `isRedexT` are the same thing
public export %inline
whnf0 : {0 isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
Definitions -> WhnfContext d n -> SQty -> tm d n -> Eff Whnf (tm d n)
whnf0 defs ctx q t = fst <$> whnf defs ctx q t
whnf0, whnfNoLog0 :
{0 isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
Definitions -> WhnfContext d n -> SQty -> tm d n -> Eff Whnf (tm d n)
whnf0 defs ctx q t = fst <$> whnf defs ctx q t
whnfNoLog0 defs ctx q t = fst <$> whnfNoLog defs ctx q t
public export
0 IsRedex, NotRedex : {isRedex : RedexTest tm} -> CanWhnf tm isRedex =>

View File

@ -4,6 +4,7 @@ import Quox.Whnf.Interface
import Quox.Whnf.ComputeElimType
import Quox.Whnf.TypeCase
import Quox.Whnf.Coercion
import Quox.Pretty
import Quox.Displace
import Data.SnocVect
@ -14,19 +15,43 @@ export covering CanWhnf Term Interface.isRedexT
export covering CanWhnf Elim Interface.isRedexE
-- the String is what to call the "s" argument in logs (maybe "s", or "e")
private %inline
whnfDefault :
{0 isRedex : RedexTest tm} ->
(CanWhnf tm isRedex, Located2 tm) =>
String ->
(forall d, n. WhnfContext d n -> tm d n -> Eff Pretty LogDoc) ->
(defs : Definitions) ->
(ctx : WhnfContext d n) ->
(sg : SQty) ->
(s : tm d n) ->
Eff Whnf (Subset (tm d n) (No . isRedex defs ctx sg))
whnfDefault name ppr defs ctx sg s = do
sayMany "whnf" s.loc
[10 :> "whnf",
95 :> hsep ["ctx =", runPretty $ prettyWhnfContext ctx],
95 :> hsep ["sg =", runPretty $ prettyQty sg.qty],
10 :> hsep [text name, "=", runPretty $ ppr ctx s]]
res <- whnfNoLog defs ctx sg s
say "whnf" 11 s.loc $ hsep ["whnf ⇝", runPretty $ ppr ctx res.fst]
pure res
covering
CanWhnf Elim Interface.isRedexE where
whnf defs ctx sg (F x u loc) with (lookupElim0 x u defs) proof eq
whnf = whnfDefault "e" $ \ctx, e => prettyElim ctx.dnames ctx.tnames e
whnfNoLog defs ctx sg (F x u loc) with (lookupElim0 x u defs) proof eq
_ | Just y = whnf defs ctx sg $ setLoc loc $ injElim ctx y
_ | Nothing = pure $ Element (F x u loc) $ rewrite eq in Ah
whnf defs ctx sg (B i loc) with (ctx.tctx !! i) proof eq1
whnfNoLog defs ctx sg (B i loc) with (ctx.tctx !! i) proof eq1
_ | l with (l.term) proof eq2
_ | Just y = whnf defs ctx sg $ Ann y l.type loc
_ | Nothing = pure $ Element (B i loc) $ rewrite eq1 in rewrite eq2 in Ah
-- ((λ x ⇒ t) ∷ (π.x : A) → B) s ⇝ t[s∷A/x] ∷ B[s∷A/x]
whnf defs ctx sg (App f s appLoc) = do
whnfNoLog defs ctx sg (App f s appLoc) = do
Element f fnf <- whnf defs ctx sg f
case nchoose $ isLamHead f of
Left _ => case f of
@ -41,7 +66,7 @@ CanWhnf Elim Interface.isRedexE where
--
-- 0 · case e return p ⇒ C of { (a, b) ⇒ u } ⇝
-- u[fst e/a, snd e/b] ∷ C[e/p]
whnf defs ctx sg (CasePair pi pair ret body caseLoc) = do
whnfNoLog defs ctx sg (CasePair pi pair ret body caseLoc) = do
Element pair pairnf <- whnf defs ctx sg pair
case nchoose $ isPairHead pair of
Left _ => case pair of
@ -64,7 +89,7 @@ CanWhnf Elim Interface.isRedexE where
(pairnf `orNo` np `orNo` notYesNo n0)
-- fst ((s, t) ∷ (x : A) × B) ⇝ s ∷ A
whnf defs ctx sg (Fst pair fstLoc) = do
whnfNoLog defs ctx sg (Fst pair fstLoc) = do
Element pair pairnf <- whnf defs ctx sg pair
case nchoose $ isPairHead pair of
Left _ => case pair of
@ -76,7 +101,7 @@ CanWhnf Elim Interface.isRedexE where
pure $ Element (Fst pair fstLoc) (pairnf `orNo` np)
-- snd ((s, t) ∷ (x : A) × B) ⇝ t ∷ B[(s ∷ A)/x]
whnf defs ctx sg (Snd pair sndLoc) = do
whnfNoLog defs ctx sg (Snd pair sndLoc) = do
Element pair pairnf <- whnf defs ctx sg pair
case nchoose $ isPairHead pair of
Left _ => case pair of
@ -89,7 +114,7 @@ CanWhnf Elim Interface.isRedexE where
-- case 'a ∷ {a,…} return p ⇒ C of { 'a ⇒ u } ⇝
-- u ∷ C['a∷{a,…}/p]
whnf defs ctx sg (CaseEnum pi tag ret arms caseLoc) = do
whnfNoLog defs ctx sg (CaseEnum pi tag ret arms caseLoc) = do
Element tag tagnf <- whnf defs ctx sg tag
case nchoose $ isTagHead tag of
Left _ => case tag of
@ -110,7 +135,7 @@ CanWhnf Elim Interface.isRedexE where
--
-- case succ n ∷ return p ⇒ C of { succ n', π.ih ⇒ u; … } ⇝
-- u[n∷/n', (case n ∷ ⋯)/ih] ∷ C[succ n ∷ /p]
whnf defs ctx sg (CaseNat pi piIH nat ret zer suc caseLoc) = do
whnfNoLog defs ctx sg (CaseNat pi piIH nat ret zer suc caseLoc) = do
Element nat natnf <- whnf defs ctx sg nat
case nchoose $ isNatHead nat of
Left _ =>
@ -137,7 +162,7 @@ CanWhnf Elim Interface.isRedexE where
-- case [t] ∷ [π.A] return p ⇒ C of { [x] ⇒ u } ⇝
-- u[t∷A/x] ∷ C[[t] ∷ [π.A]/p]
whnf defs ctx sg (CaseBox pi box ret body caseLoc) = do
whnfNoLog defs ctx sg (CaseBox pi box ret body caseLoc) = do
Element box boxnf <- whnf defs ctx sg box
case nchoose $ isBoxHead box of
Left _ => case box of
@ -153,7 +178,7 @@ CanWhnf Elim Interface.isRedexE where
-- e : Eq (𝑗 ⇒ A) t u ⊢ e @1 ⇝ u ∷ A1/𝑗
--
-- ((δ 𝑖 ⇒ s) ∷ Eq (𝑗 ⇒ A) t u) @𝑘 ⇝ s𝑘/𝑖 ∷ A𝑘/𝑗
whnf defs ctx sg (DApp f p appLoc) = do
whnfNoLog defs ctx sg (DApp f p appLoc) = do
Element f fnf <- whnf defs ctx sg f
case nchoose $ isDLamHead f of
Left _ => case f of
@ -173,7 +198,7 @@ CanWhnf Elim Interface.isRedexE where
B {} => pure $ Element (DApp f p appLoc) (fnf `orNo` ndlh `orNo` Ah)
-- e ∷ A ⇝ e
whnf defs ctx sg (Ann s a annLoc) = do
whnfNoLog defs ctx sg (Ann s a annLoc) = do
Element s snf <- whnf defs ctx sg s
case nchoose $ isE s of
Left _ => let E e = s in pure $ Element e $ noOr2 snf
@ -181,7 +206,7 @@ CanWhnf Elim Interface.isRedexE where
Element a anf <- whnf defs ctx SZero a
pure $ Element (Ann s a annLoc) (ne `orNo` snf `orNo` anf)
whnf defs ctx sg (Coe sty p q val coeLoc) =
whnfNoLog defs ctx sg (Coe sty p q val coeLoc) =
-- 𝑖 ∉ fv(A)
-- -------------------------------
-- coe (𝑖 ⇒ A) @p @q s ⇝ s ∷ A
@ -201,7 +226,7 @@ CanWhnf Elim Interface.isRedexE where
(_, Right ty) =>
whnf defs ctx sg $ Ann val ty coeLoc
whnf defs ctx sg (Comp ty p q val r zero one compLoc) =
whnfNoLog defs ctx sg (Comp ty p q val r zero one compLoc) =
case p `decEqv` q of
-- comp [A] @p @p s @r { ⋯ } ⇝ s ∷ A
Yes y => whnf defs ctx sg $ Ann val ty compLoc
@ -213,7 +238,7 @@ CanWhnf Elim Interface.isRedexE where
B {} => pure $ Element (Comp ty p q val r zero one compLoc)
(notYesNo npq `orNo` Ah)
whnf defs ctx sg (TypeCase ty ret arms def tcLoc) =
whnfNoLog defs ctx sg (TypeCase ty ret arms def tcLoc) =
case sg `decEq` SZero of
Yes Refl => do
Element ty tynf <- whnf defs ctx SZero ty
@ -226,48 +251,50 @@ CanWhnf Elim Interface.isRedexE where
No _ =>
throw $ ClashQ tcLoc sg.qty Zero
whnf defs ctx sg (CloE (Sub el th)) =
whnf defs ctx sg $ pushSubstsWith' id th el
whnf defs ctx sg (DCloE (Sub el th)) =
whnf defs ctx sg $ pushSubstsWith' th id el
whnfNoLog defs ctx sg (CloE (Sub el th)) =
whnfNoLog defs ctx sg $ pushSubstsWith' id th el
whnfNoLog defs ctx sg (DCloE (Sub el th)) =
whnfNoLog defs ctx sg $ pushSubstsWith' th id el
covering
CanWhnf Term Interface.isRedexT where
whnf _ _ _ t@(TYPE {}) = pure $ nred t
whnf _ _ _ t@(IOState {}) = pure $ nred t
whnf _ _ _ t@(Pi {}) = pure $ nred t
whnf _ _ _ t@(Lam {}) = pure $ nred t
whnf _ _ _ t@(Sig {}) = pure $ nred t
whnf _ _ _ t@(Pair {}) = pure $ nred t
whnf _ _ _ t@(Enum {}) = pure $ nred t
whnf _ _ _ t@(Tag {}) = pure $ nred t
whnf _ _ _ t@(Eq {}) = pure $ nred t
whnf _ _ _ t@(DLam {}) = pure $ nred t
whnf _ _ _ t@(NAT {}) = pure $ nred t
whnf _ _ _ t@(Nat {}) = pure $ nred t
whnf _ _ _ t@(STRING {}) = pure $ nred t
whnf _ _ _ t@(Str {}) = pure $ nred t
whnf _ _ _ t@(BOX {}) = pure $ nred t
whnf _ _ _ t@(Box {}) = pure $ nred t
whnf = whnfDefault "e" $ \ctx, s => prettyTerm ctx.dnames ctx.tnames s
whnf _ _ _ (Succ p loc) =
whnfNoLog _ _ _ t@(TYPE {}) = pure $ nred t
whnfNoLog _ _ _ t@(IOState {}) = pure $ nred t
whnfNoLog _ _ _ t@(Pi {}) = pure $ nred t
whnfNoLog _ _ _ t@(Lam {}) = pure $ nred t
whnfNoLog _ _ _ t@(Sig {}) = pure $ nred t
whnfNoLog _ _ _ t@(Pair {}) = pure $ nred t
whnfNoLog _ _ _ t@(Enum {}) = pure $ nred t
whnfNoLog _ _ _ t@(Tag {}) = pure $ nred t
whnfNoLog _ _ _ t@(Eq {}) = pure $ nred t
whnfNoLog _ _ _ t@(DLam {}) = pure $ nred t
whnfNoLog _ _ _ t@(NAT {}) = pure $ nred t
whnfNoLog _ _ _ t@(Nat {}) = pure $ nred t
whnfNoLog _ _ _ t@(STRING {}) = pure $ nred t
whnfNoLog _ _ _ t@(Str {}) = pure $ nred t
whnfNoLog _ _ _ t@(BOX {}) = pure $ nred t
whnfNoLog _ _ _ t@(Box {}) = pure $ nred t
whnfNoLog _ _ _ (Succ p loc) =
case nchoose $ isNatConst p of
Left _ => case p of
Nat p _ => pure $ nred $ Nat (S p) loc
E (Ann (Nat p _) _ _) => pure $ nred $ Nat (S p) loc
Right nc => pure $ nred $ Succ p loc
whnf defs ctx sg (Let _ rhs body _) =
whnfNoLog defs ctx sg (Let _ rhs body _) =
whnf defs ctx sg $ sub1 body rhs
-- s ∷ A ⇝ s (in term context)
whnf defs ctx sg (E e) = do
whnfNoLog defs ctx sg (E e) = do
Element e enf <- whnf defs ctx sg e
case nchoose $ isAnn e of
Left _ => let Ann {tm, _} = e in pure $ Element tm $ noOr1 $ noOr2 enf
Right na => pure $ Element (E e) $ na `orNo` enf
whnf defs ctx sg (CloT (Sub tm th)) =
whnf defs ctx sg $ pushSubstsWith' id th tm
whnf defs ctx sg (DCloT (Sub tm th)) =
whnf defs ctx sg $ pushSubstsWith' th id tm
whnfNoLog defs ctx sg (CloT (Sub tm th)) =
whnfNoLog defs ctx sg $ pushSubstsWith' id th tm
whnfNoLog defs ctx sg (DCloT (Sub tm th)) =
whnfNoLog defs ctx sg $ pushSubstsWith' th id tm

View File

@ -19,6 +19,7 @@ modules =
Quox.PrettyValExtra,
Quox.Decidable,
Quox.No,
Quox.Log,
Quox.Loc,
Quox.Var,
Quox.Scoped,

View File

@ -20,7 +20,7 @@
@article{granule,
author = {Dominic Orchard and Vilem{-}Benjamin Liepelt and Harley Eades III},
title = {Quantitative program reasoning with graded modal types},
journal = {Proceedings of the ACM on Programming Languages},
journal = {Proceedings of the {ACM} on Programming Languages},
volume = {3},
number = {{ICFP}},
pages = {110:1--110:30},
@ -198,7 +198,7 @@
number = {POPL},
url = {https://doi.org/10.1145/3498670},
doi = {10.1145/3498670},
journal = {Proc. ACM Program. Lang.},
journal = {Proc. {ACM} Program. Lang.},
month = {jan},
articleno = {9},
numpages = {31},
@ -239,7 +239,7 @@
for universe levels based on displacement algebras, for use in proof
assistant implementations.
},
journal = {Proc. ACM Program. Lang.},
journal = {Proc. {ACM} Program. Lang.},
month = {jan},
articleno = {57},
numpages = {27},
@ -362,9 +362,9 @@
date = {2019-07},
doi = {10.1145/3341711},
issn = {2475-1421},
journaltitle = {Proceedings of the ACM on Programming Languages},
journaltitle = {Proceedings of the {ACM} on Programming Languages},
keywords = {Modal types,dependent types,normalization by evaluation,type-checking},
number = {ICFP},
number = {{ICFP}},
pages = {107:1--107:29},
title = {Implementing a Modal Dependent Type Theory},
volume = {3},
@ -389,7 +389,7 @@
@article{defunc,
author = {Yulong Huang and Jeremy Yallop},
title = {Defunctionalization with Dependent Types},
journal = {Proceedings of the ACM on Programming Languages},
journal = {Proceedings of the {ACM} on Programming Languages},
volume = {7},
number = {{PLDI}},
pages = {516--538},
@ -397,3 +397,18 @@
url = {https://doi.org/10.1145/3591241},
doi = {10.1145/3591241},
}
@inproceedings{delcont-callcc,
author = {Martin Gasbichler and Michael Sperber},
editor = {Mitchell Wand and Simon L. Peyton Jones},
title = {Final shift for \texttt{call/cc}:
direct implementation of shift and reset},
journaltitle = {Proceedings of the {ACM} on Programming Languages},
number = {{ICFP}},
pages = {271--282},
publisher = {{ACM}},
year = {2002},
% url = {https://doi.org/10.1145/581478.581504},
url = {https://www.cs.tufts.edu/~nr/cs257/archive/mike-sperber/shift-reset-direct.pdf},
doi = {10.1145/581478.581504},
}

View File

@ -97,7 +97,7 @@ tests = "dimension constraints" :- [
testPrettyD iijj ZeroIsOne "𝑖, 𝑗, 0 = 1",
testPrettyD [<] new "" {label = "[empty output from empty context]"},
testPrettyD ii new "𝑖",
testPrettyD iijj (fromGround [< "𝑖", "𝑗"] [< Zero, One])
testPrettyD iijj (fromGround iijj [< Zero, One])
"𝑖, 𝑗, 𝑖 = 0, 𝑗 = 1",
testPrettyD iijj (C [< Just (^K Zero), Nothing])
"𝑖, 𝑗, 𝑖 = 0",

View File

@ -68,7 +68,7 @@ parameters {c : Bool} {auto _ : Show b}
runFromParser : {default empty defs : Definitions} ->
Eff FromParserPure a -> Either FromParser.Error a
runFromParser = map fst . fromParserPure 0 defs
runFromParser = map val . fromParserPure 0 defs initStack
export
tests : Test

View File

@ -15,7 +15,8 @@ import Control.Eff
runWhnf : Eff Whnf a -> Either Error a
runWhnf act = runSTErr $ do
runEff act [handleExcept (\e => stLeft e),
handleStateSTRef !(liftST $ newSTRef 0)]
handleStateSTRef !(liftST $ newSTRef 0),
handleLogDiscard]
parameters {0 isRedex : RedexTest tm} {auto _ : CanWhnf tm isRedex} {d, n : Nat}
{auto _ : (Eq (tm d n), Show (tm d n))}

View File

@ -22,10 +22,11 @@ ToInfo Error where
export
runEqual : Definitions -> Eff Equal a -> Either Error a
runEqual defs act = runSTErr $ do
runEff act
runEff act $ with Union.(::)
[handleExcept (\e => stLeft e),
handleReaderConst defs,
handleStateSTRef !(liftST $ newSTRef 0)]
handleStateSTRef !(liftST $ newSTRef 0),
handleLogDiscard]
export
runTC : Definitions -> Eff TC a -> Either Error a