[ performance ] Implement weak memoisation of lazy values for chez and racket (#2791)

* [ lazy ] Weakly memoise lazy expressions on chez and racket

* [ lazy ] Make weak memoisation to be controlled by codegen directive
This commit is contained in:
Denis Buzdalov 2024-08-06 17:24:57 +03:00 committed by GitHub
parent ff7c0fa3c2
commit 2482ebb432
No known key found for this signature in database
GPG Key ID: B5690EEEBB952194
37 changed files with 533 additions and 109 deletions

View File

@ -129,12 +129,20 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
* More efficient `collect-request-handler` is used.
* Add a codegen directive called `lazy=weakMemo` to make `Lazy` and `Inf` values *weakly*
memoised. That is, once accessed, they are allowed to be not re-evaluated until garbage
collector wipes them.
#### Racket
* Fixed CSE soundness bug that caused delayed expressions to sometimes be eagerly
evaluated. Now when a delayed expression is lifted by CSE, it is compiled
using Scheme's `delay` and `force` to memoize them.
* Add a codegen directive called `lazy=weakMemo` to make `Lazy` and `Inf` values *weakly*
memoised. That is, once accessed, they are allowed to be not re-evaluated until garbage
collector wipes them.
#### NodeJS Backend
* The NodeJS executable output to `build/exec/` now has its executable bit set.

View File

@ -72,6 +72,15 @@ Chez Directives
$ idris2 --codegen chez --directive extraRuntime=/path/to/extensions.scm -o main Main.idr
* ``--directive lazy=weakMemo``
Makes all non-toplevel ``Lazy`` and ``Inf`` values to be *weakly* memoised.
That is, once this expression is evaluated at runtime, it is allowed to not to be recalculated on later accesses
until memoised value is wiped by a garbage collector.
Garbage collector is allowed to collect weakly memoised values at its own discretion,
so when no free memory is available, weakly memoised values are free to be wiped.
That's why it is safer comparing to full memoisation.
Making a freestanding executable
================================

View File

@ -69,3 +69,12 @@ Racket Directives
.. code-block::
$ idris2 --codegen chez --directive extraRuntime=/path/to/extensions.scm -o main Main.idr
* ``--directive lazy=weakMemo``
Makes all non-toplevel ``Lazy`` and ``Inf`` values to be *weakly* memoised.
That is, once this expression is evaluated at runtime, it is allowed to not to be recalculated on later accesses
until memoised value is wiped by a garbage collector.
Garbage collector is allowed to collect weakly memoised values at its own discretion,
so when no free memory is available, weakly memoised values are free to be wiped.
That's why it is safer comparing to full memoisation.

View File

@ -557,6 +557,14 @@ getExtraRuntime directives
paths : List String
paths = nub $ mapMaybe getArg $ reverse directives
-- parses `--directive lazy=weakMemo` option for turning on weak memoisation of lazy values
-- (if supported by a backend).
-- This particular form of the directive string is chosen to be able to pass different variants
-- in the future (say, for strong memoisation, or turning laziness off).
export
getWeakMemoLazy : List String -> Bool
getWeakMemoLazy = elem "lazy=weakMemo"
||| Cast implementations. Values of `ConstantPrimitives` can
||| be used in a call to `castInt`, which then determines
||| the cast implementation based on the given pair of

View File

@ -133,45 +133,44 @@ export
chezString : String -> Builder
chezString cs = "\"" ++ showChezString (unpack cs) "\""
mutual
handleRet : CFType -> Builder -> Builder
handleRet CFUnit op = op ++ " " ++ mkWorld (schConstructor chezString (UN Underscore) (Just 0) [])
handleRet _ op = mkWorld op
handleRet : CFType -> Builder -> Builder
handleRet CFUnit op = op ++ " " ++ mkWorld (schConstructor chezString (UN Underscore) (Just 0) [])
handleRet _ op = mkWorld op
getFArgs : NamedCExp -> Core (List (NamedCExp, NamedCExp))
getFArgs (NmCon fc _ _ (Just 0) _) = pure []
getFArgs (NmCon fc _ _ (Just 1) [ty, val, rest]) = pure $ (ty, val) :: !(getFArgs rest)
getFArgs arg = throw (GenericMsg (getFC arg) ("Badly formed c call argument list " ++ show arg))
getFArgs : NamedCExp -> Core (List (NamedCExp, NamedCExp))
getFArgs (NmCon fc _ _ (Just 0) _) = pure []
getFArgs (NmCon fc _ _ (Just 1) [ty, val, rest]) = pure $ (ty, val) :: !(getFArgs rest)
getFArgs arg = throw (GenericMsg (getFC arg) ("Badly formed c call argument list " ++ show arg))
export
chezExtPrim : SortedSet Name -> Nat -> ExtPrim -> List NamedCExp -> Core Builder
chezExtPrim cs i GetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _]
= do structsc <- schExp cs (chezExtPrim cs) chezString 0 struct
pure $ "(ftype-ref " ++ fromString s ++ " (" ++ fromString fld ++ ") " ++ structsc ++ ")"
chezExtPrim cs i GetField [_,_,_,_,_,_]
= pure "(blodwen-error-quit \"bad getField\")"
chezExtPrim cs i SetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _, val, world]
= do structsc <- schExp cs (chezExtPrim cs) chezString 0 struct
valsc <- schExp cs (chezExtPrim cs) chezString 0 val
pure $ mkWorld $
"(ftype-set! " ++ fromString s ++ " (" ++ fromString fld ++ ") " ++ structsc ++
" " ++ valsc ++ ")"
chezExtPrim cs i SetField [_,_,_,_,_,_,_,_]
= pure "(blodwen-error-quit \"bad setField\")"
chezExtPrim cs i SysCodegen []
= pure $ "\"chez\""
chezExtPrim cs i OnCollect [_, p, c, world]
= 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 OnCollectAny [p, c, world]
= 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 prim args
= schExtCommon cs (chezExtPrim cs) chezString i prim args
export
chezExtPrim : SortedSet Name -> LazyExprProc -> Nat -> ExtPrim -> List NamedCExp -> Core Builder
chezExtPrim cs schLazy i GetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _]
= do structsc <- schExp cs (chezExtPrim cs schLazy) chezString schLazy 0 struct
pure $ "(ftype-ref " ++ fromString s ++ " (" ++ fromString fld ++ ") " ++ structsc ++ ")"
chezExtPrim cs schLazy i GetField [_,_,_,_,_,_]
= pure "(blodwen-error-quit \"bad getField\")"
chezExtPrim cs schLazy i SetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _, val, world]
= do structsc <- schExp cs (chezExtPrim cs schLazy) chezString schLazy 0 struct
valsc <- schExp cs (chezExtPrim cs schLazy) chezString schLazy 0 val
pure $ mkWorld $
"(ftype-set! " ++ fromString s ++ " (" ++ fromString fld ++ ") " ++ structsc ++
" " ++ valsc ++ ")"
chezExtPrim cs schLazy i SetField [_,_,_,_,_,_,_,_]
= pure "(blodwen-error-quit \"bad setField\")"
chezExtPrim cs schLazy i SysCodegen []
= pure $ "\"chez\""
chezExtPrim cs schLazy i OnCollect [_, p, c, world]
= do p' <- schExp cs (chezExtPrim cs schLazy) chezString schLazy 0 p
c' <- schExp cs (chezExtPrim cs schLazy) chezString schLazy 0 c
pure $ mkWorld $ "(blodwen-register-object " ++ p' ++ " " ++ c' ++ ")"
chezExtPrim cs schLazy i OnCollectAny [p, c, world]
= do p' <- schExp cs (chezExtPrim cs schLazy) chezString schLazy 0 p
c' <- schExp cs (chezExtPrim cs schLazy) chezString schLazy 0 c
pure $ mkWorld $ "(blodwen-register-object " ++ p' ++ " " ++ c' ++ ")"
chezExtPrim cs schLazy i prim args
= schExtCommon cs (chezExtPrim cs schLazy) chezString schLazy i prim args
-- Reference label for keeping track of loaded external libraries
export
@ -502,10 +501,12 @@ compileToSS c prof appdir tm outfile
fgndefs <- traverse (getFgnCall version) ndefs
loadlibs <- traverse (locateLib appdir) (mapMaybe fst fgndefs)
let schLazy = if getWeakMemoLazy ds then weakMemoLaziness else defaultLaziness
(sortedDefs, constants) <- sortDefs ndefs
compdefs <- logTime 3 "Print as scheme" $ traverse (getScheme constants (chezExtPrim constants) chezString) sortedDefs
compdefs <- logTime 3 "Print as scheme" $ traverse (getScheme constants (chezExtPrim constants schLazy) chezString schLazy) sortedDefs
let code = concat (map snd fgndefs) ++ concat compdefs
main <- schExp constants (chezExtPrim constants) chezString 0 ctm
main <- schExp constants (chezExtPrim constants schLazy) chezString schLazy 0 ctm
support <- readDataFile "chez/support.ss"
extraRuntime <- getExtraRuntime ds
let scm = concat $ the (List _)
@ -552,7 +553,7 @@ compileToSSInc c mods libs appdir tm outfile
loadlibs <- traverse (map fromString . loadLib appdir) (nub libs)
loadsos <- traverse (map fromString . loadSO appdir) (nub mods)
main <- schExp empty (chezExtPrim empty) chezString 0 ctm
main <- schExp empty (chezExtPrim empty defaultLaziness) chezString defaultLaziness 0 ctm
support <- readDataFile "chez/support.ss"
let scm = schHeader chez [] False ++
@ -692,7 +693,7 @@ incCompile c s sourceFile
version <- coreLift $ chezVersion chez
fgndefs <- traverse (getFgnCall version) ndefs
(sortedDefs, constants) <- sortDefs ndefs
compdefs <- traverse (getScheme empty (chezExtPrim empty) chezString) sortedDefs
compdefs <- traverse (getScheme empty (chezExtPrim empty defaultLaziness) chezString defaultLaziness) sortedDefs
let code = concat $ map snd fgndefs ++ compdefs
Right () <- coreLift $ writeFile ssFile $ build code
| Left err => throw (FileErr ssFile err)

View File

@ -217,7 +217,7 @@ compileToSS c chez appdir tm = do
let footer = ")"
fgndefs <- traverse (Chez.getFgnCall version) cu.definitions
compdefs <- traverse (getScheme empty (Chez.chezExtPrim empty) Chez.chezString) cu.definitions
compdefs <- traverse (getScheme empty (Chez.chezExtPrim empty defaultLaziness) Chez.chezString defaultLaziness) cu.definitions
loadlibs <- traverse (map fromString . loadLib appdir) (mapMaybe fst fgndefs)
-- write the files
@ -234,7 +234,7 @@ compileToSS c chez appdir tm = do
pure (MkChezLib chezLib hashChanged)
-- main module
main <- schExp empty (Chez.chezExtPrim empty) Chez.chezString 0 ctm
main <- schExp empty (Chez.chezExtPrim empty defaultLaziness) Chez.chezString defaultLaziness 0 ctm
Core.writeFile (appdir </> "mainprog.ss") $ build $ sepBy "\n"
[ schHeader (map snd libs) [lib.name | lib <- chezLibs]
, collectRequestHandler

View File

@ -317,9 +317,28 @@ var _ = False
getScrutineeTemp : Nat -> Builder
getScrutineeTemp i = fromString $ "sc" ++ show i
public export
record LazyExprProc where
constructor MkLazyExprProc
processDelay : Builder -> Builder
processForce : Builder -> Builder
public export
defaultLaziness : LazyExprProc
defaultLaziness = MkLazyExprProc
(\expr => "(lambda () " ++ expr ++ ")")
(\expr => "(" ++ expr ++ ")")
public export
weakMemoLaziness : LazyExprProc
weakMemoLaziness = MkLazyExprProc
(\expr => "(blodwen-delay-lazy (lambda () " ++ expr ++ "))")
(\expr => "(blodwen-force-lazy " ++ expr ++ ")")
parameters (constants : SortedSet Name,
schExtPrim : Nat -> ExtPrim -> List NamedCExp -> Core Builder,
schString : String -> Builder)
schString : String -> Builder,
schLazy : LazyExprProc)
showTag : Name -> Maybe Int -> Builder
showTag n (Just i) = showB i
showTag n Nothing = schString (show n)
@ -562,8 +581,8 @@ parameters (constants : SortedSet Name,
= schExtPrim i (toPrim p) args
schExp i (NmForce _ _ (NmApp fc x@(NmRef _ _) []))
= pure $ "(force " ++ !(schExp i x) ++ ")" -- Special version for memoized toplevel lazy definitions
schExp i (NmForce fc lr t) = pure $ "(" ++ !(schExp i t) ++ ")"
schExp i (NmDelay fc lr t) = pure $ "(lambda () " ++ !(schExp i t) ++ ")"
schExp i (NmForce fc lr t) = pure $ schLazy.processForce !(schExp i t)
schExp i (NmDelay fc lr t) = pure $ schLazy.processDelay !(schExp i t)
schExp i (NmConCase fc sc alts def)
= cond [(recordCase alts, schRecordCase i sc alts def),
(maybeCase alts, schMaybeCase i sc alts def),
@ -678,6 +697,7 @@ getScheme : {auto c : Ref Ctxt Defs} ->
(constants : SortedSet Name) ->
(schExtPrim : Nat -> ExtPrim -> List NamedCExp -> Core Builder) ->
(schString : String -> Builder) ->
(schLazy : LazyExprProc) ->
(Name, FC, NamedDef) -> Core Builder
getScheme constants schExtPrim schString (n, fc, d)
= schDef constants schExtPrim schString n d
getScheme constants schExtPrim schString schLazy (n, fc, d)
= schDef constants schExtPrim schString schLazy n d

View File

@ -74,35 +74,34 @@ showGambitString (c :: cs) acc = showGambitChar c $ showGambitString cs acc
gambitString : String -> Builder
gambitString cs = "\"" ++ showGambitString (unpack cs) "\""
mutual
handleRet : CFType -> Builder -> Builder
handleRet CFUnit op = op ++ " " ++ mkWorld (schConstructor gambitString (UN $ Basic "") (Just 0) [])
handleRet _ op = mkWorld op
handleRet : CFType -> Builder -> Builder
handleRet CFUnit op = op ++ " " ++ mkWorld (schConstructor gambitString (UN $ Basic "") (Just 0) [])
handleRet _ op = mkWorld op
getFArgs : NamedCExp -> Core (List (NamedCExp, NamedCExp))
getFArgs (NmCon fc _ _ (Just 0) _) = pure []
getFArgs (NmCon fc _ _ (Just 1) [ty, val, rest]) = pure $ (ty, val) :: !(getFArgs rest)
getFArgs arg = throw (GenericMsg (getFC arg) ("Badly formed c call argument list " ++ show arg))
getFArgs : NamedCExp -> Core (List (NamedCExp, NamedCExp))
getFArgs (NmCon fc _ _ (Just 0) _) = pure []
getFArgs (NmCon fc _ _ (Just 1) [ty, val, rest]) = pure $ (ty, val) :: !(getFArgs rest)
getFArgs arg = throw (GenericMsg (getFC arg) ("Badly formed c call argument list " ++ show arg))
gambitPrim : SortedSet Name -> Nat -> ExtPrim -> List NamedCExp -> Core Builder
gambitPrim cs i GetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _]
= do structsc <- schExp cs (gambitPrim cs) gambitString 0 struct
pure $ "(" ++ fromString s ++ "-" ++ fromString fld ++ " " ++ structsc ++ ")"
gambitPrim cs i GetField [_,_,_,_,_,_]
= pure "(error \"bad getField\")"
gambitPrim cs i SetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _, val, world]
= do structsc <- schExp cs (gambitPrim cs) gambitString 0 struct
valsc <- schExp cs (gambitPrim cs) gambitString 0 val
pure $ mkWorld $
"(" ++ fromString s ++ "-" ++ fromString fld ++ "-set! " ++ structsc ++ " " ++ valsc ++ ")"
gambitPrim cs i SetField [_,_,_,_,_,_,_,_]
= pure "(error \"bad setField\")"
gambitPrim cs i SysCodegen []
= pure $ "\"gambit\""
gambitPrim cs i prim args
= schExtCommon cs (gambitPrim cs) gambitString i prim args
gambitPrim : SortedSet Name -> LazyExprProc -> Nat -> ExtPrim -> List NamedCExp -> Core Builder
gambitPrim cs schLazy i GetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _]
= do structsc <- schExp cs (gambitPrim cs schLazy) gambitString schLazy 0 struct
pure $ "(" ++ fromString s ++ "-" ++ fromString fld ++ " " ++ structsc ++ ")"
gambitPrim cs schLazy i GetField [_,_,_,_,_,_]
= pure "(error \"bad getField\")"
gambitPrim cs schLazy i SetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _, val, world]
= do structsc <- schExp cs (gambitPrim cs schLazy) gambitString schLazy 0 struct
valsc <- schExp cs (gambitPrim cs schLazy) gambitString schLazy 0 val
pure $ mkWorld $
"(" ++ fromString s ++ "-" ++ fromString fld ++ "-set! " ++ structsc ++ " " ++ valsc ++ ")"
gambitPrim cs schLazy i SetField [_,_,_,_,_,_,_,_]
= pure "(error \"bad setField\")"
gambitPrim cs schLazy i SysCodegen []
= pure $ "\"gambit\""
gambitPrim cs schLazy i prim args
= schExtCommon cs (gambitPrim cs schLazy) gambitString schLazy i prim args
-- Reference label for keeping track of loaded external libraries
data Loaded : Type where
@ -365,16 +364,18 @@ compileToSCM c tm outfile
-- let tags = nameTags cdata
let ctm = forget (mainExpr cdata)
ds <- getDirectives Gambit
let schLazy = if getWeakMemoLazy ds then weakMemoLaziness else defaultLaziness
defs <- get Ctxt
l <- newRef {t = List String} Loaded []
s <- newRef {t = List String} Structs []
fgndefs <- traverse getFgnCall ndefs
(sortedDefs, constants) <- sortDefs ndefs
compdefs <- traverse (getScheme constants (gambitPrim constants) gambitString) ndefs
compdefs <- traverse (getScheme constants (gambitPrim constants schLazy) gambitString schLazy) ndefs
let code = concat (map snd fgndefs) ++ concat compdefs
main <- schExp constants (gambitPrim constants) gambitString 0 ctm
main <- schExp constants (gambitPrim constants schLazy) gambitString schLazy 0 ctm
support <- readDataFile "gambit/support.scm"
ds <- getDirectives Gambit
extraRuntime <- getExtraRuntime ds
foreign <- readDataFile "gambit/foreign.scm"
let scm = sepBy "\n" [schHeader, fromString support, fromString extraRuntime, fromString foreign, code, main]

View File

@ -86,34 +86,33 @@ showRacketString (c :: cs) acc = showRacketChar c $ showRacketString cs acc
racketString : String -> Builder
racketString cs = "\"" ++ showRacketString (unpack cs) "\""
mutual
racketPrim : SortedSet Name -> Nat -> ExtPrim -> List NamedCExp -> Core Builder
racketPrim cs i GetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _]
= do structsc <- schExp cs (racketPrim cs) racketString 0 struct
pure $ "(" ++ fromString s ++ "-" ++ fromString fld ++ " " ++ structsc ++ ")"
racketPrim cs i GetField [_,_,_,_,_,_]
= pure "(error \"bad getField\")"
racketPrim cs i SetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _, val, world]
= do structsc <- schExp cs (racketPrim cs) racketString 0 struct
valsc <- schExp cs (racketPrim cs) racketString 0 val
pure $ mkWorld $
"(set-" ++ fromString s ++ "-" ++ fromString fld ++ "! " ++ structsc ++ " " ++ valsc ++ ")"
racketPrim cs i SetField [_,_,_,_,_,_,_,_]
= pure "(error \"bad setField\")"
racketPrim cs i SysCodegen []
= pure $ "\"racket\""
racketPrim cs i OnCollect [_, p, c, world]
= 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 OnCollectAny [p, c, world]
= 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 prim args
= schExtCommon cs (racketPrim cs) racketString i prim args
racketPrim : SortedSet Name -> LazyExprProc -> Nat -> ExtPrim -> List NamedCExp -> Core Builder
racketPrim cs schLazy i GetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _]
= do structsc <- schExp cs (racketPrim cs schLazy) racketString schLazy 0 struct
pure $ "(" ++ fromString s ++ "-" ++ fromString fld ++ " " ++ structsc ++ ")"
racketPrim cs schLazy i GetField [_,_,_,_,_,_]
= pure "(error \"bad getField\")"
racketPrim cs schLazy i SetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _, val, world]
= do structsc <- schExp cs (racketPrim cs schLazy) racketString schLazy 0 struct
valsc <- schExp cs (racketPrim cs schLazy) racketString schLazy 0 val
pure $ mkWorld $
"(set-" ++ fromString s ++ "-" ++ fromString fld ++ "! " ++ structsc ++ " " ++ valsc ++ ")"
racketPrim cs schLazy i SetField [_,_,_,_,_,_,_,_]
= pure "(error \"bad setField\")"
racketPrim cs schLazy i SysCodegen []
= pure $ "\"racket\""
racketPrim cs schLazy i OnCollect [_, p, c, world]
= do p' <- schExp cs (racketPrim cs schLazy) racketString schLazy 0 p
c' <- schExp cs (racketPrim cs schLazy) racketString schLazy 0 c
pure $ mkWorld $ "(blodwen-register-object " ++ p' ++ " " ++ c' ++ ")"
racketPrim cs schLazy i OnCollectAny [p, c, world]
= do p' <- schExp cs (racketPrim cs schLazy) racketString schLazy 0 p
c' <- schExp cs (racketPrim cs schLazy) racketString schLazy 0 c
pure $ mkWorld $ "(blodwen-register-object " ++ p' ++ " " ++ c' ++ ")"
racketPrim cs schLazy i prim args
= schExtCommon cs (racketPrim cs schLazy) racketString schLazy i prim args
-- Reference label for keeping track of loaded external libraries
data Loaded : Type where
@ -390,17 +389,19 @@ compileToRKT c appdir tm outfile
let ndefs = namedDefs cdata
let ctm = forget (mainExpr cdata)
ds <- getDirectives Racket
let schLazy = if getWeakMemoLazy ds then weakMemoLaziness else defaultLaziness
defs <- get Ctxt
f <- newRef {t = List String} Done empty
l <- newRef {t = List String} Loaded []
s <- newRef {t = List String} Structs []
fgndefs <- traverse (getFgnCall appdir) ndefs
(sortedDefs, constants) <- sortDefs ndefs
compdefs <- traverse (getScheme constants (racketPrim constants) racketString) sortedDefs
compdefs <- traverse (getScheme constants (racketPrim constants schLazy) racketString schLazy) sortedDefs
let code = concat (map snd fgndefs) ++ concat compdefs
main <- schExp constants (racketPrim constants) racketString 0 ctm
main <- schExp constants (racketPrim constants schLazy) racketString schLazy 0 ctm
support <- readDataFile "racket/support.rkt"
ds <- getDirectives Racket
extraRuntime <- getExtraRuntime ds
let prof = profile !getSession
let runmain

View File

@ -1,3 +1,5 @@
#!chezscheme
(define (blodwen-os)
(case (machine-type)
[(i3le ti3le a6le ta6le tarm64le) "unix"] ; GNU/Linux
@ -19,6 +21,16 @@
(void))
res))))
(define (blodwen-delay-lazy f)
(weak-cons #!bwp f))
(define (blodwen-force-lazy e)
(let ((exval (car e)))
(if (bwp-object? exval)
(let ((val ((cdr e))))
(begin (set-car! e val) val))
exval)))
(define (blodwen-toSignedInt x bits)
(if (logbit? bits x)
(logor x (ash -1 bits))

View File

@ -25,6 +25,12 @@
(void))
res))))
(define (blodwen-delay-lazy f)
f)
(define (blodwen-force-lazy f)
(f))
(define (blodwen-toSignedInt x bits)
(if (bit-set? bits x)
(bitwise-ior x (arithmetic-shift -1 bits))

View File

@ -16,6 +16,18 @@
(void))
res))))
(define bwp 'bwp)
(define (blodwen-delay-lazy f)
(mcons (make-weak-box bwp) f))
(define (blodwen-force-lazy e)
(let ((exval (weak-box-value (mcar e) bwp)))
(if (eq? exval bwp)
(let ((val ((mcdr e))))
(begin (set-mcar! e (make-weak-box val)) val))
exval)))
(define (blodwen-toSignedInt x bits)
(if (bitwise-bit-set? x bits)
(bitwise-ior x (arithmetic-shift (- 1) bits))

View File

@ -0,0 +1,69 @@
-- This tests checks, whether lazy values constants are
-- properly memoized. If they are not, this test
-- will perform 10^20 additions and will therefore not
-- finish in reasonable time.
n0 : Lazy Nat
n0 = 1
n1 : Lazy Nat
n1 = n0 + n0 + n0 + n0 + n0 + n0 + n0 + n0 + n0 + n0
n2 : Lazy Nat
n2 = n1 + n1 + n1 + n1 + n1 + n1 + n1 + n1 + n1 + n1
n3 : Lazy Nat
n3 = n2 + n2 + n2 + n2 + n2 + n2 + n2 + n2 + n2 + n2
n4 : Lazy Nat
n4 = n3 + n3 + n3 + n3 + n3 + n3 + n3 + n3 + n3 + n3
n5 : Lazy Nat
n5 = n4 + n4 + n4 + n4 + n4 + n4 + n4 + n4 + n4 + n4
n6 : Lazy Nat
n6 = n5 + n5 + n5 + n5 + n5 + n5 + n5 + n5 + n5 + n5
n7 : Lazy Nat
n7 = n6 + n6 + n6 + n6 + n6 + n6 + n6 + n6 + n6 + n6
n8 : Lazy Nat
n8 = n7 + n7 + n7 + n7 + n7 + n7 + n7 + n7 + n7 + n7
n9 : Lazy Nat
n9 = n8 + n8 + n8 + n8 + n8 + n8 + n8 + n8 + n8 + n8
n10 : Lazy Nat
n10 = n9 + n9 + n9 + n9 + n9 + n9 + n9 + n9 + n9 + n9
n11 : Lazy Nat
n11 = n10 + n10 + n10 + n10 + n10 + n10 + n10 + n10 + n10 + n10
n12 : Lazy Nat
n12 = n11 + n11 + n11 + n11 + n11 + n11 + n11 + n11 + n11 + n11
n13 : Lazy Nat
n13 = n12 + n12 + n12 + n12 + n12 + n12 + n12 + n12 + n12 + n12
n14 : Lazy Nat
n14 = n13 + n13 + n13 + n13 + n13 + n13 + n13 + n13 + n13 + n13
n15 : Lazy Nat
n15 = n14 + n14 + n14 + n14 + n14 + n14 + n14 + n14 + n14 + n14
n16 : Lazy Nat
n16 = n15 + n15 + n15 + n15 + n15 + n15 + n15 + n15 + n15 + n15
n17 : Lazy Nat
n17 = n16 + n16 + n16 + n16 + n16 + n16 + n16 + n16 + n16 + n16
n18 : Lazy Nat
n18 = n17 + n17 + n17 + n17 + n17 + n17 + n17 + n17 + n17 + n17
n19 : Lazy Nat
n19 = n18 + n18 + n18 + n18 + n18 + n18 + n18 + n18 + n18 + n18
n20 : Lazy Nat
n20 = n19 + n19 + n19 + n19 + n19 + n19 + n19 + n19 + n19 + n19
main : IO ()
main = do printLn n20

View File

@ -0,0 +1 @@
100000000000000000000

View File

@ -0,0 +1,3 @@
. ../../testutils.sh
run Memo.idr

View File

@ -0,0 +1,38 @@
-- This tests checks, whether lazy values constants are
-- properly memoized. If they are not, this test
-- will perform 10^20 additions and will therefore not
-- finish in reasonable time.
-- Turn on weak memoisation
%cg chez lazy=weakMemo
%cg racket lazy=weakMemo
-- We return lazy values in a monad to avoid behaviour of common expression elimination
nx : Lazy Nat -> IO $ Lazy Nat
nx n = pure $ delay $ force n + force n + force n + force n + force n + force n + force n + force n + force n + force n
main : IO ()
main = do
n0 <- pure 1
n1 <- nx n0
n2 <- nx n1
n3 <- nx n2
n4 <- nx n3
n5 <- nx n4
n6 <- nx n5
n7 <- nx n6
n8 <- nx n7
n9 <- nx n8
n10 <- nx n9
n11 <- nx n10
n12 <- nx n11
n13 <- nx n12
n14 <- nx n13
n15 <- nx n14
n16 <- nx n15
n17 <- nx n16
n18 <- nx n17
n19 <- nx n18
n20 <- nx n19
printLn n20

View File

@ -0,0 +1 @@
100000000000000000000

View File

@ -0,0 +1,3 @@
. ../../testutils.sh
run Memo.idr

View File

@ -0,0 +1,49 @@
import Data.List.Lazy
import Data.Stream
import Debug.Trace
-- Turn on weak memoisation
%cg chez lazy=weakMemo
%cg racket lazy=weakMemo
S' : (pref : String) -> Nat -> Nat
S' pref = S . traceValBy (\n => "\{pref} \{show n}")
-- We return lazy values in a monad to avoid behaviour of common expression elimination
natsS' : IO $ Stream Nat
natsS' = pure $ iterate (S' "> s") Z
natsL' : IO $ LazyList Nat
natsL' = pure $ iterateN 200 (S' "> ll") Z
main : IO ()
main = do
natsS <- natsS'
putStrLn "\n-----------------------"
putStrLn "first take of stream (should be `s 0..9`)"
printLn $ take 10 natsS
putStrLn "\n-----------------------"
putStrLn "second take of stream (should be no `s *`)"
printLn $ take 10 natsS
natsL <- natsL'
putStrLn "\n-----------------------"
putStrLn "first take of short lazy list (should be `ll 0..9`)"
printLn $ take 10 natsL
putStrLn "\n-----------------------"
putStrLn "second take of short lazy list (should be no `ll *`)"
printLn $ take 10 natsL
putStrLn "\n-----------------------"
putStrLn "first take of longer lazy list (should be `ll 10..14`)"
printLn $ take 15 natsL
putStrLn "\n-----------------------"
putStrLn "second take of longer lazy list (should be no `ll *`)"
printLn $ take 15 natsL

View File

@ -0,0 +1,49 @@
-----------------------
first take of stream (should be `s 0..9`)
> s 0
> s 1
> s 2
> s 3
> s 4
> s 5
> s 6
> s 7
> s 8
> s 9
[0, 1, 2, 3, 4, 5, 6, 7, 8, 9]
-----------------------
second take of stream (should be no `s *`)
[0, 1, 2, 3, 4, 5, 6, 7, 8, 9]
-----------------------
first take of short lazy list (should be `ll 0..9`)
> ll 0
> ll 1
> ll 2
> ll 3
> ll 4
> ll 5
> ll 6
> ll 7
> ll 8
> ll 9
[0, 1, 2, 3, 4, 5, 6, 7, 8, 9]
-----------------------
second take of short lazy list (should be no `ll *`)
[0, 1, 2, 3, 4, 5, 6, 7, 8, 9]
-----------------------
first take of longer lazy list (should be `ll 10..14`)
> ll 10
> ll 11
> ll 12
> ll 13
> ll 14
[0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14]
-----------------------
second take of longer lazy list (should be no `ll *`)
[0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14]

View File

@ -0,0 +1,3 @@
. ../../testutils.sh
run --find-ipkg Memo.idr

View File

@ -0,0 +1,5 @@
package test
main = Memo
depends = contrib

View File

@ -0,0 +1,40 @@
import Data.List.Lazy
import Debug.Trace
-- Turn on weak memoisation
%cg chez lazy=weakMemo
%cg racket lazy=weakMemo
S' : (pref : String) -> Nat -> Nat
S' pref = S . traceValBy (\n => "\{pref} \{show n}")
-- We return lazy values in a monad to avoid behaviour of common expression elimination
natsL' : IO $ LazyList Nat
natsL' = pure $ iterateN 200 (S' "> ll") Z
%foreign "scheme,chez:collect"
"scheme,racket:collect-garbage"
prim__gc : PrimIO ()
gc : IO ()
gc = primIO prim__gc
main : IO ()
main = do
natsL <- natsL'
putStrLn "\n-----------------------"
putStrLn "first take of lazy list (should be `ll 0..9`)"
printLn $ take 10 natsL
putStrLn "\n-----------------------"
putStrLn "second take of lazy list (should be no `ll *`)"
printLn $ take 10 natsL
gc
putStrLn "\n-----------------------"
putStrLn "take of lazy list after gc (should be `ll 0..9`)"
printLn $ take 10 natsL

View File

@ -0,0 +1,32 @@
-----------------------
first take of lazy list (should be `ll 0..9`)
> ll 0
> ll 1
> ll 2
> ll 3
> ll 4
> ll 5
> ll 6
> ll 7
> ll 8
> ll 9
[0, 1, 2, 3, 4, 5, 6, 7, 8, 9]
-----------------------
second take of lazy list (should be no `ll *`)
[0, 1, 2, 3, 4, 5, 6, 7, 8, 9]
-----------------------
take of lazy list after gc (should be `ll 0..9`)
> ll 0
> ll 1
> ll 2
> ll 3
> ll 4
> ll 5
> ll 6
> ll 7
> ll 8
> ll 9
[0, 1, 2, 3, 4, 5, 6, 7, 8, 9]

View File

@ -0,0 +1,3 @@
. ../../testutils.sh
run --find-ipkg Memo.idr

View File

@ -0,0 +1,5 @@
package test
main = Memo
depends = contrib

View File

@ -8,3 +8,14 @@ This line is also printed
#1
#2
#3
----
String
Hello Foo!
Hello Bar!
Hello Idris!
Hello World!
This line is printed
This line is also printed
#1
#2
#3

View File

@ -5,4 +5,11 @@ idris2 --cg chez Futures.idr -p contrib --exec simpleIO
idris2 --cg chez Futures.idr -p contrib --exec erasureAndNonbindTest
idris2 --cg chez Futures.idr -p contrib --exec map
echo "----"
idris2 --cg chez --directive lazy=weakMemo Futures.idr -p contrib --exec constant
idris2 --cg chez --directive lazy=weakMemo Futures.idr -p contrib --exec simpleIO
idris2 --cg chez --directive lazy=weakMemo Futures.idr -p contrib --exec erasureAndNonbindTest
idris2 --cg chez --directive lazy=weakMemo Futures.idr -p contrib --exec map
rm -r build

View File

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

View File

@ -1,5 +1,7 @@
. ../../testutils.sh
idris2 --cg chez Futures.idr -p contrib --exec main
echo "----"
idris2 --cg chez --directive lazy=weakMemo Futures.idr -p contrib --exec main
rm -r build

View File

@ -2,3 +2,8 @@ String
#1
#2
#3
----
String
#1
#2
#3

View File

@ -2,3 +2,8 @@
idris2 --cg racket Futures.idr -p contrib --exec constant
idris2 --cg racket Futures.idr -p contrib --exec map
echo "----"
idris2 --cg racket --directive lazy=weakMemo Futures.idr -p contrib --exec constant
idris2 --cg racket --directive lazy=weakMemo Futures.idr -p contrib --exec map

View File

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

View File

@ -1,5 +1,7 @@
. ../../testutils.sh
idris2 --cg racket Futures.idr -p contrib --exec main
echo "----"
idris2 --cg racket --directive lazy=weakMemo Futures.idr -p contrib --exec main
rm -r build