Skip to content

Commit

Permalink
[ re #3177 ] Separate the renaming from the rest (#3248)
Browse files Browse the repository at this point in the history
  • Loading branch information
gallais committed Apr 3, 2024
1 parent 75032a7 commit fee293b
Show file tree
Hide file tree
Showing 11 changed files with 235 additions and 228 deletions.
40 changes: 20 additions & 20 deletions src/Compiler/RefC/RefC.idr
Original file line number Diff line number Diff line change
Expand Up @@ -198,7 +198,7 @@ cOp StrIndex [x, i] = "strIndex(" ++ x ++ ", " ++ i ++ ")"
cOp StrCons [x, y] = "strCons(" ++ x ++ ", " ++ y ++ ")"
cOp StrAppend [x, y] = "strAppend(" ++ x ++ ", " ++ y ++ ")"
cOp StrSubstr [x, y, z] = "strSubstr(" ++ x ++ ", " ++ y ++ ", " ++ z ++ ")"
cOp BelieveMe [_, _, x] = "newReference(" ++ x ++ ")"
cOp BelieveMe [_, _, x] = "idris2_newReference(" ++ x ++ ")"
cOp Crash [_, msg] = "idris2_crash(" ++ msg ++ ");"
cOp fn args = show fn ++ "(" ++ (showSep ", " $ toList args) ++ ")"

Expand All @@ -216,7 +216,7 @@ ReuseMap = SortedMap Name String
Owned = SortedSet AVar

||| Environment for precise reference counting.
||| If variable borrowed (that is, it is not in the owned set) when used, call a function newReference.
||| If variable borrowed (that is, it is not in the owned set) when used, call a function idris2_newReference.
||| If variable owned, then use it directly.
||| Reuse Map contains the name of the reusable constructor and variable
record Env where
Expand Down Expand Up @@ -296,25 +296,25 @@ removeVars : {auto oft : Ref OutfileText Output}
-> {auto il : Ref IndentLevel Nat}
-> List String
-> Core ()
removeVars = applyFunctionToVars "removeReference"
removeVars = applyFunctionToVars "idris2_removeReference"

dupVars : {auto oft : Ref OutfileText Output}
-> {auto il : Ref IndentLevel Nat}
-> List String
-> Core ()
dupVars = applyFunctionToVars "newReference"
dupVars = applyFunctionToVars "idris2_newReference"

removeReuseConstructors : {auto oft : Ref OutfileText Output}
-> {auto il : Ref IndentLevel Nat}
-> List String
-> Core ()
removeReuseConstructors = applyFunctionToVars "removeReuseConstructor"
removeReuseConstructors = applyFunctionToVars "idris2_removeReuseConstructor"

avarToC : Env -> AVar -> String
avarToC env var =
if contains var env.owned then varName var
-- case when the variable is borrowed
else "newReference(" ++ varName var ++ ")"
else "idris2_newReference(" ++ varName var ++ ")"

moveFromOwnedToBorrowed : Env -> SortedSet AVar -> Env
moveFromOwnedToBorrowed env vars = { owned $= (`difference` vars) } env
Expand All @@ -330,7 +330,7 @@ makeArglist missing xs = do
let arglist = "arglist_" ++ !(getNextCounter)
emit EmptyFC $ "Value_Arglist *"
++ arglist
++ " = newArglist(" ++ show missing
++ " = idris2_newArglist(" ++ show missing
++ "," ++ show (length xs + missing)
++ ");"
pushArgToArglist !(get EnvTracker) arglist xs 0
Expand Down Expand Up @@ -436,7 +436,7 @@ addReuseConstructor reuseMap sc conName conArgs consts shouldDrop actualReuseCon
let constr = "constructor_" ++ !(getNextCounter)
emit EmptyFC $ "Value_Constructor* " ++ constr ++ " = NULL;"
-- If the constructor variable is unique (has 1 reference), then assign it for reuse
emit EmptyFC $ "if (isUnique(" ++ sc ++ ")) {"
emit EmptyFC $ "if (idris2_isUnique(" ++ sc ++ ")) {"
increaseIndentation
emit EmptyFC $ constr ++ " = (Value_Constructor*)" ++ sc ++ ";"
decreaseIndentation
Expand Down Expand Up @@ -495,27 +495,27 @@ mutual
let closure_name = "closure_" ++ c
emit fc $ "Value *"
++ closure_name
++ " = (Value*)makeClosureFromArglist("
++ " = (Value*)idris2_makeClosureFromArglist("
++ f_ptr_name
++ ", "
++ arglist
++ ");"
emit fc $ ("// end " ++ cName n ++ "(" ++ showSep ", " (map (\v => varName v) args) ++ ")")
pure $ case tailPosition of
NotInTailPosition => "trampoline(\{closure_name})"
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*)makeClosureFromArglist(\{f_ptr_name}, \{arglist})"
pure "(Value*)idris2_makeClosureFromArglist(\{f_ptr_name}, \{arglist})"

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

cStatementsFromANF (ALet fc var value body) tailPosition = do
env <- get EnvTracker
Expand All @@ -527,7 +527,7 @@ mutual
let valueEnv = { reuseMap $= (`intersectionMap` usedCons) } (moveFromOwnedToBorrowed env borrowVal)
put EnvTracker valueEnv
emit fc $ "Value * var_\{show var} = \{!(cStatementsFromANF value NotInTailPosition)};"
unless (contains (ALocal var) usedVars) $ emit fc $ "removeReference(var_\{show var});"
unless (contains (ALocal var) usedVars) $ emit fc $ "idris2_removeReference(var_\{show var});"
put EnvTracker ({ owned := owned', reuseMap $= (`differenceMap` usedCons) } env)
cStatementsFromANF body tailPosition

Expand All @@ -536,7 +536,7 @@ mutual
then pure "(NULL /* \{show n} */)"
else do
env <- get EnvTracker
let createNewConstructor = " = newConstructor("
let createNewConstructor = " = idris2_newConstructor("
++ (show (length args))
++ ", " ++ maybe "-1" show tag ++ ");"

Expand Down Expand Up @@ -764,9 +764,9 @@ packCFType CFUnsigned8 varName = "idris2_mkBits8(" ++ varName ++ ")"
packCFType CFString varName = "idris2_mkString(" ++ varName ++ ")"
packCFType CFDouble varName = "idris2_mkDouble(" ++ varName ++ ")"
packCFType CFChar varName = "idris2_mkChar(" ++ varName ++ ")"
packCFType CFPtr varName = "makePointer(" ++ varName ++ ")"
packCFType CFGCPtr varName = "makePointer(" ++ varName ++ ")"
packCFType CFBuffer varName = "makeBuffer(" ++ varName ++ ")"
packCFType CFPtr varName = "idris2_makePointer(" ++ varName ++ ")"
packCFType CFGCPtr varName = "idris2_makePointer(" ++ varName ++ ")"
packCFType CFBuffer varName = "idris2_makeBuffer(" ++ varName ++ ")"
packCFType CFWorld _ = "(Value *)NULL"
packCFType (CFFun x y) varName = "makeFunction(" ++ varName ++ ")"
packCFType (CFIORes x) varName = packCFType x varName
Expand All @@ -783,7 +783,7 @@ additionalFFIStub name argTypes (CFIORes retType) = additionalFFIStub name (disc
additionalFFIStub name argTypes retType =
cTypeOfCFType retType ++
" (*" ++ cName name ++ ")(" ++
(concat $ intersperse ", " $ map cTypeOfCFType argTypes) ++ ") = (void*)missing_ffi;\n"
(concat $ intersperse ", " $ map cTypeOfCFType argTypes) ++ ") = (void*)idris2_missing_ffi;\n"

createCFunctions : {auto c : Ref Ctxt Defs}
-> {auto a : Ref ArgCounter Nat}
Expand Down Expand Up @@ -944,7 +944,7 @@ footer = do
""
}
Value *mainExprVal = __mainExpression_0();
trampoline(mainExprVal);
idris2_trampoline(mainExprVal);
return 0; // bye bye
}
"""
Expand Down
2 changes: 1 addition & 1 deletion support/refc/casts.h
Original file line number Diff line number Diff line change
Expand Up @@ -74,7 +74,7 @@ Value *idris2_cast_Int32_to_string(Value *);
(idris2_mkInt16((int16_t)idris2_vp_to_Int64(x)))
#define idris2_cast_Int64_to_Int32(x) \
(idris2_mkInt32((int32_t)idris2_vp_to_Int64(x)))
#define idris2_cast_Int64_to_Int64(x) (newReference(x))
#define idris2_cast_Int64_to_Int64(x) (idris2_newReference(x))
Value *idris2_cast_Int64_to_Integer(Value *);
#define idris2_cast_Int64_to_Double(x) \
(idris2_mkDouble((double)idris2_vp_to_Int64(x)))
Expand Down
44 changes: 23 additions & 21 deletions support/refc/memoryManagement.c
Original file line number Diff line number Diff line change
@@ -1,7 +1,7 @@
#include "refc_util.h"
#include "runtime.h"

Value *newValue(size_t size) {
Value *idris2_newValue(size_t size) {
#if defined(__STDC_VERSION__) && (__STDC_VERSION__ >= 201112L) /* C11 */
Value *retVal = (Value *)aligned_alloc(
sizeof(void *),
Expand All @@ -15,7 +15,7 @@ Value *newValue(size_t size) {
return retVal;
}

Value_Arglist *newArglist(int missing, int total) {
Value_Arglist *idris2_newArglist(int missing, int total) {
Value_Arglist *retVal = IDRIS2_NEW_VALUE(Value_Arglist);
retVal->header.tag = ARGLIST_TAG;
retVal->total = total;
Expand All @@ -25,8 +25,8 @@ Value_Arglist *newArglist(int missing, int total) {
return retVal;
}

Value_Constructor *newConstructor(int total, int tag) {
Value_Constructor *retVal = (Value_Constructor *)newValue(
Value_Constructor *idris2_newConstructor(int total, int tag) {
Value_Constructor *retVal = (Value_Constructor *)idris2_newValue(
sizeof(Value_Constructor) + sizeof(Value *) * total);
retVal->header.tag = CONSTRUCTOR_TAG;
retVal->total = total;
Expand All @@ -35,7 +35,8 @@ Value_Constructor *newConstructor(int total, int tag) {
return retVal;
}

Value_Closure *makeClosureFromArglist(fun_ptr_t f, Value_Arglist *arglist) {
Value_Closure *idris2_makeClosureFromArglist(fun_ptr_t f,
Value_Arglist *arglist) {
Value_Closure *retVal = IDRIS2_NEW_VALUE(Value_Closure);
retVal->header.tag = CLOSURE_TAG;
retVal->arglist = arglist; // (Value_Arglist *)newReference((Value*)arglist);
Expand Down Expand Up @@ -112,29 +113,30 @@ Value_String *idris2_mkString(char *s) {
return retVal;
}

Value_Pointer *makePointer(void *ptr_Raw) {
Value_Pointer *idris2_makePointer(void *ptr_Raw) {
Value_Pointer *p = IDRIS2_NEW_VALUE(Value_Pointer);
p->header.tag = POINTER_TAG;
p->p = ptr_Raw;
return p;
}

Value_GCPointer *makeGCPointer(void *ptr_Raw, Value_Closure *onCollectFct) {
Value_GCPointer *idris2_makeGCPointer(void *ptr_Raw,
Value_Closure *onCollectFct) {
Value_GCPointer *p = IDRIS2_NEW_VALUE(Value_GCPointer);
p->header.tag = GC_POINTER_TAG;
p->p = makePointer(ptr_Raw);
p->p = idris2_makePointer(ptr_Raw);
p->onCollectFct = onCollectFct;
return p;
}

Value_Buffer *makeBuffer(void *buf) {
Value_Buffer *idris2_makeBuffer(void *buf) {
Value_Buffer *b = IDRIS2_NEW_VALUE(Value_Buffer);
b->header.tag = BUFFER_TAG;
b->buffer = buf;
return b;
}

Value_Array *makeArray(int length) {
Value_Array *idris2_makeArray(int length) {
Value_Array *a = IDRIS2_NEW_VALUE(Value_Array);
a->header.tag = ARRAY_TAG;
a->capacity = length;
Expand All @@ -143,15 +145,15 @@ Value_Array *makeArray(int length) {
return a;
}

Value *newReference(Value *source) {
Value *idris2_newReference(Value *source) {
// note that we explicitly allow NULL as source (for erased arguments)
if (source && !idris2_vp_is_unboxed(source)) {
source->header.refCounter++;
}
return source;
}

void removeReference(Value *elem) {
void idris2_removeReference(Value *elem) {
if (!elem || idris2_vp_is_unboxed(elem)) {
return;
}
Expand Down Expand Up @@ -184,32 +186,32 @@ void removeReference(Value *elem) {
case CLOSURE_TAG: {
Value_Closure *cl = (Value_Closure *)elem;
Value_Arglist *al = cl->arglist;
removeReference((Value *)al);
idris2_removeReference((Value *)al);
break;
}
case COMPLETE_CLOSURE_TAG: {
Value_Closure *cl = (Value_Closure *)elem;
Value_Arglist *al = cl->arglist;
removeReference((Value *)al);
idris2_removeReference((Value *)al);
break;
}
case ARGLIST_TAG: {
Value_Arglist *al = (Value_Arglist *)elem;
for (int i = 0; i < al->filled; i++) {
removeReference(al->args[i]);
idris2_removeReference(al->args[i]);
}
free(al->args);
break;
}
case CONSTRUCTOR_TAG: {
Value_Constructor *constr = (Value_Constructor *)elem;
for (int i = 0; i < constr->total; i++) {
removeReference(constr->args[i]);
idris2_removeReference(constr->args[i]);
}
break;
}
case IOREF_TAG:
removeReference(((Value_IORef *)elem)->v);
idris2_removeReference(((Value_IORef *)elem)->v);
break;

case BUFFER_TAG: {
Expand All @@ -221,7 +223,7 @@ void removeReference(Value *elem) {
case ARRAY_TAG: {
Value_Array *a = (Value_Array *)elem;
for (int i = 0; i < a->capacity; i++) {
removeReference(a->arr[i]);
idris2_removeReference(a->arr[i]);
}
free(a->arr);
break;
Expand All @@ -234,9 +236,9 @@ void removeReference(Value *elem) {
/* maybe here we need to invoke onCollectAny */
Value_GCPointer *vPtr = (Value_GCPointer *)elem;
Value *closure1 =
apply_closure((Value *)vPtr->onCollectFct, (Value *)vPtr->p);
apply_closure(closure1, NULL);
removeReference((Value *)vPtr->p);
idris2_apply_closure((Value *)vPtr->onCollectFct, (Value *)vPtr->p);
idris2_apply_closure(closure1, NULL);
idris2_removeReference((Value *)vPtr->p);
break;
}

Expand Down
23 changes: 12 additions & 11 deletions support/refc/memoryManagement.h
Original file line number Diff line number Diff line change
Expand Up @@ -2,17 +2,17 @@

#include "cBackend.h"

Value *newValue(size_t size);
Value *newReference(Value *source);
void removeReference(Value *source);
Value *idris2_newValue(size_t size);
Value *idris2_newReference(Value *source);
void idris2_removeReference(Value *source);

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

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

// copies arglist, no pointer bending
Value_Closure *makeClosureFromArglist(fun_ptr_t f, Value_Arglist *);
Value_Closure *idris2_makeClosureFromArglist(fun_ptr_t f, Value_Arglist *);

Value *idris2_mkDouble(double d);
#define idris2_mkChar(x) \
Expand Down Expand Up @@ -62,7 +62,8 @@ Value_Integer *idris2_mkIntegerLiteral(char *i);
Value_String *idris2_mkEmptyString(size_t l);
Value_String *idris2_mkString(char *);

Value_Pointer *makePointer(void *);
Value_GCPointer *makeGCPointer(void *ptr_Raw, Value_Closure *onCollectFct);
Value_Buffer *makeBuffer(void *buf);
Value_Array *makeArray(int length);
Value_Pointer *idris2_makePointer(void *);
Value_GCPointer *idris2_makeGCPointer(void *ptr_Raw,
Value_Closure *onCollectFct);
Value_Buffer *idris2_makeBuffer(void *buf);
Value_Array *idris2_makeArray(int length);
Loading

0 comments on commit fee293b

Please sign in to comment.