mirror of
https://github.com/idris-lang/Idris2.git
synced 2025-01-01 16:12:26 +03:00
Pass Buffer as char* when using C functions in RefC
To be able to use `C` functions for both Scheme and RefC: it was not possible for `Buffer` before this PR. As an example, `writeBufferData` and `readBufferData` functions are removed: generic C backend implementations are used instead.
This commit is contained in:
parent
9cca3a7d35
commit
f81c37ea3a
@ -246,12 +246,10 @@ copyData src start len dest loc
|
||||
= primIO (prim__copyData src start len dest loc)
|
||||
|
||||
%foreign "C:idris2_readBufferData, libidris2_support, idris_file.h"
|
||||
"RefC:readBufferData"
|
||||
"node:lambda:(f,b,l,m) => require('fs').readSync(f.fd,b,l,m)"
|
||||
prim__readBufferData : FilePtr -> Buffer -> Int -> Int -> PrimIO Int
|
||||
|
||||
%foreign "C:idris2_writeBufferData, libidris2_support, idris_file.h"
|
||||
"RefC:writeBufferData"
|
||||
"node:lambda:(f,b,l,m) => require('fs').writeSync(f.fd,b,l,m)"
|
||||
prim__writeBufferData : FilePtr -> Buffer -> Int -> Int -> PrimIO Int
|
||||
|
||||
|
@ -785,30 +785,34 @@ emitFDef funcName ((varType, varName, varCFType) :: xs) = do
|
||||
decreaseIndentation
|
||||
emit EmptyFC ")"
|
||||
|
||||
extractValue : (cfType:CFType) -> (varName:String) -> String
|
||||
extractValue CFUnit varName = "void"
|
||||
extractValue CFInt varName = "((Value_Int64*)" ++ varName ++ ")->i64"
|
||||
extractValue CFInt8 varName = "((Value_Int8*)" ++ varName ++ ")->i8"
|
||||
extractValue CFInt16 varName = "((Value_Int16*)" ++ varName ++ ")->i16"
|
||||
extractValue CFInt32 varName = "((Value_Int32*)" ++ varName ++ ")->i32"
|
||||
extractValue CFInt64 varName = "((Value_Int64*)" ++ varName ++ ")->i64"
|
||||
extractValue CFUnsigned8 varName = "((Value_Bits8*)" ++ varName ++ ")->ui8"
|
||||
extractValue CFUnsigned16 varName = "((Value_Bits16*)" ++ varName ++ ")->ui16"
|
||||
extractValue CFUnsigned32 varName = "((Value_Bits32*)" ++ varName ++ ")->ui32"
|
||||
extractValue CFUnsigned64 varName = "((Value_Bits64*)" ++ varName ++ ")->ui64"
|
||||
extractValue CFString varName = "((Value_String*)" ++ varName ++ ")->str"
|
||||
extractValue CFDouble varName = "((Value_Double*)" ++ varName ++ ")->d"
|
||||
extractValue CFChar varName = "((Value_Char*)" ++ varName ++ ")->c"
|
||||
extractValue CFPtr varName = "((Value_Pointer*)" ++ varName ++ ")->p"
|
||||
extractValue CFGCPtr varName = "((Value_GCPointer*)" ++ varName ++ ")->p->p"
|
||||
extractValue CFBuffer varName = "((Value_Buffer*)" ++ varName ++ ")->buffer"
|
||||
extractValue CFWorld varName = "(Value_World*)" ++ varName
|
||||
extractValue (CFFun x y) varName = "(Value_Closure*)" ++ varName
|
||||
extractValue (CFIORes x) varName = extractValue x varName
|
||||
extractValue (CFStruct x xs) varName = assert_total $ idris_crash ("INTERNAL ERROR: Struct access not implemented: " ++ varName)
|
||||
-- Generic C parameter or RefC specific parameter
|
||||
data CLang = CLangC | CLangRefC
|
||||
|
||||
extractValue : (cLang : CLang) -> (cfType:CFType) -> (varName:String) -> String
|
||||
extractValue _ CFUnit varName = "void"
|
||||
extractValue _ CFInt varName = "((Value_Int64*)" ++ varName ++ ")->i64"
|
||||
extractValue _ CFInt8 varName = "((Value_Int8*)" ++ varName ++ ")->i8"
|
||||
extractValue _ CFInt16 varName = "((Value_Int16*)" ++ varName ++ ")->i16"
|
||||
extractValue _ CFInt32 varName = "((Value_Int32*)" ++ varName ++ ")->i32"
|
||||
extractValue _ CFInt64 varName = "((Value_Int64*)" ++ varName ++ ")->i64"
|
||||
extractValue _ CFUnsigned8 varName = "((Value_Bits8*)" ++ varName ++ ")->ui8"
|
||||
extractValue _ CFUnsigned16 varName = "((Value_Bits16*)" ++ varName ++ ")->ui16"
|
||||
extractValue _ CFUnsigned32 varName = "((Value_Bits32*)" ++ varName ++ ")->ui32"
|
||||
extractValue _ CFUnsigned64 varName = "((Value_Bits64*)" ++ varName ++ ")->ui64"
|
||||
extractValue _ CFString varName = "((Value_String*)" ++ varName ++ ")->str"
|
||||
extractValue _ CFDouble varName = "((Value_Double*)" ++ varName ++ ")->d"
|
||||
extractValue _ CFChar varName = "((Value_Char*)" ++ varName ++ ")->c"
|
||||
extractValue _ CFPtr varName = "((Value_Pointer*)" ++ varName ++ ")->p"
|
||||
extractValue _ CFGCPtr varName = "((Value_GCPointer*)" ++ varName ++ ")->p->p"
|
||||
extractValue CLangC CFBuffer varName = "((Value_Buffer*)" ++ varName ++ ")->buffer->data"
|
||||
extractValue CLangRefC CFBuffer varName = "((Value_Buffer*)" ++ varName ++ ")->buffer"
|
||||
extractValue _ CFWorld varName = "(Value_World*)" ++ varName
|
||||
extractValue _ (CFFun x y) varName = "(Value_Closure*)" ++ varName
|
||||
extractValue c (CFIORes x) varName = extractValue c x varName
|
||||
extractValue _ (CFStruct x xs) varName = assert_total $ idris_crash ("INTERNAL ERROR: Struct access not implemented: " ++ varName)
|
||||
-- not really total but this way this internal error does not contaminate everything else
|
||||
extractValue (CFUser x xs) varName = "(Value*)" ++ varName
|
||||
extractValue n _ = assert_total $ idris_crash ("INTERNAL ERROR: Unknonw FFI type in C backend: " ++ show n)
|
||||
extractValue _ (CFUser x xs) varName = "(Value*)" ++ varName
|
||||
extractValue _ n _ = assert_total $ idris_crash ("INTERNAL ERROR: Unknonw FFI type in C backend: " ++ show n)
|
||||
|
||||
packCFType : (cfType:CFType) -> (varName:String) -> String
|
||||
packCFType CFUnit varName = "NULL"
|
||||
@ -898,6 +902,9 @@ createCFunctions n (MkACon tag arity nt) = do
|
||||
createCFunctions n (MkAForeign ccs fargs ret) = do
|
||||
case parseCC (additionalFFILangs ++ ["RefC", "C"]) ccs of
|
||||
Just (lang, fctForeignName :: extLibOpts) => do
|
||||
let cLang = if lang == "RefC"
|
||||
then CLangRefC
|
||||
else CLangC
|
||||
let isStandardFFI = Prelude.elem lang ["RefC", "C"]
|
||||
let fctName = if isStandardFFI
|
||||
then UN fctForeignName
|
||||
@ -937,19 +944,19 @@ createCFunctions n (MkAForeign ccs fargs ret) = do
|
||||
CFIORes CFUnit => do
|
||||
emit EmptyFC $ cName fctName
|
||||
++ "("
|
||||
++ showSep ", " (map (\(_, vn, vt) => extractValue vt vn) (discardLastArgument typeVarNameArgList))
|
||||
++ showSep ", " (map (\(_, vn, vt) => extractValue cLang vt vn) (discardLastArgument typeVarNameArgList))
|
||||
++ ");"
|
||||
emit EmptyFC "return NULL;"
|
||||
CFIORes ret => do
|
||||
emit EmptyFC $ cTypeOfCFType ret ++ " retVal = " ++ cName fctName
|
||||
++ "("
|
||||
++ showSep ", " (map (\(_, vn, vt) => extractValue vt vn) (discardLastArgument typeVarNameArgList))
|
||||
++ showSep ", " (map (\(_, vn, vt) => extractValue cLang vt vn) (discardLastArgument typeVarNameArgList))
|
||||
++ ");"
|
||||
emit EmptyFC $ "return (Value*)" ++ packCFType ret "retVal" ++ ";"
|
||||
_ => do
|
||||
emit EmptyFC $ cTypeOfCFType ret ++ " retVal = " ++ cName fctName
|
||||
++ "("
|
||||
++ showSep ", " (map (\(_, vn, vt) => extractValue vt vn) typeVarNameArgList)
|
||||
++ showSep ", " (map (\(_, vn, vt) => extractValue cLang vt vn) typeVarNameArgList)
|
||||
++ ");"
|
||||
emit EmptyFC $ "return (Value*)" ++ packCFType ret "retVal" ++ ";"
|
||||
|
||||
|
@ -2,11 +2,6 @@
|
||||
#include <sys/stat.h>
|
||||
#include <string.h>
|
||||
|
||||
typedef struct {
|
||||
int size;
|
||||
uint8_t data[];
|
||||
} Buffer;
|
||||
|
||||
void* newBuffer(int bytes) {
|
||||
size_t size = sizeof(Buffer) + bytes*sizeof(uint8_t);
|
||||
|
||||
@ -77,11 +72,6 @@ void setBufferString(void* buffer, int loc, char* str) {
|
||||
}
|
||||
}
|
||||
|
||||
size_t writeBufferData(FILE* h, void* buffer, size_t loc, size_t len) {
|
||||
Buffer* b = buffer;
|
||||
return fwrite(b->data + loc, sizeof(uint8_t), len, h);
|
||||
}
|
||||
|
||||
uint8_t getBufferByte(void* buffer, int loc) {
|
||||
Buffer* b = buffer;
|
||||
if (loc >= 0 && loc < b->size) {
|
||||
@ -96,7 +86,7 @@ int64_t getBufferInt(void* buffer, int loc) {
|
||||
if (loc >= 0 && loc+7 < b->size) {
|
||||
int64_t result = 0;
|
||||
for (size_t i=0; i<8; i++) {
|
||||
result |= (int64_t)b->data[loc + i] << (8 * i);
|
||||
result |= (uint64_t)(uint8_t)b->data[loc + i] << (8 * i);
|
||||
}
|
||||
return result;
|
||||
} else {
|
||||
@ -129,8 +119,3 @@ char* getBufferString(void* buffer, int loc, int len) {
|
||||
rs[len] = '\0';
|
||||
return rs;
|
||||
}
|
||||
|
||||
size_t readBufferData(FILE* h, void* buffer, size_t loc, size_t max) {
|
||||
Buffer* b = buffer;
|
||||
return fread(b->data + loc, sizeof(uint8_t), max, h);
|
||||
}
|
||||
|
@ -4,6 +4,11 @@
|
||||
#include <stdlib.h>
|
||||
#include <stdint.h>
|
||||
|
||||
typedef struct {
|
||||
int size;
|
||||
char data[];
|
||||
} Buffer;
|
||||
|
||||
void* newBuffer(int bytes);
|
||||
|
||||
int getBufferSize(void* buffer);
|
||||
@ -12,7 +17,6 @@ void setBufferByte(void* buffer, int loc, int byte);
|
||||
void setBufferInt(void* buffer, int loc, int64_t val);
|
||||
void setBufferDouble(void* buffer, int loc, double val);
|
||||
void setBufferString(void* buffer, int loc, char* str);
|
||||
size_t writeBufferData(FILE* h, void* buffer, size_t loc, size_t len);
|
||||
|
||||
void copyBuffer(void* from, int start, int len,
|
||||
void* to, int loc);
|
||||
@ -21,4 +25,3 @@ uint8_t getBufferByte(void* buffer, int loc);
|
||||
int64_t getBufferInt(void* buffer, int loc);
|
||||
double getBufferDouble(void* buffer, int loc);
|
||||
char* getBufferString(void* buffer, int loc, int len);
|
||||
size_t readBufferData(FILE* h, void* buffer, size_t loc, size_t max);
|
||||
|
@ -7,6 +7,8 @@
|
||||
#include <stdint.h>
|
||||
#include <gmp.h>
|
||||
|
||||
#include "buffer.h"
|
||||
|
||||
#define NO_TAG 0
|
||||
#define BITS8_TAG 1
|
||||
#define BITS16_TAG 2
|
||||
@ -177,7 +179,7 @@ typedef struct
|
||||
{
|
||||
Value_header header;
|
||||
size_t len;
|
||||
char *buffer;
|
||||
Buffer *buffer;
|
||||
} Value_Buffer;
|
||||
|
||||
typedef struct
|
||||
|
Loading…
Reference in New Issue
Block a user