Skip to content

Commit

Permalink
[ cleanup ] Make makeFuture to be %foreign, not %extern
Browse files Browse the repository at this point in the history
  • Loading branch information
buzden committed Jul 1, 2024
1 parent 57f455d commit c94e40a
Show file tree
Hide file tree
Showing 6 changed files with 7 additions and 15 deletions.
5 changes: 3 additions & 2 deletions libs/contrib/System/Future.idr
Original file line number Diff line number Diff line change
Expand Up @@ -7,14 +7,15 @@ module System.Future
export
data Future : Type -> Type where [external]

%extern prim__makeFuture : {0 a : Type} -> Lazy a -> Future a
%foreign "scheme:blodwen-make-future"
prim__makeFuture : {0 a : Type} -> (() -> a) -> Future a

%foreign "scheme:blodwen-await-future"
prim__awaitFuture : {0 a : Type} -> Future a -> a

export %inline -- inlining is important for correct context in codegens
fork : Lazy a -> Future a
fork = prim__makeFuture
fork l = prim__makeFuture $ \_ => force l

export %inline -- inlining is important for correct context in codegens
await : Future a -> a
Expand Down
3 changes: 0 additions & 3 deletions src/Compiler/Scheme/Chez.idr
Original file line number Diff line number Diff line change
Expand Up @@ -170,9 +170,6 @@ mutual
= do p' <- schExp cs (chezExtPrim cs) chezString 0 p
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 $ 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
5 changes: 1 addition & 4 deletions src/Compiler/Scheme/Common.idr
Original file line number Diff line number Diff line change
Expand Up @@ -207,7 +207,6 @@ data ExtPrim = NewIORef | ReadIORef | WriteIORef
| SysOS | SysCodegen
| OnCollect
| OnCollectAny
| MakeFuture
| Unknown Name

export
Expand All @@ -225,7 +224,6 @@ Show ExtPrim where
show SysCodegen = "SysCodegen"
show OnCollect = "OnCollect"
show OnCollectAny = "OnCollectAny"
show MakeFuture = "MakeFuture"
show (Unknown n) = "Unknown " ++ show n

||| Match on a user given name to get the scheme primitive
Expand All @@ -243,8 +241,7 @@ toPrim pn@(NS _ n)
(n == UN (Basic "prim__os"), SysOS),
(n == UN (Basic "prim__codegen"), SysCodegen),
(n == UN (Basic "prim__onCollect"), OnCollect),
(n == UN (Basic "prim__onCollectAny"), OnCollectAny),
(n == UN (Basic "prim__makeFuture"), MakeFuture)
(n == UN (Basic "prim__onCollectAny"), OnCollectAny)
]
(Unknown pn)
toPrim pn = Unknown pn
Expand Down
3 changes: 0 additions & 3 deletions src/Compiler/Scheme/Racket.idr
Original file line number Diff line number Diff line change
Expand Up @@ -112,9 +112,6 @@ mutual
= do p' <- schExp cs (racketPrim cs) racketString 0 p
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 $ 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
4 changes: 2 additions & 2 deletions support/chez/support.ss
Original file line number Diff line number Diff line change
Expand Up @@ -454,10 +454,10 @@
;; Future

(define-record future-internal (result ready mutex signal))
(define (blodwen-make-future work)
(define (blodwen-make-future ty work)
(let ([future (make-future-internal #f #f (make-mutex) (make-condition))])
(fork-thread (lambda ()
(let ([result (work)])
(let ([result (work '())])
(with-mutex (future-internal-mutex future)
(set-future-internal-result! future result)
(set-future-internal-ready! future #t)
Expand Down
2 changes: 1 addition & 1 deletion support/racket/support.rkt
Original file line number Diff line number Diff line change
Expand Up @@ -468,7 +468,7 @@
; )


(define (blodwen-make-future work) (future work))
(define (blodwen-make-future ty work) (future (lambda () (work '()))))
(define (blodwen-await-future ty future) (touch future))

;; NB: These should *ALWAYS* be used in multi-threaded programs since Racket
Expand Down

0 comments on commit c94e40a

Please sign in to comment.