Skip to content

Commit

Permalink
- [RefC] delete unnecessary conversion to keep code simply. (#3176)
Browse files Browse the repository at this point in the history
* - [RefC] delete unnecessary conversion to keep code simply.
- [RefC] rename some C functions to confliction safe.

* make the linter feels good.

* Get me a job as a linter slave.

* [RefC] Added hard fail for unsupported primitives.

* Fix typo.
  • Loading branch information
seagull-kamome committed Jan 14, 2024
1 parent 90bf2de commit 844ea28
Show file tree
Hide file tree
Showing 6 changed files with 102 additions and 94 deletions.
59 changes: 9 additions & 50 deletions src/Compiler/RefC/RefC.idr
Original file line number Diff line number Diff line change
Expand Up @@ -223,53 +223,6 @@ cOp Crash [_, msg] = "idris2_crash(" ++ msg ++ ");"
cOp fn args = plainOp (show fn) (toList args)


data ExtPrim = NewIORef | ReadIORef | WriteIORef
| NewArray | ArrayGet | ArraySet
| GetField | SetField
| VoidElim
| SysOS | SysCodegen
| OnCollect
| OnCollectAny
| Unknown Name

export
Show ExtPrim where
show NewIORef = "newIORef"
show ReadIORef = "readIORef"
show WriteIORef = "writeIORef"
show NewArray = "newArray"
show ArrayGet = "arrayGet"
show ArraySet = "arraySet"
show GetField = "getField"
show SetField = "setField"
show VoidElim = "voidElim"
show SysOS = "sysOS"
show SysCodegen = "sysCodegen"
show OnCollect = "onCollect"
show OnCollectAny = "onCollectAny"
show (Unknown n) = "Unknown " ++ show n

||| Match on a user given name to get the scheme primitive
toPrim : Name -> ExtPrim
toPrim pn@(NS _ n)
= cond [(n == UN (Basic "prim__newIORef"), NewIORef),
(n == UN (Basic "prim__readIORef"), ReadIORef),
(n == UN (Basic "prim__writeIORef"), WriteIORef),
(n == UN (Basic "prim__newArray"), NewArray),
(n == UN (Basic "prim__arrayGet"), ArrayGet),
(n == UN (Basic "prim__arraySet"), ArraySet),
(n == UN (Basic "prim__getField"), GetField),
(n == UN (Basic "prim__setField"), SetField),
(n == UN (Basic "prim__void"), VoidElim),
(n == UN (Basic "prim__os"), SysOS),
(n == UN (Basic "prim__codegen"), SysCodegen),
(n == UN (Basic "prim__onCollect"), OnCollect),
(n == UN (Basic "prim__onCollectAny"), OnCollectAny)
]
(Unknown pn)
toPrim pn = assert_total $ idris_crash ("INTERNAL ERROR: Unknown primitive: " ++ cName pn)
-- not really total but this way this internal error does not contaminate everything else


varName : AVar -> String
varName (ALocal i) = "var_" ++ (show i)
Expand Down Expand Up @@ -468,8 +421,6 @@ const2Integer c i =





-- we return for each of the ANF a set of statements and two possible return statements
-- The first one for non-tail statements, the second one for tail statements
-- this way, we can deal with tail calls and tail recursion.
Expand Down Expand Up @@ -688,8 +639,16 @@ mutual
let opStatement = cOp op argsVec
pure $ MkRS opStatement opStatement
cStatementsFromANF (AExtPrim fc _ p args) _ = do
let prims : List String =
["prim__newIORef", "prim__readIORef", "prim__writeIORef", "prim__newArray",
"prim__arrayGet", "prim__arraySet", "prim__getField", "prim__setField",
"prim__void", "prim__os", "prim__codegen", "prim__onCollect", "prim__onCollectAny" ]
case p of
NS _ (UN (Basic pn)) =>
unless (elem pn prims) $ throw $ InternalError $ "INTERNAL ERROR: Unknown primitive: " ++ cName p
_ => throw $ InternalError $ "INTERNAL ERROR: Unknown primitive: " ++ cName p
emit fc $ "// call to external primitive " ++ cName p
let returnLine = (cCleanString (show (toPrim p)) ++ "("++ showSep ", " (map varName args) ++")")
let returnLine = "idris2_" ++ (cName p) ++ "("++ showSep ", " (map varName args) ++")"
pure $ MkRS returnLine returnLine
cStatementsFromANF (AConCase fc sc alts mDef) tailPosition = do
c <- getNextCounter
Expand Down
76 changes: 43 additions & 33 deletions support/refc/prim.c
Original file line number Diff line number Diff line change
Expand Up @@ -25,7 +25,8 @@ void doubleIORef_Storage(IORef_Storage *ior) {
ior->refs = values;
}

Value *newIORef(Value *erased, Value *input_value, Value *_world) {
Value *idris2_Data_IORef_prim__newIORef(Value *erased, Value *input_value,
Value *_world) {
// if no ioRef Storag exist, start with one
if (!global_IORef_Storage) {
global_IORef_Storage = newIORef_Storage(128);
Expand All @@ -46,13 +47,8 @@ Value *newIORef(Value *erased, Value *input_value, Value *_world) {
return (Value *)ioRef;
}

Value *readIORef(Value *erased, Value *_index, Value *_world) {
Value_IORef *index = (Value_IORef *)_index;
return newReference(global_IORef_Storage->refs[index->index]);
}

Value *writeIORef(Value *erased, Value *_index, Value *new_value,
Value *_world) {
Value *idris2_Data_IORef_prim__writeIORef(Value *erased, Value *_index,
Value *new_value, Value *_world) {
Value_IORef *index = (Value_IORef *)_index;
removeReference(global_IORef_Storage->refs[index->index]);
global_IORef_Storage->refs[index->index] = newReference(new_value);
Expand All @@ -63,31 +59,47 @@ Value *writeIORef(Value *erased, Value *_index, Value *new_value,
// System operations
// -----------------------------------

Value *sysOS(void) {
static Value *osstring = NULL;

Value *idris2_System_Info_prim__os(void) {
if (osstring == NULL) {
osstring = (Value *)makeString(
#ifdef _WIN32
return (Value *)makeString("windows");
"windows"
#elif _WIN64
return (Value *)makeString("windows");
"windows"
#elif __APPLE__ || __MACH__
return (Value *)makeString("macOS");
"macOS"
#elif __linux__
return (Value *)makeString("Linux");
"Linux"
#elif __FreeBSD__
return (Value *)makeString("FreeBSD");
"FreeBSD"
#elif __OpenBSD__
return (Value *)makeString("OpenBSD");
"OpenBSD"
#elif __NetBSD__
return (Value *)makeString("NetBSD");
"NetBSD"
#elif __DragonFly__
return (Value *)makeString("DragonFly");
"DragonFly"
#elif __unix || __unix__
return (Value *)makeString("Unix");
"Unix"
#else
return (Value *)makeString("Other");
"Other"
#endif
);
}
return newReference(osstring);
}

Value *sysCodegen(void) { return (Value *)makeString("refc"); }
// NOTE: The codegen is obviously determined at compile time,
// so the backend should optimize it by replacing it with a constant.
// It would probably also be useful for conditional compilation.
static Value *codegenstring = NULL;

Value *idris2_System_Info_prim__codegen(void) {
if (codegenstring == NULL)
codegenstring = (Value *)makeString("refc");
return newReference(codegenstring);
}

Value *idris2_crash(Value *msg) {
Value_String *str = (Value_String *)msg;
Expand All @@ -102,7 +114,8 @@ Value *idris2_crash(Value *msg) {
// // Array operations
// // -----------------------------------

Value *newArray(Value *erased, Value *_length, Value *v, Value *_word) {
Value *idris2_Data_IOArray_Prims_prim__newArray(Value *erased, Value *_length,
Value *v, Value *_word) {
int length = extractInt(_length);
Value_Array *a = makeArray(length);

Expand All @@ -113,13 +126,9 @@ Value *newArray(Value *erased, Value *_length, Value *v, Value *_word) {
return (Value *)a;
}

Value *arrayGet(Value *erased, Value *_array, Value *_index, Value *_word) {
Value_Array *a = (Value_Array *)_array;
return newReference(a->arr[((Value_Int64 *)_index)->i64]);
}

Value *arraySet(Value *erased, Value *_array, Value *_index, Value *v,
Value *_word) {
Value *idris2_Data_IOArray_Prims_prim__arraySet(Value *erased, Value *_array,
Value *_index, Value *v,
Value *_word) {
Value_Array *a = (Value_Array *)_array;
removeReference(a->arr[((Value_Int64 *)_index)->i64]);
a->arr[((Value_Int64 *)_index)->i64] = newReference(v);
Expand All @@ -131,25 +140,26 @@ Value *arraySet(Value *erased, Value *_array, Value *_index, Value *v,
// Pointer operations
// -----------------------------------

Value *onCollect(Value *_erased, Value *_anyPtr, Value *_freeingFunction,
Value *_world) {
Value *idris2_Prelude_IO_prim__onCollect(Value *_erased, Value *_anyPtr,
Value *_freeingFunction,
Value *_world) {
Value_GCPointer *retVal = IDRIS2_NEW_VALUE(Value_GCPointer);
retVal->header.tag = GC_POINTER_TAG;
retVal->p = (Value_Pointer *)newReference(_anyPtr);
retVal->onCollectFct = (Value_Closure *)newReference(_freeingFunction);
return (Value *)retVal;
}

Value *onCollectAny(Value *_anyPtr, Value *_freeingFunction, Value *_world) {
Value *idris2_Prelude_IO_prim__onCollectAny(Value *_anyPtr,
Value *_freeingFunction,
Value *_world) {
Value_GCPointer *retVal = IDRIS2_NEW_VALUE(Value_GCPointer);
retVal->header.tag = GC_POINTER_TAG;
retVal->p = (Value_Pointer *)newReference(_anyPtr);
retVal->onCollectFct = (Value_Closure *)newReference(_freeingFunction);
return (Value *)retVal;
}

Value *voidElim(Value *erased1, Value *erased2) { return NULL; }

// Threads

// %foreign "scheme:blodwen-mutex"
Expand Down
27 changes: 16 additions & 11 deletions support/refc/prim.h
Original file line number Diff line number Diff line change
Expand Up @@ -4,27 +4,32 @@

// IORef

Value *newIORef(Value *, Value *, Value *);
Value *readIORef(Value *, Value *, Value *);
Value *writeIORef(Value *, Value *, Value *, Value *);
Value *idris2_Data_IORef_prim__newIORef(Value *, Value *, Value *);
#define idris2_Data_IORef_prim__readIORef(erased, ix, world) \
(newReference(global_IORef_Storage->refs[((Value_IORef *)ix)->index]))

Value *idris2_Data_IORef_prim__writeIORef(Value *, Value *, Value *, Value *);

// Sys

Value *sysOS(void);
Value *sysCodegen(void);
Value *idris2_System_Info_prim__os(void);
Value *idris2_System_Info_prim__codegen(void);
Value *idris2_crash(Value *msg);

// Array

Value *newArray(Value *, Value *, Value *, Value *);
Value *arrayGet(Value *, Value *, Value *, Value *);
Value *arraySet(Value *, Value *, Value *, Value *, Value *);
Value *idris2_Data_IOArray_Prims_prim__newArray(Value *, Value *, Value *,
Value *);
#define idris2_Data_IOArray_Prims_prim__arrayGet(rased, array, i, word) \
(newReference(((Value_Array *)(array))->arr[((Value_Int64 *)i)->i64]))
Value *idris2_Data_IOArray_Prims_prim__arraySet(Value *, Value *, Value *,
Value *, Value *);

// Pointer
Value *onCollect(Value *, Value *, Value *, Value *);
Value *onCollectAny(Value *, Value *, Value *);
Value *idris2_Prelude_IO_prim__onCollect(Value *, Value *, Value *, Value *);
Value *idris2_Prelude_IO_prim__onCollectAny(Value *, Value *, Value *);

Value *voidElim(Value *, Value *);
#define idris2_Prelude_Uninhabited_prim__void(x, y) (NULL)

// Threads
Value *System_Concurrency_Raw_prim__mutexRelease(Value *, Value *);
Expand Down
24 changes: 24 additions & 0 deletions tests/refc/prims/Test.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,24 @@
module Test

import Data.IOArray
import Data.IORef
import System.Info

main : IO ()
main = do
do
arr <- newArray {elem=Int} 10
printLn !(readArray arr 0)
printLn !(writeArray arr 1 10)
printLn !(readArray arr 1)
do
ref <- newIORef "abcd"
printLn !(readIORef ref)
writeIORef ref "ABCD"
printLn !(readIORef ref)
do
-- printLn os
printLn codegen



6 changes: 6 additions & 0 deletions tests/refc/prims/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,6 @@
Nothing
True
Just 10
"abcd"
"ABCD"
"refc"
4 changes: 4 additions & 0 deletions tests/refc/prims/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,4 @@
. ../../testutils.sh

idris2 --cg refc -o test Test.idr > /dev/null
./build/exec/test

0 comments on commit 844ea28

Please sign in to comment.