Fix issues with use of unix sockets (#1284)

This change adds logic to set up sockaddr correctly for connect and
bind, handles the AF_UNIX case for getSockAddr and expands the existing
test to cover unix sockets.
This commit is contained in:
Raoul Hidalgo Charman 2021-04-12 11:22:45 +01:00 committed by GitHub
parent c3ff63ae5f
commit 1211f860b6
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
6 changed files with 76 additions and 28 deletions

View File

@ -40,6 +40,10 @@ prim__idrnet_sockaddr_family : (sockaddr : AnyPtr) -> PrimIO Int
export
prim__idrnet_sockaddr_ipv4 : (sockaddr : AnyPtr) -> PrimIO String
%foreign "C:idrnet_sockaddr_unix,libidris2_support"
export
prim__idrnet_sockaddr_unix : (sockaddr : AnyPtr) -> PrimIO String
%foreign "C:idrnet_sockaddr_ipv4_port,libidris2_support"
export
prim__idrnet_sockaddr_ipv4_port : (sockaddr : AnyPtr) -> PrimIO Int

View File

@ -70,6 +70,7 @@ getSockAddr (SAPtr ptr) = do
pure $ parseIPv4 ipv4_addr
Just AF_INET6 => pure IPv6Addr
Just AF_UNIX => map Hostname $ primIO (prim__idrnet_sockaddr_unix ptr)
Just AF_UNSPEC => pure InvalidAddress)
export

View File

@ -48,6 +48,14 @@ void buf_ntohl(void* buf, int len) {
}
}
struct sockaddr_un get_sockaddr_unix(char* host) {
struct sockaddr_un addr;
memset(&addr, 0, sizeof(addr));
addr.sun_family = AF_UNIX;
strcpy(addr.sun_path, host);
return addr;
}
void* idrnet_malloc(int size) {
return malloc(size);
}
@ -119,14 +127,20 @@ int idrnet_getaddrinfo(struct addrinfo** address_res, char* host, int port,
}
int idrnet_bind(int sockfd, int family, int socket_type, char* host, int port) {
struct addrinfo *address_res;
int addr_res = idrnet_getaddrinfo(&address_res, host, port, family, socket_type);
if (addr_res != 0) {
//printf("Lib err: bind getaddrinfo\n");
return -1;
}
int bind_res;
if (family == AF_UNIX) {
struct sockaddr_un addr = get_sockaddr_unix(host);
bind_res = bind(sockfd, (struct sockaddr *)&addr, sizeof(addr));
} else {
struct addrinfo *address_res;
int addr_res = idrnet_getaddrinfo(&address_res, host, port, family, socket_type);
if (addr_res != 0) {
//printf("Lib err: bind getaddrinfo\n");
return -1;
}
int bind_res = bind(sockfd, address_res->ai_addr, address_res->ai_addrlen);
bind_res = bind(sockfd, address_res->ai_addr, address_res->ai_addrlen);
}
if (bind_res == -1) {
//freeaddrinfo(address_res);
//printf("Lib err: bind\n");
@ -165,6 +179,10 @@ int idrnet_sockaddr_port(int sockfd) {
int idrnet_connect(int sockfd, int family, int socket_type, char* host, int port) {
if (family == AF_UNIX) {
struct sockaddr_un addr = get_sockaddr_unix(host);
return connect(sockfd, (struct sockaddr *)&addr, sizeof(addr));
}
struct addrinfo* remote_host;
int addr_res = idrnet_getaddrinfo(&remote_host, host, port, family, socket_type);
if (addr_res != 0) {
@ -199,6 +217,11 @@ int idrnet_sockaddr_ipv4_port(void* sockaddr) {
return ((int) ntohs(addr->sin_port));
}
char* idrnet_sockaddr_unix(void* sockaddr) {
struct sockaddr_un* addr = (struct sockaddr_un*) sockaddr;
return addr->sun_path;
}
void* idrnet_create_sockaddr() {
return malloc(sizeof(struct sockaddr_storage));
}

View File

@ -10,6 +10,7 @@
#include <unistd.h>
#include <sys/types.h>
#include <sys/socket.h>
#include <sys/un.h>
#endif
struct sockaddr_storage;
@ -28,6 +29,16 @@ typedef struct idrnet_recvfrom_result {
struct sockaddr_storage* remote_addr;
} idrnet_recvfrom_result;
#ifdef _WIN32
// mingw-w64 is currently missing the afunix.h header even though windows
// supports unix sockets so we have to define the unix sockaddr structure
// ourselves
struct sockaddr_un {
u_short sun_family;
char sun_path[108];
};
#endif
// Memory management functions
void* idrnet_malloc(int size);
void idrnet_free(void* ptr);
@ -58,6 +69,7 @@ int idrnet_connect(int sockfd, int family, int socket_type, char* host, int port
int idrnet_sockaddr_family(void* sockaddr);
char* idrnet_sockaddr_ipv4(void* sockaddr);
int idrnet_sockaddr_ipv4_port(void* sockaddr);
char* idrnet_sockaddr_unix(void *sockaddr);
void* idrnet_create_sockaddr();
int idrnet_sockaddr_port(int sockfd);

View File

@ -2,29 +2,31 @@ module Main
import System
import System.Info
import System.File
import Network.Socket
import Network.Socket.Data
import Network.Socket.Raw
runServer : IO (Either String (Port, ThreadID))
runServer = do
Right sock <- socket AF_INET Stream 0
testSocket : String
testSocket = "./idris2_test.socket"
runServer : (net : Bool) -> IO (Either String (if net then Port else ()))
runServer net = do
Right sock <- socket (if net then AF_INET else AF_UNIX) Stream 0
| Left fail => pure (Left $ "Failed to open socket: " ++ show fail)
res <- bind sock (Just (Hostname "localhost")) 0
res <- bind sock (Just (Hostname (if net then "localhost" else testSocket))) 0
if res /= 0
then pure (Left $ "Failed to bind socket with error: " ++ show res)
else do
port <- getSockPort sock
port <- if net then getSockPort sock else pure ()
res <- listen sock
if res /= 0
then pure (Left $ "Failed to listen on socket with error: " ++ show res)
else do
forked <- fork (serve port sock)
pure $ Right (port, forked)
then pure . Left $ "Failed to listen on socket with error: " ++ show res
else do ignore $ fork (serve sock)
pure $ Right port
where
serve : Port -> Socket -> IO ()
serve port sock = do
serve : Socket -> IO ()
serve sock = do
Right (s, _) <- accept sock
| Left err => putStrLn ("Failed to accept on socket with error: " ++ show err)
Right (str, _) <- recv s 1024
@ -34,15 +36,15 @@ runServer = do
| Left err => putStrLn ("Server failed to send data with error: " ++ show err)
pure ()
runClient : Port -> IO ()
runClient serverPort = do
Right sock <- socket AF_INET Stream 0
runClient : (net : Bool) -> Port -> IO ()
runClient net serverPort = do
Right sock <- socket (if net then AF_INET else AF_UNIX) Stream 0
| Left fail => putStrLn ("Failed to open socket: " ++ show fail)
res <- connect sock (Hostname "localhost") serverPort
res <- connect sock (Hostname $ if net then "localhost" else testSocket) serverPort
if res /= 0
then putStrLn ("Failed to connect client to port " ++ show serverPort ++ ": " ++ show res)
else do
Right n <- send sock ("hello world!")
Right n <- send sock ("hello world from a " ++ (if net then "ipv4" else "unix") ++ " socket!")
| Left err => putStrLn ("Client failed to send data with error: " ++ show err)
Right (str, _) <- recv sock 1024
| Left err => putStrLn ("Client failed to receive on socket with error: " ++ show err)
@ -52,6 +54,10 @@ runClient serverPort = do
main : IO ()
main = do
Right (serverPort, tid) <- runServer
Right serverPort <- runServer True
| Left err => putStrLn $ "[server] " ++ err
runClient serverPort
runClient True serverPort
Right () <- runServer False
| Left err => putStrLn $ "[server] " ++ err
runClient False 0
ignore $ removeFile testSocket

View File

@ -1,4 +1,6 @@
1/1: Building Echo (Echo.idr)
Main> Received: hello world!
Received: echo: hello world!
Main> Received: hello world from a ipv4 socket!
Received: echo: hello world from a ipv4 socket!
Received: hello world from a unix socket!
Received: echo: hello world from a unix socket!
Main> Bye for now!