Skip to content

Commit

Permalink
[RefC] Suppress arglist wrapper (#3177)
Browse files Browse the repository at this point in the history
* [RefC] Suppress code generation for unnecessary arglist wrappers.

* [RefC] cleanup dead code of arglist.

* Removed Value_Arglist to reduce Closure's allocation overhead.

* fix linter error

* [RefC] make trampoline() safety.

* [RefC] cleanup cStatementsFromANF to keep code simple.

* fix linter error

* fix linter error

* In another time, another galaxy. THE LINTER INVADORS conquaer the all humanity and make them slaves. Under 2024, a only leaved job for every humans is adjusting spaces of source code, or just type gg0vG$== in vim.

* [ test ] update golden value

* added supports 32 params on closure.

* [RefC] [Cleanup] removing duplicate codes.

* [RefC] Switch calling conventions based on the number of arguments to avoid limits on the number of arguments and to reduce stack usage.

* [RefC] Argument that are too large are placed on the heap, as are closures.

* [RefC] use idris2_malloc instead of malloc.

* [RefC] [Cleanup] Keep pure things pure.

* [RefC] Mapped some special constructors to NULL. This reduces malloc cost and generates simpler code in ConCase. But not work yet.

* [RefC] fix merge failure.

* [RefC] stringOps.c replace  NULL for NIL.

* [RefC] cleanup

* [RefC] ConstCase now generate simple if-then statements instead of using helpers. This reduces malloc/free costs.

* fix indentation

* fix whitespaces

* [RefC] The name field in Value_Constructor was restored for tycon. But changed to static const*. Hopefully the C compiler will remove the common string constants. The smartest thing to do would be to create a dummy global variable and use its address as a tag, but that would depend on the C compiler to resolve conflicts.

* [refc] a big changes of the space

* [RefC] Little tricks to reduce temporary variables

* spaces

* [RefC] fix compiler warnings

* [RefC] [test]  Perform memory leak analysis, if valgrind is installed.

* [RefC] Fix invalid memory read. Fix C compiler warnings.

* [RefC] Fix invalid memory read of strSubstr. [test] Perform memory leak analysis, if valgrind is installed.

* [test] fix junk line

* linter

* linter

* linter

* linter

* [RefC] merge with erase_trivial_constuctors

* merge w/ erase_trivial_constructors

* Revert "merge w/ erase_trivial_constructors"

This reverts commit be593a3.

* Revert "[RefC] merge with erase_trivial_constuctors"

This reverts commit 3c21eb4.

* merge w/ upstream/main

* fix merge failure

* rename

* fix renaming

* [RefC] fix merge fail

* [RefC] renamed C functions for safty.

* [RefC] cleanup

* [RefC] Fix constructor tag of UnconsResult.CHARACTER.

---------

Co-authored-by: Mathew Polzin <[email protected]>
Co-authored-by: Guillaume Allais <[email protected]>
  • Loading branch information
3 people authored Apr 17, 2024
1 parent 4799d28 commit 1dc7b74
Show file tree
Hide file tree
Showing 11 changed files with 336 additions and 756 deletions.
7 changes: 7 additions & 0 deletions CHANGELOG_NEXT.md
Original file line number Diff line number Diff line change
Expand Up @@ -75,6 +75,13 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO

* Rename C function to avoid confliction. But only a part.

* Supress code generation of _arglist wrappers to reduce code size and compilation time.

* Removed Value_Arglist to reduce Closure's allocation overhead and make code simply.

* Switch calling conventions based on the number of arguments to avoid limits on
the number of arguments and to reduce stack usage.

#### NodeJS Backend

* The NodeJS executable output to `build/exec/` now has its executable bit set.
Expand Down
186 changes: 68 additions & 118 deletions src/Compiler/RefC/RefC.idr
Original file line number Diff line number Diff line change
Expand Up @@ -241,8 +241,7 @@ getNextCounter = do
pure $ show c

getNewVarThatWillNotBeFreedAtEndOfBlock : {auto a : Ref ArgCounter Nat} -> Core String
getNewVarThatWillNotBeFreedAtEndOfBlock = do
pure $ "tmp_" ++ !(getNextCounter)
getNewVarThatWillNotBeFreedAtEndOfBlock = pure $ "tmp_" ++ !(getNextCounter)


maxLineLengthForComment : Nat
Expand Down Expand Up @@ -316,51 +315,51 @@ avarToC env var =
-- case when the variable is borrowed
else "idris2_newReference(" ++ varName var ++ ")"

avarsToC : Owned -> List AVar -> List String
avarsToC _ [] = []
avarsToC owned (v::vars) =
let v' = varName v in
if contains v owned
then v'::avarsToC (delete v owned) vars
else "idris2_newReference(\{v'})"::avarsToC owned vars -- when v is borrowed

moveFromOwnedToBorrowed : Env -> SortedSet AVar -> Env
moveFromOwnedToBorrowed env vars = { owned $= (`difference` vars) } env

makeArglist : {auto a : Ref ArgCounter Nat}
-> {auto oft : Ref OutfileText Output}
-> {auto il : Ref IndentLevel Nat}
-> {auto e : Ref EnvTracker Env}
-> Nat
-> List AVar
-> Core String
makeArglist missing xs = do
let arglist = "arglist_" ++ !(getNextCounter)
emit EmptyFC $ "Value_Arglist *"
++ arglist
++ " = idris2_newArglist(" ++ show missing
++ "," ++ show (length xs + missing)
++ ");"
pushArgToArglist !(get EnvTracker) arglist xs 0
pure arglist
where
pushArgToArglist : Env -> String -> List AVar -> Nat -> Core ()
pushArgToArglist _ arglist [] k = pure ()
pushArgToArglist env arglist (arg :: args) k = do
let ownedArg = if contains arg env.owned then singleton arg else empty
emit EmptyFC $ arglist
++ "->args[" ++ show k ++ "] = "
++ avarToC env arg ++ ";"
pushArgToArglist (moveFromOwnedToBorrowed env ownedArg) arglist args (S k)

fillConstructorArgs : {auto oft : Ref OutfileText Output}
-> {auto il : Ref IndentLevel Nat}
-> Env
-> String
-> List AVar
-> Nat
-> Core ()
fillConstructorArgs _ _ [] _ = pure ()
fillConstructorArgs env cons (v :: vars) k = do
fillArgs : {auto oft : Ref OutfileText Output}
-> {auto il : Ref IndentLevel Nat}
-> Env
-> String
-> List AVar
-> Nat
-> Core ()
fillArgs _ _ [] _ = pure ()
fillArgs env arglist (v :: vars) k = do
let ownedVars = if contains v env.owned then singleton v else empty
emit EmptyFC $ cons ++ "->args["++ show k ++ "] = " ++ avarToC env v ++ ";"
fillConstructorArgs (moveFromOwnedToBorrowed env ownedVars) cons vars (S k)

cArgsVectANF : {0 arity : Nat} -> Vect arity AVar -> Core (Vect arity String)
cArgsVectANF [] = pure []
cArgsVectANF (x :: xs) = pure $ (varName x) :: !(cArgsVectANF xs)
emit EmptyFC $ "\{arglist}[\{show k}] = \{avarToC env v};"
fillArgs (moveFromOwnedToBorrowed env ownedVars) arglist vars (S k)

makeClosure : {auto a : Ref ArgCounter Nat}
-> {auto oft : Ref OutfileText Output}
-> {auto il : Ref IndentLevel Nat}
-> {auto e : Ref EnvTracker Env}
-> FC
-> Name
-> List AVar
-> Nat
-> Core String
makeClosure fc n args missing = do
let closure = "closure_\{!(getNextCounter)}"
let nargs = length args
emit fc "Value *\{closure} = (Value *)idris2_mkClosure((Value *(*)())\{cName n}, \{show $ nargs + missing}, \{show nargs});"
fillArgs !(get EnvTracker) "((Value_Closure*)\{closure})->args" args 0
pure closure

-- When changing this number, also change idris2_dispatch_closure in runtime.c.
-- Increasing this number will worsen stack consumption and increase the codesize of idris2_dispatch_closure.
-- In C89, the maximum number of arguments is 31, so it should not be larger than 31. 127 is safe in C99, but I do not recommend it.
MaxExtractFunArgs : Nat
MaxExtractFunArgs = 16

integer_switch : List AConstAlt -> Bool
integer_switch [] = True
Expand Down Expand Up @@ -395,7 +394,6 @@ const2Integer c i =
(B64 x) => "UINT64_C(\{show x})"
_ => show i


data TailPositionStatus = InTailPosition | NotInTailPosition

||| The function takes as arguments the current ReuseMap and the constructors that will be used.
Expand Down Expand Up @@ -487,34 +485,21 @@ mutual

cStatementsFromANF (AV fc x) _ = pure $ avarToC !(get EnvTracker) x
cStatementsFromANF (AAppName fc _ n args) tailPosition = do
emit fc $ ("// start " ++ cName n ++ "(" ++ showSep ", " (map (\v => varName v) args) ++ ")")
arglist <- makeArglist 0 args
c <- getNextCounter
let f_ptr_name = "fPtr_" ++ c
emit fc $ "Value *(*"++ f_ptr_name ++ ")(Value_Arglist*) = " ++ cName n ++ "_arglist;"
let closure_name = "closure_" ++ c
emit fc $ "Value *"
++ closure_name
++ " = (Value*)idris2_makeClosureFromArglist("
++ f_ptr_name
++ ", "
++ arglist
++ ");"
emit fc $ ("// end " ++ cName n ++ "(" ++ showSep ", " (map (\v => varName v) args) ++ ")")
pure $ case tailPosition of
NotInTailPosition => "idris2_trampoline(\{closure_name})"
InTailPosition => closure_name
cStatementsFromANF (AUnderApp fc n missing args) _ = do
arglist <- makeArglist missing args
let f_ptr_name = "closure_" ++ !(getNextCounter)
let f_ptr = "Value *(*"++ f_ptr_name ++ ")(Value_Arglist*) = "++ cName n ++ "_arglist;"
emit fc f_ptr
pure "(Value*)idris2_makeClosureFromArglist(\{f_ptr_name}, \{arglist})"
let nargs = length args
case tailPosition of
InTailPosition => makeClosure fc n args 0
_ => if nargs > MaxExtractFunArgs
then pure "idris2_trampoline(\{!(makeClosure fc n args 0)})"
else do
env <- get EnvTracker
let args' = avarsToC env.owned args
pure "idris2_trampoline(\{cName n}(\{concat $ intersperse ", " args'}))"

cStatementsFromANF (AUnderApp fc n missing args) _ = makeClosure fc n args missing
cStatementsFromANF (AApp fc _ closure arg) tailPosition = do
env <- get EnvTracker
pure $ (case tailPosition of
NotInTailPosition => "idris2_apply_closure"
NotInTailPosition => "idris2_apply_closure"
InTailPosition => "idris2_tailcall_apply_closure") ++ "(\{avarToC env closure}, \{avarToC env arg})"

cStatementsFromANF (ALet fc var value body) tailPosition = do
Expand Down Expand Up @@ -554,8 +539,8 @@ mutual
emit fc $ "Value_Constructor* " ++ constr ++ createNewConstructor
when (Nothing == tag) $ emit fc "\{constr}->name = idris2_constr_\{cName n};"
pure constr
fillConstructorArgs env constr args 0
pure $ "(Value*)\{constr}"
fillArgs env "\{constr}->args" args 0
pure "(Value*)\{constr}"

cStatementsFromANF (AOp fc _ op args) _ = do
let resultVar = "primVar_" ++ !(getNextCounter)
Expand Down Expand Up @@ -658,18 +643,6 @@ addCommaToList : List String -> List String
addCommaToList [] = []
addCommaToList (x :: xs) = (" " ++ x) :: map (", " ++) xs

functionDefSignature : {auto c : Ref Ctxt Defs} -> Name -> (args:List Int) -> Core String
functionDefSignature n [] = do
let fn = (cName !(getFullName n))
pure $ "\n\nValue *" ++ fn ++ "(void)"
functionDefSignature n args = do
let argsStringList = addCommaToList (map (\i => " Value * var_" ++ (show i)) args)
let fn = (cName !(getFullName n))
pure $ "\n\nValue *" ++ fn ++ "\n(\n" ++ (showSep "\n" (argsStringList)) ++ "\n)"

functionDefSignatureArglist : {auto c : Ref Ctxt Defs} -> Name -> Core String
functionDefSignatureArglist n = pure $ "Value *" ++ (cName !(getFullName n)) ++ "_arglist(Value_Arglist* arglist)"


getArgsNrList : List ty -> Nat -> List Nat
getArgsNrList [] _ = []
Expand Down Expand Up @@ -796,37 +769,31 @@ createCFunctions : {auto c : Ref Ctxt Defs}
-> ANFDef
-> Core ()
createCFunctions n (MkAFun args anf) = do
fn <- functionDefSignature n args
fn' <- functionDefSignatureArglist n
update FunctionDefinitions $ \otherDefs => (fn ++ ";\n") :: (fn' ++ ";\n") :: otherDefs
let nargs = length args
let fn = "Value *\{cName !(getFullName n)}"
++ (if nargs == 0 then "(void)"
else if nargs > MaxExtractFunArgs then "(Value *var_arglist[\{show nargs}])"
else ("\n(\n" ++ (showSep "\n" $ addCommaToList (map (\i => " Value * var_" ++ (show i)) args))) ++ "\n)")
update FunctionDefinitions $ \otherDefs => (fn ++ ";\n") :: otherDefs

let argsVars = fromList $ ALocal <$> args
let bodyFreeVars = freeVariables anf
let shouldDrop = difference argsVars bodyFreeVars
let argsNrs = getArgsNrList args Z
emit EmptyFC fn
emit EmptyFC "{"
increaseIndentation
when (nargs > MaxExtractFunArgs) $ do
_ <- foldlC (\i, j => do
emit EmptyFC "Value *var_\{show j} = var_arglist[\{show i}];"
pure $ i + 1) 0 args
pure ()
removeVars (varName <$> SortedSet.toList shouldDrop)
_ <- newRef EnvTracker (MkEnv bodyFreeVars empty)
emit EmptyFC $ "return \{!(cStatementsFromANF anf InTailPosition)};"
decreaseIndentation
emit EmptyFC "}\n"
emit EmptyFC ""
emit EmptyFC fn'
emit EmptyFC "{"
increaseIndentation
emit EmptyFC $ "return " ++ (cName !(getFullName n))
increaseIndentation
emit EmptyFC $ "("
increaseIndentation
let commaSepArglist = addCommaToList (map (\a => "arglist->args["++ show a ++"]") argsNrs)
traverse_ (emit EmptyFC) commaSepArglist
decreaseIndentation
emit EmptyFC ");"
decreaseIndentation
decreaseIndentation
emit EmptyFC "}\n"
emit EmptyFC ""
pure ()


Expand Down Expand Up @@ -856,26 +823,9 @@ createCFunctions n (MkAForeign ccs fargs ret) = do
_ => pure ()
else emit EmptyFC $ additionalFFIStub fctName fargs ret
let fnDef = "Value *" ++ (cName n) ++ "(" ++ showSep ", " (replicate (length fargs) "Value *") ++ ");"
fn_arglist <- functionDefSignatureArglist n
update FunctionDefinitions $ \otherDefs => (fnDef ++ "\n") :: (fn_arglist ++ ";\n") :: otherDefs
update FunctionDefinitions $ \otherDefs => (fnDef ++ "\n") :: otherDefs
typeVarNameArgList <- createFFIArgList fargs

emit EmptyFC fn_arglist
emit EmptyFC "{"
increaseIndentation
emit EmptyFC $ "return " ++ (cName n)
increaseIndentation
emit EmptyFC $ "("
increaseIndentation
let commaSepArglist = addCommaToList (map (\a => "arglist->args["++ show a ++"]") (getArgsNrList fargs Z))
traverse_ (emit EmptyFC) commaSepArglist
decreaseIndentation
emit EmptyFC ");"
decreaseIndentation
decreaseIndentation
emit EmptyFC "}\n"
emit EmptyFC ""

emitFDef n typeVarNameArgList
emit EmptyFC "{"
increaseIndentation
Expand Down
18 changes: 4 additions & 14 deletions support/refc/_datatypes.h
Original file line number Diff line number Diff line change
Expand Up @@ -19,7 +19,6 @@
#define STRING_TAG 12

#define CLOSURE_TAG 15
#define ARGLIST_TAG 16
#define CONSTRUCTOR_TAG 17

#define IOREF_TAG 20
Expand All @@ -31,8 +30,6 @@
#define MUTEX_TAG 30
#define CONDITION_TAG 31

#define COMPLETE_CLOSURE_TAG 98 // for trampoline tail recursion handling

typedef struct {
int refCounter;
int tag;
Expand Down Expand Up @@ -143,17 +140,10 @@ typedef struct {

typedef struct {
Value_header header;
int32_t total;
int32_t filled;
Value **args;
} Value_Arglist;

typedef Value *(*fun_ptr_t)(Value_Arglist *);

typedef struct {
Value_header header;
fun_ptr_t f;
Value_Arglist *arglist;
Value *(*f)();
uint8_t arity;
uint8_t filled; // length of args.
Value *args[];
} Value_Closure;

typedef struct {
Expand Down
43 changes: 9 additions & 34 deletions support/refc/memoryManagement.c
Original file line number Diff line number Diff line change
Expand Up @@ -15,16 +15,6 @@ Value *idris2_newValue(size_t size) {
return retVal;
}

Value_Arglist *idris2_newArglist(int missing, int total) {
Value_Arglist *retVal = IDRIS2_NEW_VALUE(Value_Arglist);
retVal->header.tag = ARGLIST_TAG;
retVal->total = total;
retVal->filled = total - missing;
retVal->args = (Value **)malloc(sizeof(Value *) * total);
memset(retVal->args, 0, sizeof(Value *) * total);
return retVal;
}

Value_Constructor *idris2_newConstructor(int total, int tag) {
Value_Constructor *retVal = (Value_Constructor *)idris2_newValue(
sizeof(Value_Constructor) + sizeof(Value *) * total);
Expand All @@ -35,16 +25,14 @@ Value_Constructor *idris2_newConstructor(int total, int tag) {
return retVal;
}

Value_Closure *idris2_makeClosureFromArglist(fun_ptr_t f,
Value_Arglist *arglist) {
Value_Closure *retVal = IDRIS2_NEW_VALUE(Value_Closure);
Value_Closure *idris2_mkClosure(Value *(*f)(), uint8_t arity, uint8_t filled) {
Value_Closure *retVal = (Value_Closure *)idris2_newValue(
sizeof(Value_Closure) + sizeof(Value *) * filled);
retVal->header.tag = CLOSURE_TAG;
retVal->arglist = arglist; // (Value_Arglist *)newReference((Value*)arglist);
retVal->f = f;
if (retVal->arglist->filled >= retVal->arglist->total) {
retVal->header.tag = COMPLETE_CLOSURE_TAG;
}
return retVal;
retVal->arity = arity;
retVal->filled = filled;
return retVal; // caller must initialize args[].
}

Value *idris2_mkDouble(double d) {
Expand Down Expand Up @@ -185,24 +173,11 @@ void idris2_removeReference(Value *elem) {

case CLOSURE_TAG: {
Value_Closure *cl = (Value_Closure *)elem;
Value_Arglist *al = cl->arglist;
idris2_removeReference((Value *)al);
break;
}
case COMPLETE_CLOSURE_TAG: {
Value_Closure *cl = (Value_Closure *)elem;
Value_Arglist *al = cl->arglist;
idris2_removeReference((Value *)al);
break;
}
case ARGLIST_TAG: {
Value_Arglist *al = (Value_Arglist *)elem;
for (int i = 0; i < al->filled; i++) {
idris2_removeReference(al->args[i]);
}
free(al->args);
for (int i = 0; i < cl->filled; ++i)
idris2_removeReference(cl->args[i]);
break;
}

case CONSTRUCTOR_TAG: {
Value_Constructor *constr = (Value_Constructor *)elem;
for (int i = 0; i < constr->total; i++) {
Expand Down
5 changes: 1 addition & 4 deletions support/refc/memoryManagement.h
Original file line number Diff line number Diff line change
Expand Up @@ -8,11 +8,8 @@ void idris2_removeReference(Value *source);

#define IDRIS2_NEW_VALUE(t) ((t *)idris2_newValue(sizeof(t)))

Value_Arglist *idris2_newArglist(int missing, int total);
Value_Constructor *idris2_newConstructor(int total, int tag);

// copies arglist, no pointer bending
Value_Closure *idris2_makeClosureFromArglist(fun_ptr_t f, Value_Arglist *);
Value_Closure *idris2_mkClosure(Value *(*f)(), uint8_t arity, uint8_t filled);

Value *idris2_mkDouble(double d);
#define idris2_mkChar(x) \
Expand Down
Loading

0 comments on commit 1dc7b74

Please sign in to comment.