[ cleanup ] Make makeFuture to be %foreign, not %extern

This commit is contained in:
Denis Buzdalov 2024-07-01 23:05:00 +03:00 committed by G. Allais
parent 57f455d135
commit c5abf4be35
7 changed files with 12 additions and 15 deletions

View File

@ -78,6 +78,8 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
consequence, `IWithApp` appears in `TTImp` values in elaborator scripts instead
of `IApp`, as it should have been.
* `MakeFuture` primitive is removed.
### Backend changes
#### RefC Backend
@ -230,6 +232,9 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
- `Data/Morphisms/Algebra.idr`
- `Data/Nat/Algebra.idr`
* `prim__makeFuture` from `System.Future` is reimplemented as `%foreign` instead of
using now removed `MakeFuture` primitive
#### Network
* Add a missing function parameter (the flag) in the C implementation of `idrnet_recv_bytes`

View File

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

View File

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

View File

@ -207,7 +207,6 @@ data ExtPrim = NewIORef | ReadIORef | WriteIORef
| SysOS | SysCodegen
| OnCollect
| OnCollectAny
| MakeFuture
| Unknown Name
export
@ -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
@ -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

View File

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

View File

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

View File

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