mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-17 16:21:46 +03:00
Merge branch 'master' of github.com:idris-lang/Idris2 into emacs-mode-compatibility
This commit is contained in:
commit
507ace21cc
43
CHANGELOG.md
43
CHANGELOG.md
@ -3,22 +3,18 @@ Changes since Idris 2 v0.1.0
|
|||||||
|
|
||||||
The implementation is now self-hosted. To initialise the build, either use
|
The implementation is now self-hosted. To initialise the build, either use
|
||||||
the [bootstrapping version of Idris2](https://github.com/edwinb/Idris2-boot)
|
the [bootstrapping version of Idris2](https://github.com/edwinb/Idris2-boot)
|
||||||
or build fromt the generated Scheme, using `make bootstrap`.
|
or build from the generated Scheme, using `make bootstrap`.
|
||||||
|
|
||||||
Compiler updates:
|
Language changes:
|
||||||
|
|
||||||
* Data types with a single constructor, with a single unerased arguments,
|
|
||||||
are translated to just that argument, to save repeated packing and unpacking.
|
|
||||||
(c.f. `newtype` in Haskell)
|
|
||||||
+ A data type can opt out of this behaviour by specifying `noNewtype` in its
|
|
||||||
options list. `noNewtype` allows code generators to apply special handling
|
|
||||||
to the generated constructor/deconstructor, for a newtype-like data type,
|
|
||||||
that would otherwise be optimised away.
|
|
||||||
* 0-multiplicity constructor arguments are now properly erased, not just
|
|
||||||
given a placeholder null value.
|
|
||||||
|
|
||||||
Language extensions:
|
|
||||||
|
|
||||||
|
* `total`, `covering` and `partial` flags on functions now have an effect.
|
||||||
|
* Fields of records can be accessed (and updated) using the dot syntax,
|
||||||
|
such as `r.field1.field2` or `record { field1.field2 = 42 }`.
|
||||||
|
For details, see https://idris2.readthedocs.io/en/latest/reference/records.html
|
||||||
|
* New function flag `%tcinline` which means that the function should be
|
||||||
|
inlined for the purposes of totality checking (but otherwise not inlined).
|
||||||
|
This can be used as a hint for totality checking, to make the checker look
|
||||||
|
inside functions that it otherwise might not.
|
||||||
* %transform directive, for declaring transformation rules on runtime
|
* %transform directive, for declaring transformation rules on runtime
|
||||||
expressions. Transformation rules are automatically added for top level
|
expressions. Transformation rules are automatically added for top level
|
||||||
implementations of interfaces.
|
implementations of interfaces.
|
||||||
@ -33,22 +29,27 @@ Library additions:
|
|||||||
* New modules in `contrib` for JSON (`Language.JSON.*`); random numbers
|
* New modules in `contrib` for JSON (`Language.JSON.*`); random numbers
|
||||||
(`System.Random`)
|
(`System.Random`)
|
||||||
|
|
||||||
|
Compiler updates:
|
||||||
|
|
||||||
|
* Data types with a single constructor, with a single unerased arguments,
|
||||||
|
are translated to just that argument, to save repeated packing and unpacking.
|
||||||
|
(c.f. `newtype` in Haskell)
|
||||||
|
+ A data type can opt out of this behaviour by specifying `noNewtype` in its
|
||||||
|
options list. `noNewtype` allows code generators to apply special handling
|
||||||
|
to the generated constructor/deconstructor, for a newtype-like data type,
|
||||||
|
that would otherwise be optimised away.
|
||||||
|
* 0-multiplicity constructor arguments are now properly erased, not just
|
||||||
|
given a placeholder null value.
|
||||||
|
|
||||||
Other improvements:
|
Other improvements:
|
||||||
|
|
||||||
* Various performance improvements in the typechecker:
|
* Various performance improvements in the typechecker:
|
||||||
+ Noting which metavariables are blocking unification constraints, so that
|
+ Noting which metavariables are blocking unification constraints, so that
|
||||||
they only get retried if those metavariables make progress.
|
they only get retried if those metavariables make progress.
|
||||||
+ Evaluating `fromInteger` at compile time.
|
+ Evaluating `fromInteger` at compile time.
|
||||||
+ In the run-time, reuse the old heap after garbage collection if it
|
|
||||||
hasn't been resized, to avoid unnecessary system calls.
|
|
||||||
|
|
||||||
* Extend Idris2's literate mode to support reading Markdown and OrgMode files.
|
* Extend Idris2's literate mode to support reading Markdown and OrgMode files.
|
||||||
For more details see: https://idris2.readthedocs.io/en/latest/reference/literate.html
|
For more details see: https://idris2.readthedocs.io/en/latest/reference/literate.html
|
||||||
|
|
||||||
* Fields of records can be accessed (and updated) using the dot syntax,
|
|
||||||
such as `r.field1.field2` or `record { field1.field2 = 42 }`.
|
|
||||||
For details, see https://idris2.readthedocs.io/en/latest/reference/records.html
|
|
||||||
|
|
||||||
Changes since Idris 1
|
Changes since Idris 1
|
||||||
---------------------
|
---------------------
|
||||||
|
|
||||||
|
@ -29,7 +29,7 @@ data Grammar : (tok : Type) -> (consumes : Bool) -> Type -> Type where
|
|||||||
||| guaranteed to consume some input. If the first one consumes input, the
|
||| guaranteed to consume some input. If the first one consumes input, the
|
||||||
||| second is allowed to be recursive (because it means some input has been
|
||| second is allowed to be recursive (because it means some input has been
|
||||||
||| consumed and therefore the input is smaller)
|
||| consumed and therefore the input is smaller)
|
||||||
public export %inline
|
public export %inline %tcinline
|
||||||
(>>=) : {c1, c2 : Bool} ->
|
(>>=) : {c1, c2 : Bool} ->
|
||||||
Grammar tok c1 a ->
|
Grammar tok c1 a ->
|
||||||
inf c1 (a -> Grammar tok c2 b) ->
|
inf c1 (a -> Grammar tok c2 b) ->
|
||||||
|
@ -170,6 +170,9 @@ data DefFlag
|
|||||||
-- should evaluate the RHS, with reduction limits on the given names,
|
-- should evaluate the RHS, with reduction limits on the given names,
|
||||||
-- and ensure the name has made progress in doing so (i.e. has reduced
|
-- and ensure the name has made progress in doing so (i.e. has reduced
|
||||||
-- at least once)
|
-- at least once)
|
||||||
|
| AllGuarded -- safe to treat as a constructor for the purposes of
|
||||||
|
-- productivity checking. All clauses are guarded by constructors,
|
||||||
|
-- and there are no other function applications
|
||||||
|
|
||||||
export
|
export
|
||||||
Eq DefFlag where
|
Eq DefFlag where
|
||||||
@ -181,6 +184,7 @@ Eq DefFlag where
|
|||||||
(==) BlockedHint BlockedHint = True
|
(==) BlockedHint BlockedHint = True
|
||||||
(==) Macro Macro = True
|
(==) Macro Macro = True
|
||||||
(==) (PartialEval x) (PartialEval y) = x == y
|
(==) (PartialEval x) (PartialEval y) = x == y
|
||||||
|
(==) AllGuarded AllGuarded = True
|
||||||
(==) _ _ = False
|
(==) _ _ = False
|
||||||
|
|
||||||
export
|
export
|
||||||
@ -193,6 +197,7 @@ Show DefFlag where
|
|||||||
show BlockedHint = "blockedhint"
|
show BlockedHint = "blockedhint"
|
||||||
show Macro = "macro"
|
show Macro = "macro"
|
||||||
show (PartialEval _) = "partialeval"
|
show (PartialEval _) = "partialeval"
|
||||||
|
show AllGuarded = "allguarded"
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data SizeChange = Smaller | Same | Unknown
|
data SizeChange = Smaller | Same | Unknown
|
||||||
|
@ -107,7 +107,7 @@ parameters (defs : Defs, topopts : EvalOpts)
|
|||||||
eval env locs (Bind fc x (Lam r _ ty) scope) (thunk :: stk)
|
eval env locs (Bind fc x (Lam r _ ty) scope) (thunk :: stk)
|
||||||
= eval env (thunk :: locs) scope stk
|
= eval env (thunk :: locs) scope stk
|
||||||
eval env locs (Bind fc x b@(Let r val ty) scope) stk
|
eval env locs (Bind fc x b@(Let r val ty) scope) stk
|
||||||
= if holesOnly topopts || argHolesOnly topopts
|
= if holesOnly topopts || argHolesOnly topopts && not (tcInline topopts)
|
||||||
then do b' <- traverse (\tm => eval env locs tm []) b
|
then do b' <- traverse (\tm => eval env locs tm []) b
|
||||||
pure $ NBind fc x b'
|
pure $ NBind fc x b'
|
||||||
(\defs', arg => evalWithOpts defs' topopts
|
(\defs', arg => evalWithOpts defs' topopts
|
||||||
@ -595,8 +595,10 @@ mutual
|
|||||||
pure (TDelay fc r tyQ argQ)
|
pure (TDelay fc r tyQ argQ)
|
||||||
where
|
where
|
||||||
toHolesOnly : Closure vs -> Closure vs
|
toHolesOnly : Closure vs -> Closure vs
|
||||||
toHolesOnly (MkClosure _ locs env tm)
|
toHolesOnly (MkClosure opts locs env tm)
|
||||||
= MkClosure withHoles locs env tm
|
= MkClosure (record { holesOnly = True,
|
||||||
|
argHolesOnly = True } opts)
|
||||||
|
locs env tm
|
||||||
toHolesOnly c = c
|
toHolesOnly c = c
|
||||||
quoteGenNF q defs bound env (NForce fc r arg args)
|
quoteGenNF q defs bound env (NForce fc r arg args)
|
||||||
= do args' <- quoteArgs q defs bound env args
|
= do args' <- quoteArgs q defs bound env args
|
||||||
|
@ -865,6 +865,7 @@ TTC DefFlag where
|
|||||||
toBuf b BlockedHint = tag 7
|
toBuf b BlockedHint = tag 7
|
||||||
toBuf b Macro = tag 8
|
toBuf b Macro = tag 8
|
||||||
toBuf b (PartialEval x) = tag 9 -- names not useful any more
|
toBuf b (PartialEval x) = tag 9 -- names not useful any more
|
||||||
|
toBuf b AllGuarded = tag 10
|
||||||
|
|
||||||
fromBuf b
|
fromBuf b
|
||||||
= case !getTag of
|
= case !getTag of
|
||||||
@ -876,6 +877,7 @@ TTC DefFlag where
|
|||||||
7 => pure BlockedHint
|
7 => pure BlockedHint
|
||||||
8 => pure Macro
|
8 => pure Macro
|
||||||
9 => pure (PartialEval [])
|
9 => pure (PartialEval [])
|
||||||
|
10 => pure AllGuarded
|
||||||
_ => corrupt "DefFlag"
|
_ => corrupt "DefFlag"
|
||||||
|
|
||||||
export
|
export
|
||||||
|
@ -34,6 +34,55 @@ totRefsIn : {auto c : Ref Ctxt Defs} ->
|
|||||||
Defs -> Term vars -> Core Terminating
|
Defs -> Term vars -> Core Terminating
|
||||||
totRefsIn defs ty = totRefs defs (keys (getRefs (Resolved (-1)) ty))
|
totRefsIn defs ty = totRefs defs (keys (getRefs (Resolved (-1)) ty))
|
||||||
|
|
||||||
|
-- Check if all branches end up as constructor arguments, with no other
|
||||||
|
-- function application, and set 'AllGuarded' if so.
|
||||||
|
-- This is to check whether a function can be considered a constructor form
|
||||||
|
-- for the sake of termination checking (and might have other uses...)
|
||||||
|
export
|
||||||
|
checkIfGuarded : {auto c : Ref Ctxt Defs} ->
|
||||||
|
FC -> Name -> Core ()
|
||||||
|
checkIfGuarded fc n
|
||||||
|
= do defs <- get Ctxt
|
||||||
|
Just (PMDef _ _ _ _ pats) <- lookupDefExact n (gamma defs)
|
||||||
|
| _ => pure ()
|
||||||
|
t <- allGuarded pats
|
||||||
|
when t $ setFlag fc n AllGuarded
|
||||||
|
where
|
||||||
|
guardedNF : {vars : _} -> Defs -> Env Term vars -> NF vars -> Core Bool
|
||||||
|
guardedNF defs env (NDCon _ _ _ _ args) = pure True
|
||||||
|
guardedNF defs env (NApp _ (NRef _ n) args)
|
||||||
|
= do Just gdef <- lookupCtxtExact n (gamma defs)
|
||||||
|
| Nothing => pure False
|
||||||
|
pure (AllGuarded `elem` flags gdef)
|
||||||
|
guardedNF defs env _ = pure False
|
||||||
|
|
||||||
|
checkNotFn : Defs -> Name -> Core Bool
|
||||||
|
checkNotFn defs n
|
||||||
|
= do Just gdef <- lookupCtxtExact n (gamma defs)
|
||||||
|
| Nothing => pure False
|
||||||
|
case definition gdef of
|
||||||
|
DCon _ _ _ => pure True
|
||||||
|
_ => pure (multiplicity gdef == erased
|
||||||
|
|| (AllGuarded `elem` flags gdef))
|
||||||
|
|
||||||
|
guarded : {vars : _} -> Env Term vars -> Term vars -> Core Bool
|
||||||
|
guarded env tm
|
||||||
|
= do defs <- get Ctxt
|
||||||
|
empty <- clearDefs defs
|
||||||
|
tmnf <- nf empty env tm
|
||||||
|
if !(guardedNF defs env tmnf)
|
||||||
|
then do Just gdef <- lookupCtxtExact n (gamma defs)
|
||||||
|
| Nothing => pure False
|
||||||
|
allM (checkNotFn defs) (keys (refersTo gdef))
|
||||||
|
else pure False
|
||||||
|
|
||||||
|
allGuarded : List (vs ** (Env Term vs, Term vs, Term vs)) -> Core Bool
|
||||||
|
allGuarded [] = pure True
|
||||||
|
allGuarded ((_ ** (env, lhs, rhs)) :: ps)
|
||||||
|
= if !(guarded env rhs)
|
||||||
|
then allGuarded ps
|
||||||
|
else pure False
|
||||||
|
|
||||||
-- Equal for the purposes of size change means, ignoring as patterns, all
|
-- Equal for the purposes of size change means, ignoring as patterns, all
|
||||||
-- non-metavariable positions are equal
|
-- non-metavariable positions are equal
|
||||||
scEq : Term vars -> Term vars -> Bool
|
scEq : Term vars -> Term vars -> Bool
|
||||||
@ -79,8 +128,13 @@ mutual
|
|||||||
-- If we're Guarded and find a Delay, continue with the argument as InDelay
|
-- If we're Guarded and find a Delay, continue with the argument as InDelay
|
||||||
findSC defs env Guarded pats (TDelay _ _ _ tm)
|
findSC defs env Guarded pats (TDelay _ _ _ tm)
|
||||||
= findSC defs env InDelay pats tm
|
= findSC defs env InDelay pats tm
|
||||||
|
findSC defs env g pats (TDelay _ _ _ tm)
|
||||||
|
= findSC defs env g pats tm
|
||||||
findSC defs env g pats tm
|
findSC defs env g pats tm
|
||||||
= case (g, getFnArgs tm) of
|
= do let (fn, args) = getFnArgs tm
|
||||||
|
fn' <- conIfGuarded fn -- pretend it's a data constructor if
|
||||||
|
-- it has the AllGuarded flag
|
||||||
|
case (g, fn', args) of
|
||||||
-- If we're InDelay and find a constructor (or a function call which is
|
-- If we're InDelay and find a constructor (or a function call which is
|
||||||
-- guaranteed to return a constructor; AllGuarded set), continue as InDelay
|
-- guaranteed to return a constructor; AllGuarded set), continue as InDelay
|
||||||
(InDelay, Ref fc (DataCon _ _) cn, args) =>
|
(InDelay, Ref fc (DataCon _ _) cn, args) =>
|
||||||
@ -91,6 +145,9 @@ mutual
|
|||||||
(InDelay, _, args) =>
|
(InDelay, _, args) =>
|
||||||
do scs <- traverse (findSC defs env Unguarded pats) args
|
do scs <- traverse (findSC defs env Unguarded pats) args
|
||||||
pure (concat scs)
|
pure (concat scs)
|
||||||
|
(Guarded, Ref fc (DataCon _ _) cn, args) =>
|
||||||
|
do scs <- traverse (findSC defs env Guarded pats) args
|
||||||
|
pure (concat scs)
|
||||||
(Toplevel, Ref fc (DataCon _ _) cn, args) =>
|
(Toplevel, Ref fc (DataCon _ _) cn, args) =>
|
||||||
do scs <- traverse (findSC defs env Guarded pats) args
|
do scs <- traverse (findSC defs env Guarded pats) args
|
||||||
pure (concat scs)
|
pure (concat scs)
|
||||||
@ -103,6 +160,16 @@ mutual
|
|||||||
(_, f, args) =>
|
(_, f, args) =>
|
||||||
do scs <- traverse (findSC defs env Unguarded pats) args
|
do scs <- traverse (findSC defs env Unguarded pats) args
|
||||||
pure (concat scs)
|
pure (concat scs)
|
||||||
|
where
|
||||||
|
conIfGuarded : Term vars -> Core (Term vars)
|
||||||
|
conIfGuarded (Ref fc Func n)
|
||||||
|
= do defs <- get Ctxt
|
||||||
|
Just gdef <- lookupCtxtExact n (gamma defs)
|
||||||
|
| Nothing => pure $ Ref fc Func n
|
||||||
|
if AllGuarded `elem` flags gdef
|
||||||
|
then pure $ Ref fc (DataCon 0 0) n
|
||||||
|
else pure $ Ref fc Func n
|
||||||
|
conIfGuarded tm = pure tm
|
||||||
|
|
||||||
-- Expand the size change argument list with 'Nothing' to match the given
|
-- Expand the size change argument list with 'Nothing' to match the given
|
||||||
-- arity (i.e. the arity of the function we're calling) to ensure that
|
-- arity (i.e. the arity of the function we're calling) to ensure that
|
||||||
|
@ -185,7 +185,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps impln nu
|
|||||||
fns <- topMethTypes [] impName methImps impsp (params cdata)
|
fns <- topMethTypes [] impName methImps impsp (params cdata)
|
||||||
(map fst (methods cdata))
|
(map fst (methods cdata))
|
||||||
(methods cdata)
|
(methods cdata)
|
||||||
traverse (processDecl [] nest env) (map mkTopMethDecl fns)
|
traverse_ (processDecl [] nest env) (map mkTopMethDecl fns)
|
||||||
|
|
||||||
-- 3. Build the record for the implementation
|
-- 3. Build the record for the implementation
|
||||||
let mtops = map (Builtin.fst . snd) fns
|
let mtops = map (Builtin.fst . snd) fns
|
||||||
|
@ -1174,6 +1174,9 @@ fnDirectOpt fname
|
|||||||
<|> do pragma "inline"
|
<|> do pragma "inline"
|
||||||
commit
|
commit
|
||||||
pure $ IFnOpt Inline
|
pure $ IFnOpt Inline
|
||||||
|
<|> do pragma "tcinline"
|
||||||
|
commit
|
||||||
|
pure $ IFnOpt TCInline
|
||||||
<|> do pragma "extern"
|
<|> do pragma "extern"
|
||||||
pure $ IFnOpt ExternFn
|
pure $ IFnOpt ExternFn
|
||||||
<|> do pragma "macro"
|
<|> do pragma "macro"
|
||||||
|
@ -69,6 +69,7 @@ showInfo (n, idx, d)
|
|||||||
show !(traverse getFullName (keys (refersTo d))))
|
show !(traverse getFullName (keys (refersTo d))))
|
||||||
coreLift $ putStrLn ("Refers to (runtime): " ++
|
coreLift $ putStrLn ("Refers to (runtime): " ++
|
||||||
show !(traverse getFullName (keys (refersToRuntime d))))
|
show !(traverse getFullName (keys (refersToRuntime d))))
|
||||||
|
coreLift $ putStrLn ("Flags: " ++ show (flags d))
|
||||||
when (not (isNil (sizeChange d))) $
|
when (not (isNil (sizeChange d))) $
|
||||||
let scinfo = map (\s => show (fnCall s) ++ ": " ++
|
let scinfo = map (\s => show (fnCall s) ++ ": " ++
|
||||||
show (fnArgs s)) !(traverse toFullNames (sizeChange d)) in
|
show (fnArgs s)) !(traverse toFullNames (sizeChange d)) in
|
||||||
|
@ -226,7 +226,15 @@ caseBlock {vars} rigc elabinfo fc nest env scr scrtm scrty caseRig alts expected
|
|||||||
-- Start with empty nested names, since we've extended the rhs with
|
-- Start with empty nested names, since we've extended the rhs with
|
||||||
-- ICaseLocal so they'll get rebuilt with the right environment
|
-- ICaseLocal so they'll get rebuilt with the right environment
|
||||||
let nest' = MkNested []
|
let nest' = MkNested []
|
||||||
|
ust <- get UST
|
||||||
|
-- We don't want to keep rechecking delayed elaborators in the
|
||||||
|
-- case block, because they're not going to make progress until
|
||||||
|
-- we come out again, so save them
|
||||||
|
let olddelayed = delayedElab ust
|
||||||
|
put UST (record { delayedElab = [] } ust)
|
||||||
processDecl [InCase] nest' [] (IDef fc casen alts')
|
processDecl [InCase] nest' [] (IDef fc casen alts')
|
||||||
|
ust <- get UST
|
||||||
|
put UST (record { delayedElab = olddelayed } ust)
|
||||||
|
|
||||||
pure (appTm, gnf env caseretty)
|
pure (appTm, gnf env caseretty)
|
||||||
where
|
where
|
||||||
|
@ -71,7 +71,7 @@ delayOnFailure fc rig env expected pred pri elab
|
|||||||
handle (elab False)
|
handle (elab False)
|
||||||
(\err =>
|
(\err =>
|
||||||
do est <- get EST
|
do est <- get EST
|
||||||
if pred err && delayDepth est < !getAmbigLimit
|
if pred err
|
||||||
then
|
then
|
||||||
do nm <- genName "delayed"
|
do nm <- genName "delayed"
|
||||||
(ci, dtm) <- newDelayed fc linear env nm !(getTerm expected)
|
(ci, dtm) <- newDelayed fc linear env nm !(getTerm expected)
|
||||||
@ -99,19 +99,16 @@ delayElab : {vars : _} ->
|
|||||||
Core (Term vars, Glued vars)
|
Core (Term vars, Glued vars)
|
||||||
delayElab {vars} fc rig env exp pri elab
|
delayElab {vars} fc rig env exp pri elab
|
||||||
= do est <- get EST
|
= do est <- get EST
|
||||||
if delayDepth est >= !getAmbigLimit
|
nm <- genName "delayed"
|
||||||
then elab
|
expected <- mkExpected exp
|
||||||
else do
|
(ci, dtm) <- newDelayed fc linear env nm !(getTerm expected)
|
||||||
nm <- genName "delayed"
|
logGlueNF 5 ("Postponing elaborator " ++ show nm ++
|
||||||
expected <- mkExpected exp
|
" for") env expected
|
||||||
(ci, dtm) <- newDelayed fc linear env nm !(getTerm expected)
|
ust <- get UST
|
||||||
logGlueNF 5 ("Postponing elaborator " ++ show nm ++
|
put UST (record { delayedElab $=
|
||||||
" for") env expected
|
((pri, ci, mkClosedElab fc env elab) :: ) }
|
||||||
ust <- get UST
|
ust)
|
||||||
put UST (record { delayedElab $=
|
pure (dtm, expected)
|
||||||
((pri, ci, mkClosedElab fc env (deeper elab)) :: ) }
|
|
||||||
ust)
|
|
||||||
pure (dtm, expected)
|
|
||||||
where
|
where
|
||||||
mkExpected : Maybe (Glued vars) -> Core (Glued vars)
|
mkExpected : Maybe (Glued vars) -> Core (Glued vars)
|
||||||
mkExpected (Just ty) = pure ty
|
mkExpected (Just ty) = pure ty
|
||||||
@ -132,14 +129,69 @@ ambiguous (InRHS _ _ err) = ambiguous err
|
|||||||
ambiguous (WhenUnifying _ _ _ _ err) = ambiguous err
|
ambiguous (WhenUnifying _ _ _ _ err) = ambiguous err
|
||||||
ambiguous _ = False
|
ambiguous _ = False
|
||||||
|
|
||||||
|
mutual
|
||||||
|
mismatchNF : {vars : _} ->
|
||||||
|
Defs -> NF vars -> NF vars -> Core Bool
|
||||||
|
mismatchNF defs (NTCon _ xn xt _ xargs) (NTCon _ yn yt _ yargs)
|
||||||
|
= if xn /= yn
|
||||||
|
then pure True
|
||||||
|
else anyM (mismatch defs) (zip xargs yargs)
|
||||||
|
mismatchNF defs (NDCon _ _ xt _ xargs) (NDCon _ _ yt _ yargs)
|
||||||
|
= if xt /= yt
|
||||||
|
then pure True
|
||||||
|
else anyM (mismatch defs) (zip xargs yargs)
|
||||||
|
mismatchNF defs (NPrimVal _ xc) (NPrimVal _ yc) = pure (xc /= yc)
|
||||||
|
mismatchNF defs (NDelayed _ _ x) (NDelayed _ _ y) = mismatchNF defs x y
|
||||||
|
mismatchNF defs (NDelay _ _ _ x) (NDelay _ _ _ y)
|
||||||
|
= mismatchNF defs !(evalClosure defs x) !(evalClosure defs y)
|
||||||
|
mismatchNF _ _ _ = pure False
|
||||||
|
|
||||||
|
mismatch : {vars : _} ->
|
||||||
|
Defs -> (Closure vars, Closure vars) -> Core Bool
|
||||||
|
mismatch defs (x, y)
|
||||||
|
= mismatchNF defs !(evalClosure defs x) !(evalClosure defs y)
|
||||||
|
|
||||||
|
contra : {vars : _} ->
|
||||||
|
Defs -> NF vars -> NF vars -> Core Bool
|
||||||
|
-- Unlike 'impossibleOK', any mismatch indicates an unrecoverable error
|
||||||
|
contra defs (NTCon _ xn xt xa xargs) (NTCon _ yn yt ya yargs)
|
||||||
|
= if xn /= yn
|
||||||
|
then pure True
|
||||||
|
else anyM (mismatch defs) (zip xargs yargs)
|
||||||
|
contra defs (NDCon _ _ xt _ xargs) (NDCon _ _ yt _ yargs)
|
||||||
|
= if xt /= yt
|
||||||
|
then pure True
|
||||||
|
else anyM (mismatch defs) (zip xargs yargs)
|
||||||
|
contra defs (NPrimVal _ x) (NPrimVal _ y) = pure (x /= y)
|
||||||
|
contra defs (NDCon _ _ _ _ _) (NPrimVal _ _) = pure True
|
||||||
|
contra defs (NPrimVal _ _) (NDCon _ _ _ _ _) = pure True
|
||||||
|
contra defs x y = pure False
|
||||||
|
|
||||||
|
-- Errors that might be recoverable later if we try again. Generally -
|
||||||
|
-- ambiguity errors, type inference errors
|
||||||
|
export
|
||||||
|
recoverable : {auto c : Ref Ctxt Defs} ->
|
||||||
|
Error -> Core Bool
|
||||||
|
recoverable (CantConvert _ env l r)
|
||||||
|
= do defs <- get Ctxt
|
||||||
|
pure $ not !(contra defs !(nf defs env l) !(nf defs env r))
|
||||||
|
recoverable (CantSolveEq _ env l r)
|
||||||
|
= do defs <- get Ctxt
|
||||||
|
pure $ not !(contra defs !(nf defs env l) !(nf defs env r))
|
||||||
|
recoverable (UndefinedName _ _) = pure False
|
||||||
|
recoverable (InType _ _ err) = recoverable err
|
||||||
|
recoverable (InCon _ _ err) = recoverable err
|
||||||
|
recoverable (InLHS _ _ err) = recoverable err
|
||||||
|
recoverable (InRHS _ _ err) = recoverable err
|
||||||
|
recoverable (WhenUnifying _ _ _ _ err) = recoverable err
|
||||||
|
recoverable _ = pure True
|
||||||
|
|
||||||
data RetryError
|
data RetryError
|
||||||
= NoError
|
= RecoverableErrors
|
||||||
| AmbigError
|
|
||||||
| AllErrors
|
| AllErrors
|
||||||
|
|
||||||
Show RetryError where
|
Show RetryError where
|
||||||
show NoError = "NoError"
|
show RecoverableErrors = "RecoverableErrors"
|
||||||
show AmbigError = "AmbigError"
|
|
||||||
show AllErrors = "AllErrors"
|
show AllErrors = "AllErrors"
|
||||||
|
|
||||||
-- Try all the delayed elaborators. If there's a failure, we want to
|
-- Try all the delayed elaborators. If there's a failure, we want to
|
||||||
@ -178,9 +230,8 @@ retryDelayed' errmode acc (d@(_, i, elab) :: ds)
|
|||||||
(\err => do log 5 $ show errmode ++ ":Error in " ++ show !(getFullName (Resolved i))
|
(\err => do log 5 $ show errmode ++ ":Error in " ++ show !(getFullName (Resolved i))
|
||||||
++ "\n" ++ show err
|
++ "\n" ++ show err
|
||||||
case errmode of
|
case errmode of
|
||||||
NoError => retryDelayed' errmode (d :: acc) ds
|
RecoverableErrors =>
|
||||||
AmbigError =>
|
if not !(recoverable err)
|
||||||
if ambiguous err -- give up on ambiguity
|
|
||||||
then throw err
|
then throw err
|
||||||
else retryDelayed' errmode (d :: acc) ds
|
else retryDelayed' errmode (d :: acc) ds
|
||||||
AllErrors => throw err)
|
AllErrors => throw err)
|
||||||
@ -195,8 +246,7 @@ retryDelayed : {vars : _} ->
|
|||||||
Core ()
|
Core ()
|
||||||
retryDelayed ds
|
retryDelayed ds
|
||||||
= do est <- get EST
|
= do est <- get EST
|
||||||
ds <- retryDelayed' NoError [] ds -- try everything again
|
ds <- retryDelayed' RecoverableErrors [] ds -- try everything again
|
||||||
ds <- retryDelayed' AmbigError [] ds -- fail on ambiguity error
|
|
||||||
retryDelayed' AllErrors [] ds -- fail on all errors
|
retryDelayed' AllErrors [] ds -- fail on all errors
|
||||||
pure ()
|
pure ()
|
||||||
|
|
||||||
|
@ -37,7 +37,15 @@ checkLocal {vars} rig elabinfo nest env fc nestdecls scope expty
|
|||||||
-- fixes bug #115
|
-- fixes bug #115
|
||||||
let nest' = record { names $= (names' ++) } nest
|
let nest' = record { names $= (names' ++) } nest
|
||||||
let env' = dropLinear env
|
let env' = dropLinear env
|
||||||
|
-- We don't want to keep rechecking delayed elaborators in the
|
||||||
|
-- locals block, because they're not going to make progress until
|
||||||
|
-- we come out again, so save them
|
||||||
|
ust <- get UST
|
||||||
|
let olddelayed = delayedElab ust
|
||||||
|
put UST (record { delayedElab = [] } ust)
|
||||||
traverse (processDecl [] nest' env') (map (updateName nest') nestdecls)
|
traverse (processDecl [] nest' env') (map (updateName nest') nestdecls)
|
||||||
|
ust <- get UST
|
||||||
|
put UST (record { delayedElab = olddelayed } ust)
|
||||||
check rig elabinfo nest' env scope expty
|
check rig elabinfo nest' env scope expty
|
||||||
where
|
where
|
||||||
-- For the local definitions, don't allow access to linear things
|
-- For the local definitions, don't allow access to linear things
|
||||||
|
@ -7,6 +7,7 @@ import Core.Normalise
|
|||||||
import Core.TT
|
import Core.TT
|
||||||
import Core.Value
|
import Core.Value
|
||||||
|
|
||||||
|
import TTImp.Elab.Check
|
||||||
import TTImp.TTImp
|
import TTImp.TTImp
|
||||||
|
|
||||||
detagSafe : Defs -> NF [] -> Core Bool
|
detagSafe : Defs -> NF [] -> Core Bool
|
||||||
@ -64,3 +65,9 @@ updateErasable n
|
|||||||
addDef n (record { eraseArgs = es,
|
addDef n (record { eraseArgs = es,
|
||||||
safeErase = dtes } gdef)
|
safeErase = dtes } gdef)
|
||||||
pure ()
|
pure ()
|
||||||
|
export
|
||||||
|
wrapErrorC : List ElabOpt -> (Error -> Error) -> Core a -> Core a
|
||||||
|
wrapErrorC opts err
|
||||||
|
= if InCase `elem` opts
|
||||||
|
then id
|
||||||
|
else wrapError err
|
||||||
|
@ -92,7 +92,7 @@ checkCon {vars} opts nest env vis tn_in tn (MkImpTy fc cn_in ty_raw)
|
|||||||
Nothing <- lookupCtxtExact cn (gamma defs)
|
Nothing <- lookupCtxtExact cn (gamma defs)
|
||||||
| Just gdef => throw (AlreadyDefined fc cn)
|
| Just gdef => throw (AlreadyDefined fc cn)
|
||||||
ty <-
|
ty <-
|
||||||
wrapError (InCon fc cn) $
|
wrapErrorC opts (InCon fc cn) $
|
||||||
checkTerm !(resolveName cn) InType opts nest env
|
checkTerm !(resolveName cn) InType opts nest env
|
||||||
(IBindHere fc (PI erased) ty_raw)
|
(IBindHere fc (PI erased) ty_raw)
|
||||||
(gType fc)
|
(gType fc)
|
||||||
@ -255,7 +255,7 @@ processData {vars} eopts nest env fc vis (MkImpLater dfc n_in ty_raw)
|
|||||||
| Just gdef => throw (AlreadyDefined fc n)
|
| Just gdef => throw (AlreadyDefined fc n)
|
||||||
|
|
||||||
(ty, _) <-
|
(ty, _) <-
|
||||||
wrapError (InCon fc n) $
|
wrapErrorC eopts (InCon fc n) $
|
||||||
elabTerm !(resolveName n) InType eopts nest env
|
elabTerm !(resolveName n) InType eopts nest env
|
||||||
(IBindHere fc (PI erased) ty_raw)
|
(IBindHere fc (PI erased) ty_raw)
|
||||||
(Just (gType dfc))
|
(Just (gType dfc))
|
||||||
@ -289,7 +289,7 @@ processData {vars} eopts nest env fc vis (MkImpData dfc n_in ty_raw opts cons_ra
|
|||||||
log 1 $ "Processing " ++ show n
|
log 1 $ "Processing " ++ show n
|
||||||
defs <- get Ctxt
|
defs <- get Ctxt
|
||||||
(ty, _) <-
|
(ty, _) <-
|
||||||
wrapError (InCon fc n) $
|
wrapErrorC eopts (InCon fc n) $
|
||||||
elabTerm !(resolveName n) InType eopts nest env
|
elabTerm !(resolveName n) InType eopts nest env
|
||||||
(IBindHere fc (PI erased) ty_raw)
|
(IBindHere fc (PI erased) ty_raw)
|
||||||
(Just (gType dfc))
|
(Just (gType dfc))
|
||||||
|
@ -18,6 +18,7 @@ import Core.UnifyState
|
|||||||
import TTImp.BindImplicits
|
import TTImp.BindImplicits
|
||||||
import TTImp.Elab
|
import TTImp.Elab
|
||||||
import TTImp.Elab.Check
|
import TTImp.Elab.Check
|
||||||
|
import TTImp.Elab.Utils
|
||||||
import TTImp.Impossible
|
import TTImp.Impossible
|
||||||
import TTImp.PartialEval
|
import TTImp.PartialEval
|
||||||
import TTImp.TTImp
|
import TTImp.TTImp
|
||||||
@ -252,7 +253,7 @@ checkLHS {vars} trans mult hashit n opts nest env fc lhs_in
|
|||||||
then InTransform
|
then InTransform
|
||||||
else InLHS mult
|
else InLHS mult
|
||||||
(lhstm, lhstyg) <-
|
(lhstm, lhstyg) <-
|
||||||
wrapError (InLHS fc !(getFullName (Resolved n))) $
|
wrapErrorC opts (InLHS fc !(getFullName (Resolved n))) $
|
||||||
elabTerm n lhsMode opts nest env
|
elabTerm n lhsMode opts nest env
|
||||||
(IBindHere fc PATTERN lhs) Nothing
|
(IBindHere fc PATTERN lhs) Nothing
|
||||||
logTerm 5 "Checked LHS term" lhstm
|
logTerm 5 "Checked LHS term" lhstm
|
||||||
@ -387,7 +388,7 @@ checkClause {vars} mult hashit n opts nest env (PatClause fc lhs_in rhs)
|
|||||||
log 5 $ "Checking RHS " ++ show rhs
|
log 5 $ "Checking RHS " ++ show rhs
|
||||||
logEnv 5 "In env" env'
|
logEnv 5 "In env" env'
|
||||||
|
|
||||||
rhstm <- wrapError (InRHS fc !(getFullName (Resolved n))) $
|
rhstm <- wrapErrorC opts (InRHS fc !(getFullName (Resolved n))) $
|
||||||
checkTermSub n rhsMode opts nest' env' env sub' rhs (gnf env' lhsty')
|
checkTermSub n rhsMode opts nest' env' env sub' rhs (gnf env' lhsty')
|
||||||
clearHoleLHS
|
clearHoleLHS
|
||||||
|
|
||||||
@ -411,7 +412,7 @@ checkClause {vars} mult hashit n opts nest env (WithClause fc lhs_in wval_raw cs
|
|||||||
let wmode
|
let wmode
|
||||||
= if isErased mult then InType else InExpr
|
= if isErased mult then InType else InExpr
|
||||||
|
|
||||||
(wval, gwvalTy) <- wrapError (InRHS fc !(getFullName (Resolved n))) $
|
(wval, gwvalTy) <- wrapErrorC opts (InRHS fc !(getFullName (Resolved n))) $
|
||||||
elabTermSub n wmode opts nest' env' env sub' wval_raw Nothing
|
elabTermSub n wmode opts nest' env' env sub' wval_raw Nothing
|
||||||
clearHoleLHS
|
clearHoleLHS
|
||||||
|
|
||||||
@ -471,7 +472,7 @@ checkClause {vars} mult hashit n opts nest env (WithClause fc lhs_in wval_raw cs
|
|||||||
map (maybe wval_raw (\pn => IVar fc (snd pn))) wargNames)
|
map (maybe wval_raw (\pn => IVar fc (snd pn))) wargNames)
|
||||||
|
|
||||||
log 3 $ "Applying to with argument " ++ show rhs_in
|
log 3 $ "Applying to with argument " ++ show rhs_in
|
||||||
rhs <- wrapError (InRHS fc !(getFullName (Resolved n))) $
|
rhs <- wrapErrorC opts (InRHS fc !(getFullName (Resolved n))) $
|
||||||
checkTermSub n wmode opts nest' env' env sub' rhs_in
|
checkTermSub n wmode opts nest' env' env sub' rhs_in
|
||||||
(gnf env' reqty)
|
(gnf env' reqty)
|
||||||
|
|
||||||
@ -720,6 +721,7 @@ processDef opts nest env fc n_in cs_in
|
|||||||
do calcRefs False atotal (Resolved nidx)
|
do calcRefs False atotal (Resolved nidx)
|
||||||
sc <- calculateSizeChange fc n
|
sc <- calculateSizeChange fc n
|
||||||
setSizeChange fc n sc
|
setSizeChange fc n sc
|
||||||
|
checkIfGuarded fc n
|
||||||
|
|
||||||
md <- get MD -- don't need the metadata collected on the coverage check
|
md <- get MD -- don't need the metadata collected on the coverage check
|
||||||
cov <- checkCoverage nidx ty mult cs
|
cov <- checkCoverage nidx ty mult cs
|
||||||
|
@ -33,6 +33,8 @@ processFnOpt : {auto c : Ref Ctxt Defs} ->
|
|||||||
FC -> Name -> FnOpt -> Core ()
|
FC -> Name -> FnOpt -> Core ()
|
||||||
processFnOpt fc ndef Inline
|
processFnOpt fc ndef Inline
|
||||||
= setFlag fc ndef Inline
|
= setFlag fc ndef Inline
|
||||||
|
processFnOpt fc ndef TCInline
|
||||||
|
= setFlag fc ndef TCInline
|
||||||
processFnOpt fc ndef (Hint d)
|
processFnOpt fc ndef (Hint d)
|
||||||
= do defs <- get Ctxt
|
= do defs <- get Ctxt
|
||||||
Just ty <- lookupTyExact ndef (gamma defs)
|
Just ty <- lookupTyExact ndef (gamma defs)
|
||||||
@ -264,7 +266,7 @@ processType {vars} eopts nest env fc rig vis opts (MkImpTy tfc n_in ty_raw)
|
|||||||
| Just _ => throw (AlreadyDefined fc n)
|
| Just _ => throw (AlreadyDefined fc n)
|
||||||
|
|
||||||
ty <-
|
ty <-
|
||||||
wrapError (InType fc n) $
|
wrapErrorC eopts (InType fc n) $
|
||||||
checkTerm idx InType (HolesOkay :: eopts) nest env
|
checkTerm idx InType (HolesOkay :: eopts) nest env
|
||||||
(IBindHere fc (PI erased) ty_raw)
|
(IBindHere fc (PI erased) ty_raw)
|
||||||
(gType fc)
|
(gType fc)
|
||||||
|
@ -179,6 +179,7 @@ mutual
|
|||||||
public export
|
public export
|
||||||
data FnOpt : Type where
|
data FnOpt : Type where
|
||||||
Inline : FnOpt
|
Inline : FnOpt
|
||||||
|
TCInline : FnOpt
|
||||||
-- Flag means the hint is a direct hint, not a function which might
|
-- Flag means the hint is a direct hint, not a function which might
|
||||||
-- find the result (e.g. chasing parent interface dictionaries)
|
-- find the result (e.g. chasing parent interface dictionaries)
|
||||||
Hint : Bool -> FnOpt
|
Hint : Bool -> FnOpt
|
||||||
@ -196,6 +197,7 @@ mutual
|
|||||||
export
|
export
|
||||||
Show FnOpt where
|
Show FnOpt where
|
||||||
show Inline = "%inline"
|
show Inline = "%inline"
|
||||||
|
show TCInline = "%tcinline"
|
||||||
show (Hint t) = "%hint " ++ show t
|
show (Hint t) = "%hint " ++ show t
|
||||||
show (GlobalHint t) = "%globalhint " ++ show t
|
show (GlobalHint t) = "%globalhint " ++ show t
|
||||||
show ExternFn = "%extern"
|
show ExternFn = "%extern"
|
||||||
@ -210,6 +212,7 @@ mutual
|
|||||||
export
|
export
|
||||||
Eq FnOpt where
|
Eq FnOpt where
|
||||||
Inline == Inline = True
|
Inline == Inline = True
|
||||||
|
TCInline == TCInline = True
|
||||||
(Hint x) == (Hint y) = x == y
|
(Hint x) == (Hint y) = x == y
|
||||||
(GlobalHint x) == (GlobalHint y) = x == y
|
(GlobalHint x) == (GlobalHint y) = x == y
|
||||||
ExternFn == ExternFn = True
|
ExternFn == ExternFn = True
|
||||||
@ -895,6 +898,7 @@ mutual
|
|||||||
export
|
export
|
||||||
TTC FnOpt where
|
TTC FnOpt where
|
||||||
toBuf b Inline = tag 0
|
toBuf b Inline = tag 0
|
||||||
|
toBuf b TCInline = tag 11
|
||||||
toBuf b (Hint t) = do tag 1; toBuf b t
|
toBuf b (Hint t) = do tag 1; toBuf b t
|
||||||
toBuf b (GlobalHint t) = do tag 2; toBuf b t
|
toBuf b (GlobalHint t) = do tag 2; toBuf b t
|
||||||
toBuf b ExternFn = tag 3
|
toBuf b ExternFn = tag 3
|
||||||
@ -919,6 +923,7 @@ mutual
|
|||||||
8 => pure (Totality PartialOK)
|
8 => pure (Totality PartialOK)
|
||||||
9 => pure Macro
|
9 => pure Macro
|
||||||
10 => do ns <- fromBuf b; pure (SpecArgs ns)
|
10 => do ns <- fromBuf b; pure (SpecArgs ns)
|
||||||
|
11 => pure TCInline
|
||||||
_ => corrupt "FnOpt"
|
_ => corrupt "FnOpt"
|
||||||
|
|
||||||
export
|
export
|
||||||
|
@ -8,6 +8,7 @@ Inferrable args: []
|
|||||||
Compiled: Constructor tag Just 0 arity 1 (newtype by 0)
|
Compiled: Constructor tag Just 0 arity 1 (newtype by 0)
|
||||||
Refers to: []
|
Refers to: []
|
||||||
Refers to (runtime): []
|
Refers to (runtime): []
|
||||||
|
Flags: []
|
||||||
Main> Main.MkBar ==> DataCon 0 1
|
Main> Main.MkBar ==> DataCon 0 1
|
||||||
RigW
|
RigW
|
||||||
Erasable args: []
|
Erasable args: []
|
||||||
@ -17,4 +18,5 @@ Inferrable args: []
|
|||||||
Compiled: Constructor tag Just 0 arity 1
|
Compiled: Constructor tag Just 0 arity 1
|
||||||
Refers to: []
|
Refers to: []
|
||||||
Refers to (runtime): []
|
Refers to (runtime): []
|
||||||
|
Flags: []
|
||||||
Main> Bye for now!
|
Main> Bye for now!
|
||||||
|
@ -1,6 +1,6 @@
|
|||||||
1/1: Building Total (Total.idr)
|
1/1: Building Total (Total.idr)
|
||||||
Main> Main.count is total
|
Main> Main.count is total
|
||||||
Main> Main.badCount is possibly not terminating due to recursive path Main.badCount -> Prelude.Functor implementation at Prelude.idr:1010:1--1015:1 -> Prelude.map
|
Main> Main.badCount is possibly not terminating due to recursive path Main.badCount
|
||||||
Main> Main.process is total
|
Main> Main.process is total
|
||||||
Main> Main.badProcess is possibly not terminating due to recursive path Main.badProcess -> Main.badProcess -> Main.badProcess
|
Main> Main.badProcess is possibly not terminating due to recursive path Main.badProcess -> Main.badProcess -> Main.badProcess
|
||||||
Main> Main.doubleInt is total
|
Main> Main.doubleInt is total
|
||||||
|
Loading…
Reference in New Issue
Block a user