diff --git a/CHANGELOG_NEXT.md b/CHANGELOG_NEXT.md index 7791df0f8d1..bdf63b90c44 100644 --- a/CHANGELOG_NEXT.md +++ b/CHANGELOG_NEXT.md @@ -122,9 +122,9 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO * More efficient `collect-request-handler` is used. -* `Lazy` and `Inf` values are *weakly* memoised. That is, once accessed, they are allowed - to be not re-evaluated until garbage collector wipes them. - This change requires Chez with version >= 9.5.9, because of fixed bug in Chez GC. +* 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 @@ -132,8 +132,9 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO evaluated. Now when a delayed expression is lifted by CSE, it is compiled using Scheme's `delay` and `force` to memoize them. -* `Lazy` and `Inf` values are *weakly* memoised. That is, once accessed, they are allowed - to be not re-evaluated until garbage collector wipes 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 diff --git a/docs/source/backends/chez.rst b/docs/source/backends/chez.rst index 5880c615c9c..4fe4b2b9420 100644 --- a/docs/source/backends/chez.rst +++ b/docs/source/backends/chez.rst @@ -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 ================================ diff --git a/docs/source/backends/racket.rst b/docs/source/backends/racket.rst index c30f6c8d4d8..b9a69850366 100644 --- a/docs/source/backends/racket.rst +++ b/docs/source/backends/racket.rst @@ -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. diff --git a/src/Compiler/Common.idr b/src/Compiler/Common.idr index e39ca640b6a..e113a34d0af 100644 --- a/src/Compiler/Common.idr +++ b/src/Compiler/Common.idr @@ -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 diff --git a/src/Compiler/Scheme/Chez.idr b/src/Compiler/Scheme/Chez.idr index 22418c87d4e..db34eea1d2a 100644 --- a/src/Compiler/Scheme/Chez.idr +++ b/src/Compiler/Scheme/Chez.idr @@ -133,48 +133,47 @@ 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 - - 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 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 +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)) + +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 MakeFuture [_, work] + = do work' <- schExp cs (chezExtPrim cs schLazy) chezString schLazy 0 $ NmForce EmptyFC LUnknown work + pure $ "(blodwen-make-future (lambda () " ++ work' ++ "))" +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 @@ -505,10 +504,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 _) @@ -554,7 +555,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 ++ @@ -694,7 +695,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) diff --git a/src/Compiler/Scheme/ChezSep.idr b/src/Compiler/Scheme/ChezSep.idr index 202ede6d9d8..f51386f55fe 100644 --- a/src/Compiler/Scheme/ChezSep.idr +++ b/src/Compiler/Scheme/ChezSep.idr @@ -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 diff --git a/src/Compiler/Scheme/Common.idr b/src/Compiler/Scheme/Common.idr index 3df0d327095..42bc64eb693 100644 --- a/src/Compiler/Scheme/Common.idr +++ b/src/Compiler/Scheme/Common.idr @@ -320,9 +320,28 @@ var _ = False getScrutineeTemp : Nat -> Builder getScrutineeTemp i = fromString $ "sc" ++ show i +public export +record LazyExprProc where + constructor MkLazyExprProc + processForce : Builder -> Builder + processDelay : Builder -> Builder + +public export +defaultLaziness : LazyExprProc +defaultLaziness = MkLazyExprProc + (\expr => "(" ++ expr ++ ")") + (\expr => "(lambda () " ++ expr ++ ")") + +public export +weakMemoLaziness : LazyExprProc +weakMemoLaziness = MkLazyExprProc + (\expr => "(blodwen-force-lazy " ++ expr ++ ")") + (\expr => "(blodwen-delay-lazy (lambda () " ++ 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) @@ -565,8 +584,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 $ "(blodwen-force-lazy " ++ !(schExp i t) ++ ")" - schExp i (NmDelay fc lr t) = pure $ "(blodwen-delay-lazy (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), @@ -681,6 +700,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 diff --git a/src/Compiler/Scheme/Gambit.idr b/src/Compiler/Scheme/Gambit.idr index 106a70eba29..d73ca7962e5 100644 --- a/src/Compiler/Scheme/Gambit.idr +++ b/src/Compiler/Scheme/Gambit.idr @@ -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 - - 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 +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)) + +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] diff --git a/src/Compiler/Scheme/Racket.idr b/src/Compiler/Scheme/Racket.idr index c7fa06d0fee..b3d661c63c4 100644 --- a/src/Compiler/Scheme/Racket.idr +++ b/src/Compiler/Scheme/Racket.idr @@ -86,37 +86,36 @@ 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 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 +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 MakeFuture [_, work] + = do work' <- schExp cs (racketPrim cs schLazy) racketString schLazy 0 $ NmForce EmptyFC LUnknown work + pure $ mkWorld $ "(blodwen-make-future (lambda () " ++ work' ++ "))" +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 @@ -393,17 +392,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 diff --git a/tests/allschemes/memo003/Memo.idr b/tests/allschemes/memo003/Memo.idr index 4c8363c4b49..839d0558149 100644 --- a/tests/allschemes/memo003/Memo.idr +++ b/tests/allschemes/memo003/Memo.idr @@ -3,6 +3,10 @@ -- 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 diff --git a/tests/allschemes/memo004/Memo.idr b/tests/allschemes/memo004/Memo.idr index e42e616a93e..07a06a7a965 100644 --- a/tests/allschemes/memo004/Memo.idr +++ b/tests/allschemes/memo004/Memo.idr @@ -3,6 +3,10 @@ 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}") diff --git a/tests/allschemes/memo005/Memo.idr b/tests/allschemes/memo005/Memo.idr index 7f96c3bcb93..bf92afa6903 100644 --- a/tests/allschemes/memo005/Memo.idr +++ b/tests/allschemes/memo005/Memo.idr @@ -2,6 +2,10 @@ 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}") diff --git a/tests/chez/futures001/expected b/tests/chez/futures001/expected index 3ce1888704d..76ec8ae4185 100644 --- a/tests/chez/futures001/expected +++ b/tests/chez/futures001/expected @@ -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 diff --git a/tests/chez/futures001/run b/tests/chez/futures001/run index a1a267de56e..27a0b47fe2e 100644 --- a/tests/chez/futures001/run +++ b/tests/chez/futures001/run @@ -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 diff --git a/tests/chez/futures002/expected b/tests/chez/futures002/expected index e78e52e2eaa..afb104cab8d 100644 --- a/tests/chez/futures002/expected +++ b/tests/chez/futures002/expected @@ -1 +1,3 @@ top-level indeed +---- +top-level indeed diff --git a/tests/chez/futures002/run b/tests/chez/futures002/run index 843ed7e3980..e62745d3cd6 100644 --- a/tests/chez/futures002/run +++ b/tests/chez/futures002/run @@ -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 diff --git a/tests/racket/futures001/expected b/tests/racket/futures001/expected index 099413ff8f9..0da1337bd02 100644 --- a/tests/racket/futures001/expected +++ b/tests/racket/futures001/expected @@ -2,3 +2,8 @@ String #1 #2 #3 +---- +String +#1 +#2 +#3 diff --git a/tests/racket/futures001/run b/tests/racket/futures001/run index 9f97ae23c97..b362b8b38ac 100644 --- a/tests/racket/futures001/run +++ b/tests/racket/futures001/run @@ -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 diff --git a/tests/racket/futures002/expected b/tests/racket/futures002/expected index e78e52e2eaa..afb104cab8d 100644 --- a/tests/racket/futures002/expected +++ b/tests/racket/futures002/expected @@ -1 +1,3 @@ top-level indeed +---- +top-level indeed diff --git a/tests/racket/futures002/run b/tests/racket/futures002/run index 2c2753605c1..bd5fdbeb527 100644 --- a/tests/racket/futures002/run +++ b/tests/racket/futures002/run @@ -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