Skip to content

Commit

Permalink
[RefC] Remove variable only from body in let
Browse files Browse the repository at this point in the history
Co-authored-by: G. Allais <[email protected]>
  • Loading branch information
Alex1005a and gallais authored Jun 20, 2023
1 parent 90bd5a2 commit 8021fa1
Show file tree
Hide file tree
Showing 2 changed files with 7 additions and 7 deletions.
10 changes: 5 additions & 5 deletions src/Compiler/ANF.idr
Original file line number Diff line number Diff line change
Expand Up @@ -289,7 +289,7 @@ freeVariables (AAppName _ _ n args) = fromList args
freeVariables (AUnderApp _ n _ args) = fromList args
freeVariables (AApp _ _ closure arg) = fromList [closure, arg]
freeVariables (ALet _ var value body) =
delete (ALocal var) $ union (freeVariables value) (freeVariables body)
union (freeVariables value) (delete (ALocal var) $ freeVariables body)
freeVariables (ACon _ _ _ _ args) = fromList args
freeVariables (AOp _ _ _ args) = fromList $ foldl (\acc, elem => elem :: acc) [] args
freeVariables (AExtPrim _ _ _ args) = fromList args
Expand All @@ -300,13 +300,13 @@ freeVariables (AConCase _ sc alts mDef) =
let vars : List (SortedSet AVar) = case mDef of
Just anf => freeVariables anf :: altsAnf
Nothing => altsAnf in
insert sc $ foldl (\acc, var => union acc var) empty vars
insert sc $ concat vars
freeVariables (AConstCase _ sc alts mDef) =
let altsAnf = map (\(MkAConstAlt _ caseBody) => caseBody) alts in
let anfs : List ANF = case mDef of
Just anf => anf :: altsAnf
Nothing => altsAnf in
insert sc $ foldl (\acc, anf => union acc $ freeVariables anf) empty anfs
insert sc $ foldMap freeVariables anfs
freeVariables _ = empty

export
Expand All @@ -325,11 +325,11 @@ usedConstructors (AConCase _ sc alts mDef) =
let anfs : List (SortedSet Name) = case mDef of
Just anf => usedConstructors anf :: altsAnf
Nothing => altsAnf in
foldl (\acc, anf => union acc anf) empty anfs
concat anfs
usedConstructors (AConstCase _ sc alts mDef) =
let altsAnf = map (\(MkAConstAlt _ caseBody) => caseBody) alts in
let anfs : List ANF = case mDef of
Just anf => anf :: altsAnf
Nothing => altsAnf in
foldl (\acc, anf => union acc $ usedConstructors anf) empty anfs
foldMap usedConstructors anfs
usedConstructors _ = empty
4 changes: 2 additions & 2 deletions src/Compiler/RefC/RefC.idr
Original file line number Diff line number Diff line change
Expand Up @@ -392,7 +392,7 @@ avarToC : Env -> AVar -> Core String
avarToC env var = do
if contains var env.borrowed then pure $ "newReference(" ++ varName var ++ ")"
else if contains var env.owned then pure $ varName var
else coreFail $ InternalError $ "INTERNAL ERROR: variable " ++ varName var ++ " is not owned and borrowed "
else coreFail $ InternalError $ "INTERNAL ERROR: variable " ++ varName var ++ " is neither owned nor borrowed "

moveFromOwnedToBorrowed : Env -> SortedSet AVar -> Env
moveFromOwnedToBorrowed env vars = { borrowed $= (union vars), owned $= (`difference` vars) } env
Expand Down Expand Up @@ -498,7 +498,7 @@ callByPosition NotInTailPosition = nonTailCall
||| Returns constructor variables to remove and constructors to reuse.
dropUnusedReuseCons : ReuseMap -> SortedSet Name -> (List String, ReuseMap)
dropUnusedReuseCons reuseMap usedCons =
-- if there is no constructor named by that name, than the reuse constructor is deleted
-- if there is no constructor named by that name, then the reuse constructor is deleted
let dropReuseMap = differenceMap reuseMap usedCons in
let actualReuseMap = intersectionMap reuseMap usedCons in
(values dropReuseMap, actualReuseMap)
Expand Down

0 comments on commit 8021fa1

Please sign in to comment.