mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-28 11:05:17 +03:00
Small fix for incremental compilation (#1930)
* Small fix for incremental compilation In incremental mode we need to add the arity of a compiled definition to the hash, because if that changes, we need to rebuild the dependencies too to make sure the arity is correct at the call site * Fix indentation in CHANGELOG
This commit is contained in:
parent
971afa9f5d
commit
22c12046df
@ -55,6 +55,9 @@ filter p (x :: xs) with (p x)
|
||||
and reverts to whole program compilation. Incremental compilation is currently
|
||||
supported only by the Chez Scheme back end.
|
||||
This is currently supported only on Unix-like platforms (not yet Windows)
|
||||
- Note that you must set `IDRIS2_INC_CGS` when building and installing
|
||||
all libraries you plan to link with an incremental build.
|
||||
- Note also that this is experimental and not yet well tested!
|
||||
* The type checker now tries a lot harder to avoid reducing expressions where
|
||||
it is not needed. This gives a huge performance improvement in programs
|
||||
that potentially do a lot of compile time evaluation. However, sometimes
|
||||
|
@ -8,6 +8,7 @@ import Core.CompileExpr
|
||||
import Core.Context
|
||||
import Core.Context.Log
|
||||
import Core.FC
|
||||
import Core.Hash
|
||||
import Core.Options
|
||||
import Core.TT
|
||||
|
||||
@ -514,6 +515,16 @@ mergeLamDef n
|
||||
| Nothing => pure ()
|
||||
setCompiled n !(mergeLam cexpr)
|
||||
|
||||
export
|
||||
addArityHash : {auto c : Ref Ctxt Defs} ->
|
||||
Name -> Core ()
|
||||
addArityHash n
|
||||
= do defs <- get Ctxt
|
||||
Just def <- lookupCtxtExact n (gamma defs) | Nothing => pure ()
|
||||
let Just cexpr = compexpr def | Nothing => pure ()
|
||||
let MkFun args _ = cexpr | _ => pure ()
|
||||
addHash (n, length args)
|
||||
|
||||
export
|
||||
compileAndInlineAll : {auto c : Ref Ctxt Defs} ->
|
||||
Core ()
|
||||
@ -528,6 +539,11 @@ compileAndInlineAll
|
||||
-- This seems to be the point where not much useful
|
||||
-- happens any more.
|
||||
traverse_ updateCallGraph cns
|
||||
-- in incremental mode, add the arity of the definitions to the hash,
|
||||
-- because if these change we need to recompile dependencies
|
||||
-- accordingly
|
||||
when (not (isNil (incrementalCGs !getSession))) $
|
||||
traverse_ addArityHash cns
|
||||
where
|
||||
transform : Nat -> List Name -> Core ()
|
||||
transform Z cns = pure ()
|
||||
|
Loading…
Reference in New Issue
Block a user