diff -r f27ba31b7e96 -r 6f7b9039715f no_shm_selinux/Flask.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/no_shm_selinux/Flask.thy Tue Dec 17 13:30:21 2013 +0800 @@ -0,0 +1,1654 @@ +theory Flask +imports Main Flask_type OS_type_def File_renaming +begin + +(****** Policy-language-level rule match/auxiliary functions ******) + +fun patt_match :: "'a patt \ 'a \ bool" +where + "patt_match (Single x) y = (x = y)" +| "patt_match (Set s) x = (x \ s)" +| "patt_match (Tilde s) x = (x \ s)" +| "patt_match Asterisk x = True" + +(* match a target with a patt, but including the SELF special + * patt_match_self patt target source + *) +fun patt_match_self :: "(type_t patt) target_with_self_t \ type_t \ type_t \ bool" +where + "patt_match_self Self tt st = (tt = st)" +| "patt_match_self (Not_self (Single p)) tt st = (tt = p)" +| "patt_match_self (Not_self (Set p)) tt st = (tt \ p)" +| "patt_match_self (Not_self (Tilde p)) tt st = (tt \ p)" +| "patt_match_self (Not_self Asterisk) tt st = True" + +(* list lookup: find the first one *) +fun lookup :: "'a list \ ('a \ bool) \ 'a option" +where + "lookup [] P = None" +| "lookup (x # xs) P = (if (P x) then Some x else lookup xs P)" + +(* list member: if exists an element satisfy predicate P *) +fun member :: "'a list \ ('a \ bool) \ bool" +where + "member [] P = False" +| "member (x # xs) P = (if (P x) then True else member xs P)" + +(******* socket related aux functions ********) +fun is_remote_request :: "t_event \ bool" +where + "is_remote_request (Accept p fd addr port fd' i) = True" +| "is_remote_request (Connect p fd addr) = True" +| "is_remote_request (SendSock p fd) = True" +| "is_remote_request (RecvSock p fd) = True" +| "is_remote_request e = False" + +fun ip_of_addr:: "t_sock_addr \ t_inet_addr" +where + "ip_of_addr (INET_ADDR saddr sport) = saddr" + +fun port_of_addr:: "t_sock_addr \ t_socket_port" +where + "port_of_addr (INET_ADDR saddr sport) = sport" + +fun after_bind:: "t_sock_state option \ bool" +where + "after_bind None = False" +| "after_bind (Some SS_create) = False" +| "after_bind _= True" (* after bind , it should have local addr *) + +fun after_conacc:: "t_sock_state option \ bool" +where + "after_conacc None = False" +| "after_conacc (Some SS_create) = False" +| "after_conacc (Some SS_bind) = False" +| "after_conacc (Some SS_listen) = False" +| "after_conacc _ = True" (* after connect or accept, it should have remote addr *) + +(* inet_addr_compare addr1 addr2 \ addr1 \ addr2; but in \'s def,there is a \, is not a function*) +fun subnet_addr :: "t_inet_addr \ t_inet_addr \ bool" +where + "subnet_addr [] addr = True" +| "subnet_addr (x#xs) [] = False" +| "subnet_addr (x#xs) (y#ys) = ((x = y) \ subnet_addr xs ys)" + +definition port_in_range :: "t_socket_port \ t_socket_port \ t_socket_port \ bool" +where + "port_in_range minp maxp p \ ((p < maxp \ p = maxp) \ (p > minp \ p = minp))" + +(* initial value of initiate state constrains aux function *) +definition bidirect_in_init :: "'a set \ ('a \ 'b option) \ bool" +where + "bidirect_in_init S f \ (\ a. a \ S \ (\ b. f a = Some b)) \ + (\ a b. f a = Some b \ a \ S)" + +fun is_file_itag :: "t_inode_tag \ bool" +where + "is_file_itag Tag_FILE = True" +| "is_file_itag tag = False" + +fun is_dir_itag :: "t_inode_tag \ bool" +where + "is_dir_itag Tag_DIR = True" +| "is_dir_itag tag = False" + +fun is_file_dir_itag :: "t_inode_tag \ bool" +where + "is_file_dir_itag Tag_FILE = True" +| "is_file_dir_itag Tag_DIR = True" +| "is_file_dir_itag tag = False" + +fun is_tcp_itag :: "t_inode_tag \ bool" +where + "is_tcp_itag Tag_TCP_SOCK = True" +| "is_tcp_itag tag = False" + +fun is_udp_itag :: "t_inode_tag \ bool" +where + "is_udp_itag Tag_UDP_SOCK = True" +| "is_udp_itag tag = False" + +fun is_sock_itag :: "t_inode_tag \ bool" +where + "is_sock_itag Tag_TCP_SOCK = True" +| "is_sock_itag Tag_UDP_SOCK = True" +| "is_sock_itag tag = False" + + +locale init = + fixes + init_files :: "t_file set" + and init_procs :: "t_process set" + and init_nodes :: "t_node set" + and init_sockets :: "t_socket set" + and init_users :: "user_t set" + + and init_file_of_proc_fd :: "t_process \ t_fd \ t_file" + and init_fds_of_proc :: "t_process \ t_fd set" + and init_oflags_of_proc_fd:: "t_process \ t_fd \ t_open_flags" + and init_inum_of_file :: "t_file \ t_inode_num" + and init_inum_of_socket:: "t_socket \ t_inode_num" + and init_itag_of_inum:: "t_inode_num \ t_inode_tag" + and init_files_hung_by_del :: "t_file set" +(* + and init_file_at_writing_by:: "t_file \ (t_process \ t_fd) set" +*) + and init_socket_state :: "t_socket \ t_sock_state" + + and init_msgqs :: "t_msgq set" + and init_msgs_of_queue:: "t_msgq \ t_msg list" +(* + and init_shms :: "t_shm set" + and init_procs_of_shm :: "t_shm \ (t_process \ t_shm_attach_flag) set" + and init_flag_of_proc_shm :: "t_process \ t_shm \ t_shm_attach_flag" +*) + assumes + parent_file_in_init: "\ f pf. \parent f = Some pf; f \ init_files\ \ pf \ init_files" + and root_is_dir: "\ i. init_inum_of_file [] = Some i \ init_itag_of_inum i = Some Tag_DIR" + and init_file_has_inum: "\ f. f \ init_files \ \ im. init_inum_of_file f = Some im" + and inof_has_file_tag: "\ f im. init_inum_of_file f = Some im \ f \ init_files \ (\ tag. init_itag_of_inum im = Some tag \ is_file_dir_itag tag)" + and init_file_no_son: "\ f im. \init_inum_of_file f = Some im; init_itag_of_inum im = Some Tag_FILE\ \ \ (\ f' \ init_files. parent f' = Some f)" + and init_parentf_is_dir: "\ f pf im. \parent f = Some pf; f \ init_files; init_inum_of_file pf = Some im\ \ init_itag_of_inum im = Some Tag_DIR" + + and init_files_hung_valid: "\ f. f \ init_files_hung_by_del \ f \ init_files \ (\ p fd. init_file_of_proc_fd p fd = Some f)" + and init_files_hung_valid': "\ f im. \f \ init_files_hung_by_del; init_inum_of_file f = Some im\ \ + init_itag_of_inum im = Some Tag_FILE \ (init_itag_of_inum im = Some Tag_DIR \ \ (\ f'. f' \ init_files \ f \ f'))" + + and init_procfds_valid: "\ p fd. fd \ init_fds_of_proc p \ p \ init_procs \ ((\ f \ init_files. init_file_of_proc_fd p fd = Some f) \ (p, fd) \ init_sockets)" + and init_filefd_valid: "\ p fd f. init_file_of_proc_fd p fd = Some f \ fd \ init_fds_of_proc p \ (\ im. init_inum_of_file f = Some im \ init_itag_of_inum im = Some Tag_FILE) \ p \ init_procs \ (\ flags. init_oflags_of_proc_fd p fd = Some flags) \ (p, fd) \ init_sockets" + and init_fileflag_valid: "\ p fd flags. init_oflags_of_proc_fd p fd = Some flags \ \ f. init_file_of_proc_fd p fd = Some f" + +(* + and init_procs_has_shm: "\ p h flag. (p,flag) \ init_procs_of_shm h \ p \ init_procs \ h \ init_shms \ init_flag_of_proc_shm p h = Some flag" + and init_shmflag_has_proc: "\ p h flag. init_flag_of_proc_shm p h = Some flag \ (p, flag) \ init_procs_of_shm h" +*) + and init_msgs_distinct: "\ q. distinct (init_msgs_of_queue q)" + and init_msgs_valid: "\ m q. m \ set (init_msgs_of_queue q) \ q \ init_msgqs" + + and init_socket_has_inode: "\ p fd. (p, fd) \ init_sockets \ \ im. init_inum_of_socket (p, fd) = Some im \ p \ init_procs \ fd \ init_fds_of_proc p \ init_file_of_proc_fd p fd = None" + and inos_has_sock_tag: "\ s im. init_inum_of_socket s = Some im \ s \ init_sockets \ (\ tag. init_itag_of_inum im = Some tag \ is_sock_itag tag)" +(* + and init_netobj_has_state: "bidirect_in_init init_netobjs init_netobj_state" + and init_netobj_has_socktype:"bidirect_in_init init_netobjs init_netobj_socktype" + and init_netobj_has_laddr: "\ s. after_bind (init_netobj_state s) \ init_netobj_localaddr s \ None" + and init_netobj_has_raddr: "\ s. after_conacc (init_netobj_state s) \ init_netobj_remoteaddr s \ None" +*) + + and init_finite_sets: "finite init_files \ finite init_procs \ (\ p. finite (init_fds_of_proc p)) \ finite init_shms \ finite init_msgqs \ finite init_users" + +begin + +definition is_init_file:: "t_file \ bool" +where + "is_init_file f \ (case (init_inum_of_file f) of + Some i \ (case (init_itag_of_inum i) of + Some Tag_FILE \ True + | _ \ False) + | _ \ False)" + +definition is_init_dir:: "t_file \ bool" +where + "is_init_dir f \ (case (init_inum_of_file f) of + Some i \ (case (init_itag_of_inum i) of + Some Tag_DIR \ True + | _ \ False) + | _ \ False)" + +definition is_init_tcp_sock:: "t_socket \ bool" +where + "is_init_tcp_sock s \ (case (init_inum_of_socket s) of + Some i \ (case (init_itag_of_inum i) of + Some Tag_TCP_SOCK \ True + | _ \ False) + | _ \ False)" + +definition is_init_udp_sock:: "t_socket \ bool" +where + "is_init_udp_sock s \ (case (init_inum_of_socket s) of + Some i \ (case (init_itag_of_inum i) of + Some Tag_UDP_SOCK \ True + | _ \ False) + | _ \ False)" + +definition init_same_inode_files :: "t_file \ t_file set" +where + "init_same_inode_files f \ if is_init_file f then {f'. init_inum_of_file f = init_inum_of_file f'} else {}" + +fun init_alive :: "t_object \ bool" +where + "init_alive (O_proc p) = (p \ init_procs)" +| "init_alive (O_file f) = (is_init_file f)" +| "init_alive (O_dir f) = (is_init_dir f)" +| "init_alive (O_fd p fd) = (fd \ init_fds_of_proc p)" +| "init_alive (O_node n) = (n \ init_nodes)" +| "init_alive (O_tcp_sock k) = (is_init_tcp_sock k)" +| "init_alive (O_udp_sock k) = (is_init_udp_sock k)" +(* +| "init_alive (O_shm h) = (h \ init_shms)" +*) +| "init_alive (O_msgq q) = (q \ init_msgqs)" +| "init_alive (O_msg q m) = (m \ set (init_msgs_of_queue q) \ q \ init_msgqs)" + + +(************ system listeners, event-related ***********) + +fun file_of_proc_fd :: "t_state \ t_process \ t_fd \ t_file option" +where + "file_of_proc_fd [] p' fd' = (init_file_of_proc_fd p' fd')" +| "file_of_proc_fd (Open p f flags fd iopt # \) p' fd' = + (if (p' = p \ fd = fd') then Some f else file_of_proc_fd \ p' fd')" +| "file_of_proc_fd (CloseFd p fd # \) p' fd' = + (if (p = p' \ fd = fd') then None else file_of_proc_fd \ p' fd')" +(*deleted: if (f\<^isub>3 \ f\<^isub>1) then None else *) +(* +| "file_of_proc_fd (Rename p f\<^isub>2 f\<^isub>3 # \) p' fd = + (case (file_of_proc_fd \ p' fd) of + Some f\<^isub>1 \ (if (f\<^isub>2 \ f\<^isub>1) then Some (file_after_rename f\<^isub>2 f\<^isub>3 f\<^isub>1) + else Some f\<^isub>1) + | None \ None )" +*) +| "file_of_proc_fd (Execve p f fds # \) p' fd = + (if (p' = p \ fd \ fds) then file_of_proc_fd \ p fd + else if (p' = p) then None + else file_of_proc_fd \ p' fd) " +| "file_of_proc_fd (Clone p p' fds # \) p'' fd = + (if (p'' = p' \ fd \ fds) then file_of_proc_fd \ p fd + else if (p'' = p') then None + else file_of_proc_fd \ p'' fd)" +| "file_of_proc_fd (Kill p\<^isub> p\<^isub>' # \) p'' fd = + (if (p'' = p\<^isub>') then None else file_of_proc_fd \ p'' fd)" +| "file_of_proc_fd (Exit p # \) p' fd = + (if (p = p') then None else file_of_proc_fd \ p' fd)" +| "file_of_proc_fd (_ # \) p fd = file_of_proc_fd \ p fd" + +definition proc_fd_of_file :: "t_state \ t_file \ (t_process \ t_fd) set" +where + "proc_fd_of_file \ f = {(p, fd) | p fd. file_of_proc_fd \ p fd = Some f}" + +definition proc_file_fds :: "t_state \ t_process \ t_fd set" +where + "proc_file_fds s p \ {fd. \ f. file_of_proc_fd s p fd = Some f}" + +definition init_proc_file_fds:: "t_process \ t_fd set" +where + "init_proc_file_fds p \ {fd. \ f. init_file_of_proc_fd p fd = Some f}" + +(****** files and directories ******) + +fun files_hung_by_del :: "t_state \ t_file set" +where + "files_hung_by_del [] = init_files_hung_by_del" +| "files_hung_by_del (UnLink p f # \) = ( + if (proc_fd_of_file \ f = {}) then files_hung_by_del \ + else insert f (files_hung_by_del \))" +(* | files_hung_by_del (Rmdir p f) is not need, for open can only apply to file not dir *) +| "files_hung_by_del (CloseFd p fd # \) = ( + case (file_of_proc_fd \ p fd) of + Some f \ (if (proc_fd_of_file \ f = {(p, fd)}) + then files_hung_by_del \ - {f} + else files_hung_by_del \) + | None \ files_hung_by_del \ )" +(* +| "files_hung_by_del (Rename p f\<^isub>2 f\<^isub>3 # \) = {file_after_rename f\<^isub>2 f\<^isub>3 f\<^isub>1| f\<^isub>1. f\<^isub>1 \ files_hung_by_del \}" (* for rename(2) does not infect on fd of the file, see man -S 2 rename *) +*) +| "files_hung_by_del (e # \) = files_hung_by_del \" +declare files_hung_by_del.simps [simp del] + +fun inum_of_file :: "t_state \ (t_file \ t_inode_num)" +where + "inum_of_file [] = init_inum_of_file" +| "inum_of_file (Open p f flags fd (Some inum) # \) = (inum_of_file \) (f := Some inum)" +| "inum_of_file (CloseFd p fd # \) = ( + case (file_of_proc_fd \ p fd) of + Some f \ ( if ((proc_fd_of_file \ f = {(p, fd)}) \ (f \ files_hung_by_del \)) + then (inum_of_file \) (f := None) + else inum_of_file \ ) + | _ \ (inum_of_file \) )" +| "inum_of_file (UnLink p f # \) = ( + if (proc_fd_of_file \ f = {}) + then (inum_of_file \) (f := None) + else inum_of_file \ )" +| "inum_of_file (Rmdir p f # \) = (inum_of_file \) (f := None)" +| "inum_of_file (Mkdir p f ino # \) = (inum_of_file \) (f:= Some ino)" +| "inum_of_file (LinkHard p f\<^isub>1 f\<^isub>2 # \) = (inum_of_file \) (f\<^isub>2 := inum_of_file \ f\<^isub>1)" +(* +| "inum_of_file (Rename p f\<^isub>2 f\<^isub>3 # \) = (\ f. + if (f\<^isub>2 \ f) then None + else if (f\<^isub>3 \ f) then inum_of_file \ (file_before_rename f\<^isub>2 f\<^isub>3 f) + else inum_of_file \ f )" +*) +| "inum_of_file (e # \) = inum_of_file \" + +definition current_files :: "t_state \ t_file set" +where + "current_files \ = {f. \ i. inum_of_file \ f = Some i}" + +(* experimentary: for the delete obj's (file or socket) inode num, it is unnecessary for itag_of_inum to be NONE, cause this is the situation blocked by inum_of_file / inum_of_socket !!! *) +(* delete a file, just recycle its inode num, but not destroy its inode \, so it is irrelative to these events, like closeFd, Rmdir, UnLink \*) +fun itag_of_inum :: "t_state \ (t_inode_num \ t_inode_tag)" +where + "itag_of_inum [] = init_itag_of_inum" +| "itag_of_inum (Open p f flags fd (Some ino) # \) = (itag_of_inum \) (ino := Some Tag_FILE)" +| "itag_of_inum (Mkdir p f ino # \) = (itag_of_inum \) (ino := Some Tag_DIR)" +| "itag_of_inum (CreateSock p af st fd ino # \) = + (case st of + STREAM \ (itag_of_inum \) (ino := Some Tag_TCP_SOCK) + | DGRAM \ (itag_of_inum \) (ino := Some Tag_UDP_SOCK) )" +| "itag_of_inum (Accept p fd addr lport' fd' ino # \) = + (itag_of_inum \) (ino := Some Tag_TCP_SOCK)" +| "itag_of_inum (_ # \) = itag_of_inum \" + +definition is_file:: "t_state \ t_file \ bool" +where + "is_file \ f \ (case (inum_of_file \ f) of + Some i \ (case (itag_of_inum \ i) of + Some Tag_FILE \ True + | _ \ False) + | _ \ False)" + +definition is_dir:: "t_state \ t_file \ bool" +where + "is_dir \ f \ (case (inum_of_file \ f) of + Some i \ (case (itag_of_inum \ i) of + Some Tag_DIR \ True + | _ \ False) + | _ \ False)" + +definition dir_is_empty :: "t_state \ t_file \ bool" +where + "dir_is_empty \ f \ is_dir \ f \ \ (\ f'. f' \ current_files \ \ f \ f')" + +definition same_inode_files :: "t_state \ t_file \ t_file set" +where + "same_inode_files s f \ if is_file s f then {f'. inum_of_file s f = inum_of_file s f'} else {}" + +definition has_same_inode :: "t_state \ t_file \ t_file \ bool" +where + "has_same_inode s fa fb \ fb \ same_inode_files s fa" + +fun inum_of_socket :: "t_state \ (t_socket \ t_inode_num)" +where + "inum_of_socket [] = init_inum_of_socket" +| "inum_of_socket (CloseFd p fd # \) = + (inum_of_socket \) ((p, fd):= None)" +| "inum_of_socket (CreateSock p af st fd ino # \) = + (inum_of_socket \) ((p, fd) := Some ino)" +| "inum_of_socket (Accept p fd addr lport' fd' ino # \) = + (inum_of_socket \) ((p, fd') := Some ino)" +| "inum_of_socket (Clone p p' fds # \) = + (\ (p'', fd). if (p'' = p' \ fd \ fds) then inum_of_socket \ (p, fd) + else if (p'' = p') then None + else inum_of_socket \ (p'',fd))" +| "inum_of_socket (Execve p f fds # \) = + (\ (p', fd). if (p' = p \ fd \ fds) then inum_of_socket \ (p, fd) + else if (p' = p) then None + else inum_of_socket \ (p', fd))" +| "inum_of_socket (Kill p\<^isub>1 p\<^isub>2 # \) = + (\ (p, fd). if (p = p\<^isub>2) then None else inum_of_socket \ (p, fd) )" +| "inum_of_socket (Exit p # \) = + (\ (p', fd). if (p' = p) then None else inum_of_socket \ (p', fd))" +| "inum_of_socket (_ # \) = inum_of_socket \" + +definition current_sockets :: "t_state \ t_socket set" +where + "current_sockets \ = {s. \ i. inum_of_socket \ s = Some i}" + +definition is_tcp_sock :: "t_state \ t_socket \ bool" +where + "is_tcp_sock \ s \ (case (inum_of_socket \ s) of + Some i \ (case (itag_of_inum \ i) of + Some Tag_TCP_SOCK \ True + | _ \ False) + | _ \ False)" + +definition is_udp_sock :: "t_state \ t_socket \ bool" +where + "is_udp_sock \ s \ (case (inum_of_socket \ s) of + Some i \ (case (itag_of_inum \ i) of + Some Tag_UDP_SOCK \ True + | _ \ False) + | _ \ False)" + +fun current_procs :: "t_state \ t_process set" +where + "current_procs [] = init_procs" +| "current_procs (Clone p p' fds # \) = (insert p' (current_procs \))" +| "current_procs (Exit p # \) = (current_procs \ - {p})" +| "current_procs (Kill p p' # \) = (current_procs \ - {p'})" +| "current_procs (_ # \) = current_procs \" + +fun current_proc_fds :: "t_state \ t_process \ t_fd set" +where + "current_proc_fds [] = init_fds_of_proc" +| "current_proc_fds (Open p f flags fd iopt # \) = + (current_proc_fds \) (p:= insert fd (current_proc_fds \ p))" +| "current_proc_fds (CreateSock p sf st newfd newi # \) = + (current_proc_fds \) (p:= insert newfd (current_proc_fds \ p))" +| "current_proc_fds (Accept p fd raddr port fd' ino # \) = + (current_proc_fds \) (p:= insert fd' (current_proc_fds \ p))" +| "current_proc_fds (CloseFd p fd # \) = + (current_proc_fds \) (p:= (current_proc_fds \ p) - {fd})" +| "current_proc_fds (Clone p p' fds # \) = + (current_proc_fds \) (p':= fds)" +| "current_proc_fds (Execve p f fds # \) = + (current_proc_fds \) (p:= fds)" +| "current_proc_fds (Exit p # \) = (current_proc_fds \) (p := {})" +| "current_proc_fds (Kill p p' # \) = (current_proc_fds \) (p' := {})" +| "current_proc_fds (_ # \) = current_proc_fds \" + +(* +fun current_shms :: "t_state \ t_shm set" +where + "current_shms [] = init_shms" +| "current_shms (CreateShM p newshm # \) = insert newshm (current_shms \)" +| "current_shms (DeleteShM p s # \) = (current_shms \) - {s}" +| "current_shms (e # \) = current_shms \ " + +fun flag_of_proc_shm :: "t_state \ t_process \ t_shm \ t_shm_attach_flag option" +where + "flag_of_proc_shm [] = init_flag_of_proc_shm" +| "flag_of_proc_shm (Attach p h flag # s) = (\ p' h'. + if (p' = p \ h' = h) then Some flag else flag_of_proc_shm s p' h')" +| "flag_of_proc_shm (Detach p h # s) = (\ p' h'. + if (p' = p \ h' = h) then None else flag_of_proc_shm s p' h')" +| "flag_of_proc_shm (CreateShM p h # s) = (\ p' h'. + if (h' = h) then None else flag_of_proc_shm s p' h')" +| "flag_of_proc_shm (DeleteShM p h # s) = (\ p' h'. + if (h' = h) then None else flag_of_proc_shm s p' h')" +| "flag_of_proc_shm (Clone p p' fds shms # s) = (\ p'' h. + if (p'' = p' \ h \ shms) then flag_of_proc_shm s p h + else if (p'' = p') then None + else flag_of_proc_shm s p'' h)" +| "flag_of_proc_shm (Execve p f fds # s) = (\ p' h. + if (p' = p) then None else flag_of_proc_shm s p' h)" +| "flag_of_proc_shm (Kill p p' # s) = (\ p'' h. + if (p'' = p') then None else flag_of_proc_shm s p'' h)" +| "flag_of_proc_shm (Exit p # s) = (\ p' h. + if (p' = p) then None else flag_of_proc_shm s p' h)" +| "flag_of_proc_shm (e # s) = flag_of_proc_shm s" + +fun procs_of_shm :: "t_state \ t_shm \ (t_process \ t_shm_attach_flag) set" +where + "procs_of_shm [] = init_procs_of_shm" +| "procs_of_shm (CreateShM p h # \) = + (procs_of_shm \) (h := {})" (* creator may not be the sharer of the shm *) +| "procs_of_shm (Attach p h flag # \) = + (procs_of_shm \) (h := insert (p, flag) (procs_of_shm \ h))" +| "procs_of_shm (Detach p h # \) = + (procs_of_shm \) (h := (procs_of_shm \ h) - {(p,flag) | flag. (p, flag) \ procs_of_shm \ h})" +| "procs_of_shm (DeleteShM p h # \) = (procs_of_shm \) (h := {})" +| "procs_of_shm (Clone p p' fds shms # \) = + (\ h. if (h \ shms) + then (procs_of_shm \ h) \ {(p', flag) | flag. (p, flag) \ procs_of_shm \ h} + else procs_of_shm \ h)" +| "procs_of_shm (Execve p f fds # \) = + (\ h. procs_of_shm \ h - {(p, flag) | flag. (p, flag) \ procs_of_shm \ h})" +| "procs_of_shm (Kill p p' # s) = + (\ h. (procs_of_shm s h) - {(p', flag) | flag. (p', flag) \ procs_of_shm s h})" +| "procs_of_shm (Exit p # s) = + (\ h. (procs_of_shm s h) - {(p, flag) | flag. (p, flag) \ procs_of_shm s h})" +| "procs_of_shm (e # \) = procs_of_shm \" + +(* +inductive info_flow_shm :: "t_state \ t_process \ t_process \ bool" +where + "p \ current_procs s \ info_flow_shm s p p" +| "\info_flow_shm s p p'; (p', SHM_RDWR) \ procs_of_shm s h; (p'', flag) \ procs_of_shm s h; p' \ p''\ + \ info_flow_shm s p p''" +*) +definition one_flow_shm :: "t_state \ t_shm \ t_process \ t_process \ bool" +where + "one_flow_shm s h from to \ from \ to \ (from, SHM_RDWR) \ procs_of_shm s h \ + (\ flag. (to, flag) \ procs_of_shm s h)" + +fun self_shm :: "t_state \ t_process \ t_process \ bool" +where + "self_shm s p p' = (p = p' \ p' \ current_procs s)" + +definition info_flow_shm :: "t_state \ t_process \ t_process \ bool" +where + "info_flow_shm s from to \ (self_shm s from to) \ (\ h. one_flow_shm s h from to)" +*) + +fun current_msgqs :: "t_state \ t_msg set" +where + "current_msgqs [] = init_msgqs" +| "current_msgqs (CreateMsgq p q # \) = insert q (current_msgqs \)" +| "current_msgqs (RemoveMsgq p q # \) = (current_msgqs \) - {q}" +| "current_msgqs (_ # \) = current_msgqs \" + +fun msgs_of_queue :: "t_state \ t_msgq \ t_msg list" +where + "msgs_of_queue [] = init_msgs_of_queue" +| "msgs_of_queue (CreateMsgq p q # \) = (msgs_of_queue \)(q := [])" +| "msgs_of_queue (SendMsg p q m # \) = (msgs_of_queue \)(q := msgs_of_queue \ q @ [m])" +| "msgs_of_queue (RecvMsg p q m # \) = (msgs_of_queue \)(q := tl (msgs_of_queue \ q))" +| "msgs_of_queue (RemoveMsgq p q # \) = (msgs_of_queue \)(q := [])" +| "msgs_of_queue (_ # \) = msgs_of_queue \" + +definition current_file_inums :: "t_state \ t_inode_num set" +where + "current_file_inums \ \ {im. \ f. inum_of_file \ f = Some im}" + +definition current_sock_inums :: "t_state \ t_inode_num set" +where + "current_sock_inums \ = {im. \ s. inum_of_socket \ s = Some im}" + +definition current_inode_nums :: "t_state \ nat set" +where + "current_inode_nums \ = current_file_inums \ \ current_sock_inums \" + +fun flags_of_proc_fd :: "t_state \ t_process \ t_fd \ t_open_flags option" +where + "flags_of_proc_fd [] p fd = init_oflags_of_proc_fd p fd" +| "flags_of_proc_fd (Open p f flags fd iopt # \) p' fd' = + (if (p = p' \ fd = fd') then Some flags else flags_of_proc_fd \ p' fd')" +| "flags_of_proc_fd (CloseFd p fd # \) p' fd' = + (if (p = p' \ fd = fd') then None else flags_of_proc_fd \ p' fd')" +| "flags_of_proc_fd (CreateSock p af st fd ino # \) p' fd' = + (if (p = p' \ fd = fd') then None else flags_of_proc_fd \ p' fd')" +| "flags_of_proc_fd (Accept p fd addr lport' fd' ino # \) p' fd'' = + (if (p = p' \ fd' = fd'') then None else flags_of_proc_fd \ p' fd'')" +| "flags_of_proc_fd (Clone p p' fds # \) p'' fd = + (if (p' = p'' \ fd \ fds) then flags_of_proc_fd \ p fd else flags_of_proc_fd \ p'' fd)" +| "flags_of_proc_fd (Execve p f fds # \) p' fd = + (if (p' = p \ fd \ fds) then flags_of_proc_fd \ p fd else flags_of_proc_fd \ p' fd)" +| "flags_of_proc_fd (Kill p\<^isub>1 p\<^isub>2 # \) p fd = + (if (p = p\<^isub>2) then None else flags_of_proc_fd \ p fd)" +| "flags_of_proc_fd (Exit p # \) p' fd' = + (if (p = p') then None else flags_of_proc_fd \ p' fd')" +| "flags_of_proc_fd (e # \) p fd = flags_of_proc_fd \ p fd" + +(* +fun file_at_writing_by:: "t_state \ t_file \ (t_process \ t_fd) set" +where + "file_at_writing_by [] f = init_file_at_writing_by f" +| "file_at_writing_by (Open p f flags fd (Some inum) # \) f' = ( + if (f' = f) + then ( if (is_write_flag flags) + then {(p, fd)} + else {} ) + else file_at_writing_by \ f' )" +| "file_at_writing_by (Open p f flags fd None # \) f' = ( + if (f' = f \ is_write_flag flags) + then insert (p, fd) (file_at_writing_by \ f) + else file_at_writing_by \ f' )" +| "file_at_writing_by (Mkdir p f inum # \) f' = + (if (f' = f) then {} else file_at_writing_by \ f')" +| "file_at_writing_by (LinkHard p f f' # \) f'' = + (if (f'' = f') then file_at_writing_by \ f else file_at_writing_by \ f'')" +| "file_at_writing_by (Rename p f\<^isub>2 f\<^isub>3 # \) f = ( + if (f\<^isub>3 \ f) then file_at_writing_by \ (file_before_rename f\<^isub>2 f\<^isub>3 f) + else file_at_writing_by \ f )" +| "file_at_writing_by (CloseFd p fd # \) f = + (file_at_writing_by \ f - {(p, fd)})" +| "file_at_writing_by (Clone p p' fds shms # \) f = + (file_at_writing_by \ f) \ {(p',fd)| fd. fd \ fds \ (p, fd) \ file_at_writing_by \ f}" +| "file_at_writing_by (Execve p f fds # \) f' = + (if (f' = f) then {} else file_at_writing_by \ f')" +| "file_at_writing_by (Exit p # \) f = + (file_at_writing_by \ f - {(p, fd)| fd. (p, fd) \ file_at_writing_by \ f})" +| "file_at_writing_by (Kill p p' # \) f = + (file_at_writing_by \ f - {(p', fd)| fd. (p', fd) \ file_at_writing_by \ f})" +| "file_at_writing_by (e # \) f = file_at_writing_by \ f" +*) + +definition current_users :: "t_state \ user_t set" +where + "current_users \ \ init_users" + +fun update_with_shuthow :: "t_sock_state option \ t_shutdown_how \ t_sock_state option" +where + "update_with_shuthow (Some (SS_trans Trans_Recv)) SHUT_RD = Some (SS_trans Trans_No)" +| "update_with_shuthow (Some (SS_trans Trans_RS)) SHUT_RD = Some (SS_trans Trans_Send)" +| "update_with_shuthow (Some (SS_trans Trans_Send)) SHUT_RD = Some (SS_trans Trans_Send)" +| "update_with_shuthow (Some (SS_trans Trans_No)) SHUT_RD = Some (SS_trans Trans_No)" +| "update_with_shuthow (Some (SS_trans Trans_Recv)) SHUT_WR = Some (SS_trans Trans_Recv)" +| "update_with_shuthow (Some (SS_trans Trans_RS)) SHUT_WR = Some (SS_trans Trans_Recv)" +| "update_with_shuthow (Some (SS_trans Trans_Send)) SHUT_WR = Some (SS_trans Trans_No)" +| "update_with_shuthow (Some (SS_trans Trans_No)) SHUT_WR = Some (SS_trans Trans_No)" +| "update_with_shuthow (Some (SS_trans Trans_Recv)) SHUT_RDWR = Some (SS_trans Trans_No)" +| "update_with_shuthow (Some (SS_trans Trans_RS)) SHUT_RDWR = Some (SS_trans Trans_No)" +| "update_with_shuthow (Some (SS_trans Trans_Send)) SHUT_RDWR = Some (SS_trans Trans_No)" +| "update_with_shuthow (Some (SS_trans Trans_No)) SHUT_RDWR = Some (SS_trans Trans_No)" +| "update_with_shuthow opt1 how = None" + +fun socket_state :: "t_state \ t_socket \ t_sock_state option" +where + "socket_state [] = init_socket_state" +| "socket_state (CloseFd p fd # \) = + (socket_state \) ((p, fd):= None)" +| "socket_state (CreateSock p af st fd ino # \) = + (socket_state \) ((p, fd) := Some SS_create)" +| "socket_state (Bind p fd addr # \) = + (socket_state \) ((p, fd) := Some SS_bind)" +| "socket_state (Connect p fd addr # \) = + (socket_state \) ((p, fd) := Some (SS_trans Trans_RS))" +| "socket_state (Listen p fd # \) = + (socket_state \) ((p, fd) := Some SS_listen)" +| "socket_state (Accept p fd addr lport' fd' ino # \) = + (socket_state \) ((p, fd') := Some (SS_trans Trans_RS))" +| "socket_state (Shutdown p fd sh # \) = + (socket_state \) ((p, fd) := update_with_shuthow (socket_state \ (p, fd)) sh)" +| "socket_state (Clone p p' fds # \) = + (\ (p'', fd). if (p'' = p' \ fd \ fds) then socket_state \ (p, fd) + else if (p'' = p') then None + else socket_state \ (p'', fd))" +| "socket_state (Execve p f fds # \) = + (\ (p', fd). if (p' = p \ fd \ fds) then socket_state \ (p, fd) + else if (p' = p) then None + else socket_state \ (p', fd))" +| "socket_state (Kill p\<^isub>1 p\<^isub>2 # \) = + (\ (p, fd). if (p = p\<^isub>2) then None else socket_state \ (p, fd))" +| "socket_state (Exit p # \) = + (\ (p', fd'). if (p = p') then None else socket_state \ (p', fd'))" +| "socket_state (e # \) = socket_state \" + +definition socket_in_trans :: "t_state \ t_socket \ bool" +where + "socket_in_trans \ s \ (case (socket_state \ s) of + Some st \ (case st of + SS_trans para \ True + | _ \ False) + | _ \ False)" + +definition socket_in_sending :: "t_state \ t_socket \ bool" +where + "socket_in_sending \ s \ (socket_state \ s = Some (SS_trans Trans_Send)) \ (socket_state \ s = Some (SS_trans Trans_RS))" + +definition socket_in_recving :: "t_state \ t_socket \ bool" +where + "socket_in_recving \ s \ (socket_state \ s = Some (SS_trans Trans_Recv)) \ (socket_state \ s = Some (SS_trans Trans_RS))" + + +(******* admissable check by the OS *******) +fun os_grant :: "t_state \ t_event \ bool" +where + "os_grant \ (Open p f flags fd inumopt) = ( + case inumopt of + Some ino \ + (p \ current_procs \ \ f \ current_files \ \ + (\ pf. parent f = Some pf \ is_dir \ pf \ pf \ files_hung_by_del \) \ + is_creat_flag flags \ fd \ (current_proc_fds \ p) \ ino \ (current_inode_nums \)) + | None \ + (p \ current_procs \ \ \ is_creat_excl_flag flags \ is_file \ f \ + fd \ (current_proc_fds \ p)) )" +(*(if (is_dir \ f) then (is_dir_flag flags \ \ is_write_flag flags) else (\ is_dir_flag flags)), + here open is not applied to directories, readdir is for that, but it is not security-related *) +| "os_grant \ (CloseFd p fd) = + (p \ current_procs \ \ fd \ current_proc_fds \ p)" +| "os_grant \ (UnLink p f) = + (p \ current_procs \ \ is_file \ f \ f \ files_hung_by_del \)" +| "os_grant \ (Rmdir p f) = + (p \ current_procs \ \ dir_is_empty \ f \ f \ files_hung_by_del \ \ f \ [])" (* root file cannot be deleted *) +| "os_grant \ (Mkdir p f ino) = + (p \ current_procs \ \ f \ current_files \ \ + (\ pf. parent f = Some pf \ is_dir \ pf \ pf \ files_hung_by_del \) \ + ino \ (current_inode_nums \))" +| "os_grant \ (LinkHard p f\<^isub>1 f\<^isub>2) = + (p \ current_procs \ \ is_file \ f\<^isub>1 \ f\<^isub>2 \ current_files \ \ + (\ pf\<^isub>2. parent f\<^isub>2 = Some pf\<^isub>2 \ is_dir \ pf\<^isub>2 \ pf\<^isub>2 \ files_hung_by_del \))" +| "os_grant \ (Truncate p f len) = + (p \ current_procs \ \ is_file \ f)" +(* +| "os_grant \ (FTruncate p fd len) = + (p \ current_procs \ \ fd \ current_proc_fds \ p \ + (\ f. file_of_proc_fd \ p fd = Some f \ f \ current_files \ \ is_file \ f) \ + (\ flags. flags_of_proc_fd \ p fd = Some flags \ is_write_flag flags))" +*) +| "os_grant \ (ReadFile p fd) = + (p \ current_procs \ \ fd \ current_proc_fds \ p \ + (\ f. file_of_proc_fd \ p fd = Some f \ f \ current_files \) \ + (\ flags. flags_of_proc_fd \ p fd = Some flags \ is_read_flag flags))" +| "os_grant \ (WriteFile p fd) = + (p \ current_procs \ \ fd \ current_proc_fds \ p \ + (\ f. file_of_proc_fd \ p fd = Some f \ f \ current_files \) \ + (\ flags. flags_of_proc_fd \ p fd = Some flags \ is_write_flag flags))" +| "os_grant \ (Execve p f fds) = + (p \ current_procs \ \ is_file \ f \ fds \ proc_file_fds \ p)" (* file_at_writing_by \ f = {} \ fds \ current_proc_fds \ p *) +(* +| "os_grant \ (Rename p f\<^isub>2 f\<^isub>3) = + (p \ current_procs \ \ f\<^isub>2 \ current_files \ \ \(f\<^isub>2 \ f\<^isub>3) \ f\<^isub>3 \ current_files \ \ + (\ pf\<^isub>3. parent f\<^isub>3 = Some pf\<^isub>3 \ pf\<^isub>3 \ current_files \ \ is_dir \ pf\<^isub>3 \ + pf\<^isub>3 \ files_hung_by_del \) )" +*) +| "os_grant \ (CreateMsgq p q) = + (p \ current_procs \ \ q \ (current_msgqs \))" +| "os_grant \ (SendMsg p q m) = + (p \ current_procs \ \ q \ current_msgqs \ \ m \ (set (msgs_of_queue \ q)))" +| "os_grant \ (RecvMsg p q m) = + (p \ current_procs \ \ q \ current_msgqs \ \ m = hd (msgs_of_queue \ q) \ msgs_of_queue \ q \ [])" +| "os_grant \ (RemoveMsgq p q) = + (p \ current_procs \ \ q \ current_msgqs \)" +(* +| "os_grant \ (CreateShM p h) = + (p \ current_procs \ \ h \ (current_shms \))" +| "os_grant \ (Attach p h flag) = + (p \ current_procs \ \ h \ current_shms \ \ \ (\ flag. (p, flag) \ procs_of_shm \ h))" +| "os_grant \ (Detach p h) = + (p \ current_procs \ \ h \ current_shms \ \ (\ flag. (p, flag) \ procs_of_shm \ h))" +| "os_grant \ (DeleteShM p h) = + (p \ current_procs \ \ h \ current_shms \)" +*) +| "os_grant \ (CreateSock p af st fd ino) = + (p \ current_procs \ \ af = AF_INET \ fd \ (current_proc_fds \ p) \ + ino \ (current_inode_nums \))" +| "os_grant \ (Accept p fd addr port fd' ino) = + (p \ current_procs \ \ fd \ current_proc_fds \ p \ (p, fd) \ current_sockets \ \ + fd' \ (current_proc_fds \ p) \ socket_state \ (p, fd) = Some SS_listen \ + is_tcp_sock \ (p, fd) \ ino \ (current_inode_nums \))" +| "os_grant \ (Bind p fd addr) = + (p \ current_procs \ \ fd \ current_proc_fds \ p \ (p, fd) \ current_sockets \ \ + socket_state \ (p, fd) = Some SS_create)" (* ?? !! (case addr of INET_ADDR ip port \ ip \ local_ipaddrs)) *) +| "os_grant \ (Connect p fd addr) = + (p \ current_procs \ \ fd \ current_proc_fds \ p \ (p, fd) \ current_sockets \ \ + socket_state \ (p, fd) = Some SS_bind)" +| "os_grant \ (Listen p fd) = + (p \ current_procs \ \ fd \ current_proc_fds \ p \ is_tcp_sock \ (p, fd) \ + socket_state \ (p, fd) = Some SS_bind)" (* Listen and Accept should only be used for TCP sever side *) +| "os_grant \ (SendSock p fd) = + (p \ current_procs \ \ fd \ current_proc_fds \ p \ (p, fd) \ current_sockets \ \ + socket_in_sending \ (p, fd))" +| "os_grant \ (RecvSock p fd) = + (p \ current_procs \ \ fd \ current_proc_fds \ p \ (p, fd) \ current_sockets \ \ + socket_in_recving \ (p, fd))" +| "os_grant \ (Shutdown p fd sh) = + (p \ current_procs \ \ fd \ current_proc_fds \ p \ (p, fd) \ current_sockets \ \ + socket_in_trans \ (p, fd))" +(* +| "os_grant \ (ChangeOwner p u) = (p \ current_procs \ \ u \ current_users \)" +*) +| "os_grant \ (Clone p p' fds) = + (p \ current_procs \ \ p' \ (current_procs \) \ fds \ proc_file_fds \ p)" (*\ + (\ h \ shms. \ flag. (p, flag) \ procs_of_shm \ h) current_proc_fds \ proc_file_fds *) +| "os_grant \ (Kill p\<^isub>1 p\<^isub>2) = + (p\<^isub>1 \ current_procs \ \ p\<^isub>2 \ current_procs \)" (* a process can kill itself right? *) +| "os_grant \ (Exit p) = + (p \ current_procs \)" +| "os_grant \ (Ptrace p\<^isub>1 p\<^isub>2) = + (p\<^isub>1 \ current_procs \ \ p\<^isub>2 \ current_procs \)" (* a process can trace itself right? *) + +fun alive :: "t_state \ t_object \ bool" +where + "alive s (O_proc p) = (p \ current_procs s)" +| "alive s (O_file f) = (is_file s f)" +| "alive s (O_dir f) = (is_dir s f)" +| "alive s (O_fd p fd) = (fd \ current_proc_fds s p)" +| "alive s (O_node n) = (n \ init_nodes)" +| "alive s (O_tcp_sock k) = (is_tcp_sock s k)" +| "alive s (O_udp_sock k) = (is_udp_sock s k)" +(* +| "alive s (O_shm h) = (h \ current_shms s)" +*) +| "alive s (O_msgq q) = (q \ current_msgqs s)" +| "alive s (O_msg q m) = (m \ set (msgs_of_queue s q) \ q \ current_msgqs s)" + +(* deleted objects should be "taintable" objects, objects like fd should not be in consideration *) +(* actually, deleted should be renamed as "vanished", cause "exit/closefd" *) +fun deleted :: "t_object \ t_event list \ bool" +where + "deleted obj [] = False" +| "deleted obj (Kill p p' # s) = (obj = O_proc p' \ deleted obj s)" +(* +| "deleted obj (Kill p p' # s) = ((obj = O_proc p') \ (\ fd \ current_proc_fds s p'. obj = O_fd p' fd) \ + (\ fd. obj = O_tcp_sock (p', fd) \ is_tcp_sock s (p',fd)) \ + (\ fd. obj = O_udp_sock (p', fd) \ is_udp_sock s (p',fd)) \ + deleted obj s)" +| "deleted obj (CloseFd p fd # s) = ((obj = O_fd p fd) \ (obj = O_tcp_sock (p,fd) \ is_tcp_sock s (p,fd)) \ + (obj = O_udp_sock (p,fd) \ is_udp_sock s (p, fd)) \ + (\ f. obj = O_file f \ proc_fd_of_file s f = {(p, fd)} \ + f \ files_hung_by_del s) \ deleted obj s)" +| "deleted obj (Execve p f fds # s) = ((\ fd \ current_proc_fds s p. obj = O_fd p fd \ fd \ fds) \ + (\ fd. obj = O_tcp_sock (p, fd) \ is_tcp_sock s (p,fd) \ fd \ fds) \ + (\ fd. obj = O_udp_sock (p, fd) \ is_udp_sock s (p,fd) \ fd \ fds) \ + deleted obj s)" +| "deleted obj (Clone p p' fds # s) = ((\ fd \ current_proc_fds s p. obj = O_fd p' fd \ fd \ fds) \ + deleted obj s)" +*) +| "deleted obj (CloseFd p fd # s) = ((\ f. obj = O_file f \ proc_fd_of_file s f = {(p, fd)} \ + f \ files_hung_by_del s) \ deleted obj s)" +| "deleted obj (UnLink p f # s) = ((proc_fd_of_file s f = {} \ obj = O_file f) \ deleted obj s)" +| "deleted obj (Rmdir p f # s) = ((obj = O_dir f) \ deleted obj s)" +(* +| "deleted obj (Exit p # s) = ((obj = O_proc p) \ (\ fd \ current_proc_fds s p. obj = O_fd p fd) \ + (\ fd. obj = O_tcp_sock (p, fd) \ is_tcp_sock s (p,fd)) \ + (\ fd. obj = O_udp_sock (p, fd) \ is_udp_sock s (p,fd)) \ deleted obj s)" +*) +(* +| "deleted obj (Rename p f f' # s) = + (case obj of + O_file f'' \ f \ f'' \ deleted obj s + | O_dir f'' \ f \ f'' \ deleted obj s + | _ \ deleted obj s)" +*) +| "deleted obj (RemoveMsgq p q # s) = (obj = O_msgq q \ deleted obj s)" +(* +| "deleted obj (DeleteShM p h # s) = (obj = O_shm h \ deleted obj s)" + +| "deleted obj (RecvMsg p q m # s) = (obj = O_msg q m \ deleted obj s)" +*) +| "deleted obj (_ # s) = deleted obj s" + +(* actually, died should be renamed as "vanished", cause "exit/closefd" *) +fun died :: "t_object \ t_event list \ bool" +where + "died obj [] = False" +| "died obj (Kill p p' # s) = ((obj = O_proc p') \ (\ fd \ current_proc_fds s p'. obj = O_fd p' fd) \ + (\ fd. obj = O_tcp_sock (p', fd) \ is_tcp_sock s (p',fd)) \ + (\ fd. obj = O_udp_sock (p', fd) \ is_udp_sock s (p',fd)) \ + died obj s)" +| "died obj (CloseFd p fd # s) = ((obj = O_fd p fd) \ (obj = O_tcp_sock (p,fd) \ is_tcp_sock s (p,fd)) \ + (obj = O_udp_sock (p,fd) \ is_udp_sock s (p, fd)) \ + (\ f. obj = O_file f \ proc_fd_of_file s f = {(p, fd)} \ + f \ files_hung_by_del s) \ died obj s)" +| "died obj (Execve p f fds # s) = ((\ fd \ current_proc_fds s p. obj = O_fd p fd \ fd \ fds) \ + (\ fd. obj = O_tcp_sock (p, fd) \ is_tcp_sock s (p,fd) \ fd \ fds) \ + (\ fd. obj = O_udp_sock (p, fd) \ is_udp_sock s (p,fd) \ fd \ fds) \ + died obj s)" +| "died obj (Clone p p' fds # s) = ((\ fd \ current_proc_fds s p. obj = O_fd p' fd \ fd \ fds) \ + died obj s)" +| "died obj (UnLink p f # s) = ((proc_fd_of_file s f = {} \ obj = O_file f) \ died obj s)" +| "died obj (Rmdir p f # s) = ((obj = O_dir f) \ died obj s)" +| "died obj (Exit p # s) = ((obj = O_proc p) \ (\ fd \ current_proc_fds s p. obj = O_fd p fd) \ + (\ fd. obj = O_tcp_sock (p, fd) \ is_tcp_sock s (p,fd)) \ + (\ fd. obj = O_udp_sock (p, fd) \ is_udp_sock s (p,fd)) \ died obj s)" +(* +| "died obj (Rename p f f' # s) = + (case obj of + O_file f'' \ f \ f'' \ died obj s + | O_dir f'' \ f \ f'' \ died obj s + | _ \ died obj s)" +*) +| "died obj (RemoveMsgq p q # s) = (obj = O_msgq q \ (\ m. obj = O_msg q m ) \ died obj s)" +(* +| "died obj (DeleteShM p h # s) = (obj = O_shm h \ died obj s)" +*) +| "died obj (RecvMsg p q m # s) = (obj = O_msg q m \ died obj s)" +| "died obj (_ # s) = died obj s" + +end + + + +locale flask = init + + fixes + (***************** Policy-specific Parameters *************) + type_transition_rules :: "type_transition_rule_list_t" + and allowed_rules :: "allowed_rule_list_t" + and role_allowed_rules :: "role_allow_rule_list_t" + and role_declaration_rules :: "role_declarations_list_t" + and role_transition_rules:: "role_transition_rule_list_t" + and constrains_rules :: "constrains_list_t" + + and init_type_of_obj :: "t_object \ type_t" + and init_user_of_obj :: "t_object \ user_t" + and init_role_of_proc :: "t_process \ role_t" + + assumes + + init_obj_has_type: "\ obj. init_alive obj \ \ t. init_type_of_obj obj = Some t" + and init_type_has_obj: "\ obj t. init_type_of_obj obj = Some t \ init_alive obj" + and init_obj_has_user: "\ obj. init_alive obj \ \ u. init_user_of_obj obj = Some u" + and init_user_has_obj: "\ obj u. init_user_of_obj obj = Some u \ init_alive obj \ u \ init_users" + and init_proc_has_role: "bidirect_in_init init_procs init_role_of_proc" + +begin + +(*********** policy-specified computations, not-event-related ************) + +(* security_transition_sid : type_transition_rules + * It's a system-call offered by security-server? So access-control-checked too? + * according to selinux, it deals only with execve and file-creation + * is_execve: if is the execve case + * is_file: if is the creation of file/dir + * change from is_execve to is_file, because the new message by default use + * the type of the sending process too. + *) + +definition type_transition :: "type_t \ type_t \ security_class_t \ bool \ type_t" +where + "type_transition st rt tc IS_file \ ( + case (lookup type_transition_rules + (\ (stp, rtp, tc', dt). patt_match stp st \ patt_match rtp rt \ tc = tc')) of + None \ (if IS_file then rt else st) + | Some (stp,rtp,tc',dt) \ dt)" + +(* allowed_rule_list *) +definition TE_allow :: "type_t \ type_t \ security_class_t \ permission_t \ bool" +where + "TE_allow st tt tc perm \ member allowed_rules + (\ (stp, ttp, tc', pp). tc = tc' \ patt_match stp st \ patt_match_self ttp tt st \ + patt_match pp perm)" + +(* role_allow_rule_list_t *) +definition role_transition :: "role_t \ type_t \ role_t option" +where + "role_transition r t \ + (case (lookup role_transition_rules (\ (pr, ft, nr). pr = r \ t = ft)) of + None \ None + | Some (pr,ft,nr) \ Some nr)" + +definition role_decl_check :: "role_t \ type_t \ bool" +where + "role_decl_check r t \ member role_declaration_rules (\ (role, types). r = role \ t \ types)" + +definition role_allow :: "role_t \ role_t \ bool" +where + "role_allow r nr \ member role_allowed_rules (\ (from, to). r \ from \ nr \ to)" + +(* here don't use lookup, because constrains should all be satisfied, + * while transitions always choose the "first" one + *) +definition constrains_satisfied :: + "security_context_t \ security_context_t \ security_class_t \ permission_t \ bool" +where + "constrains_satisfied sctxt tctxt c p \ + (fold (\ (cset, pset, P) res. res \ (if (c \ cset \ p \ pset) + then P sctxt tctxt else True)) + constrains_rules True)" + +(* main interface for TE_allow check and constrains check *) +fun permission_check:: + "security_context_t \ security_context_t \ security_class_t \ permission_t \ bool" +where + "permission_check (su,sr,st) (tu,tr,tt) c p = + ((TE_allow st tt c p) \ (constrains_satisfied (su,sr,st) (tu,tr,tt) c p))" +declare permission_check.simps [simp del] + +(* no changeowner ??? : by adding "relabel obj sectxt" event or "login" event? + * or this explanation: the running of the sever will not grant any login but running + * of the server-programs! + *) + +fun user_of_obj :: "t_state \ t_object \ user_t option" +where + "user_of_obj [] = init_user_of_obj" +| "user_of_obj (Clone p p' fds # s) = + (\ obj. case obj of + O_proc p'' \ if (p'' = p') then user_of_obj s (O_proc p) else user_of_obj s obj + | O_fd p'' fd \ if (p'' = p' \ fd \ fds) then user_of_obj s (O_fd p fd) + else if (p'' = p') then None + else user_of_obj s obj + | O_tcp_sock (p'',fd) \ if (p'' = p' \ fd \ fds) then user_of_obj s (O_tcp_sock (p, fd)) + else if (p'' = p') then None + else user_of_obj s obj + | O_udp_sock (p'',fd) \ if (p'' = p' \ fd \ fds) then user_of_obj s (O_udp_sock (p, fd)) + else if (p'' = p') then None + else user_of_obj s obj + | _ \ user_of_obj s obj)" +| "user_of_obj (Open p f flags fd iopt # s) = + (case iopt of + None \ (user_of_obj s) ((O_fd p fd) := user_of_obj s (O_proc p)) + | _ \ ((user_of_obj s) ((O_fd p fd) := user_of_obj s (O_proc p))) + ((O_file f) := user_of_obj s (O_proc p)))" +| "user_of_obj (Mkdir p f inum # s) = + (user_of_obj s) ((O_dir f) := user_of_obj s (O_proc p))" +| "user_of_obj (LinkHard p f f' # s) = + (user_of_obj s) ((O_file f') := user_of_obj s (O_proc p))" +(* +| "user_of_obj (Rename p f2 f3 # s) = + (\ obj. case obj of + O_file f \ if (f3 \ f) then user_of_obj s (O_file (file_before_rename f2 f3 f)) + else user_of_obj s obj + | O_dir f \ if (f3 \ f) then user_of_obj s (O_dir (file_before_rename f2 f3 f)) + else user_of_obj s obj + | _ \ user_of_obj s obj)" +*) +| "user_of_obj (CreateMsgq p q # s) = + (user_of_obj s) ((O_msgq q) := user_of_obj s (O_proc p))" +| "user_of_obj (SendMsg p q m # s) = + (user_of_obj s) ((O_msg q m) := user_of_obj s (O_proc p))" +(* +| "user_of_obj (CreateShM p h # s) = + (user_of_obj s) ((O_shm h) := user_of_obj s (O_proc p))" +*) +| "user_of_obj (CreateSock p af st fd inum # s) = (\ obj. + case obj of + O_fd p' fd' \ if (p' = p \ fd' = fd) then user_of_obj s (O_proc p) + else user_of_obj s obj + | O_tcp_sock (p', fd') \ if (p' = p \ fd' = fd \ st = STREAM) then user_of_obj s (O_proc p) + else user_of_obj s obj + | O_udp_sock (p', fd') \ if (p' = p \ fd' = fd \ st = DGRAM) then user_of_obj s (O_proc p) + else user_of_obj s obj + | _ \ user_of_obj s obj )" +| "user_of_obj (Accept p fd addr port fd' inum # s) = (\ obj. + case obj of + O_fd p' fd'' \ if (p' = p \ fd'' = fd') then user_of_obj s (O_proc p) + else user_of_obj s obj + | O_tcp_sock (p', fd'') \ if (p' = p \ fd'' = fd') then user_of_obj s (O_proc p) + else user_of_obj s obj + | _ \ user_of_obj s obj )" +| "user_of_obj (_ # s) = user_of_obj s" + +fun type_of_obj :: "t_state \ (t_object \ type_t option)" +where + "type_of_obj [] = init_type_of_obj" +| "type_of_obj (Clone p p' fds # s) = (\ obj. + case obj of + O_proc p'' \ if (p'' = p') then type_of_obj s (O_proc p) else type_of_obj s obj + | O_fd p'' fd \ if (p'' = p' \ fd \ fds) then type_of_obj s (O_fd p fd) + else if (p'' = p') then None + else type_of_obj s obj + | O_tcp_sock (p'', fd) \ if (p'' = p' \ fd \ fds) then type_of_obj s (O_tcp_sock (p, fd)) + else if (p'' = p') then None + else type_of_obj s obj + | O_udp_sock (p'', fd) \ if (p'' = p' \ fd \ fds) then type_of_obj s (O_udp_sock (p, fd)) + else if (p'' = p') then None + else type_of_obj s obj + | _ \ type_of_obj s obj )" +| "type_of_obj (Execve p f fds # s) = (type_of_obj s) ((O_proc p) := + (case (type_of_obj s (O_proc p), type_of_obj s (O_file f)) of + (Some tp, Some tf) \ Some (type_transition tp tf C_process False) + | _ \ None) )" +| "type_of_obj (Open p f flags fd (Some inum) # s) = ((type_of_obj s) ((O_file f) := + (case (parent f) of + Some pf \ (case (type_of_obj s (O_proc p), type_of_obj s (O_dir pf)) of + (Some tp, Some tpf) \ Some (type_transition tp tpf C_file True) + | _ \ None) + | _ \ None) )) ((O_fd p fd) := + type_of_obj s (O_proc p))" +| "type_of_obj (Open p f flags fd None # s) = (type_of_obj s) ((O_fd p fd) := + type_of_obj s (O_proc p))" +| "type_of_obj (Mkdir p f inum # s) = (type_of_obj s) ((O_dir f) := + (case (parent f) of + Some pf \ (case (type_of_obj s (O_proc p), type_of_obj s (O_dir pf)) of + (Some tp, Some tpf) \ Some (type_transition tp tpf C_dir True) + | _ \ None) + | _ \ None) )" +| "type_of_obj (LinkHard p f f' # s) = (type_of_obj s) ((O_file f') := + (case (parent f') of + Some pf \ (case (type_of_obj s (O_proc p), type_of_obj s (O_dir pf)) of + (Some tp, Some tpf) \ Some (type_transition tp tpf C_file True) + | _ \ None) + | _ \ None) )" +(* +| "type_of_obj (Rename p f f' # s) = (\ obj. + case obj of + O_file f'' \ (if (f' \ f'') then type_of_obj s (O_file (file_before_rename f f' f'')) + else type_of_obj s (O_file f'')) + | O_dir f'' \ (if (f' \ f'') then type_of_obj s (O_file (file_before_rename f f' f'')) + else type_of_obj s (O_file f'')) + | _ \ type_of_obj s obj )" +*) +| "type_of_obj (CreateMsgq p q # s) = (type_of_obj s) ((O_msgq q) := + (case (type_of_obj s (O_proc p)) of + Some tp \ Some tp + | _ \ None) )" +| "type_of_obj (SendMsg p q m # s) = (type_of_obj s) ((O_msg q m) := + (case (type_of_obj s (O_proc p), type_of_obj s (O_msgq q)) of + (Some tp, Some tq) \ Some (type_transition tp tq C_msg False) + | _ \ None) )" +(* +| "type_of_obj (CreateShM p h # s) = (type_of_obj s) ((O_shm h) := + (case (type_of_obj s (O_proc p)) of + Some tp \ Some tp + | _ \ None) )" +*) +| "type_of_obj (CreateSock p af st fd inum # s) = (\ obj. + case obj of + O_fd p' fd' \ if (p' = p \ fd' = fd) then type_of_obj s (O_proc p) + else type_of_obj s obj + | O_tcp_sock (p', fd') \ if (p' = p \ fd' = fd \ st = STREAM) then type_of_obj s (O_proc p) + else type_of_obj s obj + | O_udp_sock (p', fd') \ if (p' = p \ fd' = fd \ st = DGRAM) then type_of_obj s (O_proc p) + else type_of_obj s obj + | _ \ type_of_obj s obj )" +| "type_of_obj (Accept p fd addr port fd' inum # s) = (\ obj. + case obj of + O_fd p' fd'' \ if (p' = p \ fd'' = fd') then type_of_obj s (O_proc p) + else type_of_obj s obj + | O_tcp_sock (p', fd'') \ if (p' = p \ fd'' = fd') then type_of_obj s (O_proc p) + else type_of_obj s obj + | _ \ type_of_obj s obj )" +| "type_of_obj (_ # s) = type_of_obj s" + +fun role_of_proc :: "t_state \ (t_process \ role_t option)" +where + "role_of_proc [] = init_role_of_proc" +| "role_of_proc (Clone p p' fds # s) = + (role_of_proc s) (p' := role_of_proc s p)" +| "role_of_proc (Execve p f fds # s) = (role_of_proc s) (p := + (case (role_of_proc s p, type_of_obj s (O_file f)) of + (Some r, Some ft) \ role_transition r ft + | _ \ None) )" +| "role_of_proc (_ # s) = role_of_proc s" + +fun role_of_obj :: "t_state \ t_object \ role_t option" +where + "role_of_obj s (O_proc p) = role_of_proc s p" +| "role_of_obj s _ = Some R_object" (* object_r *) + +definition sectxt_of_obj :: "t_state \ t_object \ security_context_t option" +where + "sectxt_of_obj s obj = ( + case (user_of_obj s obj, role_of_obj s obj, type_of_obj s obj) of + (Some u, Some r, Some t) \ Some (u, r, t) + | _ \ None)" + +fun init_role_of_obj :: "t_object \ role_t option" +where + "init_role_of_obj (O_proc p) = init_role_of_proc p" +| "init_role_of_obj _ = Some R_object" + +definition init_sectxt_of_obj :: "t_object \ security_context_t option" +where + "init_sectxt_of_obj obj \ + (case (init_user_of_obj obj, init_role_of_obj obj, init_type_of_obj obj) of + (Some u, Some r, Some t) \ Some (u,r,t) + | _ \ None)" + +definition sec_of_root :: "security_context_t" +where + "sec_of_root \ (case (init_user_of_obj (O_dir []), init_type_of_obj (O_dir [])) of + (Some u, Some t) \ (u, R_object, t))" + +(******* flask granting ******** + * involves three kinds of rules: + * 1. allow rule of types, allowed_rule_list_t, main + * 2. allow rule of roles, role_allow_rule_list_t, mainly for execve call + * 3. constrains, section 3.4.5, user-specifically defined. + * Passed all these 3, then grant Yes, else No + *) + +definition search_check_allp :: "security_context_t \ security_context_t set \ bool" +where + "search_check_allp pctxt sectxts = (\ sec \ sectxts. permission_check pctxt sec C_dir P_search)" + +definition search_check_file :: "security_context_t \ security_context_t \ bool" +where + "search_check_file sctxt fctxt \ permission_check sctxt fctxt C_file P_search" + +definition search_check_dir :: "security_context_t \ security_context_t \ bool" +where + "search_check_dir sctxt fctxt \ permission_check sctxt fctxt C_dir P_search" + +fun get_parentfs_ctxts :: "t_state \ t_file \ (security_context_t list) option" +where + "get_parentfs_ctxts s [] = + (case (sectxt_of_obj s (O_dir [])) of + Some ctxt \ Some [ctxt] + | _ \ None)" +| "get_parentfs_ctxts s (f#pf) = + (case (get_parentfs_ctxts s pf, sectxt_of_obj s (O_dir (f#pf))) of + (Some ctxts, Some ctxt) \ Some (ctxt#ctxts) + | _ \ None)" + +definition search_check_ctxt:: + "security_context_t \ security_context_t \ security_context_t set \ bool \ bool" +where + "search_check_ctxt pctxt fctxt asecs if_file \ ( + if if_file + then search_check_file pctxt fctxt \ search_check_allp pctxt asecs + else search_check_dir pctxt fctxt \ search_check_allp pctxt asecs )" + +fun search_check :: "t_state \ security_context_t \ t_file \ bool" +where + "search_check s pctxt [] = + (case (sectxt_of_obj s (O_dir [])) of + Some fctxt \ search_check_ctxt pctxt fctxt {} False + | _ \ False)" +| "search_check s pctxt (f#pf) = + (if (is_file s (f#pf)) + then (case (sectxt_of_obj s (O_file (f#pf)), get_parentfs_ctxts s pf) of + (Some fctxt, Some pfctxts) \ search_check_ctxt pctxt fctxt (set pfctxts) True + | _ \ False) + else (case (sectxt_of_obj s (O_dir (f#pf)), get_parentfs_ctxts s pf) of + (Some fctxt, Some pfctxts) \ search_check_ctxt pctxt fctxt (set pfctxts) False + | _ \ False))" + +(* this means we should prove every current fd has a security context! *) +definition sectxts_of_fds :: "t_state \ t_process \ t_fd set \ security_context_t set" +where + "sectxts_of_fds s p fds \ {ctxt. \ fd \ fds. sectxt_of_obj s (O_fd p fd) = Some ctxt}" + +definition inherit_fds_check_ctxt:: "security_context_t \ security_context_t set \ bool" +where + "inherit_fds_check_ctxt p fds \ (\ ctxt \ fds. permission_check p ctxt C_fd P_inherit)" + +definition inherit_fds_check :: "t_state \ security_context_t \ t_process \ t_fd set \ bool" +where + "inherit_fds_check s npsectxt p fds \ inherit_fds_check_ctxt npsectxt (sectxts_of_fds s p fds)" + +fun npctxt_execve :: "security_context_t \ security_context_t \ security_context_t option" +where + "npctxt_execve (pu,pr,pt) (fu,fr,ft) = + (case (role_transition pr ft) of + Some nr \ Some (pu, nr, type_transition pt ft C_process False) + | _ \ None)" + +fun nfctxt_create :: "security_context_t \ security_context_t \ security_class_t \ security_context_t" +where + "nfctxt_create (pu,pr,pt) (fu,fr,ft) cls = (pu, R_object, type_transition pt ft cls True)" + +fun grant_execve :: + "security_context_t \ security_context_t \ security_context_t \ bool" +where + "grant_execve (up,rp,tp) (uf,rf,tf) (up',nr,nt) = + (role_decl_check nr nt \ role_allow rp nr \ + permission_check (up,rp,tp) (uf,rf,tf) C_file P_execute \ + permission_check (up,nr,nt) (uf,rf,tf) C_file P_entrypoint \ + permission_check (up,rp,tp) (up,nr,nt) C_process P_transition \ + permission_check (up,nr,nt) (uf,rf,tf) C_process P_execute)" + +(* +definition sectxts_of_shms :: "t_state \ t_shm set \ security_context_t set" +where + "sectxts_of_shms s shms \ {ctxt. \ h \ shms. sectxt_of_obj s (O_shm h) = Some ctxt}" + +definition inherit_shms_check_ctxt:: "security_context_t \ security_context_t set \ bool" +where + "inherit_shms_check_ctxt p shms \ (\ ctxt \ shms. permission_check p ctxt C_shm P_inherit)" + +definition inherit_shms_check :: "t_state \ security_context_t \ t_shm set \ bool" +where + "inherit_shms_check s npsectxt shms \ inherit_shms_check_ctxt npsectxt (sectxts_of_shms s shms)" +*) + +fun perm_of_mflag :: "t_open_must_flag \ permission_t set" +where + "perm_of_mflag OF_RDONLY = {P_read}" +| "perm_of_mflag OF_WRONLY = {P_write}" +| "perm_of_mflag OF_RDWR = {P_read, P_write}" + +definition perm_of_oflag :: "t_open_option_flag set \ permission_t set" +where + "perm_of_oflag oflags \ + (case (OF_APPEND \ oflags, OF_CREAT \ oflags) of + (True, True) \ {P_append, P_create} + | (True, _ ) \ {P_append} + | (_ , True) \ {P_create} + | _ \ {})" + +definition perms_of_flags :: "t_open_flags \ permission_t set" +where + "perms_of_flags flags \ + (case flags of + (mflag,oflags) \ (perm_of_mflag mflag \ perm_of_oflag oflags))" + +(* +definition class_of_flags :: "t_open_flags \ security_class_t" +where + "class_of_flags flags \ + (case flags of + (mflag, oflags) \ if (OF_DIRECTORY \ oflags) then C_dir else C_file)" + +definition obj_of_flags :: "t_open_flags \ t_file \ t_object" +where + "obj_of_flags flags f \ + (case flags of + (mflag, oflags) \ if (OF_DIRECTORY \ oflags) then O_dir f else O_file f)" +*) + +definition oflags_check :: "t_open_flags \ security_context_t \ security_context_t \ bool" +where + "oflags_check flags pctxt fctxt \ + \ perm \ (perms_of_flags flags). permission_check pctxt fctxt C_file perm" + +fun grant :: "t_state \ t_event \ bool" +where + "grant s (Execve p f fds) = + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_file f)) of + (Some (up, rp, tp), Some (uf, rf, tf)) \ + (case (npctxt_execve (up,rp,tp) (uf,rf,tf)) of + Some (pu,nr,nt) \ ( + search_check s (up,rp,tp) f \ + grant_execve (up,rp,tp) (uf,rf,tf) (up,nr,nt) \ + inherit_fds_check s (up,nr,nt) p fds) + | _ \ False) + | _ \ False )" +| "grant s (Clone p p' fds) = + (case (sectxt_of_obj s (O_proc p)) of + Some (up, rp, tp) \ + (permission_check (up,rp,tp) (up,rp,tp) C_process P_fork \ + inherit_fds_check s (up,rp,tp) p fds) + | _ \ False )" (* \ inherit_shms_check s (up,rp,tp) shms *) +| "grant s (Kill p p') = + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_proc p')) of + (Some (su,sr,st), Some (tu,tr,tt)) \ + permission_check (su,sr,st) (tu,tr,tt) C_process P_sigkill + | _ \ False)" +| "grant s (Ptrace p p') = + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_proc p')) of + (Some (su,sr,st), Some (tu,tr,tt)) \ + permission_check (su,sr,st) (tu,tr,tt) C_process P_ptrace + | _ \ False)" +| "grant s (Exit p) = True" +| "grant s (Open p f flags fd None) = + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_file f)) of + (Some (pu,pr,pt), Some (fu,fr,ft)) \ + search_check s (pu,pr,pt) f \ + oflags_check flags (pu,pr,pt) (fu,fr,ft) \ + permission_check (pu,pr,pt) (pu,pr,pt) C_fd P_setattr + | _ \False)" +| "grant s (Open p f flags fd (Some inum)) = + (case (parent f) of + Some pf \ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of + (Some (pu,pr,pt), Some (pfu,pfr,pft)) \ + (search_check s (pu,pr,pt) pf \ + oflags_check flags (pu,pr,pt) (nfctxt_create (pu,pr,pt) (pfu,pfr,pft) C_file) \ + permission_check (pu,pr,pt) (pfu,pfr,pft) C_dir P_add_name \ + permission_check (pu,pr,pt) (pu,pr,pt) C_fd P_setattr) + | _ \ False) + | _ \ False)" +| "grant s (ReadFile p fd) = + (case (file_of_proc_fd s p fd) of + Some f \ + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_fd p fd), + sectxt_of_obj s (O_file f)) of + (Some (pu,pr,pt), Some (fdu,fdr,fdt), Some (fu,fr,ft)) \ + (permission_check (pu,pr,pt) (fdu,fdr,fdt) C_fd P_setattr \ + permission_check (pu,pr,pt) (fu,fr,ft) C_file P_read) + | _ \ False) + | _ \ False)" +| "grant s (WriteFile p fd) = + (case (file_of_proc_fd s p fd) of + Some f \ + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_fd p fd), + sectxt_of_obj s (O_file f)) of + (Some (pu,pr,pt), Some (fdu,fdr,fdt), Some (fu,fr,ft)) \ + (permission_check (pu,pr,pt) (fdu,fdr,fdt) C_fd P_setattr \ + permission_check (pu,pr,pt) (fu,fr,ft) C_file P_write) + | _ \ False) + | _ \ False)" +| "grant s (CloseFd p fd) = True" +| "grant s (UnLink p f) = + (case (parent f) of + Some pf \ + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf), + sectxt_of_obj s (O_file f)) of + (Some (pu,pr,pt), Some (du,dr,dt), Some (fu,fr,ft)) \ + (search_check s (pu,pr,pt) f \ + permission_check (pu,pr,pt) (fu,fr,ft) C_file P_unlink \ + permission_check (pu,pr,pt) (du,dr,dt) C_dir P_remove_name) + | _ \ False) + | _ \ False)" +| "grant s (Rmdir p f) = + (case (parent f) of + Some pf \ + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf), + sectxt_of_obj s (O_dir f)) of + (Some (pu,pr,pt), Some (du,dr,dt), Some (fu,fr,ft)) \ + (search_check s (pu,pr,pt) f \ + permission_check (pu,pr,pt) (fu,fr,ft) C_dir P_rmdir \ + permission_check (pu,pr,pt) (du,dr,dt) C_dir P_remove_name) + | _ \ False) + | _ \ False)" +| "grant s (Mkdir p f inum) = + (case (parent f) of + Some pf \ + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of + (Some (pu,pr,pt), Some (du,dr,dt)) \ + (search_check s (pu,pr,pt) f \ + permission_check (pu,pr,pt) (nfctxt_create (pu,pr,pt) (du,dr,dt) C_dir) C_dir P_create \ + permission_check (pu,pr,pt) (du,dr,dt) C_dir P_add_name) + | _ \ False) + | _ \ False)" +| "grant s (LinkHard p f f') = + (case (parent f') of + Some pf \ + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf), + sectxt_of_obj s (O_file f)) of + (Some (pu,pr,pt), Some (du,dr,dt), Some (fu,fr,ft)) \ + (search_check s (pu,pr,pt) f \ search_check s (pu,pr,pt) pf \ + permission_check (pu,pr,pt) (fu,fr,ft) C_file P_link \ + permission_check (pu,pr,pt) (du,dr,dt) C_dir P_add_name) + | _ \ False) + | _ \ False)" +| "grant s (Truncate p f len) = + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_file f)) of + (Some (pu,pr,pt), Some (fu,fr,ft)) \ + (search_check s (pu,pr,pt) f \ + permission_check (pu,pr,pt) (fu,fr,ft) C_file P_setattr) + | _ \ False)" +(* +| "grant s (Rename p f f') = + (case (parent f, parent f') of + (Some pf, Some pf') \ + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf), + sectxt_of_obj s (O_file f), sectxt_of_obj s (O_dir pf')) of + (Some (pu,pr,pt), Some (pfu,pfr,pft), Some (fu,fr,ft), + Some (pfu',pfr',pft')) \ + (search_check s (pu,pr,pt) f \ + permission_check (pu,pr,pt) (pfu,pfr,pft) C_dir P_remove_name \ + (if (is_file s f) + then permission_check (pu,pr,pt) (fu,fr,ft) C_file P_rename + else permission_check (pu,pr,pt) (fu,fr,ft) C_dir P_reparent) \ + search_check s (pu,pr,pt) pf' \ + permission_check (pu,pr,pt) (pfu',pfr',pft') C_dir P_add_name) + | _ \ False) + | _ \ False)" +*) +| "grant s (CreateMsgq p q) = + (case (sectxt_of_obj s (O_proc p)) of + Some (pu,pr,pt) \ + (permission_check (pu,pr,pt) (pu,pr,pt) C_msgq P_associate \ + permission_check (pu,pr,pt) (pu,pr,pt) C_msgq P_create) + | _ \ False)" +| "grant s (SendMsg p q m) = + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_msgq q)) of + (Some (pu,pr,pt), Some (qu,qr,qt)) \ + (permission_check (pu,pr,pt) (qu,qr,qt) C_msgq P_enqueue \ + permission_check (pu,pr,pt) (qu,qr,qt) C_msgq P_write \ + permission_check (pu,pr,pt) (pu,pr,pt) C_msg P_send) + | _ \ False)" +| "grant s (RecvMsg p q m) = + (case (sectxt_of_obj s (O_proc p),sectxt_of_obj s (O_msgq q),sectxt_of_obj s (O_msg q m)) of + (Some (pu,pr,pt), Some (qu,qr,qt), Some (mu,mr,mt)) \ + permission_check (pu,pr,pt) (qu,qr,qt) C_msgq P_read \ + permission_check (pu,pr,pt) (mu,mr,mt) C_msg P_receive + | _ \ False)" +| "grant s (RemoveMsgq p q) = + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_msgq q)) of + (Some (pu,pr,pt), Some (qu,qr,qt)) \ + permission_check (pu,pr,pt) (qu,qr,qt) C_msgq P_destroy + | _ \ False)" +(* +| "grant s (CreateShM p h) = + (case (sectxt_of_obj s (O_proc p)) of + Some (pu,pr,pt) \ + (permission_check (pu,pr,pt) (pu,pr,pt) C_shm P_associate \ + permission_check (pu,pr,pt) (pu,pr,pt) C_shm P_create) + | _ \ False)" +| "grant s (Attach p h flag) = + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_shm h)) of + (Some (pu,pr,pt), Some (hu,hr,ht)) \ + if flag = SHM_RDONLY + then permission_check (pu,pr,pt) (hu,hr,ht) C_shm P_read + else permission_check (pu,pr,pt) (hu,hr,ht) C_shm P_read \ + permission_check (pu,pr,pt) (hu,hr,ht) C_shm P_write + | _ \ False)" +| "grant s (Detach p h) = True" (*?*) +| "grant s (DeleteShM p h) = + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_shm h)) of + (Some (pu,pr,pt), Some (hu,hr,ht)) \ + permission_check (pu,pr,pt) (hu,hr,ht) C_shm P_destroy + | _ \ False)" +*) + +| "grant s _ = False" (* socket event is currently not in consideration *) + +inductive valid :: "t_state \ bool" +where + "valid []" +| "\valid s; os_grant s e; grant s e\ \ valid (e # s)" + +(* tobj: object that can be tainted *) +fun tobj_in_init :: "t_object \ bool" +where + "tobj_in_init (O_proc p) = (p \ init_procs)" +| "tobj_in_init (O_file f) = (is_init_file f)" +(* directory is not taintable +| "tobj_in_init (O_dir f) = (f \ init_files)" + +| "tobj_in_init (O_node n) = (n \ init_nodes)" +| "tobj_in_init (O_msg q m) = (m \ set (init_msgs_of_queue q) \ q \ init_msgqs)" +*) +| "tobj_in_init _ = False" (* other kind of obj cannot be tainted *) + +(* no use +fun no_del_event:: "t_event list \ bool" +where + "no_del_event [] = True" +| "no_del_event (Kill p p' # \) = False" +| "no_del_event (CloseFd p fd # \) = False" +| "no_del_event (UnLink p f # \) = False" +| "no_del_event (Rmdir p f # \) = False" +(* +| "no_del_event (Rename p f f' # \) = False" +*) +| "no_del_event (RemoveMsgq p q # \) = False" +| "no_del_event (RecvMsg p q m # \) = False" +| "no_del_event (_ # \) = no_del_event \" +*) + +end + +locale tainting = flask + + +fixes seeds :: "t_object set" +assumes + seeds_appropriate: "\ obj. obj \ seeds \ appropriate obj" + and same_inode_in_seeds: "\ f f'. \O_file f \ seeds; has_same_inode [] f f'\ \ O_file f' \ seeds" +(* + and flow_shm_in_seeds: "\ p p'. \O_proc p \ seeds; info_flow_shm [] p p'\ \ O_proc p' \ seeds" +*) +begin + +(* +inductive tainted :: "t_object \ t_state \ bool" ("_ \ tainted _" [100, 100] 100) +where + t_init: "obj \ seeds \ obj \ tainted []" +| t_clone: "\O_proc p \ tainted s; valid (Clone p p' fds shms # s)\ + \ O_proc p' \ tainted (Clone p p' fds # s)" +| t_execve: "\O_file f \ tainted s; valid (Execve p f fds # s)\ + \ O_proc p \ tainted (Execve p f fds # s)" +| t_ptrace: "\O_proc p \ tainted s; valid (Ptrace p p' # s); info_flow_shm s p' p''\ + \ O_proc p'' \ tainted (Ptrace p p' # s)" +| t_ptrace':"\O_proc p' \ tainted s; valid (Ptrace p p' # s); info_flow_shm s p p''\ + \ O_proc p'' \ tainted (Ptrace p p' # s)" +| t_cfile: "\O_proc p \ tainted s; valid (Open p f flags fd (Some inum) # s)\ + \ O_file f \ tainted (Open p f flags fd (Some inum) # s)" +| t_read: "\O_file f \ tainted s; valid (ReadFile p fd # s); + file_of_proc_fd s p fd = Some f; info_flow_shm s p p'\ + \ O_proc p' \ tainted (ReadFile p fd # s)" +| t_write: "\O_proc p \ tainted s; valid (WriteFile p fd # s); + file_of_proc_fd s p fd = Some f; has_same_inode s f f'\ + \ O_file f' \ tainted (WriteFile p fd # s)" +(* directory is not tainted +| t_mkdir: "\O_proc p \ tainted s; valid (Mkdir p f inum # s)\ + \ O_dir f \ tainted (Mkdir p f inum # s)" + *) +| t_link: "\O_file f \ tainted s; valid (LinkHard p f f' # s)\ + \ O_file f' \ tainted (LinkHard p f f' # s)" +| t_trunc: "\O_proc p \ tainted s; valid (Truncate p f len # s); len > 0; has_same_inode s f f'\ + \ O_file f' \ tainted (Truncate p f len # s)" +(* +| t_rename: "\O_file f'' \ tainted s; valid (Rename p f f' # s);f \ f''\ + \ O_file (file_after_rename f f' f'') \ tainted (Rename p f f' # s)" +| t_rename':"\O_dir f'' \ tainted s; valid (Rename p f f' # s);f \ f''\ + \ O_dir (file_after_rename f f' f'') \ tainted (Rename p f f' # s)" +*) +| t_attach1:"\O_proc p \ tainted s; valid (Attach p h SHM_RDWR # s); (p', flag') \ procs_of_shm s h; info_flow_shm s p' p''\ + \ O_proc p'' \ tainted (Attach p h SHM_RDWR # s)" +| t_attach2:"\O_proc p' \ tainted s; valid (Attach p h flag # s); (p', SHM_RDWR) \ procs_of_shm s h; info_flow_shm s p p''\ + \ O_proc p'' \ tainted (Attach p h flag # s)" +| t_sendmsg:"\O_proc p \ tainted s; valid (SendMsg p q m # s)\ + \ O_msg q m \ tainted (SendMsg p q m # s)" +| t_recvmsg:"\O_msg q m \ tainted s; valid (RecvMsg p q m # s); info_flow_shm s p p'\ + \ O_proc p' \ tainted (RecvMsg p q m # s)" +| t_remain: "\obj \ tainted s; valid (e # s); alive (e # s) obj\ + \ obj \ tainted (e # s)" +*) + +fun tainted :: "t_state \ t_object set" +where + "tainted [] = seeds" +| "tainted (Clone p p' fds # s) = + (if (O_proc p) \ tainted s then tainted s \ {O_proc p'} else tainted s)" +| "tainted (Execve p f fds # s) = + (if (O_file f) \ tainted s then tainted s \ {O_proc p} else tainted s)" +| "tainted (Kill p p' # s) = tainted s - {O_proc p'}" +| "tainted (Ptrace p p' # s) = + (if (O_proc p \ tainted s) + then tainted s \ {O_proc p'} + else if (O_proc p' \ tainted s) + then tainted s \ {O_proc p} + else tainted s)" +| "tainted (Exit p # s) = tainted s - {O_proc p}" +| "tainted (Open p f flags fd opt # s) = + (case opt of + Some inum \ (if (O_proc p) \ tainted s + then tainted s \ {O_file f} + else tainted s) + | _ \ tainted s)" +| "tainted (ReadFile p fd # s) = + (case (file_of_proc_fd s p fd) of + Some f \ if (O_file f) \ tainted s + then tainted s \ {O_proc p} + else tainted s + | None \ tainted s)" +| "tainted (WriteFile p fd # s) = + (case (file_of_proc_fd s p fd) of + Some f \ if (O_proc p) \ tainted s + then tainted s \ {O_file f' | f'. f' \ same_inode_files s f} + else tainted s + | None \ tainted s)" +| "tainted (CloseFd p fd # s) = + (case (file_of_proc_fd s p fd) of + Some f \ ( if ((proc_fd_of_file s f = {(p,fd)}) \ (f \ files_hung_by_del s)) + then tainted s - {O_file f} else tainted s ) + | _ \ tainted s)" +| "tainted (UnLink p f # s) = + (if (proc_fd_of_file s f = {}) then tainted s - {O_file f} else tainted s)" +| "tainted (LinkHard p f f' # s) = + (if (O_file f \ tainted s) then tainted s \ {O_file f'} else tainted s)" +| "tainted (Truncate p f len # s) = + (if (len > 0 \ O_proc p \ tainted s) + then tainted s \ {O_file f' | f'. f' \ same_inode_files s f} + else tainted s)" +| "tainted (SendMsg p q m # s) = + (if (O_proc p \ tainted s) then tainted s \ {O_msg q m} else tainted s)" +| "tainted (RecvMsg p q m # s) = + (if (O_msg q m \ tainted s) + then (tainted s \ {O_proc p}) - {O_msg q m} + else tainted s)" +| "tainted (RemoveMsgq p q # s) = tainted s - {O_msg q m| m. O_msg q m \ tainted s}" +| "tainted (e # s) = tainted s" + +(* reasonable tainted object, fd and sock and msgq and msg is excluded + * this is different from tainted, which already excluded some of them, + * tainted not excluded msg, which does not have the meaning of "tainteable", but + * but acting as a media to "pass" the virus. We normally will not say that + * a message is tableable or not. + *) +fun appropriate :: "t_object \ bool" +where + "appropriate (O_proc p) = (p \ init_procs)" +| "appropriate (O_file f) = (is_init_file f)" +| "appropriate (O_dir f) = False" +| "appropriate (O_fd p fd) = False" +| "appropriate (O_node n) = False" (* cause socket is temperary not considered *) +| "appropriate (O_tcp_sock k) = False" +| "appropriate (O_udp_sock k) = False" +(* +| "appropriate (O_shm h) = False" +*) +| "appropriate (O_msgq q) = False" +| "appropriate (O_msg q m) = False" + +definition taintable:: "t_object \ bool" +where + "taintable obj \ init_alive obj \ appropriate obj \ (\ s. obj \ tainted s)" + +definition undeletable :: "t_object \ bool" +where + "undeletable obj \ init_alive obj \ \ (\ s. valid s \ deleted obj s)" + +fun exited :: "t_object \ t_state \ bool" +where + "exited (O_proc p) s = (Exit p \ set s)" +| "exited _ s = False" + +end + +end \ No newline at end of file