# HG changeset patch # User chunhan # Date 1367565621 -3600 # Node ID 7d9c0ed02b563285b1d9d4113dc18d13fde9e237 # Parent 34d01e9a772e31908e7621edf78aa094e7c0effc thy files diff -r 34d01e9a772e -r 7d9c0ed02b56 Alive_prop.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Alive_prop.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,299 @@ +theory Alive_prop +imports Main Flask_type Flask Current_files_prop Current_sockets_prop Init_prop +begin + +context flask begin + +lemma distinct_queue_msgs: + "\q \ current_msgqs s; valid s\ \ distinct (msgs_of_queue s q)" +apply (induct s) +apply (simp add:init_msgs_distinct) +apply (frule vd_cons, frule vt_grant_os, case_tac a) +apply auto +apply (case_tac "msgs_of_queue s q", simp+) +done + +lemma received_msg_notin: + "\hd (msgs_of_queue s q) \ set (tl (msgs_of_queue s q)); q \ current_msgqs s; valid s\ \ False" +apply (drule distinct_queue_msgs, simp) +apply (case_tac "msgs_of_queue s q", auto) +done + +lemma other_msg_remains: + "\m \ hd (msgs_of_queue s q); q \ current_msgqs s; valid s\ + \ (m \ set (tl (msgs_of_queue s q))) = (m \ set (msgs_of_queue s q))" +apply (drule distinct_queue_msgs, simp) +apply (case_tac "msgs_of_queue s q", auto) +done + +lemma is_file_in_current: + "is_file s f \ f \ current_files s" +by (auto simp:is_file_def current_files_def split:option.splits) + +lemma is_dir_in_current: + "is_dir s f \ f \ current_files s" +by (auto simp:is_dir_def current_files_def split:option.splits) + +lemma is_tcp_in_current: + "is_tcp_sock \ s \ s \ current_sockets \" +by (auto simp:is_tcp_sock_def current_sockets_def split:option.splits) + +lemma is_udp_in_current: + "is_udp_sock \ s \ s \ current_sockets \" +by (auto simp:is_udp_sock_def current_sockets_def split:option.splits) + +(************ alive simpset **************) + +lemma alive_open: + "valid (Open p f flag fd opt # s) \ alive (Open p f flag fd opt # s) = ( + \ obj. case obj of + O_fd p' fd' \ if (p' = p \ fd' = fd) then True + else alive s obj + | O_file f' \ (if (opt = None) then alive s obj + else if (f = f') then True + else alive s obj) + | _ \ alive s obj)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps + dest:is_dir_in_current split:if_splits option.splits) +done + +lemma alive_execve: + "valid (Execve p f fds # s) \ alive (Execve p f fds # s) = ( + \ obj. case obj of + O_fd p' fd \ (if (p = p' \ fd \ fds) then alive s (O_fd p fd) + else if (p = p') then False + else alive s (O_fd p' fd)) + | O_tcp_sock (p', fd) \ (if (p = p' \ fd \ fds) then alive s (O_tcp_sock (p, fd)) + else if (p = p') then False + else alive s (O_tcp_sock (p', fd))) + | O_udp_sock (p', fd) \ (if (p = p' \ fd \ fds) then alive s (O_udp_sock (p, fd)) + else if (p = p') then False + else alive s (O_udp_sock (p', fd))) + | _ \ alive s obj )" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps + dest:is_dir_in_current split:if_splits option.splits) +done + +lemma alive_clone: + "valid (Clone p p' fds shms # s) \ alive (Clone p p' fds shms # s) = ( + \ obj. case obj of + O_proc p'' \ if (p'' = p') then True else alive s obj + | O_fd p'' fd \ if (p'' = p' \ fd \ fds) then True + else if (p'' = p') then False + else alive s obj + | O_tcp_sock (p'', fd) \ (if (p'' = p' \ fd \ fds) then alive s (O_tcp_sock (p, fd)) + else if (p'' = p') then False + else alive s (O_tcp_sock (p'', fd))) + | O_udp_sock (p'', fd) \ (if (p'' = p' \ fd \ fds) then alive s (O_udp_sock (p, fd)) + else if (p'' = p') then False + else alive s (O_udp_sock (p'', fd))) + | _ \ alive s obj )" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps + intro:is_tcp_in_current is_udp_in_current + dest:is_dir_in_current split:if_splits option.splits) +done + +lemma alive_kill: + "valid (Kill p p' # s) \ alive (Kill p p' # s) = ( + \ obj. case obj of + O_proc p'' \ if (p'' = p') then False else alive s obj + | O_fd p'' fd \ if (p'' = p') then False else alive s obj + | O_tcp_sock (p'', fd) \ if (p'' = p') then False else alive s obj + | O_udp_sock (p'', fd) \ if (p'' = p') then False else alive s obj + | _ \ alive s obj)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps + intro:is_tcp_in_current is_udp_in_current + dest:is_dir_in_current split:if_splits option.splits) +done + +lemma alive_exit: + "valid (Exit p' # s) \ alive (Exit p' # s) = ( + \ obj. case obj of + O_proc p'' \ if (p'' = p') then False else alive s obj + | O_fd p'' fd \ if (p'' = p') then False else alive s obj + | O_tcp_sock (p'', fd) \ if (p'' = p') then False else alive s obj + | O_udp_sock (p'', fd) \ if (p'' = p') then False else alive s obj + | _ \ alive s obj)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps + intro:is_tcp_in_current is_udp_in_current + dest:is_dir_in_current split:if_splits option.splits) +done + +lemma alive_closefd: + "valid (CloseFd p fd # s) \ alive (CloseFd p fd # s) = ( + \ obj. case obj of + O_fd p' fd' \ if (p' = p \ fd' = fd) then False else alive s obj + | O_tcp_sock (p', fd') \ if (p' = p \ fd' = fd) then False else alive s obj + | O_udp_sock (p', fd') \ if (p' = p \ fd' = fd) then False else alive s obj + | O_file f \ (case (file_of_proc_fd s p fd) of + Some f' \ (if (f = f' \ proc_fd_of_file s f = {(p, fd)} \ f \ files_hung_by_del s) + then False else alive s obj) + | _ \ alive s obj) + | _ \ alive s obj)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps + intro:is_tcp_in_current is_udp_in_current + dest:is_dir_in_current file_of_pfd_is_file' split:if_splits option.splits) +done + +lemma alive_unlink: + "valid (UnLink p f # s) \ alive (UnLink p f # s) = ( + \ obj. case obj of + O_file f' \ if (f' = f \ proc_fd_of_file s f = {}) then False else alive s obj + | _ \ alive s obj)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps + intro:is_tcp_in_current is_udp_in_current + dest:is_dir_in_current file_of_pfd_is_file' file_dir_conflict + split:if_splits option.splits) +done + +lemma alive_rmdir: + "valid (Rmdir p d # s) \ alive (Rmdir p d # s) = (alive s) (O_dir d := False)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def + intro:is_tcp_in_current is_udp_in_current + dest:is_dir_in_current file_of_pfd_is_file' file_dir_conflict + split:if_splits option.splits) +done + +lemma alive_mkdir: + "valid (Mkdir p d inum # s) \ alive (Mkdir p d inum # s) = (alive s) (O_dir d := True)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def + intro:is_tcp_in_current is_udp_in_current + dest:is_dir_in_current file_of_pfd_is_file' is_file_in_current + split:if_splits option.splits) +done + +lemma alive_linkhard: + "valid (LinkHard p f f' # s) \ alive (LinkHard p f f' # s) = (alive s) (O_file f' := True)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def + intro:is_tcp_in_current is_udp_in_current + dest:is_dir_in_current file_of_pfd_is_file' is_file_in_current + split:if_splits option.splits) +done + +lemma alive_createmsgq: + "valid (CreateMsgq p q # s) \ alive (CreateMsgq p q # s) = (alive s) (O_msgq q := True)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def) +done + +lemma alive_sendmsg: + "valid (SendMsg p q m # s) \ alive (SendMsg p q m # s) = (alive s) (O_msg q m := True)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def) +done + +lemma alive_recvmsg: + "valid (RecvMsg p q m # s) \ alive (RecvMsg p q m # s) = (alive s) (O_msg q m := False)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def other_msg_remains + dest:received_msg_notin) +done + +lemma alive_removemsgq: + "valid (RemoveMsgq p q # s) \ alive (RemoveMsgq p q # s) = ( + \ obj. case obj of + O_msgq q' \ if (q' = q) then False else alive s obj + | O_msg q' m \ if (q' = q) then False else alive s obj + | _ \ alive s obj)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def) +done + +lemma alive_createshm: + "valid (CreateShM p h # s) \ alive (CreateShM p h # s) = (alive s) (O_shm h := True)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def) +done + +lemma alive_deleteshm: + "valid (DeleteShM p h # s) \ alive (DeleteShM p h # s) = (alive s) (O_shm h := False)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def) +done + +lemma alive_createsock: + "valid (CreateSock p af st fd inum # s) \ alive (CreateSock p af st fd inum # s) = ( + \ obj. case obj of + O_fd p' fd' \ if (p' = p \ fd' = fd) then True else alive s obj + | O_tcp_sock (p', fd') \ if (p' = p \ fd' = fd \ st = STREAM) then True else alive s obj + | O_udp_sock (p', fd') \ if (p' = p \ fd' = fd \ st = DGRAM) then True else alive s obj + | _ \ alive s obj)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def + intro:is_tcp_in_current is_udp_in_current split:t_socket_type.splits) +done + +lemma alive_accept: + "valid (Accept p fd addr port fd' inum # s) \ alive (Accept p fd addr port fd' inum # s) = ( + \ obj. case obj of + O_fd p' fd'' \ if (p' = p \ fd'' = fd') then True else alive s obj + | O_tcp_sock (p', fd'') \ if (p' = p \ fd'' = fd') then True else alive s obj + | _ \ alive s obj)" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def + intro:is_tcp_in_current is_udp_in_current split:t_socket_type.splits) +done + +lemma alive_other: + "\valid (e # s); + \ p f flag fd opt. e \ Open p f flag fd opt; + \ p f fds. e \ Execve p f fds; + \ p p' fds shms. e \ Clone p p' fds shms; + \ p p'. e \ Kill p p'; + \ p. e \ Exit p; + \ p fd. e \ CloseFd p fd; + \ p f. e \ UnLink p f; + \ p d. e \ Rmdir p d; + \ p d inum. e \ Mkdir p d inum; + \ p f f'. e \ LinkHard p f f'; + \ p q. e \ CreateMsgq p q; + \ p q m. e \ SendMsg p q m; + \ p q m. e \ RecvMsg p q m; + \ p q. e \ RemoveMsgq p q; + \ p h. e \ CreateShM p h; + \ p h. e \ DeleteShM p h; + \ p af st fd inum. e \ CreateSock p af st fd inum; + \ p fd addr port fd' inum. e \ Accept p fd addr port fd' inum\ + \ alive (e # s) = alive s" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x, case_tac [!] e) +apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps + is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def + intro:is_tcp_in_current is_udp_in_current split:t_socket_type.splits) +done + +lemmas alive_simps = alive_open alive_execve alive_clone alive_kill alive_exit alive_closefd alive_unlink + alive_rmdir alive_mkdir alive_linkhard alive_createmsgq alive_removemsgq alive_createshm alive_deleteshm + alive_createsock alive_accept alive_other alive_sendmsg alive_recvmsg + + +end + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Co2sobj_prop.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Co2sobj_prop.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,494 @@ +(*<*) +theory Co2sobj_prop +imports Main Flask Flask_type Static Static_type Sectxt_prop Init_prop Current_files_prop Current_sockets_prop +begin +(*<*) + +context tainting_s begin + +(************ simpset for cf2sfile ***************) + +declare get_parentfs_ctxts.simps [simp del] + +lemma cf2sfile_open_none: + "valid (Open p f flag fd None # s) \ cf2sfile (Open p f flag fd None # s) f b = cf2sfile s f b" +apply (frule vd_cons, frule vt_grant_os) +apply (frule noroot_events) +by (auto split:if_splits option.splits dest!:current_has_sec' + simp:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps + get_parentfs_ctxts_simps) + +lemma cf2sfile_open_some1: + "\valid (Open p f flag fd (Some inum) # s); f' \ current_files s\ + \ cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'" +apply (frule vd_cons, frule vt_grant_os, frule noroot_events, rule ext) +apply (induct f') +apply (auto split:if_splits option.splits dest!:current_has_sec' + simp:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps + get_parentfs_ctxts_simps) +done + +lemma cf2sfile_open_some2: + "\valid (Open p f flag fd (Some inum) # s); is_file s f'\ + \ cf2sfile (Open p f flag fd (Some inum) # s) f' True = cf2sfile s f' True" +apply (frule vd_cons, drule is_file_in_current) +by (simp add:cf2sfile_open_some1) + +lemma cf2sfile_open_some3: + "\valid (Open p f flag fd (Some inum) # s); is_dir s f'\ + \ cf2sfile (Open p f flag fd (Some inum) # s) f' False = cf2sfile s f' False" +apply (frule vd_cons, drule is_dir_in_current) +by (simp add:cf2sfile_open_some1) + +lemma cf2sfile_open_some: + "valid (Open p f flag fd (Some inum) # s) \ cf2sfile (Open p f flag fd (Some inum) # s) f True = ( + case (parent f) of + Some pf \ (case (sectxt_of_obj (Open p f flag fd (Some inum) # s) (O_file f), sectxt_of_obj s (O_dir pf), + get_parentfs_ctxts s pf) of + (Some sec, Some psec, Some asecs) \ Some (Created, sec, Some psec, set asecs) + | _ \ None) + | None \ None)" +apply (frule vd_cons, frule vt_grant_os, frule noroot_events) +apply (auto split:if_splits option.splits dest!:current_has_sec' + simp:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps + get_parentfs_ctxts_simps) + + + +lemma cf2sfile_open: + "valid (Open p f flag fd opt # s) \ cf2sfile (Open p f flag fd opt # s) = (\ f' b. + if (opt = None) then cf2sfile s f' b + else if (f' = f \ b = True) + then (case (parent f) of + Some pf \ (case (sectxt_of_obj s (O_file f), sectxt_of_obj s (O_dir pf), + get_parentfs_ctxts s pf) of + (Some sec, Some psec, Some asecs) \ + Some (if (\ deleted (O_file f) s \ f \ init_files) + then Init f else Created, sec, Some psec, set asecs) + | _ \ None) + | None \ None) + else cf2sfile s f' b)" +apply (frule vd_cons, frule vt_grant_os, rule ext, rule ext) +apply (auto split:if_splits option.splits + simp:sectxt_of_obj_simps) + + + + + + + + + + + + +lemma cf2sfile_keep_path: "\f \ f'; \ \ vt rc_cs\ \ cf2sfile \ f \ cf2sfile \ f'" +apply (induct f', simp add:no_junior_def) +apply (case_tac "f = a # f'", simp add:cf2sfile.simps) +apply (drule no_junior_noteq, simp, simp add:cf2sfile.simps) +done + +lemma ckp'_aux: "\ f \ a # f' \ \ f \ f'" +by (auto simp:no_junior_def) + +lemma cf2sfile_keep_path': "\\ f \ f'; \ \ vt rc_cs; f \ current_files \; f' \ current_files \\ \ \ cf2sfile \ f \ cf2sfile \ f'" +apply (induct f', simp add:no_junior_def cf2sfile.simps, rule notI, drule cf2sfile_root_file, simp) +apply (case_tac "f = a # f'", simp add:cf2sfile.simps) +apply (rule notI) +apply (frule ckp'_aux, simp, frule parentf_in_current', simp+) +apply (case_tac "cf2sfile \ f = cf2sfile \ (a # f')", drule cf2sfile_inj, simp+) +apply (simp add:cf2sfile.simps) +by (drule no_junior_noteq, simp+) + +lemma cf2sfile_keep_path'': "\cf2sfile \ f \ cf2sfile \ f'; \ \ vt rc_cs; f \ current_files \; f' \ current_files \\ \ f \ f'" +using cf2sfile_keep_path' +by (auto) + +lemma cf2sfile_open_some': "\f \ current_files \; Open p f' flags fd (Some inum) # \ \ vt rc_cs\ \ cf2sfile (Open p f' flags fd (Some inum) # \) f = cf2sfile \ f" +apply (frule vt_cons') +apply (drule vt_grant_os) +apply (induct f) +apply (simp add:cf2sfile.simps)+ +apply (frule parentf_in_current', simp) +apply (auto simp:os_grant.simps index_of_file.simps split:option.splits) +done + +lemma cf2sfile_open_some: "\Open p f flags fd (Some inum) # \ \ vt rc_cs; parent f = Some pf\ + \ cf2sfile (Open p f flags fd (Some inum) # \) f = (SCrea (Suc (length \)) # (cf2sfile \ pf))" +apply (case_tac f, simp) +apply (frule vt_grant_os) +apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_open_some') +done + +lemma cf2sfile_open_none: "cf2sfile (Open p f flags fd None # \) f' = cf2sfile \ f'" +apply (induct f') +by (simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits)+ + +lemma cf2sfile_open: "\Open p f flags fd opt # \ \ vt rc_cs; f' \ current_files (Open p f flags fd opt # \)\ \ + cf2sfile (Open p f flags fd opt # \) f' = ( + if (opt = None) then cf2sfile \ f' + else if (f' = f) then (case (parent f) of + Some pf \ (SCrea (Suc (length \)) # (cf2sfile \ pf)) + | _ \ []) + else cf2sfile \ f' )" +apply (frule vt_grant_os) +by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_open_none cf2sfile_open_some cf2sfile_open_some' split:if_splits option.splits) + +lemma cf2sfile_mkdir_some': "\Mkdir p f' inum # \ \ vt rc_cs; f \ current_files \\ \ cf2sfile (Mkdir p f' inum # \) f = cf2sfile \ f" +apply (frule vt_cons', drule vt_grant_os) +apply (induct f, (simp add:cf2sfile.simps)+) +apply (frule parentf_in_current', simp) +apply (auto simp:os_grant.simps index_of_file.simps split:option.splits) +done + +lemma cf2sfile_mkdir_some: "\Mkdir p f inum # \ \ vt rc_cs; parent f = Some pf\ + \ cf2sfile (Mkdir p f inum # \) f = (SCrea (Suc (length \)) # (cf2sfile \ pf))" +apply (case_tac f, simp) +apply (frule vt_grant_os) +apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_mkdir_some') +done + +lemma cf2sfile_mkdir: "\Mkdir p f inum # \ \ vt rc_cs; f' \ current_files (Mkdir p f inum # \)\ \ + cf2sfile (Mkdir p f inum # \) f' = ( + if (f' = f) then (case (parent f) of + Some pf \ (SCrea (Suc (length \)) # (cf2sfile \ pf)) + | _ \ []) + else cf2sfile \ f' ) " +apply (frule vt_grant_os) +by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_mkdir_some cf2sfile_mkdir_some' split:if_splits option.splits) + +lemma cf2sfile_linkhard_none: "\LinkHard p f\<^isub>1 f\<^isub>2 # \ \ vt rc_cs; f \ current_files \\ \ cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \) f = cf2sfile \ f" +apply (frule vt_cons') +apply (drule vt_grant_os) +apply (induct f) +apply (simp add:cf2sfile.simps)+ +apply (frule parentf_in_current', simp) +apply (auto simp:os_grant.simps index_of_file.simps split:option.splits) +done + +lemma cf2sfile_linkhard_some: + "\LinkHard p f\<^isub>1 f\<^isub>2 # \ \ vt rc_cs; parent f\<^isub>2 = Some pf\<^isub>2\ \ cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \) f\<^isub>2 = (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>2))" +apply (case_tac f\<^isub>2, simp) +apply (frule vt_grant_os) +apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_linkhard_none) +done + +lemma cf2sfile_linkhard: "\LinkHard p f\<^isub>1 f\<^isub>2 # \ \ vt rc_cs; f \ current_files (LinkHard p f\<^isub>1 f\<^isub>2 # \)\ \ + cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \) f = ( + if (f = f\<^isub>2) then (case (parent f\<^isub>2) of + Some pf\<^isub>2 \ SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>2) + | _ \ []) + else cf2sfile \ f )" +apply (frule vt_grant_os) +by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_linkhard_none cf2sfile_linkhard_some split:if_splits option.splits) + + +lemma no_junior_aux: "\ f\<^isub>2 \ a # f \ \ f\<^isub>2 \ f" +by (auto simp:no_junior_def) + +lemma cf2sfile_rename_aux: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; f \ current_files \\ \ f \ f\<^isub>3 \ \ f\<^isub>3 \ f" +apply (frule vt_grant_os, simp add:os_grant.simps, (erule exE|erule conjE)+) +apply (rule conjI, rule notI, simp) +apply (rule notI, drule vt_cons', simp add:ancenf_in_current) +done + +lemma cf2sfile_rename'1: + "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; f \ current_files \; \ (f\<^isub>2 \ f)\ \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = cf2sfile \ f" +apply (frule vt_cons',frule vt_grant_os, induct f) +apply (simp add:cf2sfile.simps) +apply (frule parentf_in_current', simp, simp) +apply (frule no_junior_aux, simp) +apply (simp add:os_grant.simps, (erule exE|erule conjE)+) +apply (drule cf2sfile_rename_aux, simp, erule conjE) +apply (auto simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits) +done + +lemma cf2sfile_rename'2: + "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; \ (f\<^isub>3 \ f)\ \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = cf2sfile \ f" +apply (frule vt_cons', induct f) +by (auto simp add:index_of_file.simps cf2sfile.simps no_junior_def split:option.splits nat.splits) + +lemma cf2sfile_rename1: + "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; parent f\<^isub>3 = Some pf\<^isub>3\ \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f\<^isub>3 = SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)" +apply (case_tac f\<^isub>3, simp add:cf2sfile.simps) +apply (auto dest!:cf2sfile_rename'2 simp add:no_junior_def cf2sfile.simps index_of_file.simps split:option.splits) +done + +lemma index_of_file_rename1: "\f\<^isub>2 \ f; f \ f\<^isub>2\ \ index_of_file (Rename p f\<^isub>2 f\<^isub>3 # \) (file_after_rename f\<^isub>2 f\<^isub>3 f) = index_of_file \ f" +apply (clarsimp simp add:index_of_file.simps split:option.splits) +by (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop3, simp, erule conjE, simp add:file_renaming_prop5) + +lemma cf2sfile_rename2: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; (file_before_rename f\<^isub>2 f\<^isub>3 f) \ current_files \; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>3 \ f\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ (file_before_rename f\<^isub>2 f\<^isub>3 f))" +apply (induct f, simp add:no_junior_def) +apply (case_tac "a # f = f\<^isub>3", simp) +apply (drule cf2sfile_rename1, simp, simp add:file_renaming_prop0' file_renaming_prop0) + +apply (frule no_junior_noteq, simp, simp) +apply (frule_tac file_renaming_prop1') +apply (frule_tac f = f\<^isub>2 and \ = \ in cf2sfile_keep_path, simp add:vt_cons') +apply (frule_tac f\<^isub>2 = f\<^isub>2 and a = a in file_renaming_prop8') +apply (frule_tac f\<^isub>2 = f\<^isub>2 and a = a in file_renaming_prop6') +apply (frule_tac f = "file_before_rename f\<^isub>2 f\<^isub>3 (a # f)" and \ = \ and f\<^isub>2 = f\<^isub>2 and f\<^isub>3 = f\<^isub>3 and p = p in index_of_file_rename1, simp add:file_before_rename_def) +apply (drule_tac f\<^isub>3 = f\<^isub>3 and f\<^isub>1 = f and a = a and f\<^isub>2 = f\<^isub>2 in file_renaming_prop9', simp) +apply (drule parentf_in_current', simp add:vt_cons') +apply (simp add:cf2sfile.simps) +apply (erule file_renaming_prop6[THEN sym]) +done + +lemma cf2sfile_rename2': "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; f \ current_files \; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>2 \ f\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ f)" +apply (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop5) +apply (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop1) +apply (drule_tac f = "file_after_rename f\<^isub>2 f\<^isub>3 f" in cf2sfile_rename2, simp+) +done + +lemma cf2sfile_rename3: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; f \ current_files \; parent f\<^isub>3 = Some pf\<^isub>3\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ f)" +apply (case_tac "f\<^isub>2 \ f") +apply (rule cf2sfile_rename2', simp+) +apply (frule vt_grant_os) +apply (frule_tac \ = \ in cf2sfile_keep_path', simp add:vt_cons', simp add:os_grant.simps, simp) +apply (simp add:file_after_rename_def) +by (rule cf2sfile_rename'1, simp+) + +lemma cf2sfile_rename: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; f \ current_files (Rename p f\<^isub>2 f\<^isub>3 # \)\ \ + cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = ( + if (f\<^isub>3 \ f) then (case (parent f\<^isub>3) of + Some pf\<^isub>3 \ file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ (file_before_rename f\<^isub>2 f\<^isub>3 f)) + | _ \ []) + else cf2sfile \ f )" +apply (frule vt_grant_os) +apply (case_tac "f = f\<^isub>3", clarsimp simp:os_grant.simps, drule cf2sfile_rename1, simp+, simp add:file_renaming_prop0' file_renaming_prop0) + +apply (auto simp:os_grant.simps current_files_simps intro:cf2sfile_rename'2 cf2sfile_rename1 split:if_splits option.splits) +apply (rule cf2sfile_rename2, simp, drule rename_renaming_decom, simp+) +apply (drule_tac f\<^isub>2 = f\<^isub>2 and f\<^isub>1 = f\<^isub>1 and f\<^isub>3 = f\<^isub>3 in file_renaming_prop5, simp+) +done + +lemma cf2sfile_other: "\ + \ p f flags fd opt. e \ Open p f flags fd opt; + \ p f im. e \ Mkdir p f im; + \ p f\<^isub>1 f\<^isub>2. e \ LinkHard p f\<^isub>1 f\<^isub>2; + \ p f\<^isub>2 f\<^isub>3. e \ Rename p f\<^isub>2 f\<^isub>3\ \ cf2sfile (e # \) f' = cf2sfile \ f'" +apply (induct f', simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits) +apply (case_tac e, auto simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits) +done + +lemmas cf2sfile_nil = init_cf2sfile + +lemma cf2sfile_nil': "f \ current_files [] \ cf2sfile [] f = map SInit f" +by (simp add:cf2sfile_nil current_files_simps) + +lemmas cf2sfile_simps = cf2sfile_nil cf2sfile_nil' cf2sfile_open cf2sfile_mkdir cf2sfile_linkhard cf2sfile_rename cf2sfile_other + +lemmas cf2sfile_simpss = cf2sfile_nil cf2sfile_nil' cf2sfile_open_some' cf2sfile_open_some cf2sfile_open_none cf2sfile_open cf2sfile_mkdir_some' cf2sfile_mkdir_some + cf2sfile_mkdir cf2sfile_linkhard_none cf2sfile_linkhard_some cf2sfile_linkhard cf2sfile_rename'1 cf2sfile_rename'2 cf2sfile_rename1 + cf2sfile_rename2 cf2sfile_rename2' cf2sfile_rename3 cf2sfile_rename cf2sfile_other + +lemma cf2sfile_open_some'_inum: "\Open p f' flags fd (Some inum) # \ \ vt rc_cs; inum_of_file \ f = Some im\ \ cf2sfile (Open p f' flags fd (Some inum) # \) f = cf2sfile \ f" +by (simp add:cf2sfile_open_some' current_files_def) + +lemma cf2sfile_mkdir_some'_inum: "\Mkdir p f' inum # \ \ vt rc_cs; inum_of_file \ f = Some im\ \ cf2sfile (Mkdir p f' inum # \) f = cf2sfile \ f" +by (simp add:cf2sfile_mkdir_some' current_files_def) + +lemma cf2sfile_linkhard_none_inum: "\LinkHard p f\<^isub>1 f\<^isub>2 # \ \ vt rc_cs; inum_of_file \ f = Some im\ \ cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \) f = cf2sfile \ f" +by (simp add:cf2sfile_linkhard_none current_files_def) + +lemma cf2sfile_rename'1_inum: + "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; inum_of_file \ f = Some im ; \ (f\<^isub>2 \ f)\ \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = cf2sfile \ f" +by (simp add:cf2sfile_rename'1 current_files_def) + +lemma cf2sfile_rename2_inum: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; inum_of_file \ (file_before_rename f\<^isub>2 f\<^isub>3 f) = Some im; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>3 \ f\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ (file_before_rename f\<^isub>2 f\<^isub>3 f))" +by (simp add:cf2sfile_rename2 current_files_def) + +lemma cf2sfile_rename2'_inum: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; inum_of_file \ f = Some im; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>2 \ f\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ f)" +by (simp add:cf2sfile_rename2' current_files_def) + +lemma cf2sfile_rename3_inum: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; inum_of_file \ f = Some im; parent f\<^isub>3 = Some pf\<^isub>3\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ f)" +by (simp add:cf2sfile_rename3 current_files_def) + +lemma cf2sfile_nil'_inum: "inum_of_file [] f = Some im \ cf2sfile [] f = map SInit f" +by (simp add:cf2sfile_nil' current_files_def) + +lemmas cf2sfile_simps_inum = cf2sfile_nil cf2sfile_nil'_inum cf2sfile_open_some'_inum cf2sfile_open_some cf2sfile_open_none cf2sfile_open cf2sfile_mkdir_some'_inum cf2sfile_mkdir_some + cf2sfile_mkdir cf2sfile_linkhard_none_inum cf2sfile_linkhard_some cf2sfile_linkhard cf2sfile_rename'1_inum cf2sfile_rename'2 cf2sfile_rename1 + cf2sfile_rename2_inum cf2sfile_rename2'_inum cf2sfile_rename3_inum cf2sfile_rename cf2sfile_other + +(*************** cp2sproc simpset *********************) + +lemma cp2sproc_nil: "p \ init_processes \ cp2sproc [] p = SInit p" +by (simp add:cp2sproc_def index_of_proc.simps) + +lemma cp2sproc_nil': "p \ current_procs [] \ cp2sproc [] p = SInit p" +by (simp add:cp2sproc_nil current_procs.simps) + +lemma cp2sproc_clone: "cp2sproc (Clone p p' # \) p'' = ( + if (p'' = p') then SCrea (Suc (length \)) + else cp2sproc \ p'' )" +by (auto simp:cp2sproc_def index_of_proc.simps d2s_aux.simps) + +lemma cp2sproc_other: "\ p p'. e \ Clone p p' \ cp2sproc (e # \) p'' = cp2sproc \ p''" +apply (case_tac e) +by (auto simp:cp2sproc_def index_of_proc.simps d2s_aux.simps) + +lemmas cp2sproc_simps = cp2sproc_nil cp2sproc_nil' cp2sproc_clone cp2sproc_other + +(******************** ch2sshm simpset **************************) + +lemma ch2sshm_nil: "h \ init_shms \ ch2sshm [] h = SInit h" +by (simp add:ch2sshm_def index_of_shm.simps) + +lemma ch2sshm_nil': "h \ current_shms [] \ ch2sshm [] h = SInit h" +by (simp add:ch2sshm_nil current_shms.simps) + +lemma ch2sshm_createshm: "ch2sshm (CreateShM p h # \) h' = (if (h' = h) then SCrea (Suc (length \)) else ch2sshm \ h')" +by (simp add:ch2sshm_def index_of_shm.simps d2s_aux.simps) + +lemma ch2sshm_other: "\ p h. e \ CreateShM p h \ ch2sshm (e # \) h' = ch2sshm \ h'" +apply (case_tac e) +by (auto simp add:ch2sshm_def index_of_shm.simps d2s_aux.simps) + +lemmas ch2sshm_simps = ch2sshm_nil ch2sshm_nil' ch2sshm_createshm ch2sshm_other + +(********************* cm2smsg simpset ***********************) + +lemma cm2smsg_nil: "m \ init_msgs \ cm2smsg [] m = SInit m" +by (simp add:cm2smsg_def index_of_msg.simps) + +lemma cm2smsg_nil': "m \ current_msgs [] \ cm2smsg [] m = SInit m" +by (simp add:cm2smsg_nil current_msgs.simps) + +lemma cm2smsg_createmsg: "cm2smsg (CreateMsg p m # \) m' = (if (m' = m) then SCrea (Suc (length \)) else cm2smsg \ m')" +by (simp add:cm2smsg_def index_of_msg.simps d2s_aux.simps) + +lemma cm2smsg_other: "\ p m. e \ CreateMsg p m \ cm2smsg (e # \) m' = cm2smsg \ m'" +apply (case_tac e) +by (auto simp:cm2smsg_def index_of_msg.simps d2s_aux.simps) + +lemmas cm2smsg_simps = cm2smsg_nil cm2smsg_nil' cm2smsg_createmsg cm2smsg_other + +(********************** cfd2fd_s simpset ******************************) + +lemma cfd2fd_s_nil: "fd \ init_fds_of_proc p \ cfd2fd_s [] p fd = SInit fd" +by (simp add:cfd2fd_s_def index_of_fd.simps) + +lemma cfd2fd_s_nil': "fd \ current_proc_fds [] p \ cfd2fd_s [] p fd = SInit fd" +by (simp add:cfd2fd_s_nil current_proc_fds.simps) + +lemma cfd2fd_s_open: "cfd2fd_s (Open p f flags fd opt # \) p' fd' = ( + if (p = p') then (if (fd = fd') then SCrea (Suc (length \)) + else cfd2fd_s \ p' fd') + else cfd2fd_s \ p' fd' )" +by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps) + +lemma cfd2fd_s_createsock: "cfd2fd_s (CreateSock p af st fd im # \) p' fd' = ( + if (p = p') then (if (fd = fd') then SCrea (Suc (length \)) + else cfd2fd_s \ p' fd') + else cfd2fd_s \ p' fd' )" +by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps) + +lemma cfd2fd_s_accept: "cfd2fd_s (Accept p fd addr port fd' im # \) p' fd'' = ( + if (p' = p) then (if (fd'' = fd') then SCrea (Suc (length \)) + else cfd2fd_s \ p' fd'') + else cfd2fd_s \ p' fd'' )" +by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps) + +lemma cfd2fd_s_clone: "cfd2fd_s (Clone p p' # \) p'' fd = (if (p'' = p') then cfd2fd_s \ p fd else cfd2fd_s \ p'' fd)" +by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps) + +lemma cfd2fd_s_other: "\\ p f flags fd opt. e \ Open p f flags fd opt; + \ p af st fd im. e \ CreateSock p af st fd im; + \ p fd addr port fd' im. e \ Accept p fd addr port fd' im; + \ p p'. e \ Clone p p'\ \ cfd2fd_s (e # \) p'' fd'' = cfd2fd_s \ p'' fd''" +by (case_tac e, auto simp:cfd2fd_s_def index_of_fd.simps d2s_aux.simps) + +lemmas cfd2fd_s_simps = cfd2fd_s_nil cfd2fd_s_nil' cfd2fd_s_open cfd2fd_s_createsock cfd2fd_s_accept cfd2fd_s_clone cfd2fd_s_other + +(************* cim2im_s simpset **************************) + +(* no such lemma +lemma cim2im_s_nil: "init_itag_of_inum im = Some tag \ cim2im_s [] im = SInit im" +by (simp add:cim2im_s_def) +*) + +lemma cim2im_s_open: "cim2im_s (Open p f flags fd (Some im) # \) im' = (if (im' = im) then SCrea (Suc (length \)) else cim2im_s \ im')" +by (simp add:cim2im_s_def) + +lemma cim2im_s_open': "cim2im_s (Open p f flags fd None # \) im = cim2im_s \ im" +by (simp add:cim2im_s_def) + +lemma cim2im_s_mkdir: "cim2im_s (Mkdir p f im # \) im' = (if (im' = im) then SCrea (Suc (length \)) else cim2im_s \ im')" +by (simp add:cim2im_s_def) + +lemma cim2im_s_createsock: "cim2im_s (CreateSock p sf st fd im # \) im' = (if (im' = im) then SCrea (Suc (length \)) else cim2im_s \ im')" +by (simp add:cim2im_s_def) + +lemma cim2im_s_accept: "cim2im_s (Accept p fd addr port fd' im # \) im' = (if (im' = im) then SCrea (Suc (length \)) else cim2im_s \ im')" +by (simp add:cim2im_s_def) + +lemma cim2im_s_other: "\\ p f flags fd opt. e \ Open p f flags fd opt; + \ p f im. e \ Mkdir p f im; + \ p sf st fd im. e \ CreateSock p sf st fd im; + \ p fd addr port fd' im. e \ Accept p fd addr port fd' im\ \ cim2im_s (e # \) im = cim2im_s \ im" +by (case_tac e, auto simp:cim2im_s_def) + +lemmas cim2im_s_simps = cim2im_s_open cim2im_s_open' cim2im_s_mkdir cim2im_s_createsock cim2im_s_accept cim2im_s_other + + +lemma cig2ig_s_simp: "cig2ig_s (e # \) tag = cig2ig_s \ tag" +apply (case_tac tag) +by auto + +(******************* cobj2sobj no Suc (length \) ***********************) + +lemma cf2sfile_le_len: "\cf2sfile \ f = SCrea (Suc (length \)) # spf; f \ current_files \; \ \ vt rc_cs\ \ False" +apply (case_tac f, (simp add:cf2sfile.simps d2s_aux.simps)+) +apply (case_tac "index_of_file \ (a # list)", (simp add:d2s_aux.simps)+) +by (drule index_of_file_le_length', simp+) + +lemma cf2sfile_le_len': "\SCrea (Suc (length \)) # spf \ cf2sfile \ f; f \ current_files \; \ \ vt rc_cs\ \ False" +apply (induct f) +apply (simp add:no_junior_def cf2sfile.simps d2s_aux.simps) +apply (case_tac "cf2sfile \ (a # f) = SCrea (Suc (length \)) # spf") +apply (drule_tac f = "a # f" in cf2sfile_le_len, simp+) +apply (simp only:cf2sfile.simps d2s_aux.simps) +apply (drule_tac no_junior_noteq, simp+) +apply (rule impI, erule impE, simp+) +apply (drule parentf_in_current', simp+) +done + +lemma cp2sproc_le_len: "cp2sproc \ p = SCrea (Suc (length \)) \ False" +apply (simp add:cp2sproc_def, case_tac "index_of_proc \ p") +apply (simp add:d2s_aux.simps)+ +by (drule index_of_proc_le_length', simp) + +lemma ch2sshm_le_len: "ch2sshm \ h = SCrea (Suc (length \)) \ False" +apply (simp add:ch2sshm_def, case_tac "index_of_shm \ h") +apply (simp add:d2s_aux.simps)+ +by (drule index_of_shm_le_length', simp) + +lemma cm2smsg_le_len: "cm2smsg \ m = SCrea (Suc (length \)) \ False" +apply (simp add:cm2smsg_def, case_tac "index_of_msg \ m") +apply (simp add:d2s_aux.simps)+ +by (drule index_of_msg_le_length', simp) + +lemma cim2im_s_le_len: "cim2im_s \ im = SCrea (Suc (length \)) \ False" +apply (simp add:cim2im_s_def, case_tac "inum2ind \ im") +apply (simp add:d2s_aux.simps)+ +by (drule inum2ind_le_length', simp) + +lemma cfd2fd_s_le_len: "cfd2fd_s \ p fd = SCrea (Suc (length \)) \ False" +apply (simp add:cfd2fd_s_def, case_tac "index_of_fd \ p fd") +apply (simp add:d2s_aux.simps)+ +by (drule index_of_fd_le_length', simp) + +end + +(*<*) +end +(*>*) \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Current_files_prop.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Current_files_prop.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,736 @@ +(*<*) +theory Current_files_prop +imports Main Flask_type Flask My_list_prefix Init_prop Valid_prop +begin +(*<*) + +context init begin + +lemma current_files_ndef: "f \ current_files \ \ inum_of_file \ f = None" +by (simp add:current_files_def) + +(************** file_of_proc_fd vs proc_fd_of_file *****************) +lemma pfdof_simp1: "file_of_proc_fd \ p fd = Some f \ (p, fd) \ proc_fd_of_file \ f" +by (simp add:proc_fd_of_file_def) + +lemma pfdof_simp2: "(p, fd) \ proc_fd_of_file \ f \ file_of_proc_fd \ p fd = Some f" +by (simp add:proc_fd_of_file_def) + +lemma pfdof_simp3: "proc_fd_of_file \ f = {(p, fd)} \ \ p' fd'. (file_of_proc_fd \ p' fd' = Some f \ p = p' \ fd = fd')" +by (simp add:proc_fd_of_file_def, auto) + +lemma pfdof_simp4: "\file_of_proc_fd \ p' fd' = Some f; proc_fd_of_file \ f = {(p, fd)}\ \ p' = p \ fd' = fd" +by (drule pfdof_simp3, auto) + +end + +context flask begin + +(***************** inode number lemmas *************************) + +lemma iof's_im_in_cim: "inum_of_file \ f = Some im \ im \ current_inode_nums \" +by (auto simp add:current_inode_nums_def current_file_inums_def) + +lemma ios's_im_in_cim: "inum_of_socket \ s = Some im \ im \ current_inode_nums \" +by (case_tac s, auto simp add:current_inode_nums_def current_sock_inums_def) + +lemma fim_noninter_sim_aux[rule_format]: + "\ f s. inum_of_file \ f = Some im \ inum_of_socket \ s = Some im \ valid \ \ False" +apply (induct \) +apply (clarsimp simp:inum_of_file.simps inum_of_socket.simps) +apply (drule init_inum_sock_file_noninter, simp, simp) +apply ((rule allI|rule impI|erule conjE)+) +apply (frule vd_cons, frule vt_grant_os, simp, case_tac a) +apply (auto simp:inum_of_file.simps inum_of_socket.simps split:if_splits option.splits + dest:ios's_im_in_cim iof's_im_in_cim) +done + +lemma fim_noninter_sim':"\inum_of_file \ f = Some im; inum_of_socket \ s = Some im; valid \\ \ False" +by (auto intro:fim_noninter_sim_aux) + +lemma fim_noninter_sim'':"\inum_of_socket \ s = Some im; inum_of_file \ f = Some im; valid \\ \ False" +by (auto intro:fim_noninter_sim_aux) + +lemma fim_noninter_sim: "valid \ \ (current_file_inums \) \ (current_sock_inums \) = {}" +by (auto simp:current_file_inums_def current_sock_inums_def intro:fim_noninter_sim_aux[rule_format]) + +(******************* file inum has inode tag ************************) + +lemma finum_has_itag_aux[rule_format]: + "\ f im. inum_of_file \ f = Some im \ valid \ \ itag_of_inum \ im \ None" +apply (induct \) +apply (clarsimp simp:inum_of_file.simps itag_of_inum.simps init_inum_of_file_props) +apply (clarify, frule vt_grant_os, frule vd_cons, case_tac a) +apply (auto simp add:inum_of_file.simps itag_of_inum.simps os_grant.simps current_files_def + dest:fim_noninter_sim'' split:option.splits if_splits t_socket_type.splits) +done + +lemma finum_has_itag: "\inum_of_file \ f = Some im; valid \\ \ \ tag. itag_of_inum \ im = Some tag" +by (auto dest:conjI[THEN finum_has_itag_aux]) + +(*********************** file inum is file itag *************************) + +lemma finum_has_ftag_aux[rule_format]: + "\ f tag. inum_of_file \ f = Some im \ itag_of_inum \ im = Some tag \ valid \ \ is_file_dir_itag tag" +apply (induct \) +apply (clarsimp simp:inum_of_file.simps itag_of_inum.simps init_inum_of_file_props) +apply (clarify, frule vt_grant_os, frule vd_cons, case_tac a) +apply (auto simp:inum_of_file.simps os_grant.simps current_files_def itag_of_inum.simps + split:if_splits option.splits t_socket_type.splits + dest:ios's_im_in_cim iof's_im_in_cim) +done + +lemma finum_has_ftag: + "\inum_of_file \ f = Some im; itag_of_inum \ im = Some tag; valid \\ \ is_file_dir_itag tag" +by (auto intro:finum_has_ftag_aux) + +lemma finum_has_ftag': + "\inum_of_file \ f = Some im; valid \\ \ itag_of_inum \ im = Some Tag_FILE \ itag_of_inum \ im = Some Tag_DIR" +apply (frule finum_has_itag, simp, erule exE, drule finum_has_ftag, simp+) +apply (case_tac tag, auto) +done + +(******************* sock inum has inode tag ************************) + +lemma sinum_has_itag_aux[rule_format]: + "\ s im. inum_of_socket \ s = Some im \ valid \ \ itag_of_inum \ im \ None" +apply (induct \) +apply (clarsimp simp:inum_of_socket.simps itag_of_inum.simps) +apply (drule init_inumos_prop4, clarsimp) +apply (clarify, frule vt_grant_os, frule vd_cons, case_tac a) +apply (auto simp add:inum_of_socket.simps itag_of_inum.simps os_grant.simps + dest:fim_noninter_sim'' ios's_im_in_cim iof's_im_in_cim + split:option.splits if_splits t_socket_type.splits) +done + +lemma sinum_has_itag: "\inum_of_socket \ s = Some im; valid \\ \ \ tag. itag_of_inum \ im = Some tag" +by (auto dest:conjI[THEN sinum_has_itag_aux]) + +(********************** socket inum is socket itag **********************) + +lemma sinum_has_stag_aux[rule_format]: + "\ s tag. inum_of_socket \ s = Some im \ itag_of_inum \ im = Some tag \ valid \ \ is_sock_itag tag" +apply (induct \) +apply (clarsimp simp:inum_of_socket.simps itag_of_inum.simps) +apply (drule init_inumos_prop4, clarsimp) +apply (clarify, frule vt_grant_os, frule vd_cons, case_tac a) +apply (auto simp add:inum_of_socket.simps itag_of_inum.simps os_grant.simps + dest:fim_noninter_sim'' ios's_im_in_cim iof's_im_in_cim + split:option.splits if_splits t_socket_type.splits) +done + +lemma sinum_has_stag: "\inum_of_socket \ s = Some im; itag_of_inum \ im = Some tag; valid \\ \ is_sock_itag tag" +by (auto dest:conjI[THEN sinum_has_stag_aux]) + +lemma sinum_has_stag': + "\inum_of_socket \ s = Some im; valid \\ + \ itag_of_inum \ im = Some Tag_UDP_SOCK \ itag_of_inum \ im = Some Tag_TCP_SOCK" +apply (frule sinum_has_itag, simp, erule exE) +apply (drule sinum_has_stag, simp+, case_tac tag, simp+) +done + +(************************************ 4 in 1 *************************************) + +lemma file_leveling: "valid \ \ ( + (\ f. f \ files_hung_by_del \ \ inum_of_file \ f \ None) \ + (\ f pf im. parent f = Some pf \ inum_of_file \ f = Some im \ inum_of_file \ pf \ None) \ + (\ f f' im. is_file \ f \ parent f' = Some f \ inum_of_file \ f' = Some im \ False) \ + (\ f f' im. f \ files_hung_by_del \ \ inum_of_file \ f' = Some im \ parent f' = Some f \ False) )" +proof (induct \) + case Nil + show ?case + apply (auto simp:inum_of_file.simps files_hung_by_del.simps is_file_def itag_of_inum.simps parent_file_in_init split:option.splits t_inode_tag.splits) + apply (drule init_files_hung_by_del_props, simp add:init_file_has_inum) + apply (rule init_parent_file_has_inum, simp+) + apply (rule init_file_has_no_son', simp+) + apply (rule init_file_hung_has_no_son, simp+) + done +next + case (Cons a \) + assume pre: "valid \ \ + (\ f. f \ files_hung_by_del \ \ inum_of_file \ f \ None) \ + (\f pf im. parent f = Some pf \ inum_of_file \ f = Some im \ inum_of_file \ pf \ None) \ + (\f f' im. is_file \ f \ parent f' = Some f \ inum_of_file \ f' = Some im \ False) \ + (\f f' im. f \ files_hung_by_del \ \ inum_of_file \ f' = Some im \ parent f' = Some f \ False)" + show ?case + proof + assume cons:"valid (a # \)" + show "(\f. f \ files_hung_by_del (a # \) \ inum_of_file (a # \) f \ None) \ + (\f pf im. parent f = Some pf \ inum_of_file (a # \) f = Some im \ inum_of_file (a # \) pf \ None) \ + (\f f' im. is_file (a # \) f \ parent f' = Some f \ inum_of_file (a # \) f' = Some im \ False) \ + (\f f' im. f \ files_hung_by_del (a # \) \ inum_of_file (a # \) f' = Some im \ parent f' = Some f \ False)" + proof- + have vt: "valid \" using cons by (auto dest:vd_cons) + have os: "os_grant \ a" using cons by (auto dest:vt_grant_os) + have fin: "\f. f \ files_hung_by_del \ \ inum_of_file \ f \ None" using vt pre by auto + have pin: "\f pf im. parent f = Some pf \ inum_of_file \ f = Some im \ inum_of_file \ pf \ None" + using vt pre apply (erule_tac impE, simp) by ((erule_tac conjE)+, assumption) + have fns: "\f f' im. is_file \ f \ parent f' = Some f \ inum_of_file \ f' = Some im \ False" + using vt pre apply (erule_tac impE, simp) by ((erule_tac conjE)+, assumption) + have hns: "\f f' im. f \ files_hung_by_del \ \ inum_of_file \ f' = Some im \ parent f' = Some f \ False" + using vt pre apply (erule_tac impE, simp) by ((erule_tac conjE)+, assumption) + have ain: "\f' f im. f \ f' \ inum_of_file \ f' = Some im \ inum_of_file \ f \ None" + proof + fix f' + show " \f im. f \ f' \ inum_of_file \ f' = Some im \ inum_of_file \ f \ None" + proof (induct f') + case Nil show ?case by (auto simp: no_junior_def) + next + case (Cons a f') + assume pre:"\f im. f \ f' \ inum_of_file \ f' = Some im \ inum_of_file \ f \ None" + show "\f im. f \ a # f' \ inum_of_file \ (a # f') = Some im \ inum_of_file \ f \ None" + proof clarify + fix f im + assume h1: "f \ a # f'" + and h2: "inum_of_file \ (a # f') = Some im" + show "\y. inum_of_file \ f = Some y" + proof- + have h3: "\ y. inum_of_file \ f' = Some y" + proof- + have "parent (a # f') = Some f'" by simp + hence "\ y. inum_of_file \ f' = Some y" using pin h2 by blast + with h1 show ?thesis by simp + qed + from h1 have "f = a # f' \ f = f' \ f \ f'" by (induct f, auto simp:no_junior_def) + moreover have "f = a # f' \ \y. inum_of_file \ f = Some y" using h2 by simp + moreover have "f = f' \ \y. inum_of_file \ f = Some y" using h3 by simp + moreover have "f \ f' \ \y. inum_of_file \ f = Some y" using pre h3 by simp + ultimately show ?thesis by auto + qed + qed + qed + qed + + have fin': "\ f. f \ files_hung_by_del \ \ \ im. inum_of_file \ f = Some im" using fin by auto + have pin': "\ f pf im. \parent f = Some pf; inum_of_file \ f = Some im\ \ \ im'. inum_of_file \ pf = Some im'" + using pin by auto + have fns': "\ f f' im. \is_file \ f; parent f' = Some f; inum_of_file \ f' = Some im\ \ False" using fns by auto + have fns'': "\ f f' im im'. \itag_of_inum \ im = Some Tag_FILE; inum_of_file \ f = Some im; parent f' = Some f; inum_of_file \ f' = Some im'\ \ False" + by (rule_tac f = f and f' = f' in fns', auto simp:is_file_def) + have hns': "\ f f' im. \f \ files_hung_by_del \; inum_of_file \ f' = Some im; parent f' = Some f\ \ False" using hns by auto + have ain': "\ f f' im. \f \ f'; inum_of_file \ f' = Some im\ \ \ im'. inum_of_file \ f = Some im'" using ain by auto + have dns': "\ f f' im. \dir_is_empty \ f; parent f' = Some f; inum_of_file \ f' = Some im\ \ False" + apply (auto simp:dir_is_empty_def current_files_def is_dir_def split:option.splits) + by (erule_tac x = f' in allE, simp add:noJ_Anc parent_is_ancen, drule parent_is_parent, simp+) + + have "\ f. f \ files_hung_by_del (a # \) \ inum_of_file (a # \) f \ None" + apply (clarify, case_tac a) using os fin + apply (auto simp:files_hung_by_del.simps inum_of_file.simps os_grant.simps current_files_def + split:if_splits option.splits) + done + moreover + have "\f pf im. parent f = Some pf \ inum_of_file (a # \) f = Some im \ inum_of_file (a # \) pf \ None" + apply (clarify, case_tac a) + using vt os pin' + apply (auto simp:os_grant.simps current_files_def inum_of_file.simps split:if_splits option.splits) + apply (drule_tac f = pf and f' = f in hns', simp, simp, simp) + apply (drule_tac f = list and f' = f in fns', simp, simp, simp) + apply (drule_tac f = list and f' = f in dns', simp, simp, simp) + done + moreover have "\f f' im. is_file (a # \) f \ parent f' = Some f \ inum_of_file (a # \) f' = Some im \ False" + apply (clarify, case_tac a) + using vt os fns'' + apply (auto simp:os_grant.simps current_files_def inum_of_file.simps itag_of_inum.simps + is_file_def is_dir_def + dest:ios's_im_in_cim iof's_im_in_cim + split:if_splits option.splits t_inode_tag.splits t_socket_type.splits) + apply (drule_tac f = f' and pf = list in pin', simp, simp) + apply (drule_tac f = f' and pf = list2 in pin', simp, simp) + done + moreover have "\f f' im. f \ files_hung_by_del (a # \) \ inum_of_file (a # \) f' = Some im \ + parent f' = Some f \ False" + apply (clarify, case_tac a) + using vt os hns' + apply (auto simp:os_grant.simps current_files_def inum_of_file.simps files_hung_by_del.simps + split:if_splits option.splits t_sock_addr.splits) + apply (drule fns', simp+) + done + ultimately show ?thesis by blast + qed + qed +qed + +(**************** hung file in current ***********************) + +lemma hung_file_has_inum:"\f \ files_hung_by_del \; valid \\ \ inum_of_file \ f \ None" +by (drule file_leveling[rule_format], blast) + +lemma hung_file_has_inum': "\f \ files_hung_by_del \; valid \\ \ \ im. inum_of_file \ f = Some im" +by (auto dest:hung_file_has_inum) + +lemma hung_file_in_current: "\f \ files_hung_by_del \; valid \\ \ f \ current_files \" +by (clarsimp simp add:current_files_def hung_file_has_inum') + +lemma parentf_has_inum: "\parent f = Some pf; inum_of_file \ f = Some im; valid \\ \ inum_of_file \ pf \ None" +by (drule file_leveling[rule_format], blast) + +lemma parentf_has_inum': "\parent f = Some pf; inum_of_file \ f = Some im; valid \\ \ \ im'. inum_of_file \ pf = Some im'" +by (auto dest:parentf_has_inum) + +lemma parentf_in_current: "\parent f = Some pf; f \ current_files \; valid \\ \ pf \ current_files \" +by (clarsimp simp add:current_files_def parentf_has_inum') + +lemma parentf_in_current': "\a # pf \ current_files \; valid \\ \ pf \ current_files \" +apply (subgoal_tac "parent (a # pf) = Some pf") +by (erule parentf_in_current, simp+) + +lemma ancenf_has_inum_aux: "\ f im. f \ f' \ inum_of_file \ f' = Some im \ valid \ \ inum_of_file \ f \ None" +proof (induct f') + case Nil show ?case by (auto simp: no_junior_def) +next + case (Cons a f') + assume pre:"\f im. f \ f' \ inum_of_file \ f' = Some im \ valid \ \ inum_of_file \ f \ None" + show "\f im. f \ a # f' \ inum_of_file \ (a # f') = Some im \ valid \ \ inum_of_file \ f \ None" + proof clarify + fix f im + assume h1: "f \ a # f'" + and h2: "inum_of_file \ (a # f') = Some im" + and h3: "valid \" + show "\y. inum_of_file \ f = Some y" + proof- + have h4: "\ y. inum_of_file \ f' = Some y" + proof- + have "parent (a # f') = Some f'" by simp + hence "\ y. inum_of_file \ f' = Some y" using parentf_has_inum' h2 h3 by blast + with h1 show ?thesis by simp + qed + from h1 have "f = a # f' \ f = f' \ f \ f'" by (induct f, auto simp:no_junior_def) + moreover have "f = a # f' \ \y. inum_of_file \ f = Some y" using h2 by simp + moreover have "f = f' \ \y. inum_of_file \ f = Some y" using h4 by simp + moreover have "f \ f' \ \y. inum_of_file \ f = Some y" using pre h3 h4 by simp + ultimately show ?thesis by auto + qed + qed +qed + +lemma ancenf_has_inum: "\f \ f'; inum_of_file \ f' = Some im; valid \\ \ inum_of_file \ f \ None" +by (rule ancenf_has_inum_aux[rule_format], auto) + +lemma ancenf_has_inum': "\f \ f'; inum_of_file \ f' = Some im; valid \\ \ \ im'. inum_of_file \ f = Some im'" +by (auto dest:ancenf_has_inum) + +lemma ancenf_in_current: "\f \ f'; f' \ current_files \; valid \\ \ f \ current_files \" +by (simp add:current_files_def, erule exE, simp add:ancenf_has_inum') + +lemma file_has_no_son: "\is_file \ f; parent f' = Some f; inum_of_file \ f' = Some im; valid \\ \ False" +by (drule file_leveling[rule_format], blast) + +lemma file_has_no_son': "\is_file \ f; parent f' = Some f; f' \ current_files \; valid \\ \ False" +by (simp add:current_files_def, erule exE, auto intro:file_has_no_son) + +lemma hung_file_no_son: "\f \ files_hung_by_del \; valid \; inum_of_file \ f' = Some im; parent f' = Some f\ \ False" +by (drule file_leveling[rule_format], blast) + +lemma hung_file_no_son': "\f \ files_hung_by_del \; valid \; f' \ current_files \; parent f' = Some f\ \ False" +by (simp add:current_files_def, erule exE, auto intro:hung_file_no_son) + +lemma hung_file_no_des_aux: "\ f. f \ files_hung_by_del \ \ valid \ \ f' \ current_files \ \ f \ f' \ f \ f' \ False" +proof (induct f') + case Nil + show ?case + by (auto simp:files_hung_by_del.simps current_files_def inum_of_file.simps no_junior_def split:if_splits option.splits) +next + case (Cons a pf) + assume pre: "\f. f \ files_hung_by_del \ \ valid \ \ pf \ current_files \ \ f \ pf \ f \ pf\ False" + show ?case + proof clarify + fix f + assume h1: "f \ files_hung_by_del \" + and h2: "valid \" + and h3: "a # pf \ current_files \" + and h4: "f \ a # pf" + and h5: "f \ a # pf" + have h6: "parent (a # pf) = Some pf" by simp + with h2 h3 have h7: "pf \ current_files \" by (drule_tac parentf_in_current, auto) + from h4 h5 have h8: "f \ pf" by (erule_tac no_juniorE, case_tac zs, auto simp:no_junior_def) + show False + proof (cases "f = pf") + case True with h6 h2 h3 h1 + show False by (auto intro!:hung_file_no_son') + next + case False with pre h1 h2 h7 h8 + show False by blast + qed + qed +qed + +lemma hung_file_no_des: "\f \ files_hung_by_del \; valid \; f' \ current_files \; f \ f'; f \ f'\ \ False" +by (rule hung_file_no_des_aux[rule_format], blast) + +lemma hung_file_is_leaf: "\f \ files_hung_by_del \; valid \\ \ is_file \ f \ dir_is_empty \ f" +apply (frule hung_file_has_inum', simp, erule exE) +apply (auto simp add:is_file_def dir_is_empty_def is_dir_def dest:finum_has_itag finum_has_ftag split:option.splits if_splits t_inode_tag.splits) +by (simp add: noJ_Anc, auto dest:hung_file_no_des) + + + +(************** file_of_proc_fd in current ********************) + +lemma file_of_pfd_imp_inode_aux: "\ p f. file_of_proc_fd \ p fd = Some f \ valid \ \ inum_of_file \ f \ None" +apply (induct \) +apply (clarsimp simp add:file_of_proc_fd.simps inum_of_file.simps init_filefd_prop1 init_file_has_inum) +apply ((rule_tac allI|rule_tac impI|erule_tac conjE)+, frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:inum_of_file.simps file_of_proc_fd.simps os_grant.simps current_files_def + split:if_splits option.splits) +apply (simp add:pfdof_simp3)+ +apply (simp add:proc_fd_of_file_def)+ +done + +lemma file_of_pfd_imp_inode': "\file_of_proc_fd \ p fd = Some f; valid \\ \ inum_of_file \ f \ None" +by (rule file_of_pfd_imp_inode_aux[rule_format], blast) + +lemma file_of_pfd_imp_inode: "\file_of_proc_fd \ p fd = Some f; valid \\ \ \ im. inum_of_file \ f = Some im" +by (auto dest!:file_of_pfd_imp_inode') + +lemma file_of_pfd_in_current: "\file_of_proc_fd \ p fd = Some f; valid \\ \ f \ current_files \" +by (auto dest!:file_of_pfd_imp_inode' simp:current_files_def) + + +(*************** file_of_proc_fd is file *********************) + +lemma file_of_pfd_is_file_tag: + "\file_of_proc_fd \ p fd = Some f; valid \; inum_of_file \ f = Some im\ \ itag_of_inum \ im = Some Tag_FILE" +apply (induct \ arbitrary:p, simp) +apply (drule init_filefd_prop5, simp add:is_init_file_def split:option.splits t_inode_tag.splits) +apply (frule vd_cons, frule vt_grant_os, simp, case_tac a) +by (auto split:option.splits t_inode_tag.splits if_splits t_socket_type.splits + dest:file_of_pfd_imp_inode' iof's_im_in_cim + simp:is_file_def is_dir_def dir_is_empty_def current_files_def) + +lemma file_of_pfd_is_file: + "\file_of_proc_fd \ p fd = Some f; valid \\ \ is_file \ f" +apply (frule file_of_pfd_imp_inode, simp, erule exE) +apply (drule file_of_pfd_is_file_tag, simp+) +by (simp add:is_file_def) + +lemma file_of_pfd_is_file': + "\file_of_proc_fd \ p fd = Some f; is_dir \ f; valid \\ \ False" +by (drule file_of_pfd_is_file, auto simp:is_file_def is_dir_def split:option.splits t_inode_tag.splits) + +(************** parent file / ancestral file is dir *******************) + +lemma parentf_is_dir_aux: "\ f pf. parent f = Some pf \ inum_of_file \ f = Some im \ inum_of_file \ pf = Some ipm \ valid \ \ itag_of_inum \ ipm = Some Tag_DIR" +apply (induct \) +apply (clarsimp simp:inum_of_file.simps itag_of_inum.simps init_parent_file_is_dir') +apply (clarify, frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:inum_of_file.simps itag_of_inum.simps os_grant.simps + current_files_def is_dir_def is_file_def + dest: ios's_im_in_cim iof's_im_in_cim + split:if_splits option.splits t_sock_addr.splits t_inode_tag.splits t_socket_type.splits) +apply (drule parentf_has_inum', simp, simp, simp)+ +done + +lemma parentf_has_dirtag: + "\parent f = Some pf; inum_of_file \ f = Some im; inum_of_file \ pf = Some ipm; valid \\ + \ itag_of_inum \ ipm = Some Tag_DIR" +by (auto intro:parentf_is_dir_aux[rule_format]) + +lemma parentf_is_dir': "\parent f = Some pf; inum_of_file \ f = Some im; valid \\ \ is_dir \ pf" +apply (frule parentf_has_inum', simp+, erule exE, simp add:is_dir_def split:t_inode_tag.splits option.splits) +by (auto dest:parentf_has_dirtag) + +lemma parentf_is_dir: "\parent f = Some pf; f \ current_files \; valid \\ \ is_dir \ pf" +by (clarsimp simp:current_files_def parentf_is_dir') + +lemma ancenf_is_dir_aux: "\ f. f \ f' \ f \ f' \ f' \ current_files \ \ valid \ \ is_dir \ f" +proof (induct f') + case Nil show ?case + by (auto simp:current_files_def no_junior_def) +next + case (Cons a pf) + assume pre: "\f. f \ pf \ f \ pf \ pf \ current_files \ \ valid \ \ is_dir \ f" + show ?case + proof clarify + fix f + assume h1: "f \ a # pf" + and h2: "f \ a # pf" + and h3: "a # pf \ current_files \" + and h4: "valid \" + have h5: "parent (a # pf) = Some pf" by simp + with h3 h4 have h6: "pf \ current_files \" by (drule_tac parentf_in_current, auto) + from h1 h2 have h7: "f \ pf" by (erule_tac no_juniorE, case_tac zs, auto simp:no_junior_def) + show "is_dir \ f" + proof (cases "f = pf") + case True with h3 h4 h5 show "is_dir \ f" by (drule_tac parentf_is_dir, auto) + next + case False with pre h6 h7 h4 show "is_dir \ f" by blast + qed + qed +qed + +lemma ancenf_is_dir: "\f \ f'; f \ f'; f' \ current_files \; valid \\ \ is_dir \ f" +by (auto intro:ancenf_is_dir_aux[rule_format]) + +(************* rebuild current_files simpset ***********************) + +lemma current_files_nil: "current_files [] = init_files" +apply (simp add:current_files_def inum_of_file.simps) +by (auto dest:inof_has_file_tag init_file_has_inum) + +lemma current_files_open: "current_files (Open p f flags fd (Some i) # \) = insert f (current_files \)" +by (auto simp add:current_files_def inum_of_file.simps split:option.splits) + +lemma current_files_open': "current_files (Open p f flags fd None # \) = current_files \" +by (simp add:current_files_def inum_of_file.simps split:option.splits) + +lemma current_files_closefd: "current_files (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 current_files \ - {f} + else current_files \) + | _ \ current_files \ )" +by (auto simp:current_files_def inum_of_file.simps split:option.splits) + +lemma current_files_unlink: "current_files (UnLink p f # \) = (if (proc_fd_of_file \ f = {}) then (current_files \) - {f} else current_files \)" +by (auto simp:current_files_def inum_of_file.simps split:option.splits) + +lemma current_files_rmdir: "current_files (Rmdir p f # \) = (if (proc_fd_of_file \ f = {}) then current_files \ - {f} else current_files \)" +by (auto simp:current_files_def inum_of_file.simps split:option.splits) + +lemma current_files_mkdir: "current_files (Mkdir p f ino # \) = insert f (current_files \)" +by (auto simp:current_files_def inum_of_file.simps split:option.splits) + +lemma current_files_linkhard: + "valid (LinkHard p f\<^isub>1 f\<^isub>2 # \) \ current_files (LinkHard p f\<^isub>1 f\<^isub>2 # \) = insert f\<^isub>2 (current_files \)" +apply (frule vt_grant_os, frule vd_cons) +by (auto simp:current_files_def inum_of_file.simps os_grant.simps split:option.splits) + +(* +lemma rename_renaming_decom: + "\f\<^isub>3 \ file_after_rename f\<^isub>2 f\<^isub>3 f; Rename p f\<^isub>2 f\<^isub>3 # valid \; f \ current_files \\ \ f\<^isub>2 \ f" +apply (case_tac "f\<^isub>2 \ f", simp) +apply (simp add:file_after_rename_def split:if_splits) +by (frule vd_cons, frule vt_grant_os, auto simp:os_grant.simps dest!:ancenf_in_current) + +lemma rename_renaming_decom': + "\\ f\<^isub>3 \ file_after_rename f\<^isub>2 f\<^isub>3 f; Rename p f\<^isub>2 f\<^isub>3 # valid \; f \ current_files \\ \ \ f\<^isub>2 \ f" +by (case_tac "f\<^isub>2 \ f", drule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop1, simp+) + +lemma current_files_rename: "Rename p f\<^isub>2 f\<^isub>3 # valid \ \ current_files (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 \ current_files \}" +apply (frule vt_grant_os, frule vd_cons) +apply (auto simp:current_files_def inum_of_file.simps os_grant.simps split:if_splits option.splits) +apply (rule_tac x = x in exI, simp add:file_after_rename_def) +apply (frule_tac f\<^isub>2 = f\<^isub>2 in file_renaming_prop1', drule_tac f\<^isub>2 = f\<^isub>2 in file_renaming_prop5') +apply (erule_tac x = "file_before_rename f\<^isub>2 f\<^isub>3 x" in allE, simp) +apply (rule_tac x = x in exI, simp add:file_after_rename_def) +apply (drule_tac a = f\<^isub>3 and b = f\<^isub>2 in no_junior_conf, simp, simp) +apply (drule_tac f = f\<^isub>3 and f' = f\<^isub>2 in ancenf_has_inum', simp, simp, simp) +apply (drule_tac f\<^isub>2 = f\<^isub>2 in rename_renaming_decom, simp, simp add:current_files_def, simp add:file_renaming_prop5) +apply (drule_tac f\<^isub>2 = f\<^isub>2 in rename_renaming_decom', simp, simp add:current_files_def) +apply (simp add:file_after_rename_def) +apply (drule_tac f\<^isub>2 = f\<^isub>2 in rename_renaming_decom', simp, simp add:current_files_def) +apply (simp add:file_after_rename_def) +done +*) + +lemma current_files_other: + "\\ p f flag fd opt. e \ Open p f flag fd opt; + \ p fd. e \ CloseFd p fd; + \ p f. e \ UnLink p f; + \ p f. e \ Rmdir p f; + \ p f i. e \ Mkdir p f i; + \ p f f'. e \ LinkHard p f f'\ \ current_files (e # \) = current_files \" +by (case_tac e, auto simp:current_files_def inum_of_file.simps) + +lemmas current_files_simps = current_files_nil current_files_open current_files_open' + current_files_closefd current_files_unlink current_files_rmdir + current_files_mkdir current_files_linkhard current_files_other + + +(******************** is_file simpset *********************) + +lemma is_file_nil: "is_file [] = is_init_file" +by (auto simp:is_init_file_def is_file_def init_inum_of_file_props intro!:ext split:option.splits) + +lemma is_file_open: + "valid (Open p f flags fd opt # s) \ + is_file (Open p f flags fd opt # s) = (if (opt = None) then is_file s else (is_file s) (f:= True))" +apply (frule vd_cons, drule vt_grant_os, rule ext) +apply (auto dest:finum_has_itag iof's_im_in_cim + split:if_splits option.splits t_inode_tag.splits + simp:is_file_def current_files_def) +done + +lemma is_file_closefd: + "valid (CloseFd p fd # s) \ is_file (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 (is_file s) (f := False) + else is_file s) + | _ \ is_file s )" +apply (frule vd_cons, drule vt_grant_os, rule ext) +apply (auto dest:finum_has_itag iof's_im_in_cim + split:if_splits option.splits t_inode_tag.splits + simp:is_file_def) +done + +lemma is_file_unlink: + "valid (UnLink p f # s) \ is_file (UnLink p f # s) = ( + if (proc_fd_of_file s f = {}) then (is_file s) (f := False) else is_file s)" +apply (frule vd_cons, drule vt_grant_os, rule ext) +apply (auto dest:finum_has_itag iof's_im_in_cim + split:if_splits option.splits t_inode_tag.splits + simp:is_file_def) +done + +lemma is_file_linkhard: + "valid (LinkHard p f f' # s) \ is_file (LinkHard p f f' # s) = (is_file s) (f' := True)" +apply (frule vd_cons, drule vt_grant_os, rule ext) +apply (auto dest:finum_has_itag iof's_im_in_cim + split:if_splits option.splits t_inode_tag.splits + simp:is_file_def) +done + +lemma is_file_other: + "\valid (e # \); + \ p f flag fd opt. e \ Open p f flag fd opt; + \ p fd. e \ CloseFd p fd; + \ p f. e \ UnLink p f; + \ p f f'. e \ LinkHard p f f'\ \ is_file (e # \) = is_file \" +apply (frule vd_cons, drule vt_grant_os, rule_tac ext, case_tac e) +apply (auto dest:finum_has_itag iof's_im_in_cim intro!:ext + split:if_splits option.splits t_inode_tag.split t_socket_type.splits + simp:is_file_def dir_is_empty_def is_dir_def current_files_def) +done + +lemma file_dir_conflict: "\is_file s f; is_dir s f\ \ False" +by (auto simp:is_file_def is_dir_def split:option.splits t_inode_tag.splits) + +lemma is_file_not_dir: "is_file s f \ \ is_dir s f" +by (rule notI, erule file_dir_conflict, simp) + +lemma is_dir_not_file: "is_dir s f \ \ is_file s f" +by (rule notI, erule file_dir_conflict, simp) + +lemmas is_file_simps = is_file_nil is_file_open is_file_closefd is_file_unlink is_file_linkhard is_file_other + +(********* is_dir simpset **********) + +lemma is_dir_nil: "is_dir [] = is_init_dir" +by (auto simp:is_init_dir_def is_dir_def init_inum_of_file_props intro!:ext split:option.splits) + +lemma is_dir_mkdir: "valid (Mkdir p f i # s) \ is_dir (Mkdir p f i # s) = (is_dir s) (f := True)" +apply (frule vd_cons, drule vt_grant_os, rule_tac ext) +apply (auto dest:finum_has_itag iof's_im_in_cim intro!:ext + split:if_splits option.splits t_inode_tag.split t_socket_type.splits + simp:is_dir_def dir_is_empty_def is_dir_def current_files_def) +done + +lemma is_dir_rmdir: "valid (Rmdir p f # s) \ is_dir (Rmdir p f # s) = (is_dir s) (f := False)" +apply (frule vd_cons, drule vt_grant_os, rule_tac ext) +apply (auto dest:finum_has_itag iof's_im_in_cim intro!:ext + split:if_splits option.splits t_inode_tag.split t_socket_type.splits + simp:is_dir_def dir_is_empty_def is_dir_def current_files_def) +apply (drule pfdof_simp2) +apply (drule file_of_pfd_is_file, simp) +apply (simp add:is_file_def split:t_inode_tag.splits option.splits) +done + +lemma is_dir_other: + "\valid (e # s); + \ p f. e \ Rmdir p f; + \ p f i. e \ Mkdir p f i\ \ is_dir (e # s) = is_dir s" +apply (frule vd_cons, drule vt_grant_os, rule_tac ext, case_tac e) +apply (auto dest:finum_has_itag iof's_im_in_cim intro!:ext + split:if_splits option.splits t_inode_tag.split t_socket_type.splits + simp:is_file_def dir_is_empty_def is_dir_def current_files_def) +apply (drule file_of_pfd_is_file, simp) +apply (simp add:is_file_def split:t_inode_tag.splits option.splits) +done + +lemmas is_dir_simps = is_dir_nil is_dir_mkdir is_dir_rmdir is_dir_other + +(*********** no root dir involved ***********) + +lemma root_is_dir: "valid s \ is_dir s []" +apply (induct s, simp add:is_dir_nil root_is_init_dir) +apply (frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:is_dir_simps) +done + +lemma root_is_dir': "\is_file s []; valid s\ \ False" +apply (drule root_is_dir) +apply (erule file_dir_conflict, simp) +done + +lemma noroot_execve: + "valid (Execve p f fds # s) \ f \ []" +by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir') + +lemma noroot_execve': + "valid (Execve p [] fds # s) \ False" +by (drule noroot_execve, simp) + +lemma noroot_open: + "valid (Open p f flags fd opt # s) \ f \ []" +by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir' split:option.splits) + +lemma noroot_open': + "valid (Open p [] flags fd opt # s) \ False" +by (drule noroot_open, simp) + +lemma noroot_filefd': + "\file_of_proc_fd s p fd = Some []; valid s\ \ False" +apply (induct s arbitrary:p, simp) +apply (drule init_filefd_prop5, erule root_is_init_dir') +apply (frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto split:if_splits option.splits dest!:root_is_dir') +done + +lemma noroot_filefd: + "\file_of_proc_fd s p fd = Some f; valid s\ \ f \ []" +by (rule notI, simp, erule noroot_filefd', simp) + +lemma noroot_unlink: + "valid (UnLink p f # s) \ f \ []" +by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir') + +lemma noroot_unlink': + "valid (UnLink p [] # s) \ False" +by (drule noroot_unlink, simp) + +lemma noroot_rmdir: + "valid (Rmdir p f # s) \ f \ []" +by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir') + +lemma noroot_rmdir': + "valid (Rmdir p [] # s) \ False" +by (drule noroot_rmdir, simp) + +lemma noroot_mkdir: + "valid (Mkdir p f inum # s) \ f \ []" +by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir') + +lemma noroot_mkdir': + "valid (Mkdir p [] inum # s) \ False" +by (drule noroot_mkdir, simp) + +lemma noroot_linkhard: + "valid (LinkHard p f f' # s) \ f \ [] \ f' \ []" +by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir') + +lemma noroot_linkhard': + "valid (LinkHard p [] f # s) \ False" +by (drule noroot_linkhard, simp) + +lemma noroot_linkhard'': + "valid (LinkHard p f [] # s) \ False" +by (drule noroot_linkhard, simp) + +lemma noroot_truncate: + "valid (Truncate p f len # s) \ f \ []" +by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir') + +lemma noroot_truncate': + "valid (Truncate p [] len # s) \ False" +by (drule noroot_truncate, simp) + +lemmas noroot_events = noroot_execve noroot_open noroot_filefd noroot_unlink noroot_rmdir + noroot_mkdir noroot_linkhard noroot_truncate + +lemmas noroot_events' = noroot_execve' noroot_open' noroot_filefd' noroot_unlink' noroot_rmdir' + noroot_mkdir' noroot_linkhard' noroot_linkhard'' noroot_truncate' + +end + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Current_sockets_prop.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Current_sockets_prop.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,365 @@ +theory Current_sockets_prop +imports Main Flask Flask_type Current_files_prop +begin + +context flask begin + +lemma cn_in_curp: "\(p, fd) \ current_sockets \; valid \\ \ p \ current_procs \" +apply (induct \) defer +apply (frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:current_sockets_def inum_of_socket.simps current_procs.simps os_grant.simps split:if_splits option.splits + dest:inos_has_sock_tag init_socket_has_inode) +done + +lemma cn_in_curfd_aux: "\ p. (p, fd) \ current_sockets \ \ valid \ \ fd \ current_proc_fds \ p" +apply (induct \) defer +apply (rule allI, rule impI, erule conjE, frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:current_sockets_def current_proc_fds.simps os_grant.simps inum_of_socket.simps + split:if_splits option.splits t_socket_af.splits dest:inos_has_sock_tag init_socket_has_inode) +done + +lemma cn_in_curfd: "\(p, fd) \ current_sockets \; valid \\ \ fd \ current_proc_fds \ p" +by (simp add:cn_in_curfd_aux) + +lemma cn_in_curfd': "\inum_of_socket \ (p, fd) = Some im; valid \\ \ fd \ current_proc_fds \ p" +by (rule cn_in_curfd, simp_all add:current_sockets_def) + +lemma cn_in_curp': "\inum_of_socket \ (p, fd) = Some im; valid \\ \ p \ current_procs \" +apply (rule cn_in_curp) +by (auto simp:current_sockets_def) + +(***************** current_sockets simpset *****************) + +lemma current_sockets_nil: "current_sockets [] = init_sockets" +apply (simp add:current_sockets_def inum_of_socket.simps) +using init_socket_has_inode inos_has_sock_tag +by auto + +lemma current_sockets_closefd: "current_sockets (CloseFd p fd # \) = current_sockets \ - {(p, fd)}" +by (auto simp add:current_sockets_def inum_of_socket.simps) + +lemma current_sockets_createsock: "current_sockets (CreateSock p aft st fd ino # \) = insert (p, fd) (current_sockets \)" +by (auto simp add:current_sockets_def inum_of_socket.simps) + +lemma current_sockets_accept: "current_sockets (Accept p fd addr lport' fd' ino # \) = insert (p, fd') (current_sockets \)" +by (auto simp add:current_sockets_def inum_of_socket.simps) + +lemma current_sockets_clone: "valid (Clone p p' fds shms # \) + \ current_sockets (Clone p p' fds shms # \) = + current_sockets \ \ {(p', fd)| fd. (p, fd) \ current_sockets \ \ fd \ fds}" +apply (frule vd_cons, drule vt_grant_os) +apply (auto simp add:os_grant.simps current_sockets_def inum_of_socket.simps split:if_splits + dest:cn_in_curp') +done + +lemma current_sockets_execve: "valid (Execve p f fds # \) + \ current_sockets (Execve p f fds # \) = + current_sockets \ - {(p, fd)| fd. (p, fd) \ current_sockets \ \ fd \ fds}" +apply (frule vd_cons, drule vt_grant_os) +apply (auto simp add:os_grant.simps current_sockets_def inum_of_socket.simps split:if_splits + dest:cn_in_curp') +done + +lemma current_sockets_kill: "current_sockets (Kill p p' # \) = current_sockets \ - {(p', fd)| fd. (p', fd) \ current_sockets \}" +by (auto simp add:current_sockets_def inum_of_socket.simps split:if_splits) + +lemma current_sockets_exit: "current_sockets (Exit p # \) = current_sockets \ - {(p, fd)| fd. (p, fd) \ current_sockets \}" +by (auto simp add:current_sockets_def inum_of_socket.simps split:if_splits) + +lemma current_sockets_other: + "\\ p fd. e \ CloseFd p fd; + \ p aft st fd ino. e \ CreateSock p aft st fd ino; + \ p fd addr lport' fd' ino. e \ Accept p fd addr lport' fd' ino; + \ p p' fds shms. e \ Clone p p' fds shms; + \ p f fds. e \ Execve p f fds; + \ p p'. e \ Kill p p'; + \ p. e \ Exit p\ \ current_sockets (e # \) = current_sockets \" +apply (case_tac e) +by (auto simp add:current_sockets_def inum_of_socket.simps split:if_splits) + +lemmas current_sockets_simps = current_sockets_nil current_sockets_closefd current_sockets_createsock + current_sockets_accept current_sockets_execve current_sockets_clone + current_sockets_kill current_sockets_exit current_sockets_other + +(********** is_tcp_sock simpset ***************) + +lemma is_tcp_sock_nil: + "is_tcp_sock [] = is_init_tcp_sock" +apply (rule ext) +apply (auto simp add:is_init_tcp_sock_def is_tcp_sock_def init_inum_of_socket_props + dest:init_socket_has_inode split:option.splits) +done + +lemma is_tcp_sock_closefd: + "is_tcp_sock (CloseFd p fd # s) = (is_tcp_sock s) ((p, fd) := False)" +apply (rule ext) +by (auto simp add:is_tcp_sock_def inum_of_socket.simps split:option.splits t_inode_tag.splits) + +lemma is_tcp_sock_createsock: + "valid (CreateSock p af st fd ino # s) \ is_tcp_sock (CreateSock p af st fd ino # s) = ( + case st of + STREAM \ (is_tcp_sock s) ((p, fd) := True) + | _ \ is_tcp_sock s )" +apply (frule vd_cons, frule vt_grant_os, rule ext) +by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemma is_tcp_sock_accept: + "valid (Accept p fd addr lport fd' ino # s) + \ is_tcp_sock (Accept p fd addr lport fd' ino # s) = (is_tcp_sock s) ((p,fd') := True)" +apply (frule vd_cons, frule vt_grant_os, rule ext) +by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemma is_tcp_sock_execve: + "valid (Execve p f fds # s) \ is_tcp_sock (Execve p f fds # s) = ( + \ (p', fd). if (p' = p \ fd \ fds) then is_tcp_sock s (p, fd) + else if (p' = p) then False + else is_tcp_sock s (p', fd))" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemma is_tcp_sock_clone: + "valid (Clone p p' fds shms # s) \ is_tcp_sock (Clone p p' fds shms # s) = ( + \ (p'', fd). if (p'' = p' \ fd \ fds) then is_tcp_sock s (p, fd) + else if (p'' = p') then False + else is_tcp_sock s (p'', fd))" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemma is_tcp_sock_kill: + "valid (Kill p p' # s) \ is_tcp_sock (Kill p p' # s) = ( + \ (p'', fd). if (p'' = p') then False + else is_tcp_sock s (p'', fd))" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemma is_tcp_sock_exit: + "valid (Exit p # s) \ is_tcp_sock (Exit p # s) = ( + \ (p', fd). if (p' = p) then False + else is_tcp_sock s (p', fd))" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemma is_tcp_sock_other: + "\valid (e # \); + \ p fd. e \ CloseFd p fd; + \ p aft st fd ino. e \ CreateSock p aft st fd ino; + \ p fd addr lport' fd' ino. e \ Accept p fd addr lport' fd' ino; + \ p p' fds shms. e \ Clone p p' fds shms; + \ p f fds. e \ Execve p f fds; + \ p p'. e \ Kill p p'; + \ p. e \ Exit p\ \ is_tcp_sock (e # \) = is_tcp_sock \" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x, case_tac e) +prefer 6 apply (case_tac option) +by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemmas is_tcp_sock_simps = is_tcp_sock_nil is_tcp_sock_closefd is_tcp_sock_createsock is_tcp_sock_accept + is_tcp_sock_clone is_tcp_sock_execve is_tcp_sock_kill is_tcp_sock_exit is_tcp_sock_other + + +(************ is_udp_sock simpset *************) + +lemma is_udp_sock_nil: + "is_udp_sock [] = is_init_udp_sock" +apply (rule ext) +apply (auto simp add:is_init_udp_sock_def is_udp_sock_def init_inum_of_socket_props + dest:init_socket_has_inode split:option.splits) +done + +lemma is_udp_sock_closefd: + "is_udp_sock (CloseFd p fd # s) = (is_udp_sock s) ((p, fd) := False)" +apply (rule ext) +by (auto simp add:is_udp_sock_def inum_of_socket.simps split:option.splits t_inode_tag.splits) + +lemma is_udp_sock_createsock: + "valid (CreateSock p af st fd ino # s) \ is_udp_sock (CreateSock p af st fd ino # s) = ( + case st of + DGRAM \ (is_udp_sock s) ((p, fd) := True) + | _ \ is_udp_sock s )" +apply (frule vd_cons, frule vt_grant_os, rule ext) +by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemma is_udp_sock_execve: + "valid (Execve p f fds # s) \ is_udp_sock (Execve p f fds # s) = ( + \ (p', fd). if (p' = p \ fd \ fds) then is_udp_sock s (p, fd) + else if (p' = p) then False + else is_udp_sock s (p', fd))" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemma is_udp_sock_clone: + "valid (Clone p p' fds shms # s) \ is_udp_sock (Clone p p' fds shms # s) = ( + \ (p'', fd). if (p'' = p' \ fd \ fds) then is_udp_sock s (p, fd) + else if (p'' = p') then False + else is_udp_sock s (p'', fd))" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemma is_udp_sock_kill: + "valid (Kill p p' # s) \ is_udp_sock (Kill p p' # s) = ( + \ (p'', fd). if (p'' = p') then False + else is_udp_sock s (p'', fd))" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemma is_udp_sock_exit: + "valid (Exit p # s) \ is_udp_sock (Exit p # s) = ( + \ (p', fd). if (p' = p) then False + else is_udp_sock s (p', fd))" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x) +by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemma is_udp_sock_other: + "\valid (e # \); + \ p fd. e \ CloseFd p fd; + \ p aft st fd ino. e \ CreateSock p aft st fd ino; + \ p p' fds shms. e \ Clone p p' fds shms; + \ p f fds. e \ Execve p f fds; + \ p p'. e \ Kill p p'; + \ p. e \ Exit p\ \ is_udp_sock (e # \) = is_udp_sock \" +apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x, case_tac e) +prefer 6 apply (case_tac option) +by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd' + split:option.splits t_inode_tag.splits t_socket_type.splits) + +lemmas is_udp_sock_simps = is_udp_sock_nil is_udp_sock_closefd is_udp_sock_createsock + is_udp_sock_clone is_udp_sock_execve is_udp_sock_kill is_udp_sock_exit is_udp_sock_other + +lemma update_with_shuthow_someI: "socket_in_trans \ s \ \ st'. update_with_shuthow (socket_state \ s) how = Some st'" +apply (case_tac "socket_state \ s", simp add:socket_in_trans_def, case_tac a) +apply (auto simp:update_with_shuthow.simps socket_in_trans_def split:option.splits) +apply (case_tac t_sock_trans_state, case_tac [!] how) +by auto + +(***** below should be adjusted after starting static analysis of sockets ! ??? ****) + +(* +lemma cn_has_sockstate: "\(p, fd) \ current_sockets \; valid \\ \ socket_state \ (p, fd) \ None" +apply (induct \ arbitrary: p) using init_socket_has_state +apply (simp add:current_sockets_simps socket_state.simps bidirect_in_init_def) + +apply (frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:current_sockets_simps os_grant.simps socket_state.simps intro:update_with_shuthow_someI dest:cn_no_newproc' split:option.splits) +done + +lemma after_bind_update: "\after_bind (update_with_shuthow (socket_state \ s) t_shutdown_how); valid \\ \ after_bind (socket_state \ s)" +apply (case_tac "socket_state \ s", simp, case_tac a) +apply (auto simp:update_with_shuthow.simps split:option.splits) +done + +lemma after_bind_has_laddr_aux: "\ p fd. after_bind (socket_state \ (p, fd)) \ valid \ \ (\ ip port. local_addr \ (p, fd) = Some (INET_ADDR ip port))" +apply (induct \) +apply (clarsimp simp add:socket_state.simps local_addr.simps) +apply (drule init_socket_has_laddr, simp, erule exE) +apply (case_tac y, simp) + +apply (clarify, frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:socket_state.simps local_addr.simps os_grant.simps dest:after_bind_update split:if_splits option.splits t_sock_addr.splits) +apply (erule_tac x = nat1 in allE, erule_tac x = nat2 in allE, erule impE, simp+) +done + +lemma after_bind_has_laddr: "\after_bind (socket_state \ (p, fd)); valid \\ \ \ ip port. local_addr \ (p, fd) = Some (INET_ADDR ip port)" +by (rule after_bind_has_laddr_aux[rule_format], simp+) + +lemma after_bind_has_laddr': "\after_bind (socket_state \ s); valid \\ \ \ ip port. local_addr \ s = Some (INET_ADDR ip port)" +by (case_tac s, auto dest:after_bind_has_laddr) + +lemma after_bind_has_laddr_D: "\local_addr \ s = None; valid \\ \ \ after_bind (socket_state \ s)" +by (auto dest:after_bind_has_laddr') + +lemma local_addr_valid: "\local_addr \ (p, fd) = Some (INET_ADDR ip port); (p, fd) \ current_sockets \; valid \\ \ ip \ local_ipaddrs" +apply (induct \ arbitrary:p fd port) +apply (simp add:local_addr.simps current_sockets_simps local_ip_valid) + +apply (frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:current_sockets_simps local_addr.simps os_grant.simps split:if_splits option.splits t_sock_addr.splits dest:cn_no_newproc' after_bind_has_laddr_D) +done + +lemma local_addr_valid': "\local_addr \ s = Some (INET_ADDR ip port); s \ current_sockets \; valid \\ \ ip \ local_ipaddrs" +by (case_tac s, simp add:local_addr_valid) + +lemma local_ip_has_dev_D: "\local_ip2dev ip = None; local_addr \ s = Some (INET_ADDR ip port); s \ current_sockets \; valid \\ \ False" +using local_ip2dev_valid +by (drule_tac local_addr_valid', simp+) + +lemma after_bind_has_netdev: "\after_bind (socket_state \ s); s \ current_sockets \; valid \\ \ \ dev. netdev_of_socket \ s = Some dev" +apply (drule after_bind_has_laddr', simp, clarsimp) +apply (frule local_addr_valid', simp+) +apply (drule local_ip2dev_valid, simp add:netdev_of_socket_def) +done + +lemma after_bind_has_netdev_D: "\netdev_of_socket \ s = None; s \ current_sockets \; valid \\ \ \ after_bind (socket_state \ s)" +by (auto dest:after_bind_has_netdev) + +lemma after_conacc_update: "\after_conacc (update_with_shuthow (socket_state \ s) t_shutdown_how); valid \\ \ after_conacc (socket_state \ s)" +apply (case_tac "socket_state \ s", simp, case_tac a) +apply (auto simp:update_with_shuthow.simps split:option.splits) +done + +lemma after_conacc_has_raddr: "\after_conacc (socket_state \ (p, fd)); valid \\ \ \ addr. remote_addr \ (p, fd) = Some addr" +apply (induct \ arbitrary:p fd) +apply (simp add:socket_state.simps local_addr.simps) using init_socket_has_raddr apply blast + +apply (frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:socket_state.simps local_addr.simps os_grant.simps dest:after_conacc_update split:if_splits option.splits t_sock_addr.splits) +done + +lemma after_conacc_has_raddr': "\after_conacc (socket_state \ s); valid \\ \ \ addr. remote_addr \ s = Some addr" +by (case_tac s, auto dest:after_conacc_has_raddr) + +lemma after_conacc_has_raddr_D: "\remote_addr \ s = None; valid \\ \ \ after_conacc (socket_state \ s)" +by (auto dest:after_conacc_has_raddr') + + +lemma cn_has_socktype: "\(p, fd) \ current_sockets \; valid \\ \ \ st. socket_type \ (p, fd) = Some st" +apply (induct \ arbitrary:p fd) +using init_socket_has_socktype +apply (simp add:socket_type.simps current_sockets_nil bidirect_in_init_def) + +apply (frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:socket_type.simps current_sockets_simps os_grant.simps dest:cn_no_newproc' split:if_splits option.splits t_sock_addr.splits) +done + +lemma accept_has_temps: "Accept p fd\<^isub>1 addr port fd\<^isub>2 inum # valid \ \ \ temp. nettemp_of_socket (Accept p fd\<^isub>1 addr port fd\<^isub>2 inum # \) (p, fd\<^isub>2) is_remote = Some temp" +apply (frule vd_cons, frule vt_grant_os) +apply (auto dest:cn_has_socktype simp:os_grant.simps nettemp_of_socket_def netdev_of_socket_def + dest!:after_conacc_has_raddr_D after_bind_has_laddr_D local_ip_has_dev_D + simp:socket_state.simps split:if_splits option.splits t_sock_addr.splits) +done + +lemma accept_has_temps': + "Accept p fd\<^isub>1 addr port fd\<^isub>2 inum # valid \ + \ (\ ltemp. nettemp_of_socket (Accept p fd\<^isub>1 addr port fd\<^isub>2 inum # \) (p, fd\<^isub>2) False = Some ltemp) \ + (\ rtemp. nettemp_of_socket (Accept p fd\<^isub>1 addr port fd\<^isub>2 inum # \) (p, fd\<^isub>2) True = Some rtemp)" +apply (frule accept_has_temps[where is_remote = True]) +apply (frule accept_has_temps[where is_remote = False]) +by auto + +lemma bind_has_ltemp: "Bind p fd addr # valid \ \ \ temp. nettemp_of_socket (Bind p fd addr # \) (p, fd) False = Some temp" +apply (frule vd_cons, frule vt_grant_os) +apply (auto dest:cn_has_socktype simp:os_grant.simps current_sockets_simps nettemp_of_socket_def dest!:after_bind_has_netdev_D after_bind_has_laddr_D + simp:socket_state.simps split:if_splits option.splits t_sock_addr.splits) +done + +lemma connect_has_rtemp: "Connect p fd addr # valid \ \ \ temp. nettemp_of_socket (Connect p fd addr # \) (p, fd) True = Some temp" +apply (frule vd_cons, frule vt_grant_os) +apply (auto dest:cn_has_socktype simp:os_grant.simps current_sockets_simps nettemp_of_socket_def dest!:after_bind_has_netdev_D after_bind_has_laddr_D + simp:socket_state.simps split:if_splits option.splits t_sock_addr.splits) +done +*) + + +end + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Delete_prop.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Delete_prop.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,53 @@ +theory Deleted_prop +imports Main Flask Flask_type Init_prop Alive_prop +begin + +context flask begin + +lemma deleted_cons_I: "deleted obj s \ deleted obj (e # s)" +by (case_tac e, auto) + +lemma not_deleted_cons_D: "\ deleted obj (e # s) \ \ deleted obj s" +by (auto dest:deleted_cons_I) + +lemma not_deleted_imp_exists: + "\\ deleted obj s; valid s; init_alive obj\ \ alive s obj" +apply (induct s, simp add:init_alive_prop) +apply (frule vd_cons) +apply (case_tac a, auto simp:alive_simps split:t_object.splits) +pr 19 +done + +lemma cons_app_simp_aux: + "(a # b) @ c = a # (b @ c)" by auto + +lemma not_deleted_imp_exists': + "\\ deleted obj (s'@s); exists s obj\ \ exists (s'@s) obj" +apply (induct s', simp, simp only:cons_app_simp_aux) +apply (frule not_deleted_cons_D) +apply (case_tac a, case_tac [!] obj, auto) +done + +(* + +lemma nodel_imp_un_deleted: + "no_del_event s \ \ deleted obj s" +by (induct s, simp, case_tac a,auto) + +lemma nodel_exists_remains: + "\no_del_event (s'@s); exists s obj\ \ exists (s'@s) obj" +apply (drule_tac obj = obj in nodel_imp_un_deleted) +by (simp add:not_deleted_imp_exists') + +lemma nodel_imp_exists: + "\no_del_event s; exists [] obj\ \ exists s obj" +apply (drule_tac obj = obj in nodel_imp_un_deleted) +by (simp add:not_deleted_imp_exists) + +lemma no_del_event_cons_D: + "no_del_event (e # s) \ no_del_event s" +by (case_tac e, auto) +*) +end + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Dynamic2static.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Dynamic2static.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,205 @@ +theory Dynamic2static +imports Main Flask Static Init_prop Valid_prop +begin + +context tainting_s begin + +lemma d2s_main: + "valid s \ s2ss s \ static" +apply (induct s, simp add:s2ss_nil_prop s_init) +apply (frule vd_cons, simp) +apply (case_tac a, simp_all) +(* +apply +induct s, case tac e, every event analysis +*) +sorry + +lemma is_file_has_sfile: "is_file s f \ \ sf. cf2sfile s f True = Some sf" +sorry + +lemma is_dir_has_sfile: "is_dir s f \ \ sf. cf2sfile s f False = Some sf" +sorry + +lemma is_file_imp_alive: "is_file s f \ alive s (O_file f)" +sorry + + +lemma d2s_main': + "\alive s obj; co2sobj s obj= Some sobj\ \ sobj \ (s2ss s)" +apply (induct s) +apply (simp add:s2ss_def) +apply (rule_tac x = obj in exI, simp) +sorry + +lemma tainted_prop1: + "obj \ tainted s \ alive s obj" +sorry + +lemma tainted_prop2: + "obj \ tainted s \ valid s" +sorry + +lemma alive_has_sobj: + "\alive s obj; valid s\ \ \ sobj. co2sobj s obj = Some sobj" +sorry + +lemma t2ts: + "obj \ tainted s \ co2sobj s obj = Some sobj \ tainted_s (s2ss s) sobj" +apply (frule tainted_prop1, frule tainted_prop2) +apply (simp add:s2ss_def) +apply (case_tac sobj, simp_all) +apply (case_tac [!] obj, simp_all split:option.splits) +apply (rule_tac x = "O_proc nat" in exI, simp) +apply (rule_tac x = "O_file list" in exI, simp) +defer defer defer +apply (case_tac prod1, simp, case_tac prod2, clarsimp) +apply (rule conjI) +apply (rule_tac x = "O_msgq nat1" in exI, simp) +sorry (* doable, need properties about cm2smsg and cq2smsgq *) + +lemma delq_imp_delqm: + "deleted (O_msgq q) s \ deleted (O_msg q m) s" +apply (induct s, simp) +by (case_tac a, auto) + +lemma undel_init_file_remains: + "\is_init_file f; \ deleted (O_file f) s\ \ is_file s f" +sorry + + +theorem static_complete: + assumes undel: "undeletable obj" and tbl: "taintable obj" + shows "taintable_s obj" +proof- + from tbl obtain s where tainted: "obj \ tainted s" + by (auto simp:taintable_def) + hence vs: "valid s" by (simp add:tainted_prop2) + hence static: "s2ss s \ static" using d2s_main by auto + from tainted have alive: "alive s obj" + using tainted_prop1 by auto + then obtain sobj where sobj: "co2sobj s obj = Some sobj" + using vs alive_has_sobj by blast + from undel vs have "\ deleted obj s" and init_alive: "init_alive obj" + by (auto simp:undeletable_def) + with vs sobj have "init_obj_related sobj obj" + apply (case_tac obj, case_tac [!] sobj) + apply (auto split:option.splits if_splits simp:cp2sproc_def ch2sshm_def cq2smsgq_def cm2smsg_def) + apply (frule undel_init_file_remains, simp, drule is_file_has_sfile, erule exE) + apply (rule_tac x = sf in bexI) + apply (case_tac list, auto split:option.splits simp:is_init_file_props)[1] + apply (simp add:same_inode_files_def cfs2sfiles_def) + apply (rule_tac x = list in exI, simp) + apply (case_tac list, auto split:option.splits simp:is_init_dir_props delq_imp_delqm) + done + with tainted t2ts init_alive sobj static + show ?thesis unfolding taintable_s_def + apply (rule_tac x = "s2ss s" in bexI, simp) + apply (rule_tac x = "sobj" in exI, auto) + done +qed + +lemma init_deled_imp_deled_s: + "\deleted obj s; init_alive obj; sobj \ (s2ss s); valid s\ \ \ init_obj_related sobj obj" +apply (induct s, simp) +apply (frule vd_cons) +apply (case_tac a, auto) +(* need simpset for s2ss *) +sorry + +lemma deleted_imp_deletable_s: + "\deleted obj s; init_alive obj; valid s\ \ deletable_s obj" +apply (simp add:deletable_s_def) +apply (rule_tac x = "s2ss s" in bexI) +apply (clarify, simp add:init_deled_imp_deled_s) +apply (erule d2s_main) +done + +theorem undeletable_s_complete: + assumes undel_s: "undeletable_s obj" + shows "undeletable obj" +proof- + from undel_s have init_alive: "init_alive obj" + and alive_s: "\ ss \ static. \ sobj \ ss. init_obj_related sobj obj" + using undeletable_s_def by auto + have "\ (\ s. valid s \ deleted obj s)" + proof + assume "\ s. valid s \ deleted obj s" + then obtain s where vs: "valid s" and del: "deleted obj s" by auto + from vs have vss: "s2ss s \ static" by (rule d2s_main) + with alive_s obtain sobj where in_ss: "sobj \ (s2ss s)" + and related: "init_obj_related sobj obj" by auto + from init_alive del vs have "deletable_s obj" + by (auto elim:deleted_imp_deletable_s) + with alive_s + show False by (auto simp:deletable_s_def) + qed + with init_alive show ?thesis + by (simp add:undeletable_def) +qed + +theorem final_offer: + "\undeletable_s obj; \ taintable_s obj; init_alive obj\ \ \ taintable obj" +apply (erule swap) +by (simp add:static_complete undeletable_s_complete) + + + +(************** static \ dynamic ***************) + +lemma created_can_have_many: + "\valid s; alive s obj; \ init_alive obj\ \ \ s'. valid s' \ alive s' obj \ alive s' obj' \ s2ss s = s2ss s'" +sorry + +lemma s2d_main: + "ss \ static \ \ s. valid s \ s2ss s = ss" +apply (erule static.induct) +apply (rule_tac x = "[]" in exI, simp add:s2ss_nil_prop valid.intros) + +apply (erule exE|erule conjE)+ + +apply (erule exE, erule conjE)+ + +sorry + + + +lemma tainted_s_imp_tainted: + "\tainted_s ss sobj; ss \ static\ \ \ obj s. s2ss s = ss \ valid s \ co2sobj s obj = Some sobj \ obj \ tainted s" +sorry + + +theorem static_sound: + assumes tbl_s: "taintable_s obj" + shows "taintable obj" +proof- + from tbl_s obtain ss sobj where static: "ss \ static" + and sobj: "tainted_s ss sobj" and related: "init_obj_related sobj obj" + and init_alive: "init_alive obj" by (auto simp:taintable_s_def) + from static sobj tainted_s_imp_tainted + obtain s obj' where s2ss: "s2ss s = ss" and co2sobj: "co2sobj s obj' = Some sobj" + and tainted: "obj' \ tainted s" and vs: "valid s" by blast + + from co2sobj related + have eq:"obj = obj'" + apply (case_tac obj', case_tac [!] obj, case_tac [!] sobj) + apply auto + apply (auto split:option.splits if_splits) + apply (case_tac a, simp+) + apply (simp add:cp2sproc_def split:option.splits if_splits) + apply simp + sorry + with tainted vs init_alive + show ?thesis by (auto simp:taintable_def) +qed + + + +lemma ts2t: + "obj \ tainted_s ss \ \ s. obj \ tainted s" + "obj \ tainted_s ss \ \ so. so True \ ss \ so True \ ss \ \ s. valid s \ s2ss s = ss \ so True \ s2ss s \ tainted s obj. " + + + + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 File_renaming.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/File_renaming.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,221 @@ +theory File_renaming +imports Main OS_type_def My_list_prefix +begin + +fun parent :: "'a list \ ('a list) option" +where + "parent [] = None" + | "parent (n#ns) = Some ns" + +definition file_after_rename :: "'a list \ 'a list \ 'a list \ 'a list" +where + "file_after_rename oldf newf f \ (if (oldf \ f) then ((f \ oldf) @ newf) else f)" + +definition file_before_rename :: "'a list \ 'a list \ 'a list \ 'a list" +where + "file_before_rename oldf newf f \ (if (newf \ f) then ((f \ newf) @ oldf) else f)" + +lemma list_diff_rev_prop0: "f \ f = []" +by (induct f, auto simp:list_diff_rev_def) + +lemma list_diff_rev_prop0': "f \ [] = f" +apply (simp add:list_diff_rev_def) +apply (induct f rule:rev_induct, simp+) +done + +lemma list_diff_rev_prop1_aux: "a # f\<^isub>3 \ f\<^isub>3 = [a]" +by (induct f\<^isub>3, auto simp:list_diff_rev_def) + +lemma list_diff_rev_prop1: "f @ f' \ f' = f" +apply (subgoal_tac "f @ f' \ [] @ f' = f \ []") +apply (simp add:list_diff_rev_prop0') +apply (simp only: list_diff_rev_ext_left[where ys = "[]"]) +done + +lemma list_diff_rev_prop2_aux: "f \ f' \ a # f' \ f = a # (f' \ f)" +apply (erule no_juniorE) +by (simp only:list_diff_rev_prop1 append_Cons[THEN sym]) + +lemma list_diff_rev_prop2: "f \ f' \ (f' \ f) @ f = f'" +apply (induct f', simp add:no_junior_def list_diff_rev_def) +apply (case_tac "f = a # f'", simp add:list_diff_rev_prop0) +apply (drule no_junior_noteq, simp) +by (simp add:list_diff_rev_prop2_aux) + +lemma file_renaming_prop0: + "file_after_rename f\<^isub>2 f\<^isub>3 f\<^isub>2 = f\<^isub>3" +apply (simp add:file_after_rename_def list_diff_rev_def) +apply (induct f\<^isub>2, auto) +done + +lemma file_renaming_prop0': + "file_before_rename f\<^isub>2 f\<^isub>3 f\<^isub>3 = f\<^isub>2" +apply (simp add:file_before_rename_def list_diff_rev_def) +apply (induct f\<^isub>3, auto) +done + +lemma file_renaming_prop1: + "f\<^isub>2 \ f \ f\<^isub>3 \ file_after_rename f\<^isub>2 f\<^isub>3 f " +by (simp add:file_after_rename_def) + +lemma file_renaming_prop1': + "f\<^isub>3 \ f \ f\<^isub>2 \ file_before_rename f\<^isub>2 f\<^isub>3 f" +by (simp add:file_before_rename_def) + +lemma file_renaming_prop1'': + "\ f\<^isub>3 \ file_after_rename f\<^isub>2 f\<^isub>3 f \ \ f\<^isub>2 \ f" +by (case_tac "f\<^isub>2 \ f", drule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop1, simp+) + +lemma file_renaming_prop2: + "\f\<^isub>2 \ f; f \ f\<^isub>2\ \ file_after_rename f\<^isub>2 f\<^isub>3 f \ f\<^isub>3" +apply (case_tac f, simp add:no_junior_def) +apply (auto simp add:file_after_rename_def list_diff_rev_prop1_aux) +apply (erule no_juniorE, simp add:list_diff_rev_prop1) +done + +lemma file_renaming_prop2': + "\f\<^isub>3 \ f; f \ f\<^isub>3\ \ file_before_rename f\<^isub>2 f\<^isub>3 f \ f\<^isub>2" +apply (case_tac f, simp add:no_junior_def) +apply (auto simp add:file_before_rename_def list_diff_rev_prop1_aux) +apply (erule no_juniorE, simp add:list_diff_rev_prop1) +done + +lemma file_renaming_prop3: + "\f\<^isub>2 \ f; f\<^isub>2 \ f\ \ f\<^isub>3 \ file_after_rename f\<^isub>2 f\<^isub>3 f \ f\<^isub>3 \ file_after_rename f\<^isub>2 f\<^isub>3 f" +apply (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop1, simp) +by (drule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop2, simp+) + +lemma file_renaming_prop3': + "\f\<^isub>3 \ f; f\<^isub>3 \ f\ \ f\<^isub>2 \ file_before_rename f\<^isub>2 f\<^isub>3 f \ f\<^isub>2 \ file_before_rename f\<^isub>2 f\<^isub>3 f" +apply (frule_tac f\<^isub>2 = f\<^isub>2 in file_renaming_prop1', simp) +by (drule_tac f\<^isub>2 = f\<^isub>2 in file_renaming_prop2', simp+) + +lemma file_renaming_prop4: "\ f\<^isub>2 \ f \ file_after_rename f\<^isub>2 f\<^isub>3 f = f" +by (simp add:file_after_rename_def) + +lemma file_renaming_prop4': "\ f\<^isub>3 \ f \ file_before_rename f\<^isub>2 f\<^isub>3 f = f" +by (simp add:file_before_rename_def) + +lemma file_renaming_prop5: + "f\<^isub>2 \ f\<^isub>1 \ file_before_rename f\<^isub>2 f\<^isub>3 (file_after_rename f\<^isub>2 f\<^isub>3 f\<^isub>1) = f\<^isub>1" +apply (case_tac "f\<^isub>1 = f\<^isub>2") +apply (simp add:file_renaming_prop0 file_renaming_prop0') +apply (simp add:file_after_rename_def file_before_rename_def) +by (simp add:list_diff_rev_prop1 list_diff_rev_prop2) + +lemma file_renaming_prop5': + "f\<^isub>3 \ f\<^isub>1 \ file_after_rename f\<^isub>2 f\<^isub>3 (file_before_rename f\<^isub>2 f\<^isub>3 f\<^isub>1) = f\<^isub>1" +apply (case_tac "f\<^isub>1 = f\<^isub>3") +apply (simp add:file_renaming_prop0' file_renaming_prop0) +apply (simp add:file_after_rename_def file_before_rename_def) +by (simp add:list_diff_rev_prop1 list_diff_rev_prop2) + +lemma file_renaming_prop6: + "f\<^isub>2 \ f \ file_after_rename f\<^isub>2 f\<^isub>3 (a # f) = (a # (file_after_rename f\<^isub>2 f\<^isub>3 f))" +by (simp add:file_after_rename_def list_diff_rev_prop2_aux) + +lemma file_renaming_prop6': + "f\<^isub>3 \ f \ file_before_rename f\<^isub>2 f\<^isub>3 (a # f) = (a # (file_before_rename f\<^isub>2 f\<^isub>3 f))" +by (simp add:file_before_rename_def list_diff_rev_prop2_aux) + +lemma file_renaming_prop7: "file_after_rename f\<^isub>2 f\<^isub>3 (a # f\<^isub>2) = a # f\<^isub>3" +apply (auto simp:file_after_rename_def no_junior_def list_diff_rev_def) +by (induct f\<^isub>2, auto) + +lemma file_renaming_prop7': "file_before_rename f\<^isub>2 f\<^isub>3 (a # f\<^isub>3) = a # f\<^isub>2" +by (auto simp:file_before_rename_def no_junior_def list_diff_rev_prop1_aux) + +lemma file_renaming_prop8: "f\<^isub>2 \ f \ f\<^isub>3 \ (file_after_rename f\<^isub>2 f\<^isub>3 (a # f))" +by (simp add:file_after_rename_def no_junior_def) + +lemma file_renaming_prop8': "f\<^isub>3 \ f \ f\<^isub>2 \ (file_before_rename f\<^isub>2 f\<^isub>3 (a # f))" +by (simp add:file_before_rename_def no_junior_def) + +lemma no_junior_prop0: "f\ f' \ f \ (a # f')" +by (simp add:no_junior_def) + +lemma file_renaming_prop9: + "f\<^isub>2 \ f\<^isub>1 \ file_before_rename f\<^isub>2 f\<^isub>3 (file_after_rename f\<^isub>2 f\<^isub>3 (a # f\<^isub>1)) = (a # f\<^isub>1)" +by (drule_tac a = a in no_junior_prop0, simp add:file_renaming_prop5) + +lemma file_renaming_prop9': + "f\<^isub>3 \ f\<^isub>1 \ file_after_rename f\<^isub>2 f\<^isub>3 (file_before_rename f\<^isub>2 f\<^isub>3 (a # f\<^isub>1)) = (a # f\<^isub>1)" +by (drule_tac a = a in no_junior_prop0, simp add:file_renaming_prop5') + +lemma file_renaming_prop10: + "\a # f\<^isub>1 = file_after_rename fx fy f\<^isub>1'; file_after_rename fx fy f\<^isub>1' \ fy; fx \ f\<^isub>1'\ \ a # (file_before_rename fx fy f\<^isub>1) = f\<^isub>1'" +apply (frule_tac f\<^isub>3 = fy in file_renaming_prop1) +apply (subgoal_tac "fy \ f\<^isub>1") defer apply (rotate_tac 3, erule no_juniorE, case_tac zs, simp+) +apply (drule_tac f = f\<^isub>1 and f\<^isub>2 = fx and f\<^isub>3 = fy and a = a in file_renaming_prop6') +by (drule_tac f\<^isub>1 = f\<^isub>1' and f\<^isub>3 = fy in file_renaming_prop5, simp) + +lemma file_renaming_prop11: + "\f\<^isub>2 \ f; file_after_rename f\<^isub>2 f\<^isub>3 f = f'\ \ f = file_before_rename f\<^isub>2 f\<^isub>3 f'" +by (drule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop5, simp) + +lemma file_renaming_prop11': + "\f\<^isub>3 \ f; file_before_rename f\<^isub>2 f\<^isub>3 f = f'\ \ f = file_after_rename f\<^isub>2 f\<^isub>3 f'" +by (drule_tac f\<^isub>2 = f\<^isub>2 in file_renaming_prop5', simp) + +lemma file_renaming_prop12: + "\\ f\<^isub>3 \ aa; f\<^isub>3 \ f; file_after_rename f\<^isub>2 f\<^isub>3 aa = f\ \ file_before_rename f\<^isub>2 f\<^isub>3 f = aa" +apply (simp add:file_after_rename_def file_before_rename_def split:if_splits) +by (auto, simp add:list_diff_rev_prop1 list_diff_rev_prop2) + +lemma file_renaming_prop12': + "\\ f\<^isub>2 \ aa; f\<^isub>2 \ f; file_before_rename f\<^isub>2 f\<^isub>3 aa = f\ \ file_after_rename f\<^isub>2 f\<^isub>3 f = aa" +apply (simp add:file_after_rename_def file_before_rename_def split:if_splits) +by (auto, simp add:list_diff_rev_prop1 list_diff_rev_prop2) + + +lemma file_renaming_prop13: + "\f\<^isub>3 \ f; file_before_rename f\<^isub>2 f\<^isub>3 f = f\ \ f\<^isub>2 \ f" +apply (simp add:file_before_rename_def) +apply (rule_tac zs = "f \ f\<^isub>3" in no_juniorI) +by simp + +(************ something *****************) + + +lemma parent_is_parent: "parent f = Some pf \ f \ pf" +by (case_tac f, auto) + +lemma parent_is_parent': "parent f = Some f \ false" +by (drule parent_is_parent, simp) + +lemma parent_is_parent'': "parent f \ Some f" +by (case_tac f, auto) + +lemma parent_is_parent3: "parent f = Some f' \ \ f \ f'" +by (case_tac f, auto elim:no_juniorE) + +lemma parent_is_ancen: "parent f = Some f' \ f' \ f" +by (case_tac f, auto simp:no_junior_def) + +lemma parent_is_ancen': "\parent f = Some pf; f' \ pf\ \ f' \ f" +by (case_tac f, auto simp:no_junior_def) + +lemma parent_ancen: "\parent f = Some pf; f' \ f; f \ f'\ \ f' \ pf" +apply (erule no_juniorE) +apply (case_tac zs, simp) +apply (rule_tac zs = list in no_juniorI) +by auto + +lemma parent_rename_ancen: "\parent f = Some pf; f\<^isub>3 \ pf; f \ f\<^isub>3\ \ parent (file_before_rename f\<^isub>2 f\<^isub>3 f) = Some (file_before_rename f\<^isub>2 f\<^isub>3 pf)" +apply (simp add:file_before_rename_def) +apply (case_tac f, simp, simp) +by (drule_tac a = a in list_diff_rev_prop2_aux, simp) + +lemma no_junior_trans: "\f \ f'; f' \ f''\ \ f \ f''" +by (auto elim:no_juniorE) + +lemma no_junior_conf: "\a \ c; b \ c\ \ a \ b \ b \ a" +apply (simp add:no_junior_def) +apply (induct c, auto) +done + +lemma noJ_Anc: "x \ y = (x \ y \ x \ y)" +apply (simp add: no_junior_expand) +by (auto simp:is_ancestor_def) + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Finite_current.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Finite_current.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,109 @@ +theory Finite_current +imports Main Valid_prop Flask Flask_type +begin + +context flask begin + +lemma finite_cf: "valid \ \ finite (current_files \)" +apply (induct \) +apply (simp add:current_files_def inum_of_file.simps) +apply (rule_tac B = "init_files" in finite_subset) +apply (clarsimp dest!:inof_has_file_tag, simp add:init_finite_sets) + +apply (frule vt_grant_os, frule vd_cons, simp, case_tac a) + +apply (auto simp:current_files_def os_grant.simps inum_of_file.simps split:if_splits option.splits) +apply (rule_tac B = "insert list {f. \i. inum_of_file \ f = Some i}" in finite_subset, clarsimp, simp) +apply (rule_tac B = "{f. \i. inum_of_file \ f = Some i}" in finite_subset, clarsimp, simp) +apply (rule_tac B = "{f. \i. inum_of_file \ f = Some i}" in finite_subset, clarsimp, simp) +apply (rule_tac B = "{f. \i. inum_of_file \ f = Some i}" in finite_subset, clarsimp, simp) +apply (rule_tac B = "insert list {f. \i. inum_of_file \ f = Some i}" in finite_subset, clarsimp, simp) +apply (rule_tac B = "insert list2 {f. \i. inum_of_file \ f = Some i}" in finite_subset, clarsimp, simp) +done + +lemma finite_cp: "finite (current_procs \)" +apply (induct \) +apply (simp add:current_procs.simps init_finite_sets) +apply (case_tac a, auto simp:current_procs.simps) +done + +lemma finite_cfd: "valid \ \ finite (current_proc_fds \ p)" +apply (induct \ arbitrary:p) +apply (simp add:current_proc_fds.simps init_finite_sets) +apply (frule vd_cons, frule vt_grant_os, case_tac a, auto simp:current_proc_fds.simps) +apply (metis double_diff equalityE finite_Diff) +by (metis double_diff finite_Diff subset_refl) + +lemma finite_pair: "\finite A; finite B\ \ finite {(x, y). x \ A \ y \ B}" +by auto + +lemma finite_UN_I': "\finite X; \ x. x \ X \ finite (f x)\ \ finite {(x, y). x \ X \ y \ f x}" +apply (frule_tac B = f in finite_UN_I, simp) +apply (drule_tac finite_pair, simp) +apply (rule_tac B = "{(x, y). x \ X \ y \ (\a\X. f a)}" in finite_subset, auto) +done + +lemma finite_init_netobjs: "finite init_sockets" +apply (subgoal_tac "finite {(p, fd). p \ init_procs \ fd \ init_fds_of_proc p}") +apply (rule_tac B = "{(p, fd). p \ init_procs \ fd \ init_fds_of_proc p}" in finite_subset) +apply (clarsimp dest!:init_socket_has_inode, simp) +using init_finite_sets finite_UN_I' +by (metis Collect_mem_eq SetCompr_Sigma_eq internal_split_def) + +lemma finite_cn_aux: "valid \ \ finite {s. \i. inum_of_socket \ s = Some i}" +apply (induct \) +apply (rule_tac B = "init_sockets" in finite_subset) +apply (clarsimp simp:inum_of_socket.simps dest!:inos_has_sock_tag, simp add:finite_init_netobjs) + +apply (frule vd_cons, frule vt_grant_os, simp, case_tac a) +apply (auto split:option.splits if_splits) +apply (rule_tac B = "{s. \i. inum_of_socket \ s = Some i}" in finite_subset, clarsimp split:if_splits, simp) +apply (rule_tac B = "{s. \i. inum_of_socket \ s = Some i} \ {(p, fd). \ i. inum_of_socket \ (nat1, fd) = Some i \ p = nat2 \ fd \ set1}" in finite_subset, clarsimp split:if_splits) +apply (simp only:finite_Un, rule conjI, simp) +apply (rule_tac B = "{(p, fd). \ i. inum_of_socket \ (nat1, fd) = Some i \ p = nat2}" in finite_subset, clarsimp) +apply (drule_tac h = "\ (p, fd). if (p = nat1) then (nat2, fd) else (p, fd)" in finite_imageI) +apply (rule_tac B = "((\(p, fd). if p = nat1 then (nat2, fd) else (p, fd)) ` {a. \i. inum_of_socket \ a = Some i})" in finite_subset) +apply (rule subsetI,erule CollectE, case_tac x, simp, (erule exE|erule conjE)+) +unfolding image_def +apply (rule CollectI, rule_tac x = "(nat1, b)" in bexI, simp+) +apply (rule_tac B = "{s. \i. inum_of_socket \ s = Some i}" in finite_subset, clarsimp split:if_splits, simp)+ +apply (rule_tac B = "insert (nat1, nat2) {s. \i. inum_of_socket \ s = Some i}" in finite_subset, clarsimp, simp) +apply (rule_tac B = "insert (nat1, nat4) {s. \i. inum_of_socket \ s = Some i}" in finite_subset, clarsimp, simp) +done + +lemma finite_cn: "valid \ \ finite (current_sockets \)" +apply (simp add:current_sockets_def inum_of_socket.simps) +using finite_cn_aux[where \ = \] by auto + +lemma finite_ch: "finite (current_shms \)" +apply (induct \) defer +apply (case_tac a, auto simp:current_shms.simps init_finite_sets) +done + +lemma finite_cm: "finite (current_msgqs \)" +apply (induct \) defer +apply (case_tac a, auto simp: init_finite_sets) +done + + +lemma finite_option: "finite {x. \ y. f x = Some y} \ finite {y. \ x. f x = Some y}" +apply (drule_tac h = f in finite_imageI) +apply (clarsimp simp only:image_def) +apply (rule_tac f = Some in finite_imageD) +apply (rule_tac B = "{y. \x. (\y. f x = Some y) \ y = f x}" in finite_subset) +unfolding image_def +apply auto +done + +lemma finite_ci: "valid \ \ finite (current_inode_nums \)" +apply (simp add:current_inode_nums_def current_file_inums_def current_sock_inums_def) +apply (rule conjI, drule finite_cf, simp add:current_files_def, erule finite_option) +using finite_cn[where \ = \] +apply (simp add:current_sockets_def, drule_tac finite_option, simp) +done + +end + +end + + diff -r 34d01e9a772e -r 7d9c0ed02b56 Flask.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Flask.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,1410 @@ +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" + + 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)" + 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" + and init_msgs_distinct: "\ q. distinct (init_msgs_of_queue q)" + + 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" + 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)" + +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 shms # \) 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}" + + +(****** 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 # \) = ( + if (proc_fd_of_file \ f = {}) + then (inum_of_file \) (f := None) + else inum_of_file \ )" +| "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}" + +definition has_same_inode :: "t_state \ t_file \ t_file \ bool" +where + "has_same_inode \ fa fb = + (\i. inum_of_file \ fa = Some i \ inum_of_file \ fb = Some i)" + +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 shms # \) = + (\ (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}" + +(* 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 \" (* may be sth wrong with nettemp representing addr \ netdev in statical analysis ??? *) + +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 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 shms # \) = (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 shms # \) = + (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 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 (e # \) = procs_of_shm \" + +definition info_flow_shm :: "t_state \ t_process \ t_process \ bool" +where + "info_flow_shm \ from to \ (from = to) \ (\ h \ current_shms \. \ toflag. + (((from, SHM_RDWR) \ procs_of_shm \ h) \ ((to, toflag) \ procs_of_shm \ h)))" + +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 shms # \) 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 shms # \) = + (\ (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 \ pf \ current_files \ \ 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 \ \ f \ current_files \ \ \ 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 \ \ f \ current_files \ \ is_file \ f \ f \ files_hung_by_del \)" +| "os_grant \ (Rmdir p f) = + (p \ current_procs \ \ f \ current_files \ \ 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 \ pf \ current_files \ \ 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 \ \ f\<^isub>1 \ current_files \ \ f\<^isub>2 \ current_files \ \ + (\ pf\<^isub>2. parent f\<^isub>2 = Some pf\<^isub>2 \ pf\<^isub>2 \ current_files \ \ is_dir \ pf\<^isub>2 \ + pf\<^isub>2 \ files_hung_by_del \) \ is_file \ f\<^isub>1)" +| "os_grant \ (Truncate p f len) = + (p \ current_procs \ \ f \ current_files \ \ 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 \ \ f \ current_files \ \ is_file \ f \ fds \ current_proc_fds \ p)" (* file_at_writing_by \ f = {} *) +(* +| "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))" +| "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 \ (p, fd) \ current_sockets \ \ + socket_state \ (p, fd) = Some SS_bind \ is_tcp_sock \ (p, fd))" (* 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 shms) = + (p \ current_procs \ \ p' \ (current_procs \) \ fds \ current_proc_fds \ p \ + (\ h \ shms. \ flag. (p, flag) \ procs_of_shm \ h))" +| "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) = (f \ current_files s \ is_file s f)" +| "alive s (O_dir f) = (f \ current_files s \ 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) = (k \ current_sockets s \ is_tcp_sock s k)" +| "alive s (O_udp_sock k) = (k \ current_sockets s \ 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') \ (\ 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)) \ 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 (UnLink p f # s) = ((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 \ (\ m. obj = O_msg q m) \ 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" + + +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 shms # 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 shms # 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 shms # 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)" + +(******* 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)" + +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_dir pctxt fctxt + | _ \ 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_file pctxt fctxt \ search_check_allp pctxt (set pfctxts)) + | _ \ False) + else (case (sectxt_of_obj s (O_dir (f#pf)), get_parentfs_ctxts s pf) of + (Some fctxt, Some pfctxts) \ (search_check_dir pctxt fctxt \ search_check_allp pctxt (set pfctxts)) + | _ \ 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 :: "t_state \ security_context_t \ t_process \ t_fd set \ bool" +where + "inherit_fds_check s npsectxt p fds \ + (\ ctxt \ sectxts_of_fds s p fds. permission_check npsectxt ctxt C_fd P_inherit)" + +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 :: "t_state \ security_context_t \ t_shm set \ bool" +where + "inherit_shms_check s npsectxt shms \ + (\ ctxt \ sectxts_of_shms s shms. permission_check npsectxt ctxt C_shm P_inherit)" + +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 shms) = + (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 \ + inherit_shms_check s (up,rp,tp) shms) + | _ \ False )" +| "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) = (f \ init_files)" +(* 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) = (member (init_msgs_of_queue q) ((op =) m))" +| "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_in_init: "\ obj. obj \ seeds \ tobj_in_init obj" + +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 shms # 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_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)" + +definition taintable:: "t_object \ bool" +where + "taintable obj \ init_alive obj \ (\ s. obj \ tainted s)" + +definition undeletable :: "t_object \ bool" +where + "undeletable obj \ init_alive obj \ \ (\ s. valid s \ deleted obj s)" + +end + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Flask_type.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Flask_type.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,239 @@ +theory Flask_type +imports Main +begin + + +(************* types that for modeling Access-Control part of Selinux *************) + +(* Table 4-5 of "Selinux by example" *) +datatype permission_t = +(* C_file *) + P_append (* the O_append flag for open *) +| P_create (* create new file *) +| P_entrypoint (* file can be used as the entry point of new domain via a domain transition *) +(* | P_execmod execute memory-mapped files that have modified *) +| P_execute (* corresponds to x access in standard linux *) +(* | P_execute_no_trans: execute file in calller's domain, without domain transition *) +(* | P_getattr : stat, ls, ioctls, ... not-related to security *) +| P_link (* create hard link to file *) +| P_read (* corresponds to r access in standard linux *) +| P_rename (* rename a hard link *) +| P_setattr (* chmod, ioctls ... *) +| P_unlink (* remove hard link (delete *) +| P_write (* corresponds to w access in standard linux *) +(* P_lock; P_mounton; P_quotaon; P_relabelfrom; P_relabelto *) + +(* C_dir *) +| P_add_name (* add file in directory *) +| P_remove_name (* like write for file *) +| P_reparent (* like rename for file *) +| P_search +| P_rmdir (* like unlink for file *) + +(* C_fd *) +(* P_create; P_getattr; P_setattr*) +| P_inherit (* inherit across execve, +we simplify that execve would pass all fds to new "process", but flask checks the inherit right, if denied the fd is closed. *) +(* P_receive: ignored, ? I can find a reasonable explanation *) + +(* C_process *) +| P_fork +| P_ptrace +| P_transition +| P_share (* allow state sharing for execve, clone: fds and ipc + * for clone, we might have two flags, share Fd and IPC + * for execve, as default, fds are shared, IPCs are detached. + * But fds not passed access-control are closed + *) +(*| P_siginh: inheritance of signal state, pending signals, this is not needed, as +we assume all signal are delivered instantly *) +(* P_rlimitnh : inheritance of resource limits *) +(* P_dyntransition; P_execheap; P_execmem; P_execstack; get_attr : not security related; + * P_getcap; P_getsched *) +| P_sigkill (* signal *) + +(* C_msgq, C_shm *) +| P_associate +| P_destroy +| P_enqueue +| P_receive +| P_send +(* create, read, write *) + +type_synonym access_vector_t = "permission_t set" + +(************ Object Classes *************) +datatype security_class_t = + C_process + +(* file-related object classes, 4.3.1 of Book *) +| C_file +| C_fd +| C_dir +(* C_blk file C_chr_file C_fifo_file C_filesystem C_lnk_file C_sock_file *) + +(* network, 4.3.2, using only tcp/udp as a demonstration *) +| C_node (* Host by IP address or range of addresses *) +| C_netif (* network interface, eth0 *) +| C_tcp_socket +| C_udp_socket +(* netlink_socket unix_socket rawip_socket *) + +(* IPCs, 4.3.3 *) +| C_msg (* messages within a message queue *) +| C_msgq (* message queues *) +| C_shm (* shared memory segment *) +(* C_semaphores: only for sync, not tainting related *) + +(*other classes, 4.3.4*) +(* C_capability; C_security:security server; *) +(* | C_system the system itself *) + +(********* TE rules ********** + * 1. type declaration rules + * 2. type transition rules + * 3. type member rules + * 4. type change rules + * 5. access vector rules + * 6. cloning rules + * 7. access vector assertions + *) + +(* 1. type declaration rule *) +datatype type_t = + T_syslogd +| T_syslogd_exec +| T_devlog + +type_synonym attribute_t = "type_t set" + +(* for the asterisk/wildcard in the rules lists*) +datatype 'a patt = + Single "'a" +| Set "'a set" +| Tilde "'a set" (* complement *) +| Asterisk + +(* 2. type transition rule, applied only when: execve & createfile + * (creating subj_type, related obj_typ, type class, default type) + * related obj_typ: proc \ executable; file \ parent-file + * if not in the list, use subj_type for proc and parent_type for file + * if multiple in the list, use the first one (they use the last one, we can say our list is the rev of their list, for "lookup" is easy) + * if using attributes, it can be transformed into the simple form of this list, + now we use "patt" to deal with attributes + *) +type_synonym type_transition_rule_list_t = + "(type_t patt \ type_t patt \ security_class_t \ type_t) list" + + +(* for type member rules: + * ?? I don't know the ideal yet, page 13 bottom of left column + * +type_synonym type_member_rule_list_t = + "(type_t patt \ type_t patt \ security_class_t \ type_t) list" + *) + +(* for type change rules (relabeling to diff security context): + * ?? I don't know what's its meaning yet, page 13 up of the right column + * +type_synonym type_change_rule_list_t = + "(type_t patt \ type_t patt \ security_class_t \ type_t) list" + *) + +datatype 'a target_with_self_t = + Not_self 'a +| Self + +(* for access vector rules(verify a role-changed if allowed or not): + * Any role-changed should be verified by this list! + * here we only do the allow rule, audit and notify rules are ignored + * (source type, target type, class, permissionS) + * permissionS might be access_vector_t, which is already be model as a "set" of "pattern" + * Self in the target type means this rule should be applied to the source process itself + * if multiple, we search the whole list to mimic their "the union" ideal + *) +type_synonym allowed_rule_list_t = + "(type_t patt \ (type_t patt) target_with_self_t \ security_class_t \ permission_t patt) list" + +datatype role_t = + R_object +| R_user +| R_login + +(* role declarations: + * (r, ts): ts are the types/domains that can be associated with r + *) +type_synonym role_declarations_list_t = + "(role_t \ type_t set) list" + +(* role transition rules(only happens when execve call): + * (r, t, r'): the process with role r executes a file with type t, its role will be changed to r' + *) +type_synonym role_transition_rule_list_t = + "(role_t \ type_t \ role_t) list" + +(* role allow rules: + * (from_role, to_role) changes when "Execve" calls! + * What's the difference between this list and the role_transition_rule_list? + * It is: if a process execute a file, the new role will be determined by + * role-transition, but this new role should be in role-allow(if not deny), + * and the new type should be compatible with the new role, which is checked + * by role-declarations(if not deny). + *) +type_synonym role_allow_rule_list_t = + "(role_t set \ role_t set) list" + +datatype user_t = + U_system +| U_sds +| U_pal + +(* user declarations: + * (u, roles) roles are associated with user u + * here MLS are not in consideration + *) +type_synonym user_declarations_list_t = + "(user_t \ role_t set)" + +type_synonym security_context_t = + "(user_t \ role_t \ type_t)" + +(* constrains : + * (classes, perms, expression) list + * when flask do a permission check, constrains will fetch the + * associated two security-contexts (source & target), flask + * makes sure that appling expression to these two contexts, + * will get True. If False, deny the request. + *) +type_synonym constrains_list_t = + "(security_class_t set \ permission_t set \ (security_context_t \ security_context_t \ bool)) list" + +(* comments: + * 1. sid is not needed, it is a run-time index for security-context, + * what we do is policy-analysis, so no this "dynamic" sid as mid-ware; + * 2. "seeds" of tainting cannot be solid objects in the initial states, + should be extended as a "condition/predicate" to describe the "seeds" + *) + +(* "temporary unused definitions : + +datatype security_id_t = + SID nat +| SECSID_NULL +| SECSID_WILD + + +(* this list mimic the action of "security_compute_av", + * if in the list, means func returns "grant" + * if not in, means "denied" + * !!! this list is a "work result" of policy by selinux's + * pre-processing func "security_compute_av" + * in my formalisation, this is not needed, cause the specification of + * our "selinux" contains already the policy specification. + * we can directly get the access-vector for sid/sectxt based on allow-rules & constrains + *) +type_synonym avc_t = "(security_id_t \ security_id_t \ security_class_t \ access_vector_t) list" + + *) + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Init_prop.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Init_prop.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,735 @@ +(*<*) +theory Init_prop +imports Main OS_type_def Flask Flask_type Static_type Static +begin +(*>*) + +context init begin + +lemma init_files_prop1: "init_inum_of_file f = Some im \ f \ init_files" +by (simp add:inof_has_file_tag) + +lemma init_files_prop2: "finite init_files" +by (simp add:init_finite_sets) + +lemma init_files_prop3: "f \ init_files \ init_inum_of_file f \ None" +by (auto dest:init_file_has_inum) + +lemma init_files_prop4: "(f \ init_files) = (f \ current_files [])" +apply (simp add:current_files_def, rule iffI) +using init_files_prop1 init_files_prop3 by auto + +lemmas init_files_props = init_file_has_inum init_files_prop1 init_files_prop2 init_files_prop3 init_files_prop4 + +lemma init_inumof_prop1: "init_inum_of_file f = Some im \ \ tag. init_itag_of_inum im = Some tag" +by (auto dest:inof_has_file_tag) + +lemma init_inumof_prop2: "init_inum_of_file f = Some im \ init_itag_of_inum im \ None" +by (auto dest:inof_has_file_tag) + +lemma init_inumof_prop3: "\init_inum_of_file f = Some im; init_itag_of_inum im = Some tag\ \ is_file_dir_itag tag" +by (auto dest:inof_has_file_tag) + +lemmas init_inum_of_file_props = init_files_prop1 init_inumof_prop1 init_inumof_prop2 init_inumof_prop3 + +lemma init_inumos_prop1: "init_inum_of_socket s = Some im \ s \ init_sockets" +by (auto dest:inos_has_sock_tag) + +lemma init_inumos_prop2: "init_inum_of_socket s = Some im \ init_itag_of_inum im = Some Tag_TCP_SOCK \ init_itag_of_inum im = Some Tag_UDP_SOCK" +apply (auto dest!:inos_has_sock_tag) +apply (case_tac tag, simp+) +done + +lemma init_inumos_prop3: "init_inum_of_socket s = Some im \ init_itag_of_inum im \ None" +by (auto dest:inos_has_sock_tag) + +lemma init_inumos_prop4: "init_inum_of_socket s = Some im \ \ tag. init_itag_of_inum im = Some tag \ is_sock_itag tag" +by (auto dest!:inos_has_sock_tag) + +lemmas init_inum_of_socket_props = init_inumos_prop1 init_inumos_prop2 init_inumos_prop3 init_inumos_prop4 + +lemma init_sockets_prop1: "(p, fd) \ init_sockets \ p \ init_procs" +by (auto dest: init_socket_has_inode) + +lemma init_sockets_prop2: "(p, fd) \ init_sockets \ fd \ init_fds_of_proc p" +by (auto dest:init_socket_has_inode) + +lemma init_sockets_prop3: "s \ init_sockets \ \ im. init_inum_of_socket s = Some im" +by (case_tac s, auto dest:init_socket_has_inode) + +lemma init_sockets_prop4: "s \ init_sockets \ init_inum_of_socket s \ None" +by (simp add:init_sockets_prop3) + +lemma init_sockets_prop5: "s \ init_sockets = (s \ current_sockets [])" +apply (simp add:current_sockets_def, rule iffI) +using init_sockets_prop4 inos_has_sock_tag apply auto +apply (case_tac s, auto) +done + +lemmas init_sockets_props = init_sockets_prop1 init_sockets_prop2 init_sockets_prop3 init_sockets_prop4 init_sockets_prop5 + +lemma is_init_file_prop1: "is_init_file f = (f \ init_files \ is_file [] f)" +by (auto simp add:is_init_file_def is_file_def init_inum_of_file_props split:option.splits) + +lemma is_init_file_prop2: "is_init_file f = (init_alive (O_file f))" +by (auto simp add:is_init_file_def is_file_def init_inum_of_file_props split:option.splits) + +lemmas is_init_file_props = is_init_file_prop1 is_init_file_prop2 + +lemma is_init_dir_prop1: "is_init_dir f = (f \ init_files \ is_dir [] f)" +by (auto simp add:is_init_dir_def is_dir_def init_inum_of_file_props split:option.splits) + +lemma is_init_dir_prop2: "is_init_dir f = (init_alive (O_dir f))" +by (auto simp add:is_init_dir_def is_dir_def init_inum_of_file_props split:option.splits) + +lemmas is_init_dir_props = is_init_dir_prop1 is_init_dir_prop2 + +lemma is_init_udp_sock_prop1: "is_init_udp_sock s = (s \ init_sockets \ is_udp_sock [] s)" +apply (auto simp add:is_init_udp_sock_def is_udp_sock_def init_inum_of_socket_props + dest:init_socket_has_inode split:option.splits) +done + +lemma is_init_udp_sock_prop2: "is_init_udp_sock s = (init_alive (O_udp_sock s))" +apply (auto simp add:is_init_udp_sock_def is_udp_sock_def init_inum_of_socket_props + dest:init_socket_has_inode split:option.splits) +done + +lemmas is_init_udp_sock_props = is_init_udp_sock_prop1 is_init_udp_sock_prop2 + +lemma is_init_tcp_sock_prop1: "is_init_tcp_sock s = (s \ init_sockets \ is_tcp_sock [] s)" +apply (auto simp add:is_init_tcp_sock_def is_tcp_sock_def init_inum_of_socket_props + dest:init_socket_has_inode split:option.splits) +done + +lemma is_init_tcp_sock_prop2: "is_init_tcp_sock s = (init_alive (O_tcp_sock s))" +apply (auto simp add:is_init_tcp_sock_def is_tcp_sock_def init_inum_of_socket_props + dest:init_socket_has_inode split:option.splits) +done + +lemmas is_init_tcp_sock_props = is_init_tcp_sock_prop1 is_init_tcp_sock_prop2 + + +lemma init_parent_file_prop1: + "\parent f = Some pf; f \ init_files\ \ is_init_dir pf" +apply (frule parent_file_in_init, simp, frule_tac f = pf in init_files_prop3) +apply (clarsimp, drule_tac im = y in init_parentf_is_dir, simp+) +by (simp add:is_init_dir_def) + +lemma init_parent_file_prop1': + "a # f \ init_files \ is_init_dir f" +by (rule_tac pf = f in init_parent_file_prop1, auto) + +lemma init_parent_file_prop2: + "\parent f = Some pf; is_init_file f\ \ is_init_dir pf" +by (rule init_parent_file_prop1, simp, simp add: is_init_file_props) + +lemma init_parent_file_prop2': + "is_init_file (f#pf) \ is_init_dir pf" +apply (rule init_parent_file_prop2) +by auto + +lemma init_parent_file_prop3: + "\parent f = Some pf; is_init_dir f\ \ is_init_dir pf" +by (rule init_parent_file_prop1, simp, simp add: is_init_dir_props) + +lemma init_parent_file_prop3': + "is_init_dir (f#pf) \ is_init_dir pf" +apply (rule init_parent_file_prop3) +by auto + +lemma parent_file_in_init': "a # f \ init_files \ f \ init_files" +by (subgoal_tac "parent (a # f) = Some f", drule parent_file_in_init, auto) + +lemmas init_parent_file_props = parent_file_in_init init_parent_file_prop1 parent_file_in_init' init_parent_file_prop1' init_parent_file_prop2 init_parent_file_prop2' init_parent_file_prop3 init_parent_file_prop3' + +lemma root_in_filesystem: "[] \ init_files" +using init_files_prop1 root_is_dir by auto + +lemma root_is_init_dir: "is_init_dir []" +using root_is_dir +by (auto simp add:is_init_dir_def split:option.splits) + +lemma root_is_init_dir': "is_init_file [] \ False" +using root_is_dir +by (auto simp:is_init_file_def split:option.splits) + + +lemma init_files_hung_prop1: "f \ init_files_hung_by_del \ f \ init_files" +by (auto dest:init_files_hung_valid) + +lemma init_files_hung_prop2: "f \ init_files_hung_by_del \ \ p fd. init_file_of_proc_fd p fd = Some f" +by (auto dest:init_files_hung_valid) + +lemmas init_files_hung_by_del_props = init_files_hung_prop1 init_files_hung_prop2 init_files_hung_valid' + + +lemma init_fds_of_proc_prop1: "fd \ init_fds_of_proc p \ p \ init_procs" +by (auto dest!:init_procfds_valid) + +lemma init_fds_of_proc_prop2: "fd \ init_fds_of_proc p \ (\ f \ init_files. init_file_of_proc_fd p fd = Some f) \ (p, fd) \ init_sockets" +by (auto dest:init_procfds_valid) + +lemmas init_fds_of_proc_props = init_fds_of_proc_prop1 init_fds_of_proc_prop2 + +lemma init_filefd_prop1: "init_file_of_proc_fd p fd = Some f \ f \ init_files" +by (auto dest!:init_filefd_valid intro:init_files_prop1) + +lemma init_filefd_prop2: "init_file_of_proc_fd p fd = Some f \ p \ init_procs" +by (auto dest:init_filefd_valid) + +lemma init_filefd_prop3: "init_file_of_proc_fd p fd = Some f \ fd \ init_fds_of_proc p" +by (auto dest:init_filefd_valid) + +lemma init_filefd_prop4: "init_file_of_proc_fd p fd = Some f \ \ flags. init_oflags_of_proc_fd p fd = Some flags" +by (auto dest:init_filefd_valid) + +lemma init_filefd_prop5: "init_file_of_proc_fd p fd = Some f \ is_init_file f" +by (auto dest:init_filefd_valid simp:is_init_file_def) + +lemmas init_file_of_proc_fd_props = init_filefd_prop1 init_filefd_prop2 init_filefd_prop3 init_filefd_prop4 init_filefd_prop5 + +lemma init_oflags_prop1: "init_oflags_of_proc_fd p fd = Some flags \ p \ init_procs" +by (auto dest:init_fileflag_valid init_file_of_proc_fd_props) + +lemma init_oflags_prop2: "init_oflags_of_proc_fd p fd = Some flags \ fd \ init_fds_of_proc p" +by (auto dest:init_fileflag_valid init_file_of_proc_fd_props) + +lemmas init_oflags_of_proc_fd_props = init_oflags_prop1 init_oflags_prop2 init_fileflag_valid + +(* +lemma init_socketstate_prop1: "s \ init_sockets \ init_socket_state s \ None" +using init_socket_has_state +by (case_tac s, simp add:bidirect_in_init_def) + +lemma init_socketstate_prop2: "s \ init_sockets \ \ t. init_socket_state s = Some t" +using init_socket_has_state +by (case_tac s, simp add:bidirect_in_init_def) + +lemma init_socketstate_prop3: "init_socket_state s = Some t \ s \ init_sockets" +using init_socket_has_state +by (case_tac s, simp add:bidirect_in_init_def) + +lemmas init_socket_state_props = init_socketstate_prop1 init_socketstate_prop2 init_socketstate_prop3 +*) + +lemma init_inum_sock_file_noninter: "\init_inum_of_socket s = Some im; init_inum_of_file f = Some im\ \ False" +apply (frule init_inumof_prop1, erule exE, drule init_inumof_prop3, simp) +apply (frule init_inumos_prop2) +apply (case_tac tag, simp+) +done + +lemma init_parent_file_has_inum: "\parent f = Some pf; init_inum_of_file f = Some im\ \ \ im. init_inum_of_file pf = Some im" +by (drule init_files_prop1, drule parent_file_in_init, simp, simp add:init_files_props) + +lemma init_file_has_no_son': "\init_itag_of_inum im = Some Tag_FILE; init_inum_of_file f = Some im; parent f' = Some f\ \ init_inum_of_file f' = None" +apply (drule init_file_no_son, simp) +by (case_tac "init_inum_of_file f'", auto dest:init_files_prop1) + +lemma init_parent_file_is_dir': "\parent f = Some pf; init_inum_of_file f = Some im; init_inum_of_file pf = Some ipm\ \ init_itag_of_inum ipm = Some Tag_DIR" +by (drule init_parentf_is_dir, auto dest:init_files_prop1) + +lemma init_file_hung_has_no_son: "\f \ init_files_hung_by_del; parent f' = Some f; init_inum_of_file f' = Some im\ \ False" +apply (frule init_files_hung_prop1, drule init_file_has_inum, erule exE) +apply (drule init_files_hung_valid', simp) +apply (frule init_parent_file_is_dir', simp+) +apply (drule init_files_prop1) +apply (erule_tac x = f' in allE, simp) +by (case_tac f', simp_all add:no_junior_def) + + +end + +context flask begin + +lemma init_alive_prop: "init_alive obj = alive [] obj" +apply (case_tac obj, simp_all add:is_init_file_props is_init_dir_props is_init_tcp_sock_props + is_init_udp_sock_props init_files_props init_sockets_props) +done + +lemma init_alive_proc: "p \ init_procs \ init_alive (O_proc p)" by simp +lemma init_alive_file: "is_init_file f \ init_alive (O_file f)" by simp +lemma init_alive_dir: "is_init_dir f \ init_alive (O_dir f)" by simp +lemma init_alive_fd: "fd \ init_fds_of_proc p \ init_alive (O_fd p fd)" by simp +lemma init_alive_tcp: "is_init_tcp_sock s \ init_alive (O_tcp_sock s)" by simp +lemma init_alive_udp: "is_init_udp_sock s \ init_alive (O_udp_sock s)" by simp +lemma init_alive_node: "n \ init_nodes \ init_alive (O_node n)" by simp +lemma init_alive_shm: "h \ init_shms \ init_alive (O_shm h)" by simp +lemma init_alive_msgq: "q \ init_msgqs \ init_alive (O_msgq q)" by simp +lemma init_alive_msg: "\m \ set (init_msgs_of_queue q); q \ init_msgqs\ + \ init_alive (O_msg q m)" by simp + +lemmas init_alive_intros = init_alive_proc init_alive_file init_alive_dir init_alive_fd + init_alive_tcp init_alive_udp init_alive_node init_alive_shm init_alive_msgq init_alive_msg + + +lemma init_file_type_prop1: "is_init_file f \ \ t. init_type_of_obj (O_file f) = Some t" +using init_obj_has_type +by (auto simp:is_init_file_def split:option.splits) + +lemma init_file_type_prop2: "is_init_file f \ init_type_of_obj (O_file f) \ None" +by (simp add:init_file_type_prop1) + +lemma init_file_type_prop3: "init_type_of_obj (O_file f) = Some t \ f \ init_files" +apply (drule init_type_has_obj) +by (simp add:is_init_file_def init_inum_of_file_props split:option.splits) + +lemma init_file_type_prop4: "init_type_of_obj (O_file f) = Some t \ is_init_file f" +apply (drule init_type_has_obj) +by (simp add:is_init_file_def init_inum_of_file_props split:option.splits) + +lemmas init_file_types_props = init_file_type_prop1 init_file_type_prop2 init_file_type_prop3 init_file_type_prop4 + +lemma init_dir_type_prop1: "is_init_dir f \ \ t. init_type_of_obj (O_dir f) = Some t" +using init_obj_has_type +by (auto simp:is_init_dir_def split:option.splits) + +lemma init_dir_type_prop2: "is_init_dir f \ init_type_of_obj (O_dir f) \ None" +by (simp add:init_dir_type_prop1) + +lemma init_dir_type_prop3: "init_type_of_obj (O_dir f) = Some t \ f \ init_files" +apply (drule init_type_has_obj) +by (simp add:is_init_dir_def init_inum_of_file_props split:option.splits) + +lemma init_dir_type_prop4: "init_type_of_obj (O_dir f) = Some t \ is_init_dir f" +apply (drule init_type_has_obj) +by (simp add:is_init_dir_def init_inum_of_file_props split:option.splits) + +lemmas init_dir_types_props = init_dir_type_prop1 init_dir_type_prop2 init_dir_type_prop3 init_dir_type_prop4 + +lemma init_procrole_prop1: "init_role_of_proc p = Some r \ p \ init_procs" +using init_proc_has_role +by (auto simp:bidirect_in_init_def) + +lemma init_procrole_prop2: "p \ init_procs \ \ r. init_role_of_proc p = Some r" +using init_proc_has_role +by (auto simp:bidirect_in_init_def) + +lemma init_procrole_prop3: "p \ init_procs \ init_role_of_proc p \ None" +using init_proc_has_role +by (auto simp:bidirect_in_init_def) + +lemmas init_role_of_proc_props = init_procrole_prop1 init_procrole_prop2 init_procrole_prop3 + +lemma init_file_user_prop1: "is_init_file f \ \ t. init_user_of_obj (O_file f) = Some t" +apply (simp only: is_init_file_prop2) +by (drule init_obj_has_user, auto) + +lemma init_file_user_prop2: "is_init_file f \ init_user_of_obj (O_file f) \ None" +by (simp add:init_file_user_prop1) + +lemma init_file_user_prop3: "init_user_of_obj (O_file f) = Some t \ f \ init_files" +apply (drule init_user_has_obj) +by (simp add:is_init_file_def init_inum_of_file_props split:option.splits) + +lemma init_file_user_prop4: "init_user_of_obj (O_file f) = Some t \ is_init_file f" +apply (drule init_user_has_obj) +by (simp add:is_init_file_def init_inum_of_file_props split:option.splits) + +lemma init_file_user_prop5: "init_user_of_obj (O_file f) = Some u \ u \ init_users" +by (simp add:init_user_has_obj) + +lemmas init_file_users_props = init_file_user_prop1 init_file_user_prop2 init_file_user_prop3 init_file_user_prop4 init_file_user_prop5 + +lemma init_dir_user_prop1: "is_init_dir f \ \ t. init_user_of_obj (O_dir f) = Some t" +apply (simp only: is_init_dir_prop2) +by (drule init_obj_has_user, auto) + +lemma init_dir_user_prop2: "is_init_dir f \ init_user_of_obj (O_dir f) \ None" +by (simp add:init_dir_user_prop1) + +lemma init_dir_user_prop3: "init_user_of_obj (O_dir f) = Some t \ f \ init_files" +apply (drule init_user_has_obj) +by (simp add:is_init_dir_def init_inum_of_file_props split:option.splits) + +lemma init_dir_user_prop4: "init_user_of_obj (O_dir f) = Some t \ is_init_dir f" +apply (drule init_user_has_obj) +by (simp add:is_init_dir_def init_inum_of_file_props split:option.splits) + +lemma init_dir_user_prop5: "init_user_of_obj (O_dir f) = Some u \ u \ init_users" +by (simp add:init_user_has_obj) + +lemmas init_dir_users_props = init_dir_user_prop1 init_dir_user_prop2 init_dir_user_prop3 init_dir_user_prop4 init_dir_user_prop5 + +end + + +context tainting_s begin + +lemma init_file_has_ctxt: + "is_init_file f \ \ sec. init_sectxt_of_obj (O_file f) = Some sec" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +apply (rule conjI, rule init_obj_has_user, simp add:is_init_file_props) +by (simp add:init_file_types_props) + +lemma init_file_has_ctxt': + "init_sectxt_of_obj (O_file f) = None \ \ is_init_file f" +by (rule notI, drule init_file_has_ctxt, simp) + +lemma init_dir_has_ctxt: + "is_init_dir f \ \ sec. init_sectxt_of_obj (O_dir f) = Some sec" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props) +by (simp add:init_dir_types_props) + +lemma init_dir_has_ctxt': + "init_sectxt_of_obj (O_dir f) = None \ \ is_init_dir f" +by (rule notI, drule init_dir_has_ctxt, simp) + +lemma init_proc_has_ctxt: + "p \ init_procs \ \ sec. init_sectxt_of_obj (O_proc p) = Some sec" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +apply (rule conjI, rule init_obj_has_user, simp) +apply (frule init_alive_proc, drule init_obj_has_type) +by (drule init_procrole_prop2, auto) + +lemma init_proc_has_ctxt': + "init_sectxt_of_obj (O_proc p) = None \ p \ init_procs" +by (rule notI, drule init_proc_has_ctxt, simp) + +lemma init_fd_has_ctxt: + "fd \ init_fds_of_proc p \ \ sec. init_sectxt_of_obj (O_fd p fd) = Some sec" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props) +apply (drule init_alive_intros) +apply (drule init_obj_has_type, clarsimp) +done + +lemma init_fd_has_ctxt': + "init_sectxt_of_obj (O_fd p fd) = None \ fd \ init_fds_of_proc p" +by (rule notI, drule init_fd_has_ctxt, simp) + +lemma init_node_has_ctxt: + "n \ init_nodes \ \ sec. init_sectxt_of_obj (O_node n) = Some sec" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props) +apply (drule init_alive_intros) +apply (drule init_obj_has_type, clarsimp) +done + +lemma init_node_has_ctxt': + "init_sectxt_of_obj (O_node n) = None \ n \ init_nodes" +by (rule notI, drule init_node_has_ctxt, simp) + +lemma init_tcp_has_ctxt: + "is_init_tcp_sock s \ \ sec. init_sectxt_of_obj (O_tcp_sock s) = Some sec" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props) +apply (drule init_alive_intros) +apply (drule init_obj_has_type, clarsimp) +done + +lemma init_tcp_has_ctxt': + "init_sectxt_of_obj (O_tcp_sock s) = None \ \ is_init_tcp_sock s" +by (rule notI, drule init_tcp_has_ctxt, simp) + +lemma init_udp_has_ctxt: + "is_init_udp_sock s \ \ sec. init_sectxt_of_obj (O_udp_sock s) = Some sec" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props) +by (drule init_alive_intros, drule init_obj_has_type, clarsimp) + +lemma init_udp_has_ctxt': + "init_sectxt_of_obj (O_udp_sock s) = None \ \ is_init_udp_sock s" +by (rule notI, drule init_udp_has_ctxt, simp) + +lemma init_shm_has_ctxt: + "h \ init_shms \ \ sec. init_sectxt_of_obj (O_shm h) = Some sec" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props) +by (drule init_alive_intros, drule init_obj_has_type, clarsimp) + +lemma init_shm_has_ctxt': + "init_sectxt_of_obj (O_shm h) = None \ h \ init_shms" +by (rule notI, drule init_shm_has_ctxt, simp) + +lemma init_msgq_has_ctxt: + "q \ init_msgqs \ \ sec. init_sectxt_of_obj (O_msgq q) = Some sec" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props) +by (drule init_alive_intros, drule init_obj_has_type, clarsimp) + +lemma init_msgq_has_ctxt': + "init_sectxt_of_obj (O_msgq q) = None \ q \ init_msgqs" +by (rule notI, drule init_msgq_has_ctxt, simp) + +lemma init_msg_has_ctxt: + "\m \ set (init_msgs_of_queue q); q \ init_msgqs\ \ \ sec. init_sectxt_of_obj (O_msg q m) = Some sec" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props) +by (drule init_alive_intros, simp, drule init_obj_has_type, clarsimp) + +lemma init_msg_has_ctxt': + "init_sectxt_of_obj (O_msg q m) = None \ m \ set (init_msgs_of_queue q) \ q \ init_msgqs" +by (auto dest:init_msg_has_ctxt) + +lemma init_rootf_has_ctxt: + "\ sec. init_sectxt_of_obj (O_dir []) = Some sec" +apply (rule init_dir_has_ctxt, simp add:is_init_dir_def split:option.splits) +using root_is_dir by auto + +lemma init_rootf_has_ctxt': + "init_sectxt_of_obj (O_dir []) = None \ False" +using init_rootf_has_ctxt by auto + +lemmas init_has_ctxt = init_file_has_ctxt init_dir_has_ctxt init_proc_has_ctxt init_fd_has_ctxt + init_node_has_ctxt init_tcp_has_ctxt init_udp_has_ctxt init_shm_has_ctxt init_msgq_has_ctxt + init_msg_has_ctxt init_rootf_has_ctxt + +lemmas init_has_ctxt' = init_file_has_ctxt' init_dir_has_ctxt' init_proc_has_ctxt' init_fd_has_ctxt' + init_node_has_ctxt' init_tcp_has_ctxt' init_udp_has_ctxt' init_shm_has_ctxt' init_msgq_has_ctxt' + init_msg_has_ctxt' init_rootf_has_ctxt' + +lemma sec_of_root_valid: + "init_sectxt_of_obj (O_dir []) = Some sec_of_root" +using init_rootf_has_ctxt +by (auto simp:init_sectxt_of_obj_def sec_of_root_def split:option.splits) + +lemma sec_of_root_is_tuple: + "\ u t. sec_of_root = (u, R_object, t)" +using sec_of_root_valid +by (auto simp:sec_of_root_def init_sectxt_of_obj_def split:option.splits) + +lemma sroot_valid: + "init_cf2sfile [] = Some sroot" +by (simp add:init_cf2sfile_def) + +lemma sroot_valid': + "cf2sfile s [] False = Some sroot" +by (simp add:cf2sfile_def) + +lemma init_sectxt_prop: + "sectxt_of_obj [] obj = init_sectxt_of_obj obj" +apply (auto simp:init_sectxt_of_obj_def sectxt_of_obj_def split:option.splits) +apply (case_tac [!] obj, simp+) +done + +lemma init_sectxt_prop2: + "init_sectxt_of_obj obj = Some sec \ init_alive obj" +by (case_tac obj, auto simp:init_sectxt_of_obj_def split:option.splits dest:init_type_has_obj) + +lemma init_dir_has_seclist: + "is_init_dir f \ \ seclist. get_parentfs_ctxts [] f = Some seclist" +apply (induct f) +apply (simp only:get_parentfs_ctxts.simps init_sectxt_prop) +using init_rootf_has_ctxt apply (auto)[1] +apply (frule init_parent_file_prop3', simp del:get_parentfs_ctxts.simps) +apply (erule exE, drule init_dir_has_ctxt) +by (auto simp add:init_sectxt_prop) + +lemma is_init_file_dir_prop1: + "is_init_dir f \ \ is_init_file f" +by (auto simp:is_init_dir_def is_init_file_def split:option.splits t_inode_tag.splits) + +lemma is_init_file_dir_prop2: + "is_init_file f \ \ is_init_dir f" +by (auto simp:is_init_dir_def is_init_file_def split:option.splits t_inode_tag.splits) + +lemma is_init_file_dir_prop3: + "\is_init_dir f; is_init_file f\ \ False" +by (auto simp:is_init_dir_def is_init_file_def split:option.splits t_inode_tag.splits) + +lemma is_init_file_dir_prop4: + "\is_init_file f; is_init_dir f\ \ False" +by (auto simp:is_init_dir_def is_init_file_def split:option.splits t_inode_tag.splits) + +lemmas is_init_file_dir_props = is_init_file_dir_prop1 is_init_file_dir_prop2 is_init_file_dir_prop3 is_init_file_dir_prop4 + +lemma init_dir_has_sfile: + "is_init_dir f \ \ sf. init_cf2sfile f = Some sf" +apply (case_tac f) +using init_rootf_has_ctxt apply (auto)[1] +apply (simp add:sec_of_root_valid sroot_valid sroot_def) +apply (simp, frule init_parent_file_prop3') +apply (frule_tac f = list in init_dir_has_seclist) +apply (frule_tac f = list in init_dir_has_ctxt) +apply (frule_tac f = "a # list" in init_dir_has_ctxt) +apply ((erule exE)+, case_tac sec, auto simp:init_cf2sfile_def split:option.splits) +by (auto simp:is_init_file_def is_init_dir_def split:option.splits t_inode_tag.splits) + +lemma init_file_has_sfile: + "is_init_file f \ \ sf. init_cf2sfile f = Some sf" +apply (case_tac f) +apply (simp, drule root_is_init_dir', simp) +apply (simp, frule init_parent_file_prop2') +apply (frule_tac f = list in init_dir_has_seclist) +apply (frule_tac f = list in init_dir_has_ctxt) +apply (frule_tac f = "a # list" in init_file_has_ctxt) +by ((erule exE)+, case_tac sec, auto simp:init_cf2sfile_def) + +lemma init_shm_has_sshm: + "h \ init_shms \ \ sh. init_ch2sshm h = Some sh" +apply (drule init_shm_has_ctxt) +by (auto simp add:init_ch2sshm_def) + +lemma init_proc_has_sproc: + "p \ init_procs \ \ sp. init_cp2sproc p = Some sp" +apply (frule init_proc_has_ctxt, erule exE) +apply (simp add:init_cp2sproc_def) +by (case_tac sec, simp+) + +lemma init_cqm2sms_has_sms_aux: + "\ m \ set ms. init_sectxt_of_obj (O_msg q m) \ None \ (\ sms. init_cqm2sms q ms = Some sms)" +by (induct ms, auto split:option.splits simp:init_cm2smsg_def) + +lemma init_cqm2sms_has_sms: + "q \ init_msgqs \ \ sms. init_cqm2sms q (init_msgs_of_queue q) = Some sms" +apply (rule init_cqm2sms_has_sms_aux) +using init_msg_has_ctxt by auto + +lemma init_msgq_has_smsgq: + "q \ init_msgqs \ \ sq. init_cq2smsgq q = Some sq" +apply (frule init_msgq_has_ctxt, erule exE, drule init_cqm2sms_has_sms, erule exE) +apply (simp add:init_cq2smsgq_def) +by (case_tac sec, simp+) + +lemma cf2sfile_nil_prop1: + "f \ init_files \ cf2sfile [] f (is_init_file f) = init_cf2sfile f" +apply (case_tac f) +apply (simp add:init_sectxt_prop cf2sfile_def init_cf2sfile_def) +apply (rule notI, drule root_is_init_dir', simp) +apply (auto simp:init_sectxt_prop cf2sfile_def init_cf2sfile_def split:option.splits dest!:init_has_ctxt') +apply (auto simp:is_init_file_def is_init_dir_def split:option.splits t_inode_tag.splits + dest:init_file_has_inum inof_has_file_tag) +done + +lemma init_file_dir_conflict: "\is_init_file f; is_init_dir f\ \ False" +by (auto simp:is_init_file_def is_init_dir_def split:option.splits t_inode_tag.splits) + +lemma init_file_dir_conflict1: "is_init_file f \ \ is_init_dir f" +by (auto simp:is_init_file_def is_init_dir_def split:option.splits t_inode_tag.splits) + +lemma init_file_dir_conflict2: "is_init_dir f \ \ is_init_file f" +by (auto simp:is_init_file_def is_init_dir_def split:option.splits t_inode_tag.splits) + +lemma init_sec_file_dir: + "\init_sectxt_of_obj (O_file f) = Some x; init_sectxt_of_obj (O_dir f) = Some y\ \ False" +apply (drule init_sectxt_prop2)+ +apply (auto intro:init_file_dir_conflict) +done + +lemma cf2sfile_nil_prop2: + "f \ init_files \ cf2sfile [] f (\ is_init_file f) = None" +apply (case_tac f) +apply (simp add:init_sectxt_prop cf2sfile_def init_cf2sfile_def) +apply (rule notI, drule root_is_init_dir', simp) +apply (auto simp:init_sectxt_prop cf2sfile_def init_cf2sfile_def split:option.splits dest!:init_has_ctxt') +apply (auto simp:is_init_file_def is_init_dir_def split:option.splits t_inode_tag.splits + dest:init_file_has_inum inof_has_file_tag init_sec_file_dir) +done + +lemma cf2sfile_nil_prop: + "f \ init_files \ cf2sfile [] f = (\ b. if (b = is_init_file f) then init_cf2sfile f else None)" +apply (frule cf2sfile_nil_prop1, frule cf2sfile_nil_prop2) +by (rule ext, auto split:if_splits) + +lemma cf2sfile_nil_prop3: + "is_init_file f \ cf2sfile [] f True = init_cf2sfile f" +by (simp add:is_init_file_prop1 cf2sfile_nil_prop) + +lemma cf2sfile_nil_prop4: + "is_init_dir f \ cf2sfile [] f False = init_cf2sfile f" +apply (frule init_file_dir_conflict2) +by (simp add:is_init_file_prop1 is_init_dir_prop1 cf2sfile_nil_prop) + +lemma cfs2sfiles_nil_prop: + "\ f \ fs. f \ init_files \ cfs2sfiles [] fs = init_cfs2sfiles fs" +apply (simp add:cfs2sfiles_def init_cfs2sfiles_def) +using cf2sfile_nil_prop apply auto + +lemma cfd2sfd_nil_prop: + "init_file_of_proc_fd p fd = Some f \ cfd2sfd [] p fd = init_cfd2sfd p fd" +apply (simp add:cfd2sfd_def init_sectxt_prop init_cfd2sfd_def) +apply (drule init_filefd_prop1, drule cf2sfile_nil_prop) +by (auto split:option.splits) + +lemma cpfd2sfds_nil_prop: + "cpfd2sfds [] p = init_cfds2sfds p" +apply (simp only:cpfd2sfds_def init_cfds2sfds_def) +apply (rule set_eqI, rule iffI) +apply (drule CollectD, rule CollectI, (erule exE)+) +apply (rule_tac x = fd in exI, rule_tac x = sfd in exI, rule_tac x = f in exI) defer +apply (drule CollectD, rule CollectI, (erule exE)+) +apply (rule_tac x = fd in exI, rule_tac x = sfd in exI, rule_tac x = f in exI) +using cfd2sfd_nil_prop by auto + +lemma ch2sshm_nil_prop: + "h \ init_shms \ ch2sshm [] h = init_ch2sshm h" +by (simp add:ch2sshm_def init_sectxt_prop init_ch2sshm_def) + +lemma cph2spshs_nil_prop: + "cph2spshs [] p = init_cph2spshs p" +apply (auto simp add:init_cph2spshs_def cph2spshs_def init_sectxt_prop) +apply (rule_tac x = h in exI, simp) defer +apply (rule_tac x = h in exI, simp) +by (auto simp:ch2sshm_nil_prop dest:init_procs_has_shm) + +lemma cp2sproc_nil_prop: + "p \ init_procs \ cp2sproc [] p = init_cp2sproc p" +by (auto simp add:init_cp2sproc_def cp2sproc_def init_sectxt_prop cph2spshs_nil_prop cpfd2sfds_nil_prop + split:option.splits) + +lemma tainted_nil_prop: + "(x \ tainted []) = (x \ seeds)" +apply (rule iffI) +apply (erule tainted.cases, simp+) +apply (erule t_init) +done + +lemma msg_has_sec_imp_init: + "init_sectxt_of_obj (O_msg q m) = Some sec \ q \ init_msgqs \ m \ set (init_msgs_of_queue q)" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +by (drule init_type_has_obj, simp) + +lemma msgq_has_sec_imp_init: + "init_sectxt_of_obj (O_msgq q) = Some sec \ q \ init_msgqs" +apply (simp add:init_sectxt_of_obj_def split:option.splits) +by (drule init_type_has_obj, simp) + +lemma cm2smsg_nil_prop: + "cm2smsg [] q m = init_cm2smsg q m" +by (auto simp add:init_sectxt_prop cm2smsg_def init_cm2smsg_def split:option.splits + dest: msg_has_sec_imp_init intro:t_init elim:tainted.cases) + +lemma cqm2sms_nil_prop: + "cqm2sms [] q ms = init_cqm2sms q ms" +apply (induct ms, simp) +by (auto simp add:cm2smsg_def init_sectxt_prop tainted_nil_prop msg_has_sec_imp_init init_cm2smsg_def + split:option.splits) + +lemma cq2smsga_nil_prop: + "cq2smsgq [] q = init_cq2smsgq q" +by (auto simp add:cq2smsgq_def init_cq2smsgq_def init_sectxt_prop cqm2sms_nil_prop + intro:msgq_has_sec_imp_init split:option.splits) + +lemma same_inode_nil_prop: + "same_inode_files [] f = init_same_inode_files f" +by (simp add:same_inode_files_def init_same_inode_files_def) + +lemma init_same_inode_prop1: + "f \ init_files \ \ f' \ init_same_inode_files f. f' \ init_files" +apply (simp add:init_same_inode_files_def) +apply (drule init_files_prop3) +apply (auto simp:init_files_prop1) +done + +lemma co2sobj_nil_prop: + "init_alive obj \ co2sobj [] obj = init_obj2sobj obj" +apply (case_tac obj) +apply (auto simp add:cf2sfile_nil_prop cq2smsga_nil_prop cqm2sms_nil_prop tainted_nil_prop + cp2sproc_nil_prop cfs2sfiles_nil_prop is_init_dir_prop1 is_init_file_prop1 + is_init_udp_sock_prop1 is_init_tcp_sock_prop1 ch2sshm_nil_prop + same_inode_nil_prop cm2smsg_nil_prop dest:init_same_inode_prop1 + split:option.splits) +apply (rule_tac x = list in exI, simp add:init_same_inode_files_def) +by (simp add:init_files_props) + +lemma s2ss_nil_prop: + "s2ss [] = init_static_state" +using co2sobj_nil_prop init_alive_prop +by (auto simp add:s2ss_def init_static_state_def) + +end + +(*<*) +end +(*>*) \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 List_Prefix.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/List_Prefix.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,382 @@ +(* Title: HOL/Library/List_Prefix.thy + Author: Tobias Nipkow and Markus Wenzel, TU Muenchen +*) + +header {* List prefixes and postfixes *} + +theory List_Prefix +imports List Main +begin + +subsection {* Prefix order on lists *} + +instantiation list :: (type) "{order, bot}" +begin + +definition + prefix_def: "xs \ ys \ (\zs. ys = xs @ zs)" + +definition + strict_prefix_def: "xs < ys \ xs \ ys \ xs \ (ys::'a list)" + +definition + "bot = []" + +instance proof +qed (auto simp add: prefix_def strict_prefix_def bot_list_def) + +end + +lemma prefixI [intro?]: "ys = xs @ zs ==> xs \ ys" + unfolding prefix_def by blast + +lemma prefixE [elim?]: + assumes "xs \ ys" + obtains zs where "ys = xs @ zs" + using assms unfolding prefix_def by blast + +lemma strict_prefixI' [intro?]: "ys = xs @ z # zs ==> xs < ys" + unfolding strict_prefix_def prefix_def by blast + +lemma strict_prefixE' [elim?]: + assumes "xs < ys" + obtains z zs where "ys = xs @ z # zs" +proof - + from `xs < ys` obtain us where "ys = xs @ us" and "xs \ ys" + unfolding strict_prefix_def prefix_def by blast + with that show ?thesis by (auto simp add: neq_Nil_conv) +qed + +lemma strict_prefixI [intro?]: "xs \ ys ==> xs \ ys ==> xs < (ys::'a list)" + unfolding strict_prefix_def by blast + +lemma strict_prefixE [elim?]: + fixes xs ys :: "'a list" + assumes "xs < ys" + obtains "xs \ ys" and "xs \ ys" + using assms unfolding strict_prefix_def by blast + + +subsection {* Basic properties of prefixes *} + +theorem Nil_prefix [iff]: "[] \ xs" + by (simp add: prefix_def) + +theorem prefix_Nil [simp]: "(xs \ []) = (xs = [])" + by (induct xs) (simp_all add: prefix_def) + +lemma prefix_snoc [simp]: "(xs \ ys @ [y]) = (xs = ys @ [y] \ xs \ ys)" +proof + assume "xs \ ys @ [y]" + then obtain zs where zs: "ys @ [y] = xs @ zs" .. + show "xs = ys @ [y] \ xs \ ys" + by (metis append_Nil2 butlast_append butlast_snoc prefixI zs) +next + assume "xs = ys @ [y] \ xs \ ys" + then show "xs \ ys @ [y]" + by (metis order_eq_iff order_trans prefixI) +qed + +lemma Cons_prefix_Cons [simp]: "(x # xs \ y # ys) = (x = y \ xs \ ys)" + by (auto simp add: prefix_def) + +lemma less_eq_list_code [code]: + "([]\'a\{equal, ord} list) \ xs \ True" + "(x\'a\{equal, ord}) # xs \ [] \ False" + "(x\'a\{equal, ord}) # xs \ y # ys \ x = y \ xs \ ys" + by simp_all + +lemma same_prefix_prefix [simp]: "(xs @ ys \ xs @ zs) = (ys \ zs)" + by (induct xs) simp_all + +lemma same_prefix_nil [iff]: "(xs @ ys \ xs) = (ys = [])" + by (metis append_Nil2 append_self_conv order_eq_iff prefixI) + +lemma prefix_prefix [simp]: "xs \ ys ==> xs \ ys @ zs" + by (metis order_le_less_trans prefixI strict_prefixE strict_prefixI) + +lemma append_prefixD: "xs @ ys \ zs \ xs \ zs" + by (auto simp add: prefix_def) + +theorem prefix_Cons: "(xs \ y # ys) = (xs = [] \ (\zs. xs = y # zs \ zs \ ys))" + by (cases xs) (auto simp add: prefix_def) + +theorem prefix_append: + "(xs \ ys @ zs) = (xs \ ys \ (\us. xs = ys @ us \ us \ zs))" + apply (induct zs rule: rev_induct) + apply force + apply (simp del: append_assoc add: append_assoc [symmetric]) + apply (metis append_eq_appendI) + done + +lemma append_one_prefix: + "xs \ ys ==> length xs < length ys ==> xs @ [ys ! length xs] \ ys" + unfolding prefix_def + by (metis Cons_eq_appendI append_eq_appendI append_eq_conv_conj + eq_Nil_appendI nth_drop') + +theorem prefix_length_le: "xs \ ys ==> length xs \ length ys" + by (auto simp add: prefix_def) + +lemma prefix_same_cases: + "(xs\<^isub>1::'a list) \ ys \ xs\<^isub>2 \ ys \ xs\<^isub>1 \ xs\<^isub>2 \ xs\<^isub>2 \ xs\<^isub>1" + unfolding prefix_def by (metis append_eq_append_conv2) + +lemma set_mono_prefix: "xs \ ys \ set xs \ set ys" + by (auto simp add: prefix_def) + +lemma take_is_prefix: "take n xs \ xs" + unfolding prefix_def by (metis append_take_drop_id) + +lemma map_prefixI: "xs \ ys \ map f xs \ map f ys" + by (auto simp: prefix_def) + +lemma prefix_length_less: "xs < ys \ length xs < length ys" + by (auto simp: strict_prefix_def prefix_def) + +lemma strict_prefix_simps [simp, code]: + "xs < [] \ False" + "[] < x # xs \ True" + "x # xs < y # ys \ x = y \ xs < ys" + by (simp_all add: strict_prefix_def cong: conj_cong) + +lemma take_strict_prefix: "xs < ys \ take n xs < ys" + apply (induct n arbitrary: xs ys) + apply (case_tac ys, simp_all)[1] + apply (metis order_less_trans strict_prefixI take_is_prefix) + done + +lemma not_prefix_cases: + assumes pfx: "\ ps \ ls" + obtains + (c1) "ps \ []" and "ls = []" + | (c2) a as x xs where "ps = a#as" and "ls = x#xs" and "x = a" and "\ as \ xs" + | (c3) a as x xs where "ps = a#as" and "ls = x#xs" and "x \ a" +proof (cases ps) + case Nil then show ?thesis using pfx by simp +next + case (Cons a as) + note c = `ps = a#as` + show ?thesis + proof (cases ls) + case Nil then show ?thesis by (metis append_Nil2 pfx c1 same_prefix_nil) + next + case (Cons x xs) + show ?thesis + proof (cases "x = a") + case True + have "\ as \ xs" using pfx c Cons True by simp + with c Cons True show ?thesis by (rule c2) + next + case False + with c Cons show ?thesis by (rule c3) + qed + qed +qed + +lemma not_prefix_induct [consumes 1, case_names Nil Neq Eq]: + assumes np: "\ ps \ ls" + and base: "\x xs. P (x#xs) []" + and r1: "\x xs y ys. x \ y \ P (x#xs) (y#ys)" + and r2: "\x xs y ys. \ x = y; \ xs \ ys; P xs ys \ \ P (x#xs) (y#ys)" + shows "P ps ls" using np +proof (induct ls arbitrary: ps) + case Nil then show ?case + by (auto simp: neq_Nil_conv elim!: not_prefix_cases intro!: base) +next + case (Cons y ys) + then have npfx: "\ ps \ (y # ys)" by simp + then obtain x xs where pv: "ps = x # xs" + by (rule not_prefix_cases) auto + show ?case by (metis Cons.hyps Cons_prefix_Cons npfx pv r1 r2) +qed + + +subsection {* Parallel lists *} + +definition + parallel :: "'a list => 'a list => bool" (infixl "\" 50) where + "(xs \ ys) = (\ xs \ ys \ \ ys \ xs)" + +lemma parallelI [intro]: "\ xs \ ys ==> \ ys \ xs ==> xs \ ys" + unfolding parallel_def by blast + +lemma parallelE [elim]: + assumes "xs \ ys" + obtains "\ xs \ ys \ \ ys \ xs" + using assms unfolding parallel_def by blast + +theorem prefix_cases: + obtains "xs \ ys" | "ys < xs" | "xs \ ys" + unfolding parallel_def strict_prefix_def by blast + +theorem parallel_decomp: + "xs \ ys ==> \as b bs c cs. b \ c \ xs = as @ b # bs \ ys = as @ c # cs" +proof (induct xs rule: rev_induct) + case Nil + then have False by auto + then show ?case .. +next + case (snoc x xs) + show ?case + proof (rule prefix_cases) + assume le: "xs \ ys" + then obtain ys' where ys: "ys = xs @ ys'" .. + show ?thesis + proof (cases ys') + assume "ys' = []" + then show ?thesis by (metis append_Nil2 parallelE prefixI snoc.prems ys) + next + fix c cs assume ys': "ys' = c # cs" + then show ?thesis + by (metis Cons_eq_appendI eq_Nil_appendI parallelE prefixI + same_prefix_prefix snoc.prems ys) + qed + next + assume "ys < xs" then have "ys \ xs @ [x]" by (simp add: strict_prefix_def) + with snoc have False by blast + then show ?thesis .. + next + assume "xs \ ys" + with snoc obtain as b bs c cs where neq: "(b::'a) \ c" + and xs: "xs = as @ b # bs" and ys: "ys = as @ c # cs" + by blast + from xs have "xs @ [x] = as @ b # (bs @ [x])" by simp + with neq ys show ?thesis by blast + qed +qed + +lemma parallel_append: "a \ b \ a @ c \ b @ d" + apply (rule parallelI) + apply (erule parallelE, erule conjE, + induct rule: not_prefix_induct, simp+)+ + done + +lemma parallel_appendI: "xs \ ys \ x = xs @ xs' \ y = ys @ ys' \ x \ y" + by (simp add: parallel_append) + +lemma parallel_commute: "a \ b \ b \ a" + unfolding parallel_def by auto + + +subsection {* Postfix order on lists *} + +definition + postfix :: "'a list => 'a list => bool" ("(_/ >>= _)" [51, 50] 50) where + "(xs >>= ys) = (\zs. xs = zs @ ys)" + +lemma postfixI [intro?]: "xs = zs @ ys ==> xs >>= ys" + unfolding postfix_def by blast + +lemma postfixE [elim?]: + assumes "xs >>= ys" + obtains zs where "xs = zs @ ys" + using assms unfolding postfix_def by blast + +lemma postfix_refl [iff]: "xs >>= xs" + by (auto simp add: postfix_def) +lemma postfix_trans: "\xs >>= ys; ys >>= zs\ \ xs >>= zs" + by (auto simp add: postfix_def) +lemma postfix_antisym: "\xs >>= ys; ys >>= xs\ \ xs = ys" + by (auto simp add: postfix_def) + +lemma Nil_postfix [iff]: "xs >>= []" + by (simp add: postfix_def) +lemma postfix_Nil [simp]: "([] >>= xs) = (xs = [])" + by (auto simp add: postfix_def) + +lemma postfix_ConsI: "xs >>= ys \ x#xs >>= ys" + by (auto simp add: postfix_def) +lemma postfix_ConsD: "xs >>= y#ys \ xs >>= ys" + by (auto simp add: postfix_def) + +lemma postfix_appendI: "xs >>= ys \ zs @ xs >>= ys" + by (auto simp add: postfix_def) +lemma postfix_appendD: "xs >>= zs @ ys \ xs >>= ys" + by (auto simp add: postfix_def) + +lemma postfix_is_subset: "xs >>= ys ==> set ys \ set xs" +proof - + assume "xs >>= ys" + then obtain zs where "xs = zs @ ys" .. + then show ?thesis by (induct zs) auto +qed + +lemma postfix_ConsD2: "x#xs >>= y#ys ==> xs >>= ys" +proof - + assume "x#xs >>= y#ys" + then obtain zs where "x#xs = zs @ y#ys" .. + then show ?thesis + by (induct zs) (auto intro!: postfix_appendI postfix_ConsI) +qed + +lemma postfix_to_prefix [code]: "xs >>= ys \ rev ys \ rev xs" +proof + assume "xs >>= ys" + then obtain zs where "xs = zs @ ys" .. + then have "rev xs = rev ys @ rev zs" by simp + then show "rev ys <= rev xs" .. +next + assume "rev ys <= rev xs" + then obtain zs where "rev xs = rev ys @ zs" .. + then have "rev (rev xs) = rev zs @ rev (rev ys)" by simp + then have "xs = rev zs @ ys" by simp + then show "xs >>= ys" .. +qed + +lemma distinct_postfix: "distinct xs \ xs >>= ys \ distinct ys" + by (clarsimp elim!: postfixE) + +lemma postfix_map: "xs >>= ys \ map f xs >>= map f ys" + by (auto elim!: postfixE intro: postfixI) + +lemma postfix_drop: "as >>= drop n as" + unfolding postfix_def + apply (rule exI [where x = "take n as"]) + apply simp + done + +lemma postfix_take: "xs >>= ys \ xs = take (length xs - length ys) xs @ ys" + by (clarsimp elim!: postfixE) + +lemma parallelD1: "x \ y \ \ x \ y" + by blast + +lemma parallelD2: "x \ y \ \ y \ x" + by blast + +lemma parallel_Nil1 [simp]: "\ x \ []" + unfolding parallel_def by simp + +lemma parallel_Nil2 [simp]: "\ [] \ x" + unfolding parallel_def by simp + +lemma Cons_parallelI1: "a \ b \ a # as \ b # bs" + by auto + +lemma Cons_parallelI2: "\ a = b; as \ bs \ \ a # as \ b # bs" + by (metis Cons_prefix_Cons parallelE parallelI) + +lemma not_equal_is_parallel: + assumes neq: "xs \ ys" + and len: "length xs = length ys" + shows "xs \ ys" + using len neq +proof (induct rule: list_induct2) + case Nil + then show ?case by simp +next + case (Cons a as b bs) + have ih: "as \ bs \ as \ bs" by fact + show ?case + proof (cases "a = b") + case True + then have "as \ bs" using Cons by simp + then show ?thesis by (rule Cons_parallelI2 [OF True ih]) + next + case False + then show ?thesis by (rule Cons_parallelI1) + qed +qed + +end diff -r 34d01e9a772e -r 7d9c0ed02b56 My_list_prefix.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/My_list_prefix.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,402 @@ +(*<*) +theory My_list_prefix +imports List_Prefix +begin +(*>*) + +(* cmp:: 1:complete equal; 2:less; 3:greater; 4: len equal,but ele no equal *) +fun cmp :: "'a list \ 'a list \ nat" +where + "cmp [] [] = 1" | + "cmp [] (e#es) = 2" | + "cmp (e#es) [] = 3" | + "cmp (e#es) (a#as) = (let r = cmp es as in + if (e = a) then r else 4)" + +(* list_com:: fetch the same ele of the same left order into a new list*) +fun list_com :: "'a list \ 'a list \ 'a list" +where + "list_com [] ys = []" | + "list_com xs [] = []" | + "list_com (x#xs) (y#ys) = (if x = y then x#(list_com xs ys) else [])" + +(* list_com_rev:: by the right order of list_com *) +definition list_com_rev :: "'a list \ 'a list \ 'a list" (infix "\" 50) +where + "xs \ ys \ rev (list_com (rev xs) (rev ys))" + +(* list_diff:: list substract, once different return tailer *) +fun list_diff :: "'a list \ 'a list \ 'a list" +where + "list_diff [] xs = []" | + "list_diff (x#xs) [] = x#xs" | + "list_diff (x#xs) (y#ys) = (if x = y then list_diff xs ys else (x#xs))" + +(* list_diff_rev:: list substract with rev order*) +definition list_diff_rev :: "'a list \ 'a list \ 'a list" (infix "\" 51) +where + "xs \ ys \ rev (list_diff (rev xs) (rev ys))" + +(* xs <= ys:: \zs. ys = xs @ zs *) +(* no_junior:: xs is ys' tail,or equal *) +definition no_junior :: "'a list \ 'a list \ bool" (infix "\" 50) +where + "xs \ ys \ rev xs \ rev ys" + +(* < :: xs <= ys \ xs \ ys *) +(* is_ancestor:: xs is ys' tail, but no equal *) +definition is_ancestor :: "'a list \ 'a list \ bool" (infix "\" 50) +where + "xs \ ys \ rev xs < rev ys" + +lemma list_com_diff [simp]: "(list_com xs ys) @ (list_diff xs ys) = xs" (is "?P xs ys") + by (rule_tac P = ?P in cmp.induct, simp+) + +lemma list_com_diff_rev [simp]: "(xs \ ys) @ (xs \ ys) = xs" + apply (simp only:list_com_rev_def list_diff_rev_def) + by (fold rev_append, simp) + +lemma list_com_commute: "list_com xs ys = list_com ys xs" (is "?P xs ys") + by (rule_tac P = ?P in cmp.induct, simp+) + +lemma list_com_ido: "xs \ ys \ list_com xs ys = xs" (is "?P xs ys") + by (rule_tac P = ?P in cmp.induct, auto+) + +lemma list_com_rev_ido [simp]: "xs \ ys \ xs \ ys = xs" + by (cut_tac list_com_ido, auto simp: no_junior_def list_com_rev_def) + +lemma list_com_rev_commute [iff]: "(xs \ ys) = (ys \ xs)" + by (simp only:list_com_rev_def list_com_commute) + +lemma list_com_rev_ido1 [simp]: "xs \ ys \ ys \ xs = xs" + by simp + +lemma list_diff_le: "(list_diff xs ys = []) = (xs \ ys)" (is "?P xs ys") + by (rule_tac P = ?P in cmp.induct, simp+) + + +lemma list_diff_rev_le: "(xs \ ys = []) = (xs \ ys)" + by (auto simp:list_diff_rev_def no_junior_def list_diff_le) + +lemma list_diff_lt: "(list_diff xs ys = [] \ list_diff ys xs \ []) = (xs < ys)" (is "?P xs ys") + by (rule_tac P = ?P in cmp.induct, simp+) + +lemma list_diff_rev_lt: "(xs \ ys = [] \ ys \ xs \ []) = (xs \ ys)" + by (auto simp: list_diff_rev_def list_diff_lt is_ancestor_def) + + +(* xs diff ys result not [] \ \ e \ xs. a \ ys. e \ a *) +lemma list_diff_neq: + "\ e es a as. list_diff xs ys = (e#es) \ list_diff ys xs = (a#as) \ e \ a" (is "?P xs ys") + by (rule_tac P = ?P in cmp.induct, simp+) + +lemma list_diff_rev_neq_pre: "\ e es a as. xs \ ys = rev (e#es) \ ys \ xs = rev (a#as) \ e \ a" + apply (simp only:list_diff_rev_def, clarify) + apply (insert list_diff_neq, atomize) + by (erule_tac x = "rev xs" in allE, erule_tac x = "rev ys" in allE, blast) + +lemma list_diff_rev_neq: "\ e es a as. xs \ ys = es @ [e] \ ys \ xs = as @ [a] \ e \ a" + apply (rule_tac allI)+ + apply (insert list_diff_rev_neq_pre, atomize) + apply (erule_tac x = "xs" in allE) + apply (erule_tac x = "ys" in allE) + apply (erule_tac x = "e" in allE) + apply (erule_tac x = "rev es" in allE) + apply (erule_tac x = "a" in allE) + apply (erule_tac x = "rev as" in allE) + by auto + +lemma list_com_self [simp]: "list_com zs zs = zs" + by (induct_tac zs, simp+) + +lemma list_com_rev_self [simp]: "zs \ zs = zs" + by (simp add:list_com_rev_def) + +lemma list_com_append [simp]: "(list_com (zs @ xs) (zs @ ys)) = (zs @ (list_com xs ys))" + by (induct_tac zs, simp+) + +lemma list_inter_append [simp]: "((xs @ zs) \ (ys @ zs)) = ((xs \ ys) @ zs)" + by (simp add:list_com_rev_def) + +lemma list_diff_djoin_pre: + "\ e es a as. list_diff xs ys = e#es \ list_diff ys xs = a#as \ (\ zs zs'. (list_diff (xs @ zs) (ys @ zs') = [e]@es@zs))" + (is "?P xs ys") + by (rule_tac P = ?P in cmp.induct, simp+) + +lemma list_diff_djoin_rev_pre: + "\ e es a as. xs \ ys = rev (e#es) \ ys \ xs = rev (a#as) \ (\ zs zs'. ((zs @ xs) \ (zs' @ ys) = rev ([e]@es@rev zs)))" + apply (simp only: list_diff_rev_def, clarify) + apply (insert list_diff_djoin_pre, atomize) + apply (erule_tac x = "rev xs" in allE) + apply (erule_tac x = "rev ys" in allE) + apply (erule_tac x = "e" in allE) + apply (erule_tac x = "es" in allE) + apply (erule_tac x = "a" in allE) + apply (erule_tac x = "as" in allE) + by simp + +lemma list_diff_djoin_rev: + "xs \ ys = es @ [e] \ ys \ xs = as @ [a] \ zs @ xs \ zs' @ ys = zs @ es @ [e]" + apply (insert list_diff_djoin_rev_pre [rule_format, simplified]) + apply (clarsimp, atomize) + apply (erule_tac x = "xs" in allE) + apply (erule_tac x = "ys" in allE) + apply (erule_tac x = "rev es" in allE) + apply (erule_tac x = "e" in allE) + apply (erule_tac x = "rev as" in allE) + apply (erule_tac x = "a" in allE) + by auto + +lemmas list_diff_djoin_rev_simplified = conjI [THEN list_diff_djoin_rev, simp] + +lemmas list_diff_djoin = conjI [THEN list_diff_djoin_pre [rule_format], simp] + +lemma list_diff_ext_left [simp]: "(list_diff (zs @ xs) (zs @ ys) = (list_diff xs ys))" + by (induct_tac zs, simp+) + +lemma list_diff_rev_ext_left [simp]: "((xs @ zs \ ys @ zs) = (xs \ ys))" + by (auto simp: list_diff_rev_def) + +declare no_junior_def [simp] + +lemma no_juniorE: "\xs \ ys; \ zs. ys = zs @ xs \ R\ \ R" +proof - + assume h: "xs \ ys" + and h1: "\ zs. ys = zs @ xs \ R" + show "R" + proof - + from h have "rev xs \ rev ys" by (simp) + from this obtain zs where eq_rev: "rev ys = rev xs @ zs" by (auto simp:prefix_def) + show R + proof(rule h1 [where zs = "rev zs"]) + from rev_rev_ident and eq_rev have "rev (rev (ys)) = rev zs @ rev (rev xs)" + by simp + thus "ys = rev zs @ xs" by simp + qed + qed +qed + +lemma no_juniorI: "\ys = zs @ xs\ \ xs \ ys" + by simp + +lemma no_junior_ident [simp]: "xs \ xs" + by simp + +lemma no_junior_expand: "xs \ ys = ((xs \ ys) \ xs = ys)" + by (simp only:no_junior_def is_ancestor_def strict_prefix_def, blast) + +lemma no_junior_same_prefix: " e # \ \ e' # \' \ \ \ \'" +apply (simp add:no_junior_def ) +apply (erule disjE, simp) +apply (simp only:prefix_def) +by (erule exE, rule_tac x = "[e] @ zs" in exI, auto) + +lemma no_junior_noteq: "\\ \ a # \'; \ \ a # \'\ \ \ \ \'" +apply (erule no_juniorE) +by (case_tac zs, simp+) + +lemma is_ancestor_app [simp]: "xs \ ys \ xs \ zs @ ys" + by (auto simp:is_ancestor_def strict_prefix_def) + +lemma is_ancestor_cons [simp]: "xs \ ys \ xs \ a # ys" + by (auto simp:is_ancestor_def strict_prefix_def) + +lemma no_junior_app [simp]: "xs \ ys \ xs \ zs @ ys" + by simp + +lemma is_ancestor_no_junior [simp]: "xs \ ys \ xs \ ys" + by (simp add:is_ancestor_def) + +lemma is_ancestor_y [simp]: "ys \ y#ys" + by (simp add:is_ancestor_def strict_prefix_def) + +lemma no_junior_cons [simp]: "xs \ ys \ xs \ (y#ys)" + by (unfold no_junior_expand, auto) + +lemma no_junior_anti_sym: "\xs \ ys; ys \ xs\ \ xs = ys" + by simp + +declare no_junior_def [simp del] + +(* djoin:: xs and ys is not the other's tail, not equal either *) +definition djoin :: "'a list \ 'a list \ bool" (infix "\" 50) +where + "xs \ ys \ \ (xs \ ys \ ys \ xs)" + +(* dinj:: function f's returning list is not tailing when paras not equal *) +definition dinj :: "('a \ 'b list) \ bool" +where + "dinj f \ (\ a b. a \ b \ f a \ f b)" + + +(* list_cmp:: list comparison: one is other's prefix or no equal at some position *) +lemma list_cmp: "xs \ ys \ ys \ xs \ (\ zs x y a b. xs = zs @ [a] @ x \ ys = zs @ [b] @ y \ a \ b)" +proof(cases "list_diff xs ys") + assume " list_diff xs ys = []" with list_diff_le show ?thesis by blast +next + fix e es + assume h: "list_diff xs ys = e # es" + show ?thesis + proof(cases "list_diff ys xs") + assume " list_diff ys xs = []" with list_diff_le show ?thesis by blast + next + fix a as assume h1: "list_diff ys xs = (a # as)" + have "xs = (list_com xs ys) @ [e] @ es \ ys = (list_com xs ys) @ [a] @ as \ e \ a" + apply (simp, fold h1, fold h) + apply (simp,subst list_com_commute, simp) + apply (rule_tac list_diff_neq[rule_format]) + by (insert h1, insert h, blast) + thus ?thesis by blast + qed +qed + +(* In fact, this is a case split *) +lemma list_diff_ind: "\list_diff xs ys = [] \ R; list_diff ys xs = [] \ R; + \ e es a as. \list_diff xs ys = e#es; list_diff ys xs = a#as; e \ a\ \ R\ \ R" +proof - + assume h1: "list_diff xs ys = [] \ R" + and h2: "list_diff ys xs = [] \ R" + and h3: "\ e es a as. \list_diff xs ys = e#es; list_diff ys xs = a#as; e \ a\ \ R" + show R + proof(cases "list_diff xs ys") + assume "list_diff xs ys = []" from h1 [OF this] show R . + next + fix e es + assume he: "list_diff xs ys = e#es" + show R + proof(cases "list_diff ys xs") + assume "list_diff ys xs = []" from h2 [OF this] show R . + next + fix a as + assume ha: "list_diff ys xs = a#as" show R + proof(rule h3 [OF he ha]) + from list_diff_neq [rule_format, OF conjI [OF he ha ]] + show "e \ a" . + qed + qed + qed +qed + +lemma list_diff_rev_ind: + "\xs \ ys = [] \ R; ys \ xs = [] \ R; \ e es a as. \xs \ ys = es@[e]; ys \ xs = as@[a]; e \ a\ \ R\ \ R" +proof - + fix xs ys R + assume h1: "xs \ ys = [] \ R" + and h2: "ys \ xs = [] \ R" + and h3: "\ e es a as. \xs \ ys = es@[e]; ys \ xs = as@[a]; e \ a\ \ R" + show R + proof (rule list_diff_ind [where xs = "rev xs" and ys = "rev ys"]) + assume "list_diff (rev xs) (rev ys) = []" thus R by (auto intro:h1 simp:list_diff_rev_def) + next + assume "list_diff (rev ys) (rev xs) = []" thus R by (auto intro:h2 simp:list_diff_rev_def) + next + fix e es a as + assume "list_diff (rev xs) (rev ys) = e # es" + and "list_diff (rev ys) (rev xs) = a # as" + and " e \ a" + thus R by (auto intro:h3 simp:list_diff_rev_def) + qed +qed + +lemma djoin_diff_iff: "(xs \ ys) = (\ e es a as. list_diff (rev xs) (rev ys) = e#es \ list_diff (rev ys) (rev xs) = a#as \ a \ e)" +proof (rule list_diff_ind [where xs = "rev xs" and ys = "rev ys"]) + assume "list_diff (rev xs) (rev ys) = []" + hence "xs \ ys" by (unfold no_junior_def, simp add:list_diff_le) + thus ?thesis + apply (auto simp:djoin_def no_junior_def) + by (fold list_diff_le, simp) +next + assume "list_diff (rev ys) (rev xs) = []" + hence "ys \ xs" by (unfold no_junior_def, simp add:list_diff_le) + thus ?thesis + apply (auto simp:djoin_def no_junior_def) + by (fold list_diff_le, simp) +next + fix e es a as + assume he: "list_diff (rev xs) (rev ys) = e # es" + and ha: "list_diff (rev ys) (rev xs) = a # as" + and hn: "e \ a" + show ?thesis + proof + from he ha hn + show + "\e es a as. list_diff (rev xs) (rev ys) = e # es \ list_diff (rev ys) (rev xs) = a # as \ a \ e" + by blast + next + from he ha hn + show "xs \ ys" + by (auto simp:djoin_def no_junior_def, fold list_diff_le, simp+) + qed +qed + +lemma djoin_diff_rev_iff: "(xs \ ys) = (\ e es a as. xs \ ys = es@[e] \ ys \ xs = as@[a] \ a \ e)" + apply (auto simp:djoin_diff_iff list_diff_rev_def) + apply (rule_tac x = e in exI, safe) + apply (rule_tac x = "rev es" in exI) + apply (rule_tac injD[where f = rev], simp+) + apply (rule_tac x = "a" in exI, safe) + apply (rule_tac x = "rev as" in exI) + apply (rule_tac injD[where f = rev], simp+) + done + +lemma djoin_revE: "\xs \ ys; \e es a as. \xs \ ys = es@[e]; ys \ xs = as@[a]; a \ e\ \ R\ \ R" + by (unfold djoin_diff_rev_iff, blast) + +lemma djoin_append_left[simp, intro]: "xs \ ys \ (zs' @ xs) \ (zs @ ys)" + by (auto simp:djoin_diff_iff intro:list_diff_djoin[simplified]) + +lemma djoin_cons_left[simp]: "xs \ ys \ (e # xs) \ (a # ys)" + by (drule_tac zs' = "[e]" and zs = "[a]" in djoin_append_left, simp) + +lemma djoin_simp_1 [simp]: "xs \ ys \ xs \ (zs @ ys)" + by (drule_tac djoin_append_left [where zs' = "[]"], simp) + +lemma djoin_simp_2 [simp]: "xs \ ys \ (zs' @ xs) \ ys" + by (drule_tac djoin_append_left [where zs = "[]"], simp) + +lemma djoin_append_right[simp]: "xs \ ys \ (xs @ zs) \ (ys @ zs)" + by (simp add:djoin_diff_iff) + +lemma djoin_cons_append[simp]: "xs \ ys \ (x # xs) \ (zs @ ys)" + by (subgoal_tac "[x] @ xs \ zs @ ys", simp, blast) + +lemma djoin_append_cons[simp]: "xs \ ys \ (zs @ xs) \ (y # ys)" + by (subgoal_tac "zs @ xs \ [y] @ ys", simp, blast) + +lemma djoin_neq [simp]: "xs \ ys \ xs \ ys" + by (simp only:djoin_diff_iff, clarsimp) + +lemma djoin_cons [simp]: "e \ a \ e # xs \ a # xs" + by (unfold djoin_diff_iff, simp) + +lemma djoin_append_e [simp]: "e \ a \ (zs @ [e] @ xs) \ (zs' @ [a] @ xs)" + by (unfold djoin_diff_iff, simp) + +lemma djoin_mono [simp]: "\xs \ ys; xs \ xs'; ys \ ys'\ \ xs' \ ys'" +proof(erule_tac djoin_revE,unfold djoin_diff_rev_iff) + fix e es a as + assume hx: "xs \ xs'" + and hy: "ys \ ys'" + and hmx: "xs \ ys = es @ [e]" + and hmy: "ys \ xs = as @ [a]" + and neq: "a \ e" + have "xs' \ ys' = ((xs' \ xs) @ es) @ [e] \ ys' \ xs' = ((ys' \ ys) @ as) @ [a] \ a \ e" + proof - + from hx have heqx: "(xs' \ xs) @ xs = xs'" + by (cut_tac list_com_diff_rev [of xs' xs], subgoal_tac "xs' \ xs = xs", simp+) + moreover from hy have heqy: "(ys' \ ys) @ ys = ys'" + by (cut_tac list_com_diff_rev [of ys' ys], subgoal_tac "ys' \ ys = ys", simp+) + moreover from list_diff_djoin_rev_simplified [OF hmx hmy] + have "((xs' \ xs) @ xs) \ ((ys' \ ys) @ ys) = (xs' \ xs) @ es @ [e]" by simp + moreover from list_diff_djoin_rev_simplified [OF hmy hmx] + have "((ys' \ ys) @ ys) \ ((xs' \ xs) @ xs) = (ys' \ ys) @ as @ [a]" by simp + ultimately show ?thesis by (simp add:neq) + qed + thus "\e es a as. xs' \ ys' = es @ [e] \ ys' \ xs' = as @ [a] \ a \ e" by blast +qed + +lemmas djoin_append_e_simplified [simp] = djoin_append_e [simplified] + +(*<*) +end +(*>*) \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 New_obj_prop.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/New_obj_prop.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,92 @@ +theory New_obj_prop +imports Main Finite_current Flask_type Flask Static +begin + +context tainting_s begin + +lemma nn_notin_aux: "finite s \ \ a \ s. Max s \ a " +apply (erule finite.induct, simp) +apply (rule ballI) +apply (case_tac "aa = a", simp+) +done + +lemma nn_notin: "finite s \ next_nat s \ s" +apply (drule nn_notin_aux) +apply (simp add:next_nat_def) +by (auto) + + +(* lemmas of new created obj *) + +lemma np_notin_curp: "valid \ \ new_proc \ \ current_procs \" using finite_cp +by (simp add:new_proc_def nn_notin) + +lemma np_notin_curp': "\new_proc \ \ current_procs \; valid \\ \ False" +apply (drule np_notin_curp, simp) +done + +lemma ni_notin_curi: "valid \ \ new_inode_num \ \ current_inode_nums \" +apply (drule finite_ci) +by (simp add:new_inode_num_def nn_notin) + +lemma ni_notin_curi': "\new_inode_num \ \ current_inode_nums \; valid \\ \ False" +by (drule ni_notin_curi, simp) + +lemma nm_notin_curm: "valid \ \ new_msgq \ \ current_msgqs \" using finite_cm +by (simp add:new_msgq_def nn_notin) + +lemma nm_notin_curm': "\new_msgq \ \ current_msgqs \; valid \\ \ False" +by (drule nm_notin_curm, simp) + +lemma nh_notin_curh: "valid \ \ new_shm \ \ current_shms \" using finite_ch +by (simp add:new_shm_def nn_notin) + +lemma nh_notin_curh': "\new_shm \ \ current_shms \; valid \\ \ False" +by (drule nh_notin_curh, simp) + +lemma nfd_notin_curfd: "valid \ \ new_proc_fd \ p \ current_proc_fds \ p" +using finite_cfd[where p = p] +apply (simp add:new_proc_fd_def nn_notin) +done + +lemma nfd_notin_curfd': "\new_proc_fd \ p \ current_proc_fds \ p; valid \\ \ False" +by (drule nfd_notin_curfd[where p = p], simp) + +lemma nim_notin_curim: "valid \ \ new_inode_num \ \ current_inode_nums \" +by (drule finite_ci, simp add:new_inode_num_def nn_notin) + +lemma nim_notin_curim': "\new_inode_num \ \ current_inode_nums \; valid \\ \ False" +by (drule nim_notin_curim, simp) + +lemma len_fname_all: "length (fname_all_a len) = len" +by (induct len, auto simp:fname_all_a.simps) + +lemma ncf_notin_curf: "valid \ \ new_childf f \ \ current_files \" +apply (drule finite_cf) +apply (simp add:new_childf_def next_fname_def all_fname_under_dir_def) +apply (rule notI) +apply (subgoal_tac "(CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \ current_files \}))) \ {fn. fn # f \ current_files \}") +defer apply simp +apply (subgoal_tac "length (CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \ current_files \}))) \ fname_length_set {fn. fn # f \ current_files \}") +defer apply (auto simp:fname_length_set_def image_def)[1] +apply (subgoal_tac "finite (fname_length_set {fn. fn # f \ current_files \})") +defer +apply (simp add:fname_length_set_def) +apply (rule finite_imageI) +apply (drule_tac h = "\ f'. case f' of [] \ '''' | fn # pf' \ if (pf' = f) then fn else ''''" in finite_imageI) +apply (rule_tac B = "(list_case [] (\fn pf'. if pf' = f then fn else []) ` current_files \)" in finite_subset) +unfolding image_def +apply (clarsimp split:if_splits) +apply (rule_tac x = "x # f" in bexI, simp+) + +apply (drule_tac s = "(fname_length_set {fn. fn # f \ current_files \})" in nn_notin_aux) +apply (erule_tac x = "length (CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \ current_files \})))" in ballE) +apply (simp add:len_fname_all, simp) +done + +lemma ncf_parent: "valid \ \ parent (new_childf f \) = Some f" +by (simp add:new_childf_def) + +end + +end diff -r 34d01e9a772e -r 7d9c0ed02b56 OS_type_def.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/OS_type_def.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,276 @@ +theory OS_type_def +imports Main Flask_type +begin + +(****************************** os files ***************************) + +datatype t_open_must_flag = + OF_RDONLY +| OF_WRONLY +| OF_RDWR + +datatype t_open_option_flag = + OF_APPEND +| OF_CREAT +(* +| OF_TRUNC +*) +| OF_EXCL +(* +| OF_DIRECTORY + +| OF_OTHER +*) + +type_synonym t_open_flags = "t_open_must_flag \ t_open_option_flag set" + +definition is_read_flag :: "t_open_flags \ bool" +where + "is_read_flag flag \ let (mflag, oflag) = flag in (mflag = OF_RDONLY \ mflag = OF_RDWR)" + +definition is_write_flag :: "t_open_flags \ bool" +where + "is_write_flag flag \ let (mflag, oflag) = flag in (mflag = OF_WRONLY \ mflag = OF_RDWR)" + +definition is_excl_flag :: "t_open_flags \ bool" +where + "is_excl_flag flag \ let (mflag, oflag) = flag in (OF_EXCL \ oflag)" + +definition is_creat_flag :: "t_open_flags \ bool" +where + "is_creat_flag flag \ let (mflag, oflag) = flag in (OF_CREAT \ oflag)" + +definition is_creat_excl_flag :: "t_open_flags \ bool" +where + "is_creat_excl_flag flag \ let (mflag, oflag) = flag in (OF_CREAT \ oflag \ OF_EXCL \ oflag)" +(* +definition is_trunc_flag :: "t_open_flags \ bool" +where + "is_trunc_flag flag \ let (mflag, oflag) = flag in (OF_TRUNC \ oflag \ is_write_flag flag)" + +definition is_cloexec_flag :: "t_open_flags \ bool" +where + "is_cloexec_flag flag \ let (mflag, oflag) = flag in (OF_CLOEXEC \ oflag)" +*) +definition is_append_flag :: "t_open_flags \ bool" +where + "is_append_flag flag \ let (mflag, oflag) = flag in (OF_APPEND \ oflag)" + +(* +definition is_dir_flag :: "t_open_flags \ bool" +where + "is_dir_flag flag \ let (mflag, oflag) = flag in (OF_DIRECTORY \ oflag)" +*) + + +type_synonym t_fname = string + +(* listing of file path: + * a#b#c:: b is a's parent and c is b's parent; + * [] represents the root file + *) +type_synonym t_file = "t_fname list" + +type_synonym t_fd = "nat" (* given a fd, must has funs to tell its corresponding: "t_file \ t_open_flags \ t_inode" *) + + +(* need not consideration, for rdonly transfer virus too +datatype t_sharememo_flag = + SHM_RDONLY +| SHM_OTHER + +type_synonym t_sharememo_flags = "t_sharememo_flag set" +*) + + + +(****************************** os sockets ***************************) + +datatype t_shutdown_how = + SHUT_RD +| SHUT_WR +| SHUT_RDWR + +datatype t_zero_one = + Z +| O + +type_synonym t_inet_addr = "t_zero_one list" + +type_synonym t_socket_port = nat + +type_synonym t_netif = string + +(* +datatype t_socket_af = + AF_UNIX "t_file \ t_file" +| AF_INET "t_inet_addr \ t_socket_port \ t_inet_addr \ t_socket_port \ t_netdev" +*) + +datatype t_socket_af = +(* AF_UNIX +| *)AF_INET + +datatype t_socket_type = + STREAM +| DGRAM + +datatype t_sock_addr = + INET_ADDR t_inet_addr t_socket_port(* +| UNIX_ADDR t_file*) + +type_synonym t_node = t_inet_addr (* a endpoint in the netword *) + +datatype t_sock_trans_state = + Trans_Recv +| Trans_Send +| Trans_RS +| Trans_No + +datatype t_sock_state = + SS_create +| SS_bind +| SS_connect +| SS_listen +| SS_accepted +| SS_trans "t_sock_trans_state" + + +(****************************** os other IPCs ***************************) + +datatype t_shm_attach_flag = + SHM_RDONLY +| SHM_RDWR + +type_synonym t_shm = nat + + +(* message queue is indexed by a nat, in OS + * but logically it is a list, created as [], updated as a queue + * so, it type should not be nat, but a t_msg list + *) +type_synonym t_msgq = nat + +(* a message of ID nat, but it has meaning under "in what message queue" *) +type_synonym t_msg = "nat" + +(****************************** os inodes ***************************) + +type_synonym t_inode_num = nat + +(* +datatype t_sock_af = + Tag_AF_UNIX "(t_file option \ t_file option)" "t_socket_type" +| Tag_AF_INET "(t_inet_addr \ t_socket_port) option \ (t_inet_addr \ t_socket_port) option \ t_netdev option" "t_socket_type" +*) + +datatype t_inode_tag = + Tag_FILE (* this inode is a file, and its file is: "t_file" :: no need to do that and another reason is link may add two or more files to the inode, so it should be a list, can the maintainment of this list is tough *) +| Tag_DIR (* do not contain chs anymore, for more infomation see svn log + "t_inode_num list" *) (* Tag_DIR inode number list: the inode number list of subchild files (direct son the file) , means: new created child file should add to this field. *) +| Tag_TCP_SOCK (* information below is turn to seperative listener functions +"(t_inet_addr \ t_socket_port) option \ (t_inet_addr \ t_socket_port) option \ t_netdev option" "t_socket_type" *) +| Tag_UDP_SOCK + +(* type_synonym t_inode = "nat \ t_inode_tag" (* inode (id, tag) contains file internal id and directory tag attributes *) + // not needed any more: for inode num and its tag now are seperated by two funcs. +*) + + +type_synonym t_process = nat + +type_synonym t_trunc_len = nat + +type_synonym t_socket = "t_process \ t_fd" + +(* each constructor is defined according to each security-class *) +datatype t_object = + O_proc "t_process" + +| O_file "t_file" +| O_fd "t_process" "t_fd" +| O_dir "t_file" + +| O_node "t_node" +| O_tcp_sock "t_socket" +| O_udp_sock "t_socket" + +| O_shm "t_shm" +| O_msgq "t_msgq" +| O_msg "t_msgq" "t_msg" + +(* +datatype t_tainted_obj = + TO_proc "t_process" +| TO_file "t_file" +| TO_msg "t_msgq" "t_msg" +*) + + +(* not needed anymore +datatype t_clone_share_flag = +| CF_share_ipc +| CF_share_fd +| CF_share_both +| CF_no_share *) + +(* unformalised system calls: + * 1. process management + * 1.1 wait: nothing to do with tainting relation + * 1.2 getpriority/setpriority/getsched/setsched: nothing to do with security + * 1.3 getcap/setcap: "at this time that will not be done" + * 1.4 getpgid/setpgid: ... + *) +datatype t_event = + (* process management *) +(* trace can be done trough tags of cloning and execving, but I ignore this. *) + Execve "t_process" "t_file" "t_fd set" +(* | Uselib "t_process" "t_file" *) +| Clone "t_process" "t_process" "t_fd set" "t_shm set" (* "t_clone_share_flag" is not needed, as event should record the fds and shms that passed to childprocess *) +| Kill "t_process" "t_process" +| Ptrace "t_process" "t_process" +| Exit "t_process" + +| Open "t_process" "t_file" "t_open_flags" "t_fd" "nat option" (* the last nat option : if create file, the is Some inode_num else None *) +| ReadFile "t_process" "t_fd" +| WriteFile "t_process" "t_fd" +(* | Sendfile "t_process" "t_fd" "t_fd" + can be modeled as readfile then writefile *) +| CloseFd "t_process" "t_fd" +| UnLink "t_process" "t_file" +| Rmdir "t_process" "t_file" +| Mkdir "t_process" "t_file" "nat" (* last nat: inode_num *) +| LinkHard "t_process" "t_file" "t_file" +| Truncate "t_process" "t_file" "t_trunc_len" +(* | FTruncate "t_process" "t_fd" "t_trunc_len" *) +(* +| Rename "t_process" "t_file" "t_file" +*) + +(* the message is in the msgq, so we must build a cache(list) + * to store the unreceived msg yet. And the os_grant should + * check if a message queue is empty before RecvMsg. + *) +| CreateMsgq "t_process" "t_msgq" +| SendMsg "t_process" "t_msgq" "t_msg" +| RecvMsg "t_process" "t_msgq" "t_msg" +| RemoveMsgq "t_process" "t_msgq" + +| CreateShM "t_process" "t_shm" +| Attach "t_process" "t_shm" "t_shm_attach_flag" +| Detach "t_process" "t_shm" +| DeleteShM "t_process" "t_shm" + +| CreateSock "t_process" "t_socket_af" "t_socket_type" "t_fd" "nat" (* last nat: inode_num *) +| Bind "t_process" "t_fd" "t_sock_addr" +| Connect "t_process" "t_fd" "t_sock_addr" +| Listen "t_process" "t_fd" +| Accept "t_process" "t_fd" "t_sock_addr" "nat" "t_fd" "nat" (* addr: remote socket address; nat: port num; last nat: inode_num *) +| SendSock "t_process" "t_fd" +| RecvSock "t_process" "t_fd" +| Shutdown "t_process" "t_fd" "t_shutdown_how" + +(* the state of the system is decided by the initial state and the event trace*) +type_synonym t_state = "t_event list" + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Proc_fd_of_file_prop.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Proc_fd_of_file_prop.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,115 @@ +theory Proc_fd_of_file_prop +imports Main Flask Flask_type Valid_prop Current_files_prop +begin + +context flask begin + +lemma proc_fd_in_procs: "\file_of_proc_fd \ p fd = Some f; valid \\ \ p \ current_procs \" +apply (induct \ arbitrary: f) defer +apply (frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:file_of_proc_fd.simps current_procs.simps os_grant.simps split:if_splits option.splits) +by (drule init_filefd_valid, simp) + +lemma proc_fd_in_fds_aux: "\ p f. file_of_proc_fd \ p fd = Some f \ valid \ \ fd \ current_proc_fds \ p" +apply (induct \) +apply (simp add:file_of_proc_fd.simps current_proc_fds.simps) +apply (clarify, drule init_filefd_valid, simp) +apply (clarify, frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto simp:file_of_proc_fd.simps current_proc_fds.simps split:if_splits option.splits t_sock_addr.splits) +done + +lemma proc_fd_in_fds: "\file_of_proc_fd \ p fd = Some f; valid \\ \ fd \ current_proc_fds \ p" +by (rule proc_fd_in_fds_aux[rule_format], simp+) + +lemma proc_fd_file_in_cur: "\(p, fd) \ proc_fd_of_file \ f; valid \\ \ f \ current_files \" +by (auto simp:proc_fd_of_file_def intro:file_of_pfd_in_current) + +lemma proc_fd_file_in_cur': "\proc_fd_of_file \ f \ {}; valid \\ \ f \ current_files \" +by (auto simp:proc_fd_file_in_cur) + +lemma proc_fd_file_in_cur'': "\proc_fd_of_file \ f = {(p,fd)}; valid \\ \ f \ current_files \" +by (auto simp:proc_fd_file_in_cur') + +lemma procfd_of_file_imp_fpfd: "proc_fd_of_file \ f = {(p, fd)} \ file_of_proc_fd \ p fd = Some f" +by (auto simp:proc_fd_of_file_def) + +lemma procfd_of_file_imp_fpfd': "proc_fd_of_file \ f = {(p, fd)} \ file_of_proc_fd \ p fd \ None" +by (auto simp:proc_fd_of_file_def) + +lemma procfd_of_file_eq_fpfd'': "(p, fd) \ proc_fd_of_file \ f = (file_of_proc_fd \ p fd = Some f)" +by (auto simp:proc_fd_of_file_def) + +lemma procfd_of_file_non_empty: "file_of_proc_fd \ p fd = Some f \ proc_fd_of_file \ f \ {}" +by (auto simp:proc_fd_of_file_def) + +lemma file_of_proc_fd_in_curf: "\file_of_proc_fd \ p fd = Some f; valid \\ \ f \ current_files \" +by (drule procfd_of_file_non_empty, simp add:proc_fd_file_in_cur') + + +(******************* rebuild proc_fd_of_file simpset ***********************) +(* +lemma proc_fd_of_file_open: "Open p f flags fd iopt # valid \ \ + proc_fd_of_file (Open p f flags fd iopt # \) f' = (if (f' = f) then insert (p, fd) (proc_fd_of_file \ f') else proc_fd_of_file \ f')" +apply (auto simp:proc_fd_of_file_def file_of_proc_fd.simps split:if_splits) +apply (frule vd_cons, drule vt_grant_os, case_tac iopt) +apply (drule proc_fd_in_fds, simp, simp add:os_grant.simps nfd_notin_curfd)+ +done + +lemma proc_fd_of_file_closefd: "proc_fd_of_file (CloseFd p fd # \) f = (if (file_of_proc_fd \ p fd = Some f) then (proc_fd_of_file \ f - {(p,fd)}) else proc_fd_of_file \ f) " +by (auto simp:proc_fd_of_file_def file_of_proc_fd.simps split:if_splits) + +lemma proc_fd_of_file_rename: "\Rename p f\<^isub>2 f\<^isub>3 # valid \; f \ current_files (Rename p f\<^isub>2 f\<^isub>3 # \)\ \ + proc_fd_of_file (Rename p f\<^isub>2 f\<^isub>3 # \) f = (if (f\<^isub>3 \ f) then proc_fd_of_file \ (file_before_rename f\<^isub>2 f\<^isub>3 f) else proc_fd_of_file \ f)" +apply (frule vt_grant_os, frule vd_cons) +apply (case_tac "f\<^isub>3 \ f") +apply (subgoal_tac "f \ current_files \") prefer 2 apply (rule notI) +apply (clarsimp simp:os_grant.simps, drule_tac f = f\<^isub>3 and f' = f in ancenf_in_current, simp, simp, simp) +apply (auto simp add:proc_fd_of_file_def)[1] + +apply (simp add:file_of_proc_fd.simps split:option.splits if_splits) +apply (drule_tac f\<^isub>3 = f\<^isub>3 and f\<^isub>1 = aa and f\<^isub>2 = f\<^isub>2 in file_renaming_prop5, simp) +apply (drule file_of_pfd_in_current, simp+) +apply (simp add:file_of_proc_fd.simps) +apply (rule conjI, rule impI, simp add:file_renaming_prop5') +apply (rule impI, simp add:file_before_rename_def) + +apply (simp add:proc_fd_of_file_def split:if_splits) +apply auto +apply (simp add:file_of_proc_fd.simps split:option.splits if_splits) +apply (drule_tac f\<^isub>3 = f\<^isub>3 and f\<^isub>2 = f\<^isub>2 and f = aa in file_renaming_prop1, simp) +apply (simp add:current_files_simps) +apply (erule exE| erule conjE)+ +apply (simp add:file_of_proc_fd.simps split:option.splits if_splits) +apply (drule_tac f = f\<^isub>1 in rename_renaming_decom', simp+) +apply (simp add:file_after_rename_def) +done + + +lemma proc_fd_of_file_kill: "proc_fd_of_file (Kill p\<^isub>1 p\<^isub>2 # \) f = {(p, fd). (p, fd) \ proc_fd_of_file \ f \ p \ p\<^isub>2}" +by (auto simp:proc_fd_of_file_def file_of_proc_fd.simps) + +lemma proc_fd_of_file_exit: "proc_fd_of_file (Exit p' # \) f = {(p, fd). (p, fd) \ proc_fd_of_file \ f \ p \ p'}" +by (auto simp:proc_fd_of_file_def file_of_proc_fd.simps) + +lemma proc_fd_of_file_clone: "Clone p\<^isub>1 p\<^isub>2 # valid \ \ proc_fd_of_file (Clone p\<^isub>1 p\<^isub>2 # \) f = proc_fd_of_file \ f \ {(p\<^isub>2, fd)| fd. (p\<^isub>1, fd) \ proc_fd_of_file \ f}" +apply (auto simp:proc_fd_of_file_def file_of_proc_fd.simps) +apply (frule vd_cons, drule vt_grant_os) +apply (drule proc_fd_in_procs, (simp add:os_grant.simps np_notin_curp)+) +done + +lemma proc_fd_of_file_other: "\e # valid \; + \ p f flags fd opt. e \ Open p f flags fd opt; + \ p fd. e \ CloseFd p fd; + \ p f f'. e \ Rename p f f'; + \ p p'. e \ Kill p p'; + \ p. e \ Exit p; + \ p p'. e \ Clone p p'\ \ proc_fd_of_file (e # \) f = proc_fd_of_file \ f" +apply (case_tac e, auto simp:proc_fd_of_file_def file_of_proc_fd.simps) +done + +lemmas proc_fd_of_file_simps = proc_fd_of_file_open proc_fd_of_file_closefd proc_fd_of_file_rename proc_fd_of_file_kill proc_fd_of_file_exit proc_fd_of_file_clone proc_fd_of_file_other +*) +end + + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 ROOT --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ROOT Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,22 @@ +session "dynamic" = "HOL" + + options [document = false] + theories + List_Prefix + My_list_prefix + File_renaming + Flask_type + OS_type_def + Flask + + +session "alive" = "dynamic" + + options [document = false] + theories + Valid_prop + Init_prop + Current_files_prop + Current_sockets_prop + Alive_prop + Proc_fd_of_file_prop + Finite_current + New_obj_prop \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Sectxt_prop.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Sectxt_prop.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,627 @@ +theory Sectxt_prop +imports Main Flask Flask_type Current_files_prop Current_sockets_prop Alive_prop +begin + +context flask begin + +lemma alive_obj_has_type: + assumes alive: "alive s obj" and vs: "valid s" + shows "\ t. type_of_obj s obj = Some t" +using alive vs +proof (induct s arbitrary:obj) + case Nil + thus ?case + by (simp add:init_alive_prop[THEN sym] init_obj_has_type) +next + case (Cons e s) + hence a2: "alive (e # s) obj" and a3: "valid (e # s)" by auto + hence a4: "valid s" and a5: "os_grant s e" and a6: "grant s e" + by (auto intro:vd_cons vt_grant_os vt_grant) + hence a1: "\ obj. alive s obj \ \t. type_of_obj s obj = Some t" using Cons by auto + have a1': "\ obj. type_of_obj s obj = None \ \ alive s obj" by (rule_tac notI, auto dest:a1) + have p1: "\ p. p \ current_procs s \ \ t. type_of_obj s (O_proc p) = Some t" by (auto intro:a1) + have p2: "\ f. is_file s f \ \ t. type_of_obj s (O_file f) = Some t" + by (auto intro!:a1 simp:is_file_in_current) + have p3: "\ f. is_dir s f \ \ t. type_of_obj s (O_dir f) = Some t" + by (auto intro!:a1 simp:is_dir_in_current) + have p4: "\ sock. is_tcp_sock s sock \ \ t. type_of_obj s (O_tcp_sock sock) = Some t" + by (auto intro!:a1 simp:is_tcp_in_current) + have p5: "\ sock. is_udp_sock s sock \ \ t. type_of_obj s (O_udp_sock sock) = Some t" + by (auto intro!:a1 simp:is_udp_in_current) + have p6: "\ p fd. fd \ current_proc_fds s p \ \ t. type_of_obj s (O_fd p fd) = Some t" + by (auto intro:a1) + have p7: "\ n. n \ init_nodes \ \ t. type_of_obj s (O_node n) = Some t" by (auto intro:a1) + have p8: "\ h. h \ current_shms s \ \ t. type_of_obj s (O_shm h) = Some t" by (auto intro:a1) + have p9: "\ q. q \ current_msgqs s \ \ t. type_of_obj s (O_msgq q) = Some t" by (auto intro:a1) + have p10: "\ q m. \m \ set (msgs_of_queue s q); q \ current_msgqs s\ + \ \ t. type_of_obj s (O_msg q m) = Some t" by (auto intro:a1) + show ?case using a5 a2 a4 a3 + apply (case_tac e) + apply (auto split:option.splits t_object.splits if_splits t_socket_type.splits + dest!:a1' intro:a1 intro:p1 p2 p3 p4 p5 p6 p7 p8 p9 p10 simp:alive_simps) + done +qed + +lemma alive_obj_has_user: + assumes alive: "alive s obj" and vs: "valid s" + shows "\ t. user_of_obj s obj = Some t" +using alive vs +proof (induct s arbitrary:obj) + case Nil + thus ?case + by (simp add:init_alive_prop[THEN sym] init_obj_has_user) +next + case (Cons e s) + hence a2: "alive (e # s) obj" and a3: "valid (e # s)" by auto + hence a4: "valid s" and a5: "os_grant s e" and a6: "grant s e" + by (auto intro:vd_cons vt_grant_os vt_grant) + hence a1: "\ obj. alive s obj \ \t. user_of_obj s obj = Some t" using Cons by auto + have a1': "\ obj. user_of_obj s obj = None \ \ alive s obj" by (rule_tac notI, auto dest:a1) + have p1: "\ p. p \ current_procs s \ \ t. user_of_obj s (O_proc p) = Some t" by (auto intro:a1) + have p2: "\ f. is_file s f \ \ t. user_of_obj s (O_file f) = Some t" + by (auto intro!:a1 simp:is_file_in_current) + have p3: "\ f. is_dir s f \ \ t. user_of_obj s (O_dir f) = Some t" + by (auto intro!:a1 simp:is_dir_in_current) + have p4: "\ sock. is_tcp_sock s sock \ \ t. user_of_obj s (O_tcp_sock sock) = Some t" + by (auto intro!:a1 simp:is_tcp_in_current) + have p5: "\ sock. is_udp_sock s sock \ \ t. user_of_obj s (O_udp_sock sock) = Some t" + by (auto intro!:a1 simp:is_udp_in_current) + have p6: "\ p fd. fd \ current_proc_fds s p \ \ t. user_of_obj s (O_fd p fd) = Some t" + by (auto intro:a1) + have p7: "\ n. n \ init_nodes \ \ t. user_of_obj s (O_node n) = Some t" by (auto intro:a1) + have p8: "\ h. h \ current_shms s \ \ t. user_of_obj s (O_shm h) = Some t" by (auto intro:a1) + have p9: "\ q. q \ current_msgqs s \ \ t. user_of_obj s (O_msgq q) = Some t" by (auto intro:a1) + have p10: "\ q m. \m \ set (msgs_of_queue s q); q \ current_msgqs s\ + \ \ t. user_of_obj s (O_msg q m) = Some t" by (auto intro:a1) + show ?case using a5 a2 a4 a3 + apply (case_tac e) + apply (auto split:option.splits t_object.splits if_splits t_socket_type.splits + dest!:a1' intro:a1 intro:p1 p2 p3 p4 p5 p6 p7 p8 p9 p10 simp:alive_simps) + done +qed + +lemma alive_obj_has_type': "\type_of_obj s obj = None; valid s\ \ \ alive s obj" +by (rule_tac notI, auto dest:alive_obj_has_type) + +lemma current_proc_has_type: + "\p \ current_procs s; valid s\ \ \ t. type_of_obj s (O_proc p) = Some t" +by (auto intro:alive_obj_has_type) + +lemma current_proc_has_type': + "\type_of_obj s (O_proc p) = None; valid s\ \ p \ current_procs s" +by (rule notI, auto dest:current_proc_has_type) + +lemma is_file_has_type: "\is_file s f; valid s\ \ \ t. type_of_obj s (O_file f) = Some t" +by (auto intro:alive_obj_has_type simp:is_file_in_current) + +lemma is_file_has_type': "\type_of_obj s (O_file f) = None; valid s\ \ \ is_file s f" +by (auto intro:notI dest:is_file_has_type) + +lemma is_dir_has_type: "\is_dir s f; valid s\ \ \ t. type_of_obj s (O_dir f) = Some t" +by (auto intro:alive_obj_has_type simp:is_dir_in_current) + +lemma is_dir_has_type': "\type_of_obj s (O_dir f) = None; valid s\ \ \ is_dir s f" +by (auto intro:notI dest:is_dir_has_type) + +lemma is_tcp_has_type: "\is_tcp_sock s sock; valid s\ \ \ t. type_of_obj s (O_tcp_sock sock) = Some t" +by (auto intro:alive_obj_has_type simp:is_tcp_in_current) + +lemma is_tcp_has_type': "\type_of_obj s (O_tcp_sock sock) = None; valid s\ \ \ is_tcp_sock s sock" +by (auto intro:notI dest:is_tcp_has_type) + +lemma is_udp_has_type: "\is_udp_sock s sock; valid s\ \ \ t. type_of_obj s (O_udp_sock sock) = Some t" +by (auto intro:alive_obj_has_type simp:is_udp_in_current) + +lemma is_udp_has_type': "\type_of_obj s (O_udp_sock sock) = None; valid s\ \ \ is_udp_sock s sock" +by (auto intro:notI dest:is_udp_has_type) + +lemma current_fd_has_type: "\fd \ current_proc_fds s p; valid s\ \ \ t. type_of_obj s (O_fd p fd) = Some t" +by (auto intro:alive_obj_has_type) + +lemma current_fd_has_type': "\type_of_obj s (O_fd p fd) = None; valid s\ \ fd \ current_proc_fds s p" +by (auto intro:notI dest:current_fd_has_type) + +lemma init_node_has_type: "\n \ init_nodes; valid s\ \ \ t. type_of_obj s (O_node n) = Some t" +by (auto intro: alive_obj_has_type) + +lemma init_node_has_type': "\type_of_obj s (O_node n) = None; valid s\ \ n \ init_nodes" +by (auto intro:notI dest:init_node_has_type) + +lemma current_shm_has_type: "\h \ current_shms s; valid s\ \ \ t. type_of_obj s (O_shm h) = Some t" +by (auto intro:alive_obj_has_type) + +lemma current_shm_has_type': "\type_of_obj s (O_shm h) = None; valid s\ \ h \ current_shms s" +by (auto intro:notI dest:current_shm_has_type) + +lemma current_msgq_has_type: "\q \ current_msgqs s; valid s\ \ \ t. type_of_obj s (O_msgq q) = Some t" +by (auto intro:alive_obj_has_type) + +lemma current_msgq_has_type': "\type_of_obj s (O_msgq q) = None; valid s\ \ q \ current_msgqs s" +by (auto intro:notI dest:current_msgq_has_type) + +lemma current_msg_has_type: "\m \ set (msgs_of_queue s q); q \ current_msgqs s; valid s\ + \ \ t. type_of_obj s (O_msg q m) = Some t" +by (auto intro:alive_obj_has_type) + +lemma current_msg_has_type': "\type_of_obj s (O_msg q m) = None; valid s\ + \ m \ set (msgs_of_queue s q) \ q \ current_msgqs s" +by (auto dest!:current_msg_has_type) + +lemmas current_has_type = alive_obj_has_type current_proc_has_type is_file_has_type is_dir_has_type + is_tcp_has_type is_udp_has_type current_fd_has_type init_node_has_type current_shm_has_type + current_msgq_has_type current_msg_has_type + +lemmas current_has_type' = alive_obj_has_type' current_proc_has_type' is_file_has_type' is_dir_has_type' + is_tcp_has_type' is_udp_has_type' current_fd_has_type' init_node_has_type' current_shm_has_type' + current_msgq_has_type' current_msg_has_type' + +lemma alive_obj_has_user': "\user_of_obj s obj = None; valid s\ \ \ alive s obj" +by (rule_tac notI, auto dest:alive_obj_has_user) + +lemma current_proc_has_user: + "\p \ current_procs s; valid s\ \ \ t. user_of_obj s (O_proc p) = Some t" +by (auto intro:alive_obj_has_user) + +lemma current_proc_has_user': + "\user_of_obj s (O_proc p) = None; valid s\ \ p \ current_procs s" +by (rule notI, auto dest:current_proc_has_user) + +lemma is_file_has_user: "\is_file s f; valid s\ \ \ t. user_of_obj s (O_file f) = Some t" +by (auto intro:alive_obj_has_user simp:is_file_in_current) + +lemma is_file_has_user': "\user_of_obj s (O_file f) = None; valid s\ \ \ is_file s f" +by (auto intro:notI dest:is_file_has_user) + +lemma is_dir_has_user: "\is_dir s f; valid s\ \ \ t. user_of_obj s (O_dir f) = Some t" +by (auto intro:alive_obj_has_user simp:is_dir_in_current) + +lemma is_dir_has_user': "\user_of_obj s (O_dir f) = None; valid s\ \ \ is_dir s f" +by (auto intro:notI dest:is_dir_has_user) + +lemma is_tcp_has_user: "\is_tcp_sock s sock; valid s\ \ \ t. user_of_obj s (O_tcp_sock sock) = Some t" +by (auto intro:alive_obj_has_user simp:is_tcp_in_current) + +lemma is_tcp_has_user': "\user_of_obj s (O_tcp_sock sock) = None; valid s\ \ \ is_tcp_sock s sock" +by (auto intro:notI dest:is_tcp_has_user) + +lemma is_udp_has_user: "\is_udp_sock s sock; valid s\ \ \ t. user_of_obj s (O_udp_sock sock) = Some t" +by (auto intro:alive_obj_has_user simp:is_udp_in_current) + +lemma is_udp_has_user': "\user_of_obj s (O_udp_sock sock) = None; valid s\ \ \ is_udp_sock s sock" +by (auto intro:notI dest:is_udp_has_user) + +lemma current_fd_has_user: "\fd \ current_proc_fds s p; valid s\ \ \ t. user_of_obj s (O_fd p fd) = Some t" +by (auto intro:alive_obj_has_user) + +lemma current_fd_has_user': "\user_of_obj s (O_fd p fd) = None; valid s\ \ fd \ current_proc_fds s p" +by (auto intro:notI dest:current_fd_has_user) + +lemma init_node_has_user: "\n \ init_nodes; valid s\ \ \ t. user_of_obj s (O_node n) = Some t" +by (auto intro: alive_obj_has_user) + +lemma init_node_has_user': "\user_of_obj s (O_node n) = None; valid s\ \ n \ init_nodes" +by (auto intro:notI dest:init_node_has_user) + +lemma current_shm_has_user: "\h \ current_shms s; valid s\ \ \ t. user_of_obj s (O_shm h) = Some t" +by (auto intro:alive_obj_has_user) + +lemma current_shm_has_user': "\user_of_obj s (O_shm h) = None; valid s\ \ h \ current_shms s" +by (auto intro:notI dest:current_shm_has_user) + +lemma current_msgq_has_user: "\q \ current_msgqs s; valid s\ \ \ t. user_of_obj s (O_msgq q) = Some t" +by (auto intro:alive_obj_has_user) + +lemma current_msgq_has_user': "\user_of_obj s (O_msgq q) = None; valid s\ \ q \ current_msgqs s" +by (auto intro:notI dest:current_msgq_has_user) + +lemma current_msg_has_user: "\m \ set (msgs_of_queue s q); q \ current_msgqs s; valid s\ + \ \ t. user_of_obj s (O_msg q m) = Some t" +by (auto intro:alive_obj_has_user) + +lemma current_msg_has_user': "\user_of_obj s (O_msg q m) = None; valid s\ + \ m \ set (msgs_of_queue s q) \ q \ current_msgqs s" +by (auto dest!:current_msg_has_user) + +lemmas current_has_user = alive_obj_has_user current_proc_has_user is_file_has_user is_dir_has_user + is_tcp_has_user is_udp_has_user current_fd_has_user init_node_has_user current_shm_has_user + current_msgq_has_user current_msg_has_user + +lemmas current_has_user' = alive_obj_has_user' current_proc_has_user' is_file_has_user' is_dir_has_user' + is_tcp_has_user' is_udp_has_user' current_fd_has_user' init_node_has_user' current_shm_has_user' + current_msgq_has_user' current_msg_has_user' + +lemma current_proc_has_role: + "\p \ current_procs s; valid s\ \ \ r. role_of_proc s p = Some r" +apply (induct s arbitrary:p, simp add:init_procrole_prop2) +apply (frule vd_cons, frule vt_grant_os, frule vt_grant, case_tac a) +apply (auto split:option.splits dest:current_has_type' simp:sectxt_of_obj_def) +done + +lemma current_proc_has_role': + "\role_of_proc s p = None; valid s\ \ p \ current_procs s" +by (auto dest:current_proc_has_role) + +lemma alive_obj_has_role: + "\alive s obj; valid s\ \ \ r. role_of_obj s obj = Some r" +by (case_tac obj, auto intro:current_proc_has_role) + +lemma alive_obj_has_sec: + "\alive s obj; valid s\ \ \ sec. sectxt_of_obj s obj = Some sec" +apply (frule alive_obj_has_type, simp) +apply (frule alive_obj_has_role, simp) +apply (frule alive_obj_has_user, simp) +apply (auto split:option.splits simp:sectxt_of_obj_def) +done + +lemma alive_obj_has_sec': + "\sectxt_of_obj s obj = None; valid s\ \ \ alive s obj" +by (auto dest:alive_obj_has_sec) + +lemma alive_obj_has_sec'': + "\alive s obj; valid s\ \ \ u r t. sectxt_of_obj s obj = Some (u,r,t)" +by (auto dest:alive_obj_has_sec) + +lemma current_proc_has_sec: + "\p \ current_procs s; valid s\ \ \ sec. sectxt_of_obj s (O_proc p) = Some sec" +by (rule alive_obj_has_sec, simp+) + +lemma current_proc_has_sec': + "\sectxt_of_obj s (O_proc p) = None; valid s\ \ p \ current_procs s" +by (rule notI, auto dest:current_proc_has_sec) + +lemma current_proc_has_sec'': + "\p \ current_procs s; valid s\ \ \ u r t. sectxt_of_obj s (O_proc p) = Some (u,r,t)" +by (drule current_proc_has_sec, auto) + +lemma is_file_has_sec: "\is_file s f; valid s\ \ \ sec. sectxt_of_obj s (O_file f) = Some sec" +by (rule alive_obj_has_sec, simp_all add:is_file_in_current) + +lemma is_file_has_sec': "\sectxt_of_obj s (O_file f) = None; valid s\ \ \ is_file s f" +by (auto intro:notI dest:is_file_has_sec) + +lemma is_file_has_sec'': "\is_file s f; valid s\ \ \ u r t. sectxt_of_obj s (O_file f) = Some (u,r,t)" +by (drule is_file_has_sec, auto) + +lemma is_dir_has_sec: "\is_dir s f; valid s\ \ \ sec. sectxt_of_obj s (O_dir f) = Some sec" +by (rule alive_obj_has_sec, simp_all add:is_dir_in_current) + +lemma is_dir_has_sec': "\sectxt_of_obj s (O_dir f) = None; valid s\ \ \ is_dir s f" +by (auto intro:notI dest:is_dir_has_sec) + +lemma is_dir_has_sec'': "\is_dir s f; valid s\ \ \ u r t. sectxt_of_obj s (O_dir f) = Some (u,r,t)" +by (drule is_dir_has_sec, auto) + +lemma is_tcp_has_sec: "\is_tcp_sock s sock; valid s\ \ \ sec. sectxt_of_obj s (O_tcp_sock sock) = Some sec" +by (rule alive_obj_has_sec, simp_all add:is_tcp_in_current) + +lemma is_tcp_has_sec': "\sectxt_of_obj s (O_tcp_sock sock) = None; valid s\ \ \ is_tcp_sock s sock" +by (auto intro:notI dest:is_tcp_has_sec) + +lemma is_tcp_has_sec'': "\is_tcp_sock s sock; valid s\ + \ \ u r t. sectxt_of_obj s (O_tcp_sock sock) = Some (u,r,t)" +by (drule is_tcp_has_sec, simp_all add:is_tcp_in_current) + +lemma is_udp_has_sec: "\is_udp_sock s sock; valid s\ \ \ sec. sectxt_of_obj s (O_udp_sock sock) = Some sec" +by (rule alive_obj_has_sec, simp_all add:is_udp_in_current) + +lemma is_udp_has_sec': "\sectxt_of_obj s (O_udp_sock sock) = None; valid s\ \ \ is_udp_sock s sock" +by (auto intro:notI dest:is_udp_has_sec) + +lemma is_udp_has_sec'': "\is_udp_sock s sock; valid s\ + \ \ u r t. sectxt_of_obj s (O_udp_sock sock) = Some (u,r,t)" +by (drule is_udp_has_sec, simp_all add:is_udp_in_current) + +lemma current_fd_has_sec: "\fd \ current_proc_fds s p; valid s\ \ \ sec. sectxt_of_obj s (O_fd p fd) = Some sec" +by (rule alive_obj_has_sec, simp+) + +lemma current_fd_has_sec': "\sectxt_of_obj s (O_fd p fd) = None; valid s\ \ fd \ current_proc_fds s p" +by (auto intro:notI dest:current_fd_has_sec) + +lemma current_fd_has_sec'': "\fd \ current_proc_fds s p; valid s\ + \ \ u r t. sectxt_of_obj s (O_fd p fd) = Some (u, r, t)" +by (drule current_fd_has_sec, simp+) + +lemma init_node_has_sec: "\n \ init_nodes; valid s\ \ \ sec. sectxt_of_obj s (O_node n) = Some sec" +by (rule alive_obj_has_sec, simp+) + +lemma init_node_has_sec': "\sectxt_of_obj s (O_node n) = None; valid s\ \ n \ init_nodes" +by (auto intro:notI dest:init_node_has_sec) + +lemma init_node_has_sec'': "\n \ init_nodes; valid s\ \ \ u r t. sectxt_of_obj s (O_node n) = Some (u,r,t)" +by (drule init_node_has_sec, simp+) + +lemma current_shm_has_sec: "\h \ current_shms s; valid s\ \ \ sec. sectxt_of_obj s (O_shm h) = Some sec" +by (rule alive_obj_has_sec, simp+) + +lemma current_shm_has_sec': "\sectxt_of_obj s (O_shm h) = None; valid s\ \ h \ current_shms s" +by (auto intro:notI dest:current_shm_has_sec) + +lemma current_shm_has_sec'': "\h \ current_shms s; valid s\ \ \ u r t. sectxt_of_obj s (O_shm h) = Some (u,r,t)" +by (drule current_shm_has_sec, simp+) + +lemma current_msgq_has_sec: "\q \ current_msgqs s; valid s\ \ \ sec. sectxt_of_obj s (O_msgq q) = Some sec" +by (rule alive_obj_has_sec, simp+) + +lemma current_msgq_has_sec': "\sectxt_of_obj s (O_msgq q) = None; valid s\ \ q \ current_msgqs s" +by (auto intro:notI dest:current_msgq_has_sec) + +lemma current_msgq_has_sec'': "\q \ current_msgqs s; valid s\ + \ \ u r t. sectxt_of_obj s (O_msgq q) = Some (u,r,t)" +by (drule current_msgq_has_sec, simp+) + +lemma current_msg_has_sec: "\m \ set (msgs_of_queue s q); q \ current_msgqs s; valid s\ + \ \ sec. sectxt_of_obj s (O_msg q m) = Some sec" +by (rule alive_obj_has_sec, simp+) + +lemma current_msg_has_sec': "\sectxt_of_obj s (O_msg q m) = None; valid s\ + \ m \ set (msgs_of_queue s q) \ q \ current_msgqs s" +by (auto dest!:current_msg_has_sec) + +lemma current_msg_has_sec'': "\m \ set (msgs_of_queue s q); q \ current_msgqs s; valid s\ + \ \ u r t. sectxt_of_obj s (O_msg q m) = Some (u, r, t)" +by (drule current_msg_has_sec, simp+) + +lemmas current_has_sec = alive_obj_has_sec current_proc_has_sec is_file_has_sec is_dir_has_sec + is_tcp_has_sec is_udp_has_sec current_fd_has_sec init_node_has_sec current_shm_has_sec + current_msgq_has_sec current_msg_has_sec + +lemmas current_has_sec' = alive_obj_has_sec' current_proc_has_sec' is_file_has_sec' is_dir_has_sec' + is_tcp_has_sec' is_udp_has_sec' current_fd_has_sec' init_node_has_sec' current_shm_has_sec' + current_msgq_has_sec' current_msg_has_sec' + +lemmas current_has_sec'' = alive_obj_has_sec'' current_proc_has_sec'' is_file_has_sec'' is_dir_has_sec'' + is_tcp_has_sec'' is_udp_has_sec'' current_fd_has_sec'' init_node_has_sec'' current_shm_has_sec'' + current_msgq_has_sec'' current_msg_has_sec'' + +(*************** sectxt_of_obj simpset ****************) + +lemma sec_execve: + "valid (Execve p f fds # s) \ sectxt_of_obj (Execve p f fds # s) = + (sectxt_of_obj s) (O_proc p := ( + case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_file f)) of + (Some psec, Some fsec) \ npctxt_execve psec fsec + | _ \ None))" +apply (rule ext, frule vd_cons, frule vt_grant_os, frule vt_grant,case_tac x) +apply (auto simp:sectxt_of_obj_def split:option.splits + dest!:current_has_type' current_proc_has_role' current_has_user') +done + +lemma sec_clone: + "valid (Clone p p' fds shms # s) \ sectxt_of_obj (Clone p p' fds shms # s) = (\ obj. + case obj of + O_proc p'' \ if (p'' = p') then sectxt_of_obj s (O_proc p) + else sectxt_of_obj s obj + | O_fd p'' fd \ if (p'' = p' \ fd \ fds) then sectxt_of_obj s (O_fd p fd) + else if (p'' = p') then None + else sectxt_of_obj s obj + | O_tcp_sock (p'', fd) \ if (p'' = p' \ fd \ fds) then sectxt_of_obj s (O_tcp_sock (p, fd)) + else if (p'' = p') then None + else sectxt_of_obj s obj + | O_udp_sock (p'', fd) \ if (p'' = p' \ fd \ fds) then sectxt_of_obj s (O_udp_sock (p, fd)) + else if (p'' = p') then None + else sectxt_of_obj s obj + | _ \ sectxt_of_obj s obj )" +apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac obj) +apply (auto simp:sectxt_of_obj_def split:option.splits if_splits + dest!:current_has_type' current_proc_has_role' current_has_user') +done + +lemma sec_open: + "valid (Open p f flags fd opt # s) \ sectxt_of_obj (Open p f flags fd opt # s) = (\ obj. + case obj of + O_file f' \ if (f' = f \ opt \ None) then + (case (parent f) of + None \ None + | Some pf \ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of + (Some psec, Some pfsec) \ Some (fst psec, R_object, + type_transition ((snd o snd) psec) ((snd o snd) pfsec) C_file True) + | _ \ None)) + else sectxt_of_obj s obj + | O_fd p' fd' \ if (p' = p \ fd' = fd) then + (case (sectxt_of_obj s (O_proc p)) of + Some psec \ Some (fst psec, R_object, (snd o snd) psec) + | _ \ None) + else sectxt_of_obj s obj + | _ \ sectxt_of_obj s obj)" +apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac obj) +apply (auto simp:sectxt_of_obj_def split:option.splits if_splits + dest!:current_has_type' current_proc_has_role' current_has_user') +done + +lemma sec_mkdir: + "valid (Mkdir p f inum # s) \ sectxt_of_obj (Mkdir p f inum # s) = + (sectxt_of_obj s) (O_dir f := + (case parent f of + None \ None + | Some pf \ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of + (Some psec, Some pfsec) \ Some (fst psec, R_object, + type_transition ((snd o snd) psec) ((snd o snd) pfsec) C_dir True) + | _ \ None)))" +apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac x) +apply (auto simp:sectxt_of_obj_def split:option.splits if_splits + dest!:current_has_type' current_proc_has_role' current_has_user') +done + +lemma sec_linkhard: + "valid (LinkHard p f f' # s) \ sectxt_of_obj (LinkHard p f f' # s) = + (sectxt_of_obj s) (O_file f' := + (case parent f' of + None \ None + | Some pf \ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of + (Some psec, Some pfsec) \ Some (fst psec, R_object, + type_transition ((snd o snd) psec) ((snd o snd) pfsec) C_file True) + | _ \ None)))" +apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac x) +apply (auto simp:sectxt_of_obj_def split:option.splits if_splits + dest!:current_has_type' current_proc_has_role' current_has_user') +done + +lemma sec_createmsgq: + "valid (CreateMsgq p q # s) \ sectxt_of_obj (CreateMsgq p q # s) = (sectxt_of_obj s) (O_msgq q := + (case (sectxt_of_obj s (O_proc p)) of + Some psec \ Some (fst psec, R_object, (snd o snd) psec) + | _ \ None))" +apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac x) +apply (auto simp:sectxt_of_obj_def split:option.splits if_splits + dest!:current_has_type' current_proc_has_role' current_has_user') +done + +lemma sec_sendmsg: + "valid (SendMsg p q m # s) \ sectxt_of_obj (SendMsg p q m # s) = (sectxt_of_obj s) (O_msg q m := + (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_msgq q)) of + (Some psec, Some qsec) \ Some (fst psec, R_object, + type_transition ((snd o snd) psec) ((snd o snd) qsec) C_msg False) + | _ \ None))" +apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac x) +apply (auto simp:sectxt_of_obj_def split:option.splits if_splits + dest!:current_has_type' current_proc_has_role' current_has_user') +done + +lemma sec_createshm: + "valid (CreateShM p h # s) \ sectxt_of_obj (CreateShM p h # s) = (sectxt_of_obj s) (O_shm h := + case (sectxt_of_obj s (O_proc p)) of + Some psec \ Some (fst psec, R_object, (snd o snd) psec) + | _ \ None)" +apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac x) +apply (auto simp:sectxt_of_obj_def split:option.splits if_splits + dest!:current_has_type' current_proc_has_role' current_has_user') +done + +lemma sec_createsock: + "valid (CreateSock p af st fd inum # s) \ sectxt_of_obj (CreateSock p af st fd inum # s) = (\ obj. + case obj of + O_fd p' fd' \ if (p' = p \ fd' = fd) then + (case (sectxt_of_obj s (O_proc p)) of + Some psec \ Some (fst psec, R_object, (snd o snd) psec) + | _ \ None) + else sectxt_of_obj s obj + | O_tcp_sock (p', fd') \ if (p' = p \ fd' = fd \ st = STREAM) then + (case (sectxt_of_obj s (O_proc p)) of + Some psec \ Some (fst psec, R_object, (snd o snd) psec) + | _ \ None) + else sectxt_of_obj s obj + | O_udp_sock (p', fd') \ if (p' = p \ fd' = fd \ st = DGRAM) then + (case (sectxt_of_obj s (O_proc p)) of + Some psec \ Some (fst psec, R_object, (snd o snd) psec) + | _ \ None) + else sectxt_of_obj s obj + | _ \ sectxt_of_obj s obj )" +apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac obj) +apply (auto simp:sectxt_of_obj_def split:option.splits if_splits + dest!:current_has_type' current_proc_has_role' current_has_user') +done + +lemma sec_accept: + "valid (Accept p fd addr port fd' inum # s) \ sectxt_of_obj (Accept p fd addr port fd' inum # s) = (\ obj. + case obj of + O_fd p' fd'' \ if (p' = p \ fd'' = fd') then + (case (sectxt_of_obj s (O_proc p)) of + Some psec \ Some (fst psec, R_object, (snd o snd) psec) + | _ \ None) + else sectxt_of_obj s obj + | O_tcp_sock (p', fd'') \ if (p' = p \ fd'' = fd') then + (case (sectxt_of_obj s (O_proc p)) of + Some psec \ Some (fst psec, R_object, (snd o snd) psec) + | _ \ None) + else sectxt_of_obj s obj + | _ \ sectxt_of_obj s obj )" +apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac obj) +apply (auto simp:sectxt_of_obj_def split:option.splits if_splits + dest!:current_has_type' current_proc_has_role' current_has_user') +done + +lemma sec_others : + "\valid (e # s); + \ p p' fds shms. e \ Clone p p' fds shms; + \ p f fds. e \ Execve p f fds; + \ p f flags fd opt. e \ Open p f flags fd opt; + \ p f inum. e \ Mkdir p f inum; + \ p f f'. e \ LinkHard p f f'; + \ p q. e \ CreateMsgq p q; + \ p q m. e \ SendMsg p q m; + \ p h. e \ CreateShM p h; + \ p af st fd inum. e \ CreateSock p af st fd inum; + \ p fd addr port fd' inum. e \ Accept p fd addr port fd' inum + \ \ sectxt_of_obj (e # s) = sectxt_of_obj s" +apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac e, case_tac[!] x) +apply (auto simp:sectxt_of_obj_def split:option.splits if_splits + dest!:current_has_type' current_proc_has_role' current_has_user') +done + +lemmas sectxt_of_obj_simps = sec_execve sec_open sec_mkdir sec_linkhard sec_createmsgq sec_sendmsg + sec_createshm sec_createsock sec_accept sec_clone sec_others (* init_sectxt_prop *) + +(**************** get_parentfs_ctxts simpset **************) + +lemma parentf_is_dir_prop1: "\is_dir s (x # pf); valid s\ \ is_dir s pf" +apply (rule_tac f = "x # pf" in parentf_is_dir) +by (auto simp:is_dir_def current_files_def split:option.splits t_inode_tag.splits) + +lemma parentf_is_dir_prop2: "\is_file s (x # pf); valid s\ \ is_dir s pf" +apply (rule_tac f = "x # pf" in parentf_is_dir) +by (auto simp:is_dir_def is_file_def current_files_def split:option.splits t_inode_tag.splits) + +lemma parentf_is_dir_prop3: "\(x # pf) \ current_files s; valid s\ \ is_dir s pf" +apply (rule_tac f = "x # pf" in parentf_is_dir) +by (auto simp:is_dir_def current_files_def split:option.splits t_inode_tag.splits) + +lemma get_pfs_secs_prop: + "\get_parentfs_ctxts s f = None; valid s\ \ \ is_dir s f" +apply (induct f) +by (auto split:option.splits dest:current_has_sec' parentf_is_dir_prop1) + +lemma get_pfs_secs_open: + "valid (Open p f flags fd opt # s) \ get_parentfs_ctxts (Open p f flags fd opt # s) = get_parentfs_ctxts s" +apply (frule noroot_events, frule vd_cons, frule vt_grant_os) +apply (rule ext, induct_tac x) +by (auto split:option.splits simp:sectxt_of_obj_simps) + +lemma get_pfs_secs_other: + "\valid (e # s); \ p f inum. e \ Mkdir p f inum\ + \ get_parentfs_ctxts (e # s) = get_parentfs_ctxts s" +apply (frule vd_cons, frule vt_grant_os, rule ext, induct_tac x) +apply (case_tac [!] e) +apply (auto split:option.splits if_splits simp:sectxt_of_obj_simps) +done + +lemma get_pfs_secs_mkdir1: + assumes mkdir: "valid (Mkdir p f inum # s)" and noteq: "is_dir s f'" + shows "get_parentfs_ctxts (Mkdir p f inum # s) f' = get_parentfs_ctxts s f'" +proof- + from mkdir have vd: "valid s" and os: "os_grant s (Mkdir p f inum)" + by (frule_tac vd_cons, simp, frule_tac vt_grant_os, simp) + from mkdir have notroot: "f \ []" by (auto intro!:noroot_mkdir) + show ?thesis using noteq + proof (induct f') + case Nil + show ?case using notroot mkdir by (simp add:sectxt_of_obj_simps) + next + case (Cons a f') + hence p1: "is_dir s f'" using vd by (simp add:parentf_is_dir_prop1) + from Cons have p2: "is_dir s (a # f')" by simp + from Cons p1 have p3: "get_parentfs_ctxts (Mkdir p f inum # s) f' = get_parentfs_ctxts s f'" by simp + from p2 os have p4: "a # f' \ f" by (auto simp:is_dir_in_current) + from p1 os have p5: "f' \ f" by (auto simp:is_dir_in_current) + show ?case using mkdir vd os p4 p5 p1 + by (auto simp:sectxt_of_obj_simps is_dir_simps p3 + split:option.splits dest:current_has_sec' get_pfs_secs_prop) + qed +qed + +lemma get_pfs_secs_mkdir2: + "valid (Mkdir p f inum # s) \ get_parentfs_ctxts (Mkdir p f inum # s) f = ( + case f of + [] \ get_parentfs_ctxts s [] + | x#pf \ (case (get_parentfs_ctxts s pf, sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of + (Some pfsecs, Some psec, Some pfsec) \ Some ((fst psec, R_object, type_transition ((snd o snd) psec) ((snd o snd) pfsec) C_dir True) # pfsecs) + | _ \ None))" +apply (frule vd_cons, frule vt_grant_os) +apply (frule noroot_events, case_tac f) +by (auto split:option.splits dest:current_has_sec' get_pfs_secs_prop get_pfs_secs_mkdir1 + simp:sectxt_of_obj_simps is_dir_simps) + +lemmas get_parentfs_ctxts_simps = get_pfs_secs_other get_pfs_secs_mkdir1 get_pfs_secs_mkdir2 + +end + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Static.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Static.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,871 @@ +theory Static +imports Static_type OS_type_def Flask_type Flask +begin + +locale tainting_s = tainting + +begin + +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 init_sectxt_of_file:: "t_file \ security_context_t option" +where + "init_sectxt_of_file f \ + if (is_init_file f) then init_sectxt_of_obj (O_file f) + else if (is_init_dir f) then init_sectxt_of_obj (O_dir f) + else 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))" + +definition sroot :: "t_sfile" +where + "sroot \ (Init [], sec_of_root, None, {})" + +definition init_cf2sfile :: "t_file \ t_sfile option" +where + "init_cf2sfile f \ + case (parent f) of + None \ Some sroot + | Some pf \ if (is_init_file f) then + (case (init_sectxt_of_obj (O_file f), init_sectxt_of_obj (O_dir pf), get_parentfs_ctxts [] pf) of + (Some sec, Some psec, Some aseclist) \ Some (Init f, sec, Some psec, set aseclist) + | _ \ None) else + (case (init_sectxt_of_obj (O_dir f), init_sectxt_of_obj (O_dir pf), get_parentfs_ctxts [] pf) of + (Some sec, Some psec, Some aseclist) \ Some (Init f, sec, Some psec, set aseclist) + | _ \ None)" + +definition init_cfs2sfiles :: "t_file set \ t_sfile set" +where + "init_cfs2sfiles fs \ {sf. \f \ fs. init_cf2sfile f = Some sf}" + +definition init_cfd2sfd :: "t_process \ t_fd \ (security_context_t \ t_open_flags \ t_sfile) option" +where + "init_cfd2sfd p fd = + (case (init_file_of_proc_fd p fd, init_oflags_of_proc_fd p fd, init_sectxt_of_obj (O_fd p fd)) of + (Some f, Some flags, Some sec) \ (case (init_cf2sfile f) of + Some sf \ Some (sec, flags, sf) + | _ \ None) + | _ \ None)" + +definition init_cfds2sfds :: "t_process \ (security_context_t \ t_open_flags \ t_sfile) set" +where + "init_cfds2sfds p \ { sfd. \ fd sfd f. init_file_of_proc_fd p fd = Some f \ init_cfd2sfd p fd = Some sfd}" + +definition init_ch2sshm :: "t_shm \ t_sshm option" +where + "init_ch2sshm h \ (case (init_sectxt_of_obj (O_shm h)) of + Some sec \ Some (Init h, sec) + | _ \ None)" + +definition init_cph2spshs + :: "t_process \ (t_sshm \ t_shm_attach_flag) set" +where + "init_cph2spshs p \ {(sh, flag)| sh flag h. (p, flag) \ init_procs_of_shm h \ + init_ch2sshm h = Some sh}" + +definition init_cp2sproc :: "t_process \ t_sproc option" +where + "init_cp2sproc p \ (case (init_sectxt_of_obj (O_proc p)) of + Some sec \ Some (Init p, sec, (init_cfds2sfds p), (init_cph2spshs p)) + | None \ None)" + +(* acient files of init-file +definition init_o2s_afs :: "t_file \ security_context_t set" +where + "init_o2s_afs f \ {sec. \ f'. f' \ f \ init_sectxt_of_obj (O_dir f') = Some sec}" *) + +definition init_cm2smsg :: "t_msgq \ t_msg \ t_smsg option" +where + "init_cm2smsg q m \ (case (init_sectxt_of_obj (O_msg q m)) of + Some sec \ Some (Init m, sec, (O_msg q m) \ seeds) + | _ \ None)" + +fun init_cqm2sms :: "t_msgq \ t_msg list \ (t_smsg list) option" +where + "init_cqm2sms q [] = Some []" +| "init_cqm2sms q (m#ms) = + (case (init_cqm2sms q ms, init_cm2smsg q m) of + (Some sms, Some sm) \ Some (sm # sms) + | _ \ None)" + +definition init_cq2smsgq :: "t_msgq \ t_smsgq option" +where + "init_cq2smsgq q \ (case (init_sectxt_of_obj (O_msgq q), init_cqm2sms q (init_msgs_of_queue q)) of + (Some sec, Some sms) \ Some (Init q, sec, sms) + | _ \ None)" + +definition init_same_inode_files :: "t_file \ t_file set" +where + "init_same_inode_files f \ {f'. init_inum_of_file f = init_inum_of_file f'}" + +fun init_obj2sobj :: "t_object \ t_sobject option" +where + "init_obj2sobj (O_proc p) = + (case (init_cp2sproc p) of + Some sp \ Some (S_proc sp (O_proc p \ seeds)) + | _ \ None)" +| "init_obj2sobj (O_file f) = + (if (f \ init_files \ is_init_file f) + then Some (S_file (init_cfs2sfiles (init_same_inode_files f)) + (\ f'. f' \ (init_same_inode_files f) \ O_file f \ seeds)) + else None)" +| "init_obj2sobj (O_dir f) = + (if (is_init_dir f) then + (case (init_cf2sfile f) of + Some sf \ Some (S_dir sf) + | _ \ None) + else None)" +| "init_obj2sobj (O_msgq q) = + (case (init_cq2smsgq q) of + Some sq \ Some (S_msgq sq) + | _ \ None)" +| "init_obj2sobj (O_shm h) = + (case (init_ch2sshm h) of + Some sh \ Some (S_shm sh) + | _ \ None)" +| "init_obj2sobj (O_msg q m) = + (case (init_cq2smsgq q, init_cm2smsg q m) of + (Some sq, Some sm) \ Some (S_msg sq sm) + | _ \ None)" +| "init_obj2sobj _ = None" + +definition + "init_static_state \ {sobj. \ obj. init_alive obj \ init_obj2sobj obj = Some sobj}" + +fun is_init_sfile :: "t_sfile \ bool" +where + "is_init_sfile (Init _, sec, psec,asec) = True" +| "is_init_sfile _ = False" + +fun is_many_sfile :: "t_sfile \ bool" +where + "is_many_sfile (Created, sec, psec, asec) = True" +| "is_many_sfile _ = False" + +fun is_init_sproc :: "t_sproc \ bool" +where + "is_init_sproc (Init p, sec, fds, shms) = True" +| "is_init_sproc _ = False" + +fun is_many_sproc :: "t_sproc \ bool" +where + "is_many_sproc (Created, sec,fds,shms) = True" +| "is_many_sproc _ = False" + +fun is_many_smsg :: "t_smsg \ bool" +where + "is_many_smsg (Created,sec,tag) = True" +| "is_many_smsg _ = False" + +fun is_many_smsgq :: "t_smsgq \ bool" +where + "is_many_smsgq (Created,sec,sms) = (True \ (\ sm \ (set sms). is_many_smsg sm))" +| "is_many_smsgq _ = False" + +fun is_many_sshm :: "t_sshm \ bool" +where + "is_many_sshm (Created, sec) = True" +| "is_many_sshm _ = False" + +(* +fun is_init_sobj :: "t_sobject \ bool" +where + "is_init_sobj (S_proc (popt, sec, fds, shms) tag) = (popt \ None)" +| "is_init_sobj (S_file sf tag) = (is_init_sfile sf)" +| "is_init_sobj (S_dir sf tag) = (is_init_sfile sf)" +| "is_init_sobj (S_msg" *) + +fun is_many :: "t_sobject \ bool" +where + "is_many (S_proc (Many, sec, fds, shms) tag) = True" +| "is_many (S_file sfs tag) = (\ sf \ sfs. is_many_sfile sf)" +| "is_many (S_dir sf ) = is_many_sfile sf" +| "is_many (S_msgq sq ) = is_many_smsgq sq" +| "is_many (S_shm sh ) = is_many_sshm sh" + +(* +fun update_ss_sp:: "t_static_state \ t_sobject \ t_sobject \ t_static_state" +where + "update_ss_sp ss (S_proc sp tag) (S_proc sp' tag') = + (if (is_many_sproc sp) then ss \ {S_proc sp' tag'} + else (ss - {S_proc sp tag}) \ {S_proc sp' tag'})" + +fun update_ss_sd:: "t_static_state \ t_sobject \ t_sobject \ t_static_state" +where + "update_ss_sd ss (S_dir sf tag) (S_dir sf' tag') = + (if (is_many_sfile sf) then ss \ {S_dir sf' tag'} + else (ss - {S_dir sf tag}) \ {S_dir sf' tag'})" +*) + +definition update_ss :: "t_static_state \ t_sobject \ t_sobject \ t_static_state" +where + "update_ss ss so so' \ if (is_many so) then ss \ {so'} else (ss - {so}) \ {so'}" + +definition add_ss :: "t_static_state \ t_sobject \ t_static_state" +where + "add_ss ss so \ ss \ {so}" + +definition del_ss :: "t_static_state \ t_sobject \ t_static_state" +where + "del_ss ss so \ if (is_many so) then ss else ss - {so}" + +(* +fun sparent :: "t_sfile \ t_sfile option" +where + "sparent (Sroot si sec) = None" +| "sparent (Sfile si sec spf) = Some spf" + +inductive is_ancesf :: "t_sfile \ t_sfile \ bool" +where + "is_ancesf sf sf" +| "sparent sf = Some spf \ is_ancesf spf sf" +| "\sparent sf = Some spf; is_ancesf saf spf\ \ is_ancesf saf sf" + +definition sfile_reparent :: "t_sfile \ t_sfile \ t_sfile" +where + "sfile_reparent (Sroot)" +*) + +(* +(* sfds rename aux definitions *) +definition sfds_rename_notrelated + :: "t_sfd set \ t_sfile \ t_sfile \ t_sfd set \ bool" +where + "sfds_rename_notrelated sfds from to sfds' \ + (\ sec flag sf. (sec,flag,sf) \ sfds \ (\ from \ sf) \ (sec,flag,sf) \ sfds')" + +definition sfds_rename_renamed + :: "t_sfd set \ t_sfile \ t_sfile \ t_sfile \ t_sfd set \ bool" +where + "sfds_rename_renamed sfds sf from to sfds' \ + (\ sec flag sf'. (sec,flag,sf) \ sfds \ (sf \ sf') \ + (sec, flag, file_after_rename sf' from to) \ sfds' \ (sec,flag,sf) \ sfds')" + +definition sfds_rename_remain + :: "t_sfd set \ t_sfile \ t_sfile \ t_sfile \ t_sfd set \ bool" +where + "sfds_rename_remain sfds sf from to sfds' \ + (\ sec flag sf'. (sec,flag,sf') \ sfds \ (sf \ sf') \ + (sec, flag, sf') \ sfds' \ (sec,flag, file_after_rename sf' from to) \ sfds')" + +(* for not many, choose on renamed or not *) +definition sfds_rename_choices + :: "t_sfd set \ t_sfile \ t_sfile \ t_sfile \ t_sfd set \ bool" +where + "sfds_rename_choices sfds sf from to sfds' \ + sfds_rename_remain sfds sf from to sfds' \ sfds_rename_renamed sfds sf from to sfds'" + +(* for many, merge renamed with not renamed *) +definition sfds_rename_both + :: "t_sfd set \ t_sfile \ t_sfile \ t_sfile \ t_sfd set \ bool" +where + "sfds_rename_both sfds sf from to sfds' \ + (\ sec flag sf'. (sec,flag,sf') \ sfds \ (sf \ sf') \ + (sec, flag, sf') \ sfds' \ (sec,flag, file_after_rename sf' from to) \ sfds')" + +(* added to the new sfds, are those only under the new sfile *) +definition sfds_rename_added + :: "t_sfd set \ t_sfile \ t_sfile \ t_sfd set \ bool" +where + "sfds_rename_added sfds from to sfds' \ + (\ sec' flag' sf' sec flag. (sec',flag',sf') \ sfds' \ (sec,flag,sf') \ sfds \ + (\ sf. (sec,flag,sf) \ sfds \ sf' = file_after_rename from to sf))" + +definition sproc_sfds_renamed + :: "t_static_state \ t_sfile \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "sproc_sfds_renamed ss sf from to ss' \ + (\ pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \ ss \ + (\ sfds'. S_proc (pi,sec,sfds',sshms) tagp \ ss \ sfds_rename_renamed sfds sf from to sfds'))" + +definition sproc_sfds_remain + :: "t_static_state \ t_sfile \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "sproc_sfds_remain ss sf from to ss' \ + (\ pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \ ss \ + (\ sfds'. S_proc (pi,sec,sfds',sshms) tagp \ ss \ sfds_rename_remain sfds sf from to sfds'))" + +(* for not many, choose on renamed or not *) +definition sproc_sfds_choices + :: "t_static_state \ t_sfile \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "sproc_sfds_choices ss sf from to ss' \ + (\ pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \ ss \ + (\ sfds'. S_proc (pi,sec,sfds',sshms) tagp \ ss \ sfds_rename_choices sfds sf from to sfds'))" + +(* for many, merge renamed with not renamed *) +definition sproc_sfds_both + :: "t_static_state \ t_sfile \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "sproc_sfds_both ss sf from to ss' \ + (\ pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \ ss \ + (\ sfds'. S_proc (pi,sec,sfds',sshms) tagp \ ss \ sfds_rename_both sfds sf from to sfds'))" + +(* remove (\ sp tag. S_proc sp tag \ ss \ S_proc sp tag \ ss'), + * cause sfds contains sfs informations *) +definition ss_rename_notrelated + :: "t_static_state \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "ss_rename_notrelated ss sf sf' ss' \ + (\ sq. S_msgq sq \ ss \ S_msgq sq \ ss') \ + (\ pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \ ss \ (\ sfds'. + S_proc (pi,sec,sfds',sshms) tagp \ ss'\ sfds_rename_notrelated sfds sf sf' sfds')) \ + (\ sfs sf'' tag. S_file sfs tag \ ss \ sf'' \ sfs \ (\ sf \ sf'') \ (\ sfs'. + S_file sfs tag \ ss' \ sf'' \ sfs')) \ + (\ sf'' tag. S_dir sf'' tag \ ss \ (\ sf \ sf'') \ S_dir sf'' tag \ ss')" + +(* rename from to, from should definited renamed if not many *) +definition all_descendant_sf_renamed + :: "t_static_state \ t_sfile \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "all_descendant_sf_renamed ss sf from to ss' \ + (\ sfs sf' tagf. sf \ sf' \ S_file sfs tagf \ ss \ sf' \ sfs \ (\ sfs'. + S_file sfs' tagf \ ss' \ file_after_rename from to sf' \ sfs' \ sf' \ sfs')) \ + (\ sf' tagf. sf \ sf' \ S_dir sf' tagf \ ss \ S_dir (file_after_rename from to sf') tagf \ ss' \ S_dir sf' tagf \ ss') \ + sproc_sfds_renamed ss sf from to ss'" + +(* not renamed *) +definition all_descendant_sf_remain + :: "t_static_state \ t_sfile \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "all_descendant_sf_remain ss sf from to ss' \ + (\ sfs sf' tag'. sf \ sf' \ S_file sfs tag' \ ss \ sf' \ sfs \ (\ sfs'. + S_file sfs' tag' \ ss \ file_after_rename from to sf' \ sfs' \ sf' \ sfs')) \ + (\ sf' tag'. sf \ sf' \ S_dir sf' tag' \ ss \ S_dir (file_after_rename from to sf') tag' \ ss' \ S_dir sf' tag' \ ss') \ + sproc_sfds_remain ss sf from to ss'" + +definition all_descendant_sf_choices + :: "t_static_state \ t_sfile \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "all_descendant_sf_choices ss sf from to ss' \ + all_descendant_sf_renamed ss sf from to ss' \ all_descendant_sf_remain ss sf from to ss'" + +definition all_descendant_sf_both + :: "t_static_state \ t_sfile \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "all_descendant_sf_both ss sf from to ss' \ + (\ sfs sf' tag'. sf \ sf' \ S_file sfs tag' \ ss \ sf' \ sfs \ (\ sfs'. + S_file sfs' tag' \ ss \ file_after_rename from to sf' \ sfs' \ sf' \ sfs')) \ + (\ sf' tag'. sf \ sf' \ S_dir sf' tag' \ ss \ + S_dir (file_after_rename from to sf') tag' \ ss' \ S_dir sf' tag' \ ss') \ + sproc_sfds_both ss sf from to ss'" + +definition ss_renamed_file + :: "t_static_state \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "ss_renamed_file ss sf sf' ss' \ + (if (is_many_sfile sf) + then all_descendant_sf_choices ss sf sf sf' ss' + else all_descendant_sf_renamed ss sf sf sf' ss')" + +(* added to the new sfs, are those only under the new sfile *) +definition sfs_rename_added + :: "t_sfile set \ t_sfile \ t_sfile \ t_sfile set \ bool" +where + "sfs_rename_added sfs from to sfs' \ + (\ sf'. sf' \ sfs' \ sf' \ sfs \ (\ sf. sf \ sfs \ sf' = file_after_rename from to sf))" + +(* added to the new sfile, are those only under the new sfile *) +definition ss_rename_added + :: "t_static_state \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "ss_rename_added ss from to ss' \ + (\ pi sec fds fds' shms tagp. S_proc (pi, sec, fds,shms) tagp \ ss \ + S_proc (pi,sec,fds',shms) tagp \ ss' \ sfds_rename_added fds from to fds') \ + (\ sq. S_msgq sq \ ss' \ S_msgq sq \ ss) \ + (\ sfs sfs' tagf. S_file sfs' tagf \ ss' \ S_file sfs tagf \ ss \ + sfs_rename_added sfs from to sfs') \ + (\ sf' tagf. S_dir sf' tagf \ ss' \ S_dir sf' tagf \ ss \ + (\ sf. S_dir sf tagf \ ss \ sf' = file_after_rename from to sf))" + +definition sfile_alive :: "t_static_state \ t_sfile \ bool" +where + "sfile_alive ss sf \ (\ sfs tagf. sf \ sfs \ S_file sfs tagf \ ss)" + +definition sf_alive :: "t_static_state \ t_sfile \ bool" +where + "sf_alive ss sf \ (\ tagd. S_dir sf tagd \ ss) \ sfile_alive ss sf" + +(* constrains that the new static state should satisfy *) +definition ss_rename:: "t_static_state \ t_sfile \ t_sfile \ t_static_state \ bool" +where + "ss_rename ss sf sf' ss' \ + ss_rename_notrelated ss sf sf' ss' \ + ss_renamed_file ss sf sf' ss' \ ss_rename_added ss sf sf' ss' \ + (\ sf''. sf_alive ss sf'' \ sf \ sf'' \ + (if (is_many_sfile sf'') + then all_descendant_sf_choices ss sf'' sf sf' ss' + else all_descendant_sf_both ss sf'' sf sf' ss'))" + +(* two sfile, the last fname should not be equal *) +fun sfile_same_fname:: "t_sfile \ t_sfile \ bool" +where + "sfile_same_fname ((Init n, sec)#spf) ((Init n', sec')#spf') = (n = n')" +| "sfile_same_fname _ _ = False" + +(* no same init sfile/only sfile in the target-to spf, should be in static_state addmissble check *) +definition ss_rename_no_same_fname:: "t_static_state \ t_sfile \ t_sfile \ bool" +where + "ss_rename_no_same_fname ss from spf \ + \ (\ to. sfile_same_fname from to \ parent to = Some spf \ sf_alive ss to)" + +(* is not a function, is a relation (one 2 many) +definition update_ss_rename :: "t_static_state \ t_sfile \ t_sfile \ t_static_state" +where + "update_ss_rename ss sf sf' \ if (is_many_sfile sf) + then (ss \ {S_file (file_after_rename sf sf' sf'') tag | sf'' tag. sf \ sf'' \ S_file sf'' tag \ ss} + \ {S_dir (file_after_rename sf sf' sf'') tag | sf'' tag. sf \ sf'' \ S_dir sf'' tag \ ss}) + else (ss - {S_file sf'' tag | sf'' tag. sf \ sf'' \ S_file sf'' tag \ ss} + - {S_dir sf'' tag | sf'' tag. sf \ sf'' \ S_dir sf'' tag \ ss}) + \ {S_file (file_after_rename sf sf' sf'') tag | sf'' tag. sf \ sf'' \ S_file sf'' tag \ ss} + \ {S_dir (file_after_rename sf sf' sf'') tag | sf'' tag. sf \ sf'' \ S_dir sf'' tag \ ss}" +*) +*) + +fun sectxt_of_sproc :: "t_sproc \ security_context_t" +where + "sectxt_of_sproc (pi,sec,fds,shms) = sec" + +fun sectxt_of_sfile :: "t_sfile \ security_context_t" +where + "sectxt_of_sfile (fi,sec,psec,asecs) = sec" + +fun asecs_of_sfile :: "t_sfile \ security_context_t set" +where + "asecs_of_sfile (fi,sec,psec,asecs) = asecs" + +definition search_check_s :: "security_context_t \ t_sfile \ bool \ bool" +where + "search_check_s pctxt sf if_file = + (if if_file + then search_check_file pctxt (sectxt_of_sfile sf) \ search_check_allp pctxt (asecs_of_sfile sf) + else search_check_dir pctxt (sectxt_of_sfile sf) \ search_check_allp pctxt (asecs_of_sfile sf))" + +definition sectxts_of_sfds :: "t_sfd set \ security_context_t set" +where + "sectxts_of_sfds sfds \ {ctxt. \ flag sf. (ctxt, flag, sf) \ sfds}" + +definition inherit_fds_check_s :: "security_context_t \ t_sfd set \ bool" +where + "inherit_fds_check_s pctxt sfds \ + (\ ctxt \ sectxts_of_sfds sfds. permission_check pctxt ctxt C_fd P_inherit)" + +definition sectxts_of_sproc_sshms :: "t_sproc_sshm set \ security_context_t set" +where + "sectxts_of_sproc_sshms sshms \ {ctxt. \ hi flag. ((hi, ctxt),flag) \ sshms}" + +definition inherit_shms_check_s :: "security_context_t \ t_sproc_sshm set \ bool" +where + "inherit_shms_check_s pctxt sshms \ + (\ ctxt \ sectxts_of_sproc_sshms sshms. permission_check pctxt ctxt C_shm P_inherit)" + +(* +fun info_flow_sshm :: "t_sproc \ t_sproc \ bool" +where + "info_flow_sshm (pi,sec,fds,shms) (pi',sec',fds',shms') = + (\ sh flag'. (sh, SHM_RDWR) \ shms \ (sh, flag') \ shms')" +*) +definition info_flow_sproc_sshms :: "t_sproc_sshm set \ t_sproc_sshm set \ bool" +where + "info_flow_sproc_sshms shms shms' \ (\ sh flag'. (sh, SHM_RDWR) \ shms \ (sh, flag') \ shms')" + +fun info_flow_sshm :: "t_sproc \ t_sproc \ bool" +where + "info_flow_sshm (pi,sec,fds,shms) (pi',sec',fds',shms') = info_flow_sproc_sshms shms shms'" + +definition update_ss_shms :: "t_static_state \ t_sproc \ bool \ t_static_state" +where + "update_ss_shms ss spfrom tag \ {sobj. \ sobj' \ ss. + (case sobj' of + S_proc sp tagp \ if (info_flow_sshm spfrom sp) + then if (is_many_sproc sp) + then (sobj = S_proc sp tagp \ sobj = S_proc sp (tagp \ tag)) + else (sobj = S_proc sp (tagp \ tag)) + else (sobj = S_proc sp tag) + | _ \ sobj = sobj')}" + + + +(* all reachable static states(sobjects set) *) +inductive_set static :: "t_static_state set" +where + s_init: "init_static_state \ static" +| s_execve: "\ss \ static; S_proc (pi, pctxt, fds, shms) tagp \ ss; S_file sfs tagf \ ss; + (fi,fsec,pfsec,asecs) \ sfs; npctxt_execve pctxt fsec = Some pctxt'; + grant_execve pctxt fsec pctxt'; search_check_s pctxt (fi,fsec,pfsec,asecs) True; + inherit_fds_check_s pctxt' fds'; fds' \ fds\ + \ (update_ss ss (S_proc (pi, pctxt, fds, shms) tagp) + (S_proc (pi, pctxt', fds', {}) (tagp \ tagf))) \ static" +| s_clone: "\ss \ static; S_proc (pi, pctxt, fds, shms) tagp \ ss; + permission_check pctxt pctxt C_process P_fork; + inherit_fds_check_s pctxt fds'; fds' \ fds; + inherit_shms_check_s pctxt shms'; shms' \ shms\ + \ (add_ss ss (S_proc (Created, pctxt, fds', shms') tagp)) \ static" +| s_kill: "\ss \ static; S_proc (pi, pctxt, fds, shms) tagp \ ss; + S_proc (pi', pctxt', fds', shms') tagp' \ ss; + permission_check pctxt pctxt' C_process P_sigkill\ + \ (del_ss ss (S_proc (pi', pctxt', fds', shms') tagp')) \ static" +| s_ptrace: "\ss \ static; S_proc sp tagp \ ss; S_proc sp' tagp' \ ss; + permission_check (sextxt_of_sproc sp) (sectxt_of_sproc sp') C_process P_ptrace\ + \ (update_ss_shms (update_ss_shms ss sp tagp) sp' tagp') \ static" +| s_exit: "\ss \ static; S_proc sp tagp \ ss\ \ (del_ss ss (S_proc sp tagp)) \ static" +| s_open: "\ss \ static; S_proc (pi, pctxt, fds, shms) tagp \ ss; S_file sfs tagf \ ss; sf \ sfs; + search_check_s pctxt sf True; \ is_creat_excl_flag flags; + oflags_check flags pctxt (sectxt_of_sfile sf); permission_check pctxt pctxt C_fd P_setattr\ + \ (update_ss ss (S_proc (pi, pctxt, fds, shms) tagp) + (S_proc (pi, pctxt, fds \ {(pctxt,flags,sf)}, shms) tagp)) \ static" +| s_open': "\ss \ static; S_proc (pi, pctxt, fds, shms) tagp \ ss; is_creat_excl_flag flags; + S_dir (pfi,fsec,pfsec,asecs) \ ss; search_check_s pctxt (pfi,fsec,pfsec,asecs) False; + nfsec = nfctxt_create pctxt fsec C_file; oflags_check flags pctxt nfsec; + permission_check pctxt fsec C_dir P_add_rename; permission_check pctxt pctxt C_fd P_setattr\ + \ (update_ss (add_ss ss (S_file {(Created, nfsec, Some fsec, asecs \ {fsec})} tagp)) + (S_proc (pi, pctxt, fds, shms) tagp) + (S_proc (pi, pctxt, fds \ {(pctxt, flags, (Created, nfsec, Some fsec, asecs \ {fsec}))}, shms) tagp) + ) \ static" +| S_readf: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; (fdctxt,flags,sf) \ fds; + permission_check pctxt fdctxt C_fd P_setattr; S_file sfs tagf \ ss; sf \ sfs; + permission_check pctxt (sectxt_of_sfile sf) C_file P_read; is_read_flag flags\ + \ (update_ss_shms ss (pi, pctxt,fds,shms) (tagp \ tagf)) \ static" +| S_writef: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; (fdctxt,flags,sf) \ fds; + permission_check pctxt fdctxt C_fd P_setattr; sf \ sfs; S_file sfs tagf \ ss; + permission_check pctxt (sectxt_of_sfile sf) C_file P_write; is_write_flag flags\ + \ (update_ss ss (S_file sfs tagf) (S_file sfs (tagp \ tagf))) \ static" +| S_unlink: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_file sfs tagf \ ss; + (Init f,fsec,Some pfsec,asecs) \ sfs; + search_check_s pctxt (Init f,fsec,Some pfsec,asecs) True; + permission_check pctxt fsec C_file P_unlink; + permission_check pctxt pfsec C_dir P_remove_name\ + \ ((ss - {S_file sfs tagf}) \ {S_file (sfs - {(Init f,fsec,Some pfsec,asecs)}) tagf}) \ static" +| S_rmdir: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; + S_dir (fi,fsec,Some pfsec,asecs) \ ss; + search_check_s pctxt (fi,fsec,Some pfsec,asecs) False; + permission_check pctxt fsec C_dir P_rmdir; + permission_check pctxt pfsec C_dir P_remove_name\ + \ (del_ss ss (S_dir (fi,fsec,Some pfsec,asecs))) \ static" +| S_mkdir: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_dir (fi,fsec,pfsec,asecs) \ ss; + search_check_s pctxt (fi,fsec,pfsec,asecs) False; + permission_check pctxt (nfctxt_create pctxt fsec C_dir) C_dir P_create; + permission_check pctxt fsec C_dir P_add_name\ + \ (add_ss ss (S_dir (Created,nfctxt_create pctxt fsec C_dir,Some fsec,asecs \ {fsec}))) \ static" +| s_link: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_dir (pfi,pfsec,ppfsec,asecs) \ ss; + S_file sfs tagf \ ss; sf \ sfs; nfsec = nfctxt_create pctxt pfsec C_file; + search_check_s pctxt (pfi,pfsec,ppfsec,asecs) False; search_check_s pctxt sf True; + permission_check pctxt (sectxt_of_sfile sf) C_file P_link; + permission_check pctxt pfsec C_dir P_add_name\ + \ (update_ss ss (S_file sfs tagf) + (S_file (sfs \ {(Created,nfsec,Some pfsec, asecs \ {pfsec})}) tagf)) \ static" +| s_trunc: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_file sfs tagf \ ss; sf \ sfs; + search_check_s pctxt sf True; permission_check pctxt (sectxt_of_sfile sf) C_file P_setattr\ + \ (update_ss ss (S_file sfs tagf) (S_file sfs (tagf \ tagp))) \ static" +(* +| s_rename: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_file sfs tagf \ ss; + (sf#spf') \ sfs; S_dir spf tagpf \ ss; \((sf#spf') \ (sf#spf)); + search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True; + sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt; + permission_check pctxt fctxt C_file P_rename; + permission_check pctxt pfctxt C_dir P_add_name; + ss_rename ss (sf#spf') (sf#spf) ss'; + ss_rename_no_same_fname ss (sf#spf') (sf#spf)\ + \ ss' \ static" +| s_rename': "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_dir (sf#spf') tagf \ ss; + S_dir spf tagpf \ ss; \((sf#spf') \ (sf#spf)); + search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True; + sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt; + permission_check pctxt fctxt C_dir P_reparent; + permission_check pctxt pfctxt C_dir P_add_name; + ss_rename ss (sf#spf') (sf#spf) ss'; + ss_rename_no_same_fname ss (sf#spf') (sf#spf)\ + \ ss' \ static" +*) +| s_createq: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; + permission_check pctxt pctxt C_msgq P_associate; + permission_check pctxt pctxt C_msgq P_create\ + \ (add_ss ss (S_msgq (Created,pctxt,[]))) \ static" +| s_sendmsg: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_msgq (qi,qctxt,sms) \ ss; + permission_check pctxt qctxt C_msgq P_enqueue; + permission_check pctxt qctxt C_msgq P_write; + permission_check pctxt pctxt C_msg P_create\ + \ (update_ss ss (S_msgq (qi,qctxt,sms)) + (S_msgq (qi,qctxt,sms @ [(Created, pctxt, tagp)]))) \ static" +| s_recvmsg: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; + S_msgq (qi,qctxt,(mi,mctxt,tagm)#sms) \ ss; + permission_check pctxt qctxt C_msgq P_read; + permission_check pctxt mctxt C_msg P_receive\ + \ (update_ss_shms ss (pi,pctxt,fds,shms) (tagp \ tagm)) \ static" +| s_removeq: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_msgq (qi,qctxt,sms) \ ss; + permission_check pctxt qctxt C_msgq P_destroy\ + \ (del_ss ss (S_msgq (qi,qctxt,sms))) \ static" +| s_createh: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; + permission_check pctxt pctxt C_shm P_associate; + permission_check pctxt pctxt C_shm P_create\ + \ (add_ss ss (S_shm (Created, pctxt))) \ static" +| s_attach: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_shm (hi,hctxt) \ ss; + if flag = SHM_RDONLY then permission_check pctxt hctxt C_shm P_read + else (permission_check pctxt hctxt C_shm P_read \ + permission_check pctxt hctxt C_shm P_write)\ + \ (update_ss ss (S_proc (pi,pctxt,fds,shms) tagp) + (S_proc (pi,pctxt,fds,shms \ {((hi,hctxt),flag)}) tagp)) \ static" +| s_detach: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_shm sh \ ss; + (sh,flag) \ shms; \ is_many_sshm sh\ + \ (update_ss ss (S_proc (pi,pctxt,fds,shms) tagp) + (S_proc (pi,pctxt,fds,shms - {(sh,flag)}) tagp)) \ static" +| s_deleteh: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_shm (hi,hctxt) \ ss; + permission_check pctxt hctxt C_shm P_destroy; \ is_many_sshm sh\ + \ (remove_sproc_sshm (del_ss ss (S_shm (hi,hctxt))) (hi,hctxt)) \ static" + +(* +fun smsg_related :: "t_msg \ t_smsg list \ bool" +where + "smsg_related m [] = False" +| "smsg_related m ((mi, sec, tag)#sms) = + (if (mi = Init m) then True else smsg_related m sms)" + +fun smsgq_smsg_related :: "t_msgq \ t_msg \ t_smsgq \ bool" +where + "smsgq_smsg_related q m (qi, sec, smsgslist) = ((qi = Init q) \ (smsg_related m smsgslist))" + +fun smsg_relatainted :: "t_msg \ t_smsg list \ bool" +where + "smsg_relatainted m [] = False" +| "smsg_relatainted m ((mi, sec, tag)#sms) = + (if (mi = Init m \ tag = True) then True else smsg_relatainted m sms)" + +fun smsgq_smsg_relatainted :: "t_msgq \ t_msg \ t_smsgq \ bool" +where + "smsgq_smsg_relatainted q m (qi, sec, smsgslist) = + ((qi = Init q) \ (smsg_relatainted m smsgslist))" +*) + +fun sfile_related :: "t_sfile \ t_file \ bool" +where + "sfile_related (fi,sec,psec,asecs) f = (fi = Init f)" +(* +fun sproc_related :: "t_process \ t_sproc \ bool" +where + "sproc_related p (pi, sec, fds, shms) = (pi = Init p)" +*) +fun init_obj_related :: "t_sobject \ t_object \ bool" +where + "init_obj_related (S_proc (Init p, sec, fds, shms) tag) (O_proc p') = (p = p')" +| "init_obj_related (S_file sfs tag) (O_file f) = (\ sf \ sfs. sfile_related sf f)" +| "init_obj_related (S_dir sf) (O_dir f) = (sfile_related sf f)" +| "init_obj_related (S_msgq (Init q, sec, sms)) (O_msgq q') = (q = q')" +| "init_obj_related (S_shm (Init h, sec)) (O_shm h') = (h = h')" +| "init_obj_related (S_msg (Init q, sec, sms) (Init m, secm, tagm)) (O_msg q' m') = (q = q' \ m = m')" +| "init_obj_related _ _ = False" + +fun tainted_s :: "t_static_state \ t_sobject \ bool" +where + "tainted_s ss (S_proc sp tag) = (S_proc sp tag \ ss \ tag = True)" +| "tainted_s ss (S_file sfs tag) = (S_file sfs tag \ ss \ tag = True)" +| "tainted_s ss (S_msg (qi, sec, sms) (mi, secm, tag)) = + (S_msgq (qi, sec, sms) \ ss \ (mi,sec,tag) \ set sms \ tag = True)" +| "tainted_s ss _ = False" + +(* +fun tainted_s :: "t_object \ t_static_state \ bool" +where + "tainted_s (O_proc p) ss = (\ sp. S_proc sp True \ ss \ sproc_related p sp)" +| "tainted_s (O_file f) ss = (\ sfs sf. S_file sfs True \ ss \ sf \ sfs \ sfile_related f sf)" +| "tainted_s (O_msg q m) ss = (\ sq. S_msgq sq \ ss \ smsgq_smsg_relatainted q m sq)" +| "tainted_s _ ss = False" +*) + +definition taintable_s :: "t_object \ bool" +where + "taintable_s obj \ \ ss \ static. \ sobj. tainted_s ss sobj \ init_obj_related sobj obj \ init_alive obj" + +definition deletable_s :: "t_object \ bool" +where + "deletable_s obj \ init_alive obj \ (\ ss \ static. \ sobj \ ss. \ init_obj_related sobj obj)" + +definition undeletable_s :: "t_object \ bool" +where + "undeletable_s obj \ init_alive obj \ (\ ss \ static. \ sobj \ ss. init_obj_related sobj obj)" + + +(**************** translation from dynamic to static *******************) + +definition cf2sfile :: "t_state \ t_file \ bool \ t_sfile option" +where + "cf2sfile s f Is_file \ + case (parent f) of + None \ if Is_file then None else Some sroot + | Some pf \ if Is_file + then (case (sectxt_of_obj s (O_file f), sectxt_of_obj s (O_dir pf), get_parentfs_ctxts s pf) of + (Some sec, Some psec, Some asecs) \ + Some (if (\ deleted (O_file f) s \ is_init_file f) then Init f else Created, sec, Some psec, set asecs) + | _ \ None) + else (case (sectxt_of_obj s (O_dir f), sectxt_of_obj s (O_dir pf), get_parentfs_ctxts s pf) of + (Some sec, Some psec, Some asecs) \ + Some (if (\ deleted (O_dir f) s \ is_init_dir f) then Init f else Created, sec, Some psec, set asecs) + | _ \ None)" + +definition cfs2sfiles :: "t_state \ t_file set \ t_sfile set" +where + "cfs2sfiles s fs \ {sf. \ f \ fs. cf2sfile s f True = Some sf}" + +definition same_inode_files :: "t_state \ t_file \ t_file set" +where + "same_inode_files s f \ {f'. inum_of_file s f = inum_of_file s f'}" + +(* here cf2sfile is passed with True, because, process' fds are only for files not dirs *) +definition cfd2sfd :: "t_state \ t_process \ t_fd \ t_sfd option" +where + "cfd2sfd s p fd \ + (case (file_of_proc_fd s p fd, flags_of_proc_fd s p fd, sectxt_of_obj s (O_fd p fd)) of + (Some f, Some flags, Some sec) \ (case (cf2sfile s f True) of + Some sf \ Some (sec, flags, sf) + | _ \ None) + | _ \ None)" + + +definition cpfd2sfds :: "t_state \ t_process \ t_sfd set" +where + "cpfd2sfds s p \ {sfd. \ fd sfd f. file_of_proc_fd s p fd = Some f \ cfd2sfd s p fd = Some sfd}" + +definition ch2sshm :: "t_state \ t_shm \ t_sshm option" +where + "ch2sshm s h \ (case (sectxt_of_obj s (O_shm h)) of + Some sec \ + Some (if (\ deleted (O_shm h) s \ h \ init_shms) then Init h else Created, sec) + | _ \ None)" + +definition cph2spshs :: "t_state \ t_process \ t_sproc_sshm set" +where + "cph2spshs s p \ {(sh, flag) | sh flag h. (p, flag) \ procs_of_shm s h \ ch2sshm s h = Some sh}" + +definition cp2sproc :: "t_state \ t_process \ t_sproc option" +where + "cp2sproc s p \ (case (sectxt_of_obj s (O_proc p)) of + Some sec \ + Some (if (\ deleted (O_proc p) s \ p \ init_procs) then Init p else Created, sec, + cpfd2sfds s p, cph2spshs s p) + | _ \ None)" + +definition cm2smsg :: "t_state \ t_msgq \ t_msg \ t_smsg option" +where + "cm2smsg s q m \ (case (sectxt_of_obj s (O_msg q m)) of + Some sec \ + Some (if (\ deleted (O_msg q m) s \ m \ set (init_msgs_of_queue q)) then Init m else Created, + sec, O_msg q m \ tainted s) + | _ \ None)" + +fun cqm2sms:: "t_state \ t_msgq \ t_msg list \ (t_smsg list) option" +where + "cqm2sms s q [] = Some []" +| "cqm2sms s q (m#ms) = + (case (cqm2sms s q ms, cm2smsg s q m) of + (Some sms, Some sm) \ Some (sm#sms) + | _ \ None)" + +definition cq2smsgq :: "t_state \ t_msgq \ t_smsgq option" +where + "cq2smsgq s q \ (case (sectxt_of_obj s (O_msgq q), cqm2sms s q (msgs_of_queue s q)) of + (Some sec, Some sms) \ + Some (if (\ deleted (O_msgq q) s \ q \ init_msgqs) then Init q else Created, + sec, sms) + | _ \ None)" + +fun co2sobj :: "t_state \ t_object \ t_sobject option" +where + "co2sobj s (O_proc p) = + (case (cp2sproc s p) of + Some sp \ Some (S_proc sp (O_proc p \ tainted s)) + | _ \ None)" +| "co2sobj s (O_file f) = + (if (f \ current_files s) then + Some (S_file (cfs2sfiles s (same_inode_files s f)) (O_file f \ tainted s)) + else None)" +| "co2sobj s (O_dir f) = + (case (cf2sfile s f False) of + Some sf \ Some (S_dir sf) + | _ \ None)" +| "co2sobj s (O_msgq q) = + (case (cq2smsgq s q) of + Some sq \ Some (S_msgq sq) + | _ \ None)" +| "co2sobj s (O_shm h) = + (case (ch2sshm s h) of + Some sh \ Some (S_shm sh) + | _ \ None)" +| "co2sobj s (O_msg q m) = + (case (cq2smsgq s q, cm2smsg s q m) of + (Some sq, Some sm) \ Some (S_msg sq sm) + | _ \ None)" +| "co2sobj s _ = None" + + +(***************** for backward proof when Instancing static objects ******************) + +definition next_nat :: "nat set \ nat" +where + "next_nat nset = (Max nset) + 1" + +definition new_proc :: "t_state \ t_process" +where + "new_proc \ = next_nat (current_procs \)" + +definition new_inode_num :: "t_state \ nat" +where + "new_inode_num \ = next_nat (current_inode_nums \)" + +definition new_msgq :: "t_state \ t_msgq" +where + "new_msgq s = next_nat (current_msgqs s)" + +definition new_msg :: "t_state \ t_msgq \ nat" +where + "new_msg s q = next_nat (set (msgs_of_queue s q))" + +definition new_shm :: "t_state \ nat" +where + "new_shm \ = next_nat (current_shms \)" + +definition new_proc_fd :: "t_state \ t_process \ t_fd" +where + "new_proc_fd \ p = next_nat (current_proc_fds \ p)" + +definition all_fname_under_dir:: "t_file \ t_state \ t_fname set" +where + "all_fname_under_dir d \ = {fn. \ f. fn # d = f \ f \ current_files \}" + +fun fname_all_a:: "nat \ t_fname" +where + "fname_all_a 0 = []" | + "fname_all_a (Suc n) = ''a''@(fname_all_a n)" + +definition fname_length_set :: "t_fname set \ nat set" +where + "fname_length_set fns = length`fns" + +definition next_fname:: "t_file \ t_state \ t_fname" +where + "next_fname pf \ = fname_all_a ((Max (fname_length_set (all_fname_under_dir pf \))) + 1)" + +definition new_childf:: "t_file \ t_state \ t_file" +where + "new_childf pf \ = next_fname pf \ # pf" + +definition s2ss :: "t_state \ t_static_state" +where + "s2ss s \ {sobj. \ obj. alive s obj \ co2sobj s obj = Some sobj}" + +end + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Static_type.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Static_type.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,95 @@ +theory Static_type +imports Main OS_type_def Flask_type OS_type_def +begin + +(* option: if some \ initial else none \ new created + * the constructor each is one-to-one according to security_class_t + * but dynamic-special objects are not included, such as: + * 1. file-descriptor: C_fd + * 2. network node : C_node + * 3. + *) +(* Init \ exists in the initial state; Only: the only static object(cannot be cloned, e.g.) + * Many \ Many instantiation of this static object in the dynamic world + * tainted under: + * Init \ this init obj is tainted + * Only \ this new created is tainted + * Many \ at least one of them is tainted or all of them can be tainted + *) +datatype 'a Instance = + Init 'a +| Created (* new created *) +(* +| Only (* for process, cannot clone(that means no childprocesses in ss!); + for file, renamed init-file(still keeps Init fname!)) *) +| Many +*) + +(* (fi, file sectxt, parent file's sectxt, ancestral files' sectxts) *) +type_synonym t_sfile = "(t_file Instance) \ security_context_t \ (security_context_t option) \ (security_context_t set)" + +(* +type_synonym t_sfn = "(t_fname Instance) \ security_context_t" +type_synonym t_sfile = "t_sfn list" (* apparently, "sfile set" is not a finite set! *) +*) +(* +(* Sroot si sectxt : root sfile + * Sfile si sectxt Sdir : normal sfile *) +datatype t_sfile = + Sroot "t_file Instance" "security_context_t" +| Sfile "t_file Instance" "security_context_t" "t_sfile" +*) + + +(* This is not finite set! think it as nature numbers +lemma finite_sfile_set: "finite (UNIV::(t_sfile set))" +apply (rule finite.intros) +unfolding UNIV_def +apply (induct x rule:t_sfile.induct) +apply (simp add:UNIV_def) + +done *) + +(* the flags tells the sfile is a file or dir: NO! + * here sfile means sfd is opened to a sfile that has that "type", it is associated with all dynamic + * files that according to this sfile. When we need to give such a file instance, we can give any dynamic + * file in that state which according to this sfd's "creating static_state", so we might record ss in fds?? + * that means our static is a "(static-state list) set"?? *) +type_synonym t_sfd = "(security_context_t \ t_open_flags \ t_sfile)" + +type_synonym t_sshm = "(t_shm Instance \ security_context_t)" + +type_synonym t_sproc_sshm = "(t_sshm \ t_shm_attach_flag)" + +(* (si, sectxt, sectxts of fds, sectxt of attached shms) *) +type_synonym t_sproc = + "t_process Instance \ security_context_t \ (t_sfd set) \ (t_sproc_sshm set)" + +(* here is a exception, the tainted-flag is in this type, case, msgq will not be-tainted *) +type_synonym t_smsg = "t_msg Instance \ security_context_t \ bool" + +(* (qmi, sectxt, sectxt of queue) *) +type_synonym t_smsgq = "t_msgq Instance \ security_context_t \ (t_smsg list)" + +(* O_fd, O_*_sock, O_shm, O_msgq are all that can not be tainted *) +datatype t_sobject = + S_proc "t_sproc" "bool" +| S_file "t_sfile set" "bool" (* here sfile set is for "linking", tainted is the inode *) +| S_dir "t_sfile" +| S_msgq "t_smsgq" +| S_shm "t_sshm" +| S_msg "t_smsgq" "t_smsg" + +(* +datatype t_tainted_sobj = + TS_proc "t_sproc" "bool" +| TS_file "t_sfile set" "bool" +| TS_msg "t_smsgq" "t_smsg" +*) + +(* the static state is the current static objects + * But, how to record the static fds ??? fd is a dynamic concept !!! + *) +type_synonym t_static_state = "t_sobject set" + +end \ No newline at end of file diff -r 34d01e9a772e -r 7d9c0ed02b56 Valid_prop.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Valid_prop.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,33 @@ +theory Valid_prop +imports Main Flask Flask_type My_list_prefix +begin + +context flask begin + +lemma vd_cons'[rule_format]: "valid s' \ e # s = s' \ valid s" +by (erule valid.induct, auto) + +lemma vd_cons: "valid (e # s) \ valid s" +by (simp only:vd_cons') + +lemma vd_appd: " valid (\' @ \) \ valid \" +apply (induct \') +by (auto simp:vd_cons) + +lemma vd_preceq: "\\' \ \; valid \\ \ valid \'" +apply (erule no_juniorE) +by (simp only:vd_appd) + +lemma vd_prec: "\\' \ \; valid \\ \ valid \'" +apply (drule is_ancestor_no_junior) +by (simp only:vd_preceq) + +lemma vt_grant_os: "valid (e # \) \ os_grant \ e" +by (erule valid.cases, simp+) + +lemma vt_grant: "valid (e # \) \ grant \ e" +by (erule valid.cases, simp+) + +end + +end \ No newline at end of file