2020-07-21 14:30:33 +03:00
|
|
|
||| FFI binding to the low-Level C Sockets bindings for Idris.
|
|
|
|
|||
|
2020-05-20 22:09:56 +03:00
|
|
|
||| Modified (C) The Idris Community, 2020
|
|
|
|
module Network.FFI
|
|
|
|
|
|
|
|
import Network.Socket.Data
|
|
|
|
|
|
|
|
-- From idris_net.h
|
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_socket, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_socket : (domain, type, protocol : Int) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-07-21 16:35:21 +03:00
|
|
|
%foreign "C:idrnet_close, libidris2_support, idris_net.h"
|
|
|
|
export
|
|
|
|
prim__idrnet_close : (sockdes : SocketDescriptor) -> PrimIO Int
|
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_bind, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_bind : (sockfd : SocketDescriptor) -> (family, socket_type : Int) ->
|
|
|
|
(host : String) -> (port : Port) -> PrimIO Int
|
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_connect, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_connect : (sockfd : SocketDescriptor) -> (family, socket_type : Int) ->
|
|
|
|
(host : String) -> (port : Port) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-07-21 16:35:21 +03:00
|
|
|
%foreign "C:idrnet_listen, libidris2_support, idris_net.h"
|
|
|
|
export
|
|
|
|
prim__idrnet_listen : (sockfd : SocketDescriptor) -> (backlog : Int) -> PrimIO Int
|
|
|
|
|
|
|
|
%foreign "C:idrnet_fdopen, libidris2_support, idris_net.h"
|
|
|
|
export
|
|
|
|
prim__idrnet_fdopen : Int -> String -> PrimIO AnyPtr
|
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_sockaddr_family, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_sockaddr_family : (sockaddr : AnyPtr) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_sockaddr_ipv4, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_sockaddr_ipv4 : (sockaddr : AnyPtr) -> PrimIO String
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_sockaddr_unix, libidris2_support, idris_net.h"
|
2021-04-12 13:22:45 +03:00
|
|
|
export
|
|
|
|
prim__idrnet_sockaddr_unix : (sockaddr : AnyPtr) -> PrimIO String
|
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_sockaddr_ipv4_port, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_sockaddr_ipv4_port : (sockaddr : AnyPtr) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_sockaddr_port, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_sockaddr_port : (sockfd : SocketDescriptor) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_create_sockaddr, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_create_sockaddr : PrimIO AnyPtr
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-06-15 16:31:31 +03:00
|
|
|
%foreign "C__collect_safe:idrnet_accept, libidris2_support, idris_net.h"
|
|
|
|
"C:idrnet_accept, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_accept : (sockfd : SocketDescriptor) -> (sockaddr : AnyPtr) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_send, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_send : (sockfd : SocketDescriptor) -> (dataString : String) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_send_buf, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_send_buf : (sockfd : SocketDescriptor) -> (dataBuffer : AnyPtr) -> (len : Int) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
|
|
|
|
2021-06-15 16:31:31 +03:00
|
|
|
%foreign "C__collect_safe:idrnet_recv, libidris2_support, idris_net.h"
|
|
|
|
"C:idrnet_recv, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_recv : (sockfd : SocketDescriptor) -> (len : Int) -> PrimIO AnyPtr
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-06-15 16:31:31 +03:00
|
|
|
%foreign "C__collect_safe:idrnet_recv_buf, libidris2_support, idris_net.h"
|
|
|
|
"C:idrnet_recv_buf, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_recv_buf : (sockfd : SocketDescriptor) -> (buf : AnyPtr) -> (len : Int) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_sendto, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_sendto : (sockfd : SocketDescriptor) -> (dataString,host : String) ->
|
|
|
|
(port : Port) -> (family : Int) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_sendto_buf, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_sendto_buf : (sockfd : SocketDescriptor) -> (dataBuf : AnyPtr) ->
|
|
|
|
(buf_len : Int) -> (host : String) -> (port : Port) ->
|
|
|
|
(family : Int) -> PrimIO Int
|
|
|
|
|
2021-06-15 16:31:31 +03:00
|
|
|
%foreign "C__collect_safe:idrnet_recvfrom, libidris2_support, idris_net.h"
|
|
|
|
"C:idrnet_recvfrom, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_recvfrom : (sockfd : SocketDescriptor) -> (len : Int) -> PrimIO AnyPtr
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-06-15 16:31:31 +03:00
|
|
|
%foreign "C__collect_safe:idrnet_recvfrom_buf, libidris2_support, idris_net.h"
|
|
|
|
"C:idrnet_recvfrom_buf, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_recvfrom_buf : (sockfd : SocketDescriptor) -> (buf : AnyPtr) -> (len : Int) -> PrimIO AnyPtr
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_get_recv_res, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_get_recv_res : (res_struct : AnyPtr) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_get_recv_payload, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_get_recv_payload : (res_struct : AnyPtr) -> PrimIO String
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_free_recv_struct, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_free_recv_struct : (res_struct : AnyPtr) -> PrimIO ()
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_get_recvfrom_res, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_get_recvfrom_res : (res_struct : AnyPtr) -> PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_get_recvfrom_payload, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_get_recvfrom_payload : (res_struct : AnyPtr) -> PrimIO String
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_get_recvfrom_sockaddr, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_get_recvfrom_sockaddr : (res_struct : AnyPtr) -> PrimIO AnyPtr
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_free_recvfrom_struct, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_free_recvfrom_struct : (res_struct : AnyPtr) -> PrimIO ()
|
2020-05-20 22:09:56 +03:00
|
|
|
|
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_geteagain, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_geteagain : PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_errno, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_errno : PrimIO Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_peek, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_peek : (ptr : AnyPtr) -> (offset : {-Unsigned-} Int) -> PrimIO {-Unsigned-} Int
|
2020-05-20 22:09:56 +03:00
|
|
|
|
2021-04-23 13:09:31 +03:00
|
|
|
%foreign "C:idrnet_poke, libidris2_support, idris_net.h"
|
2020-05-20 22:09:56 +03:00
|
|
|
export
|
2020-07-21 14:30:33 +03:00
|
|
|
prim__idrnet_poke : (ptr : AnyPtr) -> (offset : {-Unsigned-} Int) ->
|
|
|
|
(val : Int {- should be Char? -}) -> PrimIO ()
|