[ cleanup ] Do not wrap the mainExpression into lazy in ES backends

This commit is contained in:
Denis Buzdalov 2024-01-22 23:46:24 +03:00 committed by Mathew Polzin
parent 2e1c3fbf2e
commit 381438533c
5 changed files with 50 additions and 4 deletions

View File

@ -749,14 +749,14 @@ def (MkFunction n as body) = do
mde <- mode <$> get ESs
b <- stmt Returns body >>= stmt
let cmt = comment $ hsep (shown n :: toList ((":" <++>) <$> mty))
case args of
if null args && n /= mainExpr
-- zero argument toplevel functions are converted to
-- lazily evaluated constants.
[] => pure $ printDoc mde $ vcat
-- lazily evaluated constants (except the main expression).
then pure $ printDoc mde $ vcat
[ cmt
, constant (var !(get NoMangleMap) ref)
("__lazy(" <+> function neutral [] b <+> ")") ]
_ => pure $ printDoc mde $ vcat
else pure $ printDoc mde $ vcat
[ cmt
, function (var !(get NoMangleMap) ref)
(map (var !(get NoMangleMap)) args) b ]

View File

@ -0,0 +1,4 @@
module HelloWorld
main : IO ()
main = putStrLn "Hello, world!"

View File

@ -0,0 +1,7 @@
module LazyIsStillThere
topLevelConst : Nat
topLevelConst = 16 + 8
main : IO ()
main = putStrLn "Top-level constant: \{show topLevelConst}"

View File

@ -0,0 +1,15 @@
# Top-level constants are lazily evaluated and strongly memoised.
# This is implemented by wrapping them to the function called `__lazy`.
# The only top-level function that should not be treated so is the expression for `main : IO ()`.
# In this test we check this.
--------------
# Running an example without any top-level constants...
# We expect no usages of `__lazy` to be present, maybe only a definition.
Hello, world!
function __lazy(thunk) {
--------------
# Running an example with some top-level constant...
# We expect `__lazy` to be used on the RHS for the top-level constant called `topLevelConst`.
Top-level constant: 24
function __lazy(thunk) {
const LazyIsStillThere_topLevelConst = __lazy(function () {

20
tests/node/node028/run Normal file
View File

@ -0,0 +1,20 @@
. ../../testutils.sh
echo '# Top-level constants are lazily evaluated and strongly memoised.'
echo '# This is implemented by wrapping them to the function called `__lazy`.'
echo '# The only top-level function that should not be treated so is the expression for `main : IO ()`.'
echo '# In this test we check this.'
echo '--------------'
echo '# Running an example without any top-level constants...'
echo '# We expect no usages of `__lazy` to be present, maybe only a definition.'
run --cg node -o hw.js HelloWorld.idr
grep '__lazy' build/exec/hw.js
echo '--------------'
echo '# Running an example with some top-level constant...'
echo '# We expect `__lazy` to be used on the RHS for the top-level constant called `topLevelConst`.'
run --cg node -o lsth.js LazyIsStillThere.idr
grep '__lazy' build/exec/lsth.js