Skip to content

Commit

Permalink
[ fixup #2939 ] Make futures not interfere with optimisations
Browse files Browse the repository at this point in the history
  • Loading branch information
buzden authored and gallais committed Jun 3, 2024
1 parent 10b0cc3 commit a6c5cf5
Show file tree
Hide file tree
Showing 10 changed files with 47 additions and 7 deletions.
4 changes: 2 additions & 2 deletions libs/contrib/System/Future.idr
Original file line number Diff line number Diff line change
Expand Up @@ -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

Expand Down
4 changes: 2 additions & 2 deletions src/Compiler/Scheme/Chez.idr
Original file line number Diff line number Diff line change
Expand Up @@ -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

Expand Down
4 changes: 2 additions & 2 deletions src/Compiler/Scheme/Racket.idr
Original file line number Diff line number Diff line change
Expand Up @@ -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

Expand Down
14 changes: 14 additions & 0 deletions tests/chez/futures002/Futures.idr
Original file line number Diff line number Diff line change
@@ -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 change: 1 addition & 0 deletions tests/chez/futures002/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
top-level indeed
5 changes: 5 additions & 0 deletions tests/chez/futures002/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,5 @@
. ../../testutils.sh

idris2 --cg chez Futures.idr -p contrib --exec main

rm -r build
2 changes: 1 addition & 1 deletion tests/racket/barrier001/run
Original file line number Diff line number Diff line change
@@ -1,3 +1,3 @@
. ../../testutils.sh

run --cg chez Main.idr
run --cg racket Main.idr
14 changes: 14 additions & 0 deletions tests/racket/futures002/Futures.idr
Original file line number Diff line number Diff line change
@@ -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 change: 1 addition & 0 deletions tests/racket/futures002/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
top-level indeed
5 changes: 5 additions & 0 deletions tests/racket/futures002/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,5 @@
. ../../testutils.sh

idris2 --cg racket Futures.idr -p contrib --exec main

rm -r build

0 comments on commit a6c5cf5

Please sign in to comment.