mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-23 22:22:07 +03:00
[ fixup #2939 ] Make futures not interfere with optimisations
This commit is contained in:
parent
10b0cc3240
commit
a6c5cf5af0
@ -12,11 +12,11 @@ data Future : Type -> Type where [external]
|
||||
%foreign "scheme:blodwen-await-future"
|
||||
prim__awaitFuture : {0 a : Type} -> Future a -> a
|
||||
|
||||
export
|
||||
export %inline -- inlining is important for correct context in codegens
|
||||
fork : Lazy a -> Future a
|
||||
fork = prim__makeFuture
|
||||
|
||||
export
|
||||
export %inline -- inlining is important for correct context in codegens
|
||||
await : Future a -> a
|
||||
await f = prim__awaitFuture f
|
||||
|
||||
|
@ -171,8 +171,8 @@ mutual
|
||||
c' <- schExp cs (chezExtPrim cs) chezString 0 c
|
||||
pure $ mkWorld $ "(blodwen-register-object " ++ p' ++ " " ++ c' ++ ")"
|
||||
chezExtPrim cs i MakeFuture [_, work]
|
||||
= do work' <- schExp cs (chezExtPrim cs) chezString 0 work
|
||||
pure $ "(blodwen-make-future " ++ work' ++ ")"
|
||||
= do work' <- schExp cs (chezExtPrim cs) chezString 0 $ NmForce EmptyFC LUnknown work
|
||||
pure $ "(blodwen-make-future (lambda () " ++ work' ++ "))"
|
||||
chezExtPrim cs i prim args
|
||||
= schExtCommon cs (chezExtPrim cs) chezString i prim args
|
||||
|
||||
|
@ -113,8 +113,8 @@ mutual
|
||||
c' <- schExp cs (racketPrim cs) racketString 0 c
|
||||
pure $ mkWorld $ "(blodwen-register-object " ++ p' ++ " " ++ c' ++ ")"
|
||||
racketPrim cs i MakeFuture [_, work]
|
||||
= do work' <- schExp cs (racketPrim cs) racketString 0 work
|
||||
pure $ mkWorld $ "(blodwen-make-future " ++ work' ++ ")"
|
||||
= do work' <- schExp cs (racketPrim cs) racketString 0 $ NmForce EmptyFC LUnknown work
|
||||
pure $ mkWorld $ "(blodwen-make-future (lambda () " ++ work' ++ "))"
|
||||
racketPrim cs i prim args
|
||||
= schExtCommon cs (racketPrim cs) racketString i prim args
|
||||
|
||||
|
14
tests/chez/futures002/Futures.idr
Normal file
14
tests/chez/futures002/Futures.idr
Normal file
@ -0,0 +1,14 @@
|
||||
module Futures
|
||||
|
||||
import System.Future
|
||||
|
||||
-- Checks the interference between CSE optimisations and de-optimisations
|
||||
-- and management of lazy values
|
||||
|
||||
topLevelConstant : Lazy String
|
||||
topLevelConstant = "top-level indeed"
|
||||
|
||||
main : IO ()
|
||||
main = do
|
||||
let a = await $ fork topLevelConstant
|
||||
putStrLn a
|
1
tests/chez/futures002/expected
Normal file
1
tests/chez/futures002/expected
Normal file
@ -0,0 +1 @@
|
||||
top-level indeed
|
5
tests/chez/futures002/run
Normal file
5
tests/chez/futures002/run
Normal file
@ -0,0 +1,5 @@
|
||||
. ../../testutils.sh
|
||||
|
||||
idris2 --cg chez Futures.idr -p contrib --exec main
|
||||
|
||||
rm -r build
|
@ -1,3 +1,3 @@
|
||||
. ../../testutils.sh
|
||||
|
||||
run --cg chez Main.idr
|
||||
run --cg racket Main.idr
|
||||
|
14
tests/racket/futures002/Futures.idr
Normal file
14
tests/racket/futures002/Futures.idr
Normal file
@ -0,0 +1,14 @@
|
||||
module Futures
|
||||
|
||||
import System.Future
|
||||
|
||||
-- Checks the interference between CSE optimisations and de-optimisations
|
||||
-- and management of lazy values
|
||||
|
||||
topLevelConstant : Lazy String
|
||||
topLevelConstant = "top-level indeed"
|
||||
|
||||
main : IO ()
|
||||
main = do
|
||||
let a = await $ fork topLevelConstant
|
||||
putStrLn a
|
1
tests/racket/futures002/expected
Normal file
1
tests/racket/futures002/expected
Normal file
@ -0,0 +1 @@
|
||||
top-level indeed
|
5
tests/racket/futures002/run
Normal file
5
tests/racket/futures002/run
Normal file
@ -0,0 +1,5 @@
|
||||
. ../../testutils.sh
|
||||
|
||||
idris2 --cg racket Futures.idr -p contrib --exec main
|
||||
|
||||
rm -r build
|
Loading…
Reference in New Issue
Block a user