Skip to content

Commit

Permalink
[ perf ] Use delay and force to memoize lazy toplevel definitions
Browse files Browse the repository at this point in the history
  • Loading branch information
AlgebraicWolf committed Oct 9, 2023
1 parent fa54392 commit cc30795
Show file tree
Hide file tree
Showing 2 changed files with 11 additions and 1 deletion.
11 changes: 10 additions & 1 deletion src/Compiler/Scheme/Common.idr
Original file line number Diff line number Diff line change
Expand Up @@ -535,10 +535,11 @@ parameters (constants : SortedSet Name,
= do val' <- schExp i val
sc' <- schExp i sc
pure $ "(let ((" ++ schName x ++ " " ++ val' ++ ")) " ++ sc' ++ ")"
schExp i (NmApp fc x@(NmRef _ n) []) =
schExp i (NmApp fc x@(NmRef exp n) []) =
if contains n constants
then schExp i x
else pure $ "(" ++ !(schExp i x) ++ ")"

schExp i (NmApp fc x args)
= pure $ "(" ++ !(schExp i x) ++ " " ++ sepBy " " !(traverse (schExp i) args) ++ ")"
schExp i (NmCon fc _ NIL tag []) = pure $ "'()"
Expand All @@ -562,6 +563,8 @@ parameters (constants : SortedSet Name,
= schOp op !(schArgs i args)
schExp i (NmExtPrim fc p args)
= 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 (NmConCase fc sc alts def)
Expand Down Expand Up @@ -652,11 +655,17 @@ parameters (constants : SortedSet Name,

schDef : {auto c : Ref Ctxt Defs} ->
Name -> NamedDef -> Core Builder

schDef n (MkNmFun [] (NmDelay _ _ exp))
= pure $ "(define " ++ schName !(getFullName n) ++ "(delay "
++ !(schExp 0 exp) ++ "))\n" -- Special version for memoized toplevel lazy definitions

schDef n (MkNmFun [] exp)
= if contains n constants
then pure $ "(define " ++ schName !(getFullName n) ++ " " ++ !(schExp 0 exp) ++ ")\n"
else pure $ "(define " ++ schName !(getFullName n) ++ " (lambda () " ++ !(schExp 0 exp) ++ "))\n"


schDef n (MkNmFun args exp)
= pure $ "(define " ++ schName !(getFullName n) ++ " (lambda (" ++ schArglist args ++ ") "
++ !(schExp 0 exp) ++ "))\n"
Expand Down
1 change: 1 addition & 0 deletions src/Compiler/Scheme/Racket.idr
Original file line number Diff line number Diff line change
Expand Up @@ -48,6 +48,7 @@ schHeader prof libs = fromString """
(require racket/async-channel) ; for asynchronous channels
(require racket/future) ; for parallelism/concurrency
(require racket/math) ; for math ops
(require racket/promise) ; for delay/force in toplevel defs
(require racket/system) ; for system
(require racket/unsafe/ops) ; for fast fixnum ops
(require rnrs/bytevectors-6) ; for buffers
Expand Down

0 comments on commit cc30795

Please sign in to comment.