2020-05-18 15:59:07 +03:00
|
|
|
||| Low-Level C Sockets bindings for Idris. Used by higher-level, cleverer things.
|
|
|
|
||| Type-unsafe parts. Use Network.Socket for a safe variant.
|
|
|
|
|||
|
|
|
|
||| Original (C) SimonJF, MIT Licensed, 2014
|
|
|
|
||| Modified (C) The Idris Community, 2015, 2016
|
|
|
|
module Network.Socket.Raw
|
|
|
|
|
|
|
|
import public Network.Socket.Data
|
|
|
|
|
2020-05-20 22:09:56 +03:00
|
|
|
import Network.FFI
|
|
|
|
|
2020-05-18 15:59:07 +03:00
|
|
|
-- ---------------------------------------------------------------- [ Pointers ]
|
|
|
|
|
|
|
|
public export
|
|
|
|
data RecvStructPtr = RSPtr AnyPtr
|
|
|
|
|
|
|
|
public export
|
|
|
|
data RecvfromStructPtr = RFPtr AnyPtr
|
|
|
|
|
|
|
|
public export
|
|
|
|
data BufPtr = BPtr AnyPtr
|
|
|
|
|
|
|
|
public export
|
|
|
|
data SockaddrPtr = SAPtr AnyPtr
|
|
|
|
|
|
|
|
-- ---------------------------------------------------------- [ Socket Utilies ]
|
|
|
|
|
|
|
|
||| Put a value in a buffer
|
|
|
|
export
|
2020-06-21 17:25:40 +03:00
|
|
|
sock_poke : HasIO io => BufPtr -> Int -> Int -> io ()
|
2020-05-20 22:09:56 +03:00
|
|
|
sock_poke (BPtr ptr) offset val = primIO $ idrnet_poke ptr offset val
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Take a value from a buffer
|
|
|
|
export
|
2020-06-21 17:25:40 +03:00
|
|
|
sock_peek : HasIO io => BufPtr -> Int -> io Int
|
2020-05-20 22:09:56 +03:00
|
|
|
sock_peek (BPtr ptr) offset = primIO $ idrnet_peek ptr offset
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Frees a given pointer
|
|
|
|
export
|
2020-06-21 17:25:40 +03:00
|
|
|
sock_free : HasIO io => BufPtr -> io ()
|
2020-05-20 22:09:56 +03:00
|
|
|
sock_free (BPtr ptr) = primIO $ idrnet_free ptr
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
export
|
2020-06-21 17:25:40 +03:00
|
|
|
sockaddr_free : HasIO io => SockaddrPtr -> io ()
|
2020-05-20 22:09:56 +03:00
|
|
|
sockaddr_free (SAPtr ptr) = primIO $ idrnet_free ptr
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Allocates an amount of memory given by the ByteLength parameter.
|
|
|
|
|||
|
|
|
|
||| Used to allocate a mutable pointer to be given to the Recv functions.
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
sock_alloc : HasIO io => ByteLength -> io BufPtr
|
2020-05-20 22:09:56 +03:00
|
|
|
sock_alloc bl = map BPtr $ primIO $ idrnet_malloc bl
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Retrieves the port the given socket is bound to
|
|
|
|
export
|
2020-06-21 17:25:40 +03:00
|
|
|
getSockPort : HasIO io => Socket -> io Port
|
2020-05-20 22:09:56 +03:00
|
|
|
getSockPort sock = primIO $ idrnet_sockaddr_port $ descriptor sock
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
|
|
|
|
||| Retrieves a socket address from a sockaddr pointer
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
getSockAddr : HasIO io => SockaddrPtr -> io SocketAddress
|
2020-05-18 15:59:07 +03:00
|
|
|
getSockAddr (SAPtr ptr) = do
|
2020-05-20 22:09:56 +03:00
|
|
|
addr_family_int <- primIO $ idrnet_sockaddr_family ptr
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
-- ASSUMPTION: Foreign call returns a valid int
|
|
|
|
assert_total (case getSocketFamily addr_family_int of
|
|
|
|
Just AF_INET => do
|
2020-05-20 22:09:56 +03:00
|
|
|
ipv4_addr <- primIO $ idrnet_sockaddr_ipv4 ptr
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
pure $ parseIPv4 ipv4_addr
|
|
|
|
Just AF_INET6 => pure IPv6Addr
|
|
|
|
Just AF_UNSPEC => pure InvalidAddress)
|
|
|
|
|
|
|
|
export
|
2020-06-21 17:25:40 +03:00
|
|
|
freeRecvStruct : HasIO io => RecvStructPtr -> io ()
|
2020-05-20 22:09:56 +03:00
|
|
|
freeRecvStruct (RSPtr p) = primIO $ idrnet_free_recv_struct p
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Utility to extract data.
|
|
|
|
export
|
2020-06-21 17:25:40 +03:00
|
|
|
freeRecvfromStruct : HasIO io => RecvfromStructPtr -> io ()
|
2020-05-20 22:09:56 +03:00
|
|
|
freeRecvfromStruct (RFPtr p) = primIO $ idrnet_free_recvfrom_struct p
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Sends the data in a given memory location
|
|
|
|
|||
|
|
|
|
||| Returns on failure a `SocketError`
|
|
|
|
||| Returns on success the `ResultCode`
|
|
|
|
|||
|
|
|
|
||| @sock The socket on which to send the message.
|
|
|
|
||| @ptr The location containing the data to send.
|
|
|
|
||| @len How much of the data to send.
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
sendBuf : HasIO io
|
2020-06-21 17:25:40 +03:00
|
|
|
=> (sock : Socket)
|
2020-05-18 15:59:07 +03:00
|
|
|
-> (ptr : BufPtr)
|
|
|
|
-> (len : ByteLength)
|
2020-06-21 17:25:40 +03:00
|
|
|
-> io (Either SocketError ResultCode)
|
2020-05-18 15:59:07 +03:00
|
|
|
sendBuf sock (BPtr ptr) len = do
|
2020-05-20 22:09:56 +03:00
|
|
|
send_res <- primIO $ idrnet_send_buf (descriptor sock) ptr len
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
if send_res == (-1)
|
|
|
|
then map Left getErrno
|
|
|
|
else pure $ Right send_res
|
|
|
|
|
|
|
|
||| Receive data from a given memory location.
|
|
|
|
|||
|
|
|
|
||| Returns on failure a `SocketError`
|
|
|
|
||| Returns on success the `ResultCode`
|
|
|
|
|||
|
|
|
|
||| @sock The socket on which to receive the message.
|
|
|
|
||| @ptr The location containing the data to receive.
|
|
|
|
||| @len How much of the data to receive.
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
recvBuf : HasIO io
|
2020-06-21 17:25:40 +03:00
|
|
|
=> (sock : Socket)
|
2020-05-18 15:59:07 +03:00
|
|
|
-> (ptr : BufPtr)
|
|
|
|
-> (len : ByteLength)
|
2020-06-21 17:25:40 +03:00
|
|
|
-> io (Either SocketError ResultCode)
|
2020-05-18 15:59:07 +03:00
|
|
|
recvBuf sock (BPtr ptr) len = do
|
2020-05-20 22:09:56 +03:00
|
|
|
recv_res <- primIO $ idrnet_recv_buf (descriptor sock) ptr len
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
if (recv_res == (-1))
|
|
|
|
then map Left getErrno
|
|
|
|
else pure $ Right recv_res
|
|
|
|
|
|
|
|
||| Send a message stored in some buffer.
|
|
|
|
|||
|
|
|
|
||| Returns on failure a `SocketError`
|
|
|
|
||| Returns on success the `ResultCode`
|
|
|
|
|||
|
|
|
|
||| @sock The socket on which to send the message.
|
|
|
|
||| @addr Address of the recipient.
|
|
|
|
||| @port The port on which to send the message.
|
|
|
|
||| @ptr A Pointer to the buffer containing the message.
|
|
|
|
||| @len The size of the message.
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
sendToBuf : HasIO io
|
2020-06-21 17:25:40 +03:00
|
|
|
=> (sock : Socket)
|
2020-05-18 15:59:07 +03:00
|
|
|
-> (addr : SocketAddress)
|
|
|
|
-> (port : Port)
|
|
|
|
-> (ptr : BufPtr)
|
|
|
|
-> (len : ByteLength)
|
2020-06-21 17:25:40 +03:00
|
|
|
-> io (Either SocketError ResultCode)
|
2020-05-18 15:59:07 +03:00
|
|
|
sendToBuf sock addr p (BPtr dat) len = do
|
2020-05-20 22:09:56 +03:00
|
|
|
sendto_res <- primIO $ idrnet_sendto_buf
|
|
|
|
(descriptor sock) dat len (show addr) p (toCode $ family sock)
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
if sendto_res == (-1)
|
|
|
|
then map Left getErrno
|
|
|
|
else pure $ Right sendto_res
|
|
|
|
|
|
|
|
||| Utility function to get the payload of the sent message as a `String`.
|
|
|
|
export
|
2020-06-21 17:25:40 +03:00
|
|
|
foreignGetRecvfromPayload : HasIO io => RecvfromStructPtr -> io String
|
2020-05-20 22:09:56 +03:00
|
|
|
foreignGetRecvfromPayload (RFPtr p) = primIO $ idrnet_get_recvfrom_payload p
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
||| Utility function to return senders socket address.
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
foreignGetRecvfromAddr : HasIO io => RecvfromStructPtr -> io SocketAddress
|
2020-05-18 15:59:07 +03:00
|
|
|
foreignGetRecvfromAddr (RFPtr p) = do
|
2020-05-20 22:09:56 +03:00
|
|
|
sockaddr_ptr <- map SAPtr $ primIO $ idrnet_get_recvfrom_sockaddr p
|
2020-05-18 15:59:07 +03:00
|
|
|
getSockAddr sockaddr_ptr
|
|
|
|
|
|
|
|
||| Utility function to return sender's IPV4 port.
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
foreignGetRecvfromPort : HasIO io => RecvfromStructPtr -> io Port
|
2020-05-18 15:59:07 +03:00
|
|
|
foreignGetRecvfromPort (RFPtr p) = do
|
2020-05-20 22:09:56 +03:00
|
|
|
sockaddr_ptr <- primIO $ idrnet_get_recvfrom_sockaddr p
|
|
|
|
port <- primIO $ idrnet_sockaddr_ipv4_port sockaddr_ptr
|
2020-05-18 15:59:07 +03:00
|
|
|
pure port
|
|
|
|
|
|
|
|
||| Receive a message placed on a 'known' buffer.
|
|
|
|
|||
|
|
|
|
||| Returns on failure a `SocketError`.
|
|
|
|
||| Returns on success a pair of
|
|
|
|
||| + `UDPAddrInfo` :: The address of the sender.
|
|
|
|
||| + `Int` :: Result value from underlying function.
|
|
|
|
|||
|
|
|
|
||| @sock The channel on which to receive.
|
|
|
|
||| @ptr Pointer to the buffer to place the message.
|
|
|
|
||| @len Size of the expected message.
|
|
|
|
|||
|
|
|
|
export
|
Back to HasIO, remove MonadIO
Following a fairly detailed discussion on slack, the feeling is
generally that it's better to have a single interface. While precision
is nice, it doesn't appear to buy us anything here. If that turns out to
be wrong, or limiting somehow, we can revisit it later. Also:
- it's easier for backend authors if the type of IO operations is
slightly less restrictive. For example, if it's in HasIO, that limits
alternative implementations, which might be awkward for some
alternative back ends.
- it's one less extra detail to learn. This is minor, but there needs to
be a clear advantage if there's more detail to learn.
- It is difficult to think of an underlying type that can't have a Monad
instance (I have personally never encountered one - if they turns out
to exist, again, we can revisit!)
2020-06-21 21:21:22 +03:00
|
|
|
recvFromBuf : HasIO io
|
2020-06-21 17:25:40 +03:00
|
|
|
=> (sock : Socket)
|
2020-05-18 15:59:07 +03:00
|
|
|
-> (ptr : BufPtr)
|
|
|
|
-> (len : ByteLength)
|
2020-06-21 17:25:40 +03:00
|
|
|
-> io (Either SocketError (UDPAddrInfo, ResultCode))
|
2020-05-18 15:59:07 +03:00
|
|
|
recvFromBuf sock (BPtr ptr) bl = do
|
2020-05-20 22:09:56 +03:00
|
|
|
recv_ptr <- primIO $ idrnet_recvfrom_buf (descriptor sock) ptr bl
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
let recv_ptr' = RFPtr recv_ptr
|
|
|
|
|
|
|
|
isnull <- nullPtr recv_ptr
|
|
|
|
|
|
|
|
if isnull
|
|
|
|
then map Left getErrno
|
|
|
|
else do
|
2020-05-20 22:09:56 +03:00
|
|
|
result <- primIO $ idrnet_get_recvfrom_res recv_ptr
|
2020-05-18 15:59:07 +03:00
|
|
|
if result == -1
|
|
|
|
then do
|
|
|
|
freeRecvfromStruct recv_ptr'
|
|
|
|
map Left getErrno
|
|
|
|
else do
|
|
|
|
port <- foreignGetRecvfromPort recv_ptr'
|
|
|
|
addr <- foreignGetRecvfromAddr recv_ptr'
|
|
|
|
freeRecvfromStruct recv_ptr'
|
|
|
|
pure $ Right (MkUDPAddrInfo addr port, result + 1)
|