diff --git a/CHANGELOG_NEXT.md b/CHANGELOG_NEXT.md index c24735d6bb..db1651740f 100644 --- a/CHANGELOG_NEXT.md +++ b/CHANGELOG_NEXT.md @@ -127,12 +127,19 @@ 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. + #### 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. +* `Lazy` and `Inf` values are *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. diff --git a/src/Compiler/Scheme/Common.idr b/src/Compiler/Scheme/Common.idr index 662ae136d0..140df845dc 100644 --- a/src/Compiler/Scheme/Common.idr +++ b/src/Compiler/Scheme/Common.idr @@ -562,8 +562,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 $ "(blodwen-force-lazy " ++ !(schExp i t) ++ ")" + schExp i (NmDelay fc lr t) = pure $ "(blodwen-delay-lazy (lambda () " ++ !(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), diff --git a/support/chez/support.ss b/support/chez/support.ss index b37b5a0270..8162585b87 100644 --- a/support/chez/support.ss +++ b/support/chez/support.ss @@ -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)) diff --git a/support/gambit/support.scm b/support/gambit/support.scm index 13e1d9b5f3..898c536ddd 100644 --- a/support/gambit/support.scm +++ b/support/gambit/support.scm @@ -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)) diff --git a/support/racket/support.rkt b/support/racket/support.rkt index db3b0c2340..a5d4bac2a1 100644 --- a/support/racket/support.rkt +++ b/support/racket/support.rkt @@ -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)) diff --git a/tests/chez/memo/Memo.idr b/tests/allschemes/memo001/Memo.idr similarity index 100% rename from tests/chez/memo/Memo.idr rename to tests/allschemes/memo001/Memo.idr diff --git a/tests/chez/memo/expected b/tests/allschemes/memo001/expected similarity index 100% rename from tests/chez/memo/expected rename to tests/allschemes/memo001/expected diff --git a/tests/chez/memo/run b/tests/allschemes/memo001/run similarity index 100% rename from tests/chez/memo/run rename to tests/allschemes/memo001/run diff --git a/tests/allschemes/memo002/Memo.idr b/tests/allschemes/memo002/Memo.idr new file mode 100644 index 0000000000..d70a4934aa --- /dev/null +++ b/tests/allschemes/memo002/Memo.idr @@ -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 diff --git a/tests/allschemes/memo002/expected b/tests/allschemes/memo002/expected new file mode 100644 index 0000000000..493d3b8b14 --- /dev/null +++ b/tests/allschemes/memo002/expected @@ -0,0 +1 @@ +100000000000000000000 diff --git a/tests/allschemes/memo002/run b/tests/allschemes/memo002/run new file mode 100644 index 0000000000..f8843a026d --- /dev/null +++ b/tests/allschemes/memo002/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +run Memo.idr diff --git a/tests/allschemes/memo003/Memo.idr b/tests/allschemes/memo003/Memo.idr new file mode 100644 index 0000000000..4c8363c4b4 --- /dev/null +++ b/tests/allschemes/memo003/Memo.idr @@ -0,0 +1,34 @@ +-- 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. + +-- 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 diff --git a/tests/allschemes/memo003/expected b/tests/allschemes/memo003/expected new file mode 100644 index 0000000000..493d3b8b14 --- /dev/null +++ b/tests/allschemes/memo003/expected @@ -0,0 +1 @@ +100000000000000000000 diff --git a/tests/allschemes/memo003/run b/tests/allschemes/memo003/run new file mode 100644 index 0000000000..f8843a026d --- /dev/null +++ b/tests/allschemes/memo003/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +run Memo.idr diff --git a/tests/allschemes/memo004/Memo.idr b/tests/allschemes/memo004/Memo.idr new file mode 100644 index 0000000000..e42e616a93 --- /dev/null +++ b/tests/allschemes/memo004/Memo.idr @@ -0,0 +1,45 @@ +import Data.List.Lazy +import Data.Stream + +import Debug.Trace + +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 diff --git a/tests/allschemes/memo004/expected b/tests/allschemes/memo004/expected new file mode 100644 index 0000000000..27757b8ec5 --- /dev/null +++ b/tests/allschemes/memo004/expected @@ -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] diff --git a/tests/allschemes/memo004/run b/tests/allschemes/memo004/run new file mode 100644 index 0000000000..2faac00933 --- /dev/null +++ b/tests/allschemes/memo004/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +run --find-ipkg Memo.idr diff --git a/tests/allschemes/memo004/test.ipkg b/tests/allschemes/memo004/test.ipkg new file mode 100644 index 0000000000..e45ed0c98d --- /dev/null +++ b/tests/allschemes/memo004/test.ipkg @@ -0,0 +1,5 @@ +package test + +main = Memo + +depends = contrib diff --git a/tests/allschemes/memo005/Memo.idr b/tests/allschemes/memo005/Memo.idr new file mode 100644 index 0000000000..7f96c3bcb9 --- /dev/null +++ b/tests/allschemes/memo005/Memo.idr @@ -0,0 +1,36 @@ +import Data.List.Lazy + +import Debug.Trace + +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 diff --git a/tests/allschemes/memo005/expected b/tests/allschemes/memo005/expected new file mode 100644 index 0000000000..f97ff55014 --- /dev/null +++ b/tests/allschemes/memo005/expected @@ -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] diff --git a/tests/allschemes/memo005/run b/tests/allschemes/memo005/run new file mode 100644 index 0000000000..2faac00933 --- /dev/null +++ b/tests/allschemes/memo005/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +run --find-ipkg Memo.idr diff --git a/tests/allschemes/memo005/test.ipkg b/tests/allschemes/memo005/test.ipkg new file mode 100644 index 0000000000..e45ed0c98d --- /dev/null +++ b/tests/allschemes/memo005/test.ipkg @@ -0,0 +1,5 @@ +package test + +main = Memo + +depends = contrib