[ fixup #2939 ] Make futures not interfere with optimisations

This commit is contained in:
Denis Buzdalov 2024-05-31 15:22:53 +03:00 committed by G. Allais
parent 10b0cc3240
commit a6c5cf5af0
10 changed files with 47 additions and 7 deletions

View File

@ -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

View File

@ -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

View File

@ -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

View 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

View File

@ -0,0 +1 @@
top-level indeed

View File

@ -0,0 +1,5 @@
. ../../testutils.sh
idris2 --cg chez Futures.idr -p contrib --exec main
rm -r build

View File

@ -1,3 +1,3 @@
. ../../testutils.sh
run --cg chez Main.idr
run --cg racket Main.idr

View 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

View File

@ -0,0 +1 @@
top-level indeed

View File

@ -0,0 +1,5 @@
. ../../testutils.sh
idris2 --cg racket Futures.idr -p contrib --exec main
rm -r build