Skip to content

Commit

Permalink
[RefC] Use locally function as combinator
Browse files Browse the repository at this point in the history
  • Loading branch information
Alex1005a committed Jun 30, 2023
1 parent 5f5b491 commit 7cb98c2
Showing 1 changed file with 25 additions and 22 deletions.
47 changes: 25 additions & 22 deletions src/Compiler/RefC/RefC.idr
Original file line number Diff line number Diff line change
Expand Up @@ -509,6 +509,13 @@ dropUnusedOwnedVars owned usedVars =
let shouldDrop = difference owned actualOwned in
(varName <$> SortedSet.toList shouldDrop, actualOwned)

locally : {auto t : Ref EnvTracker Env} -> Env -> Core () -> Core ()
locally newEnv act = do
oldEnv <- get EnvTracker
put EnvTracker newEnv
act
put EnvTracker oldEnv

mutual
copyConstructors : {auto a : Ref ArgCounter Nat}
-> {auto oft : Ref OutfileText Output}
Expand Down Expand Up @@ -556,13 +563,12 @@ mutual
(shouldDrop, actualReuseMap) <- addReuseConstructor env.reuseMap conArgs usedCons shouldDrop actualReuseMap
removeVars shouldDrop
removeReuseConstructors dropReuseCons
put EnvTracker ({owned := actualOwned, reuseMap := actualReuseMap} env)
assignment <- cStatementsFromANF body tailStatus
emit EmptyFC $ retValVar ++ " = " ++ callByPosition tailStatus assignment ++ ";"
emit EmptyFC $ "break;"
decreaseIndentation
emit EmptyFC $ " }"
put EnvTracker env
locally ({owned := actualOwned, reuseMap := actualReuseMap} env) $ do
assignment <- cStatementsFromANF body tailStatus
emit EmptyFC $ retValVar ++ " = " ++ callByPosition tailStatus assignment ++ ";"
emit EmptyFC $ "break;"
decreaseIndentation
emit EmptyFC $ " }"
conBlocks sc xs retValVar (S k) tailStatus
where
-- if the constructor is unique use it, otherwise add it to should drop vars and create null constructor
Expand Down Expand Up @@ -628,13 +634,12 @@ mutual
increaseIndentation
removeReuseConstructors dropReuseCons
removeVars shouldDrop
put EnvTracker ({owned := actualOwned, reuseMap := actualReuseMap} env)
assignment <- cStatementsFromANF caseBody tailStatus
emit EmptyFC $ retValVar ++ " = " ++ callByPosition tailStatus assignment ++ ";"
emit EmptyFC "break;"
decreaseIndentation
emit EmptyFC " }"
put EnvTracker env
locally ({owned := actualOwned, reuseMap := actualReuseMap} env) $ do
assignment <- cStatementsFromANF caseBody tailStatus
emit EmptyFC $ retValVar ++ " = " ++ callByPosition tailStatus assignment ++ ";"
emit EmptyFC "break;"
decreaseIndentation
emit EmptyFC " }"
constBlockSwitch alts retValVar (i+1) tailStatus


Expand All @@ -658,12 +663,11 @@ mutual
increaseIndentation
removeReuseConstructors dropReuseCons
removeVars shouldDrop
put EnvTracker ({owned := actualOwned, reuseMap := actualReuseMap} env)
assignment <- cStatementsFromANF defaultBody tailStatus
emit EmptyFC $ retValVar ++ " = " ++ callByPosition tailStatus assignment ++ ";"
decreaseIndentation
emit EmptyFC " }"
put EnvTracker env
locally ({owned := actualOwned, reuseMap := actualReuseMap} env) $ do
assignment <- cStatementsFromANF defaultBody tailStatus
emit EmptyFC $ retValVar ++ " = " ++ callByPosition tailStatus assignment ++ ";"
decreaseIndentation
emit EmptyFC " }"



Expand Down Expand Up @@ -792,7 +796,6 @@ mutual
let returnLine = (cCleanString (show (toPrim p)) ++ "("++ showSep ", " (map varName args) ++")")
pure $ MkRS returnLine returnLine
cStatementsFromANF (AConCase fc sc alts mDef) tailPosition = do
env <- get EnvTracker
c <- getNextCounter
switchReturnVar <- getNewVarThatWillNotBeFreedAtEndOfBlock
let newValueLine = "Value * " ++ switchReturnVar ++ " = NULL;"
Expand All @@ -817,6 +820,7 @@ mutual
emit EmptyFC $ "free(" ++ constructorField ++ ");"
pure $ MkRS switchReturnVar switchReturnVar
(Just d) => do
env <- get EnvTracker
let (shouldDrop, actualOwned) = dropUnusedOwnedVars env.owned (freeVariables d)
let usedCons = usedConstructors d
let (dropReuseCons, actualReuseMap) = dropUnusedReuseCons env.reuseMap usedCons
Expand All @@ -833,7 +837,6 @@ mutual
emit EmptyFC $ "free(" ++ constructorField ++ ");"
pure $ MkRS switchReturnVar switchReturnVar
cStatementsFromANF (AConstCase fc sc alts def) tailPosition = do
env <- get EnvTracker
switchReturnVar <- getNewVarThatWillNotBeFreedAtEndOfBlock
let newValueLine = "Value * " ++ switchReturnVar ++ " = NULL;"
emit fc newValueLine
Expand Down

0 comments on commit 7cb98c2

Please sign in to comment.