# HG changeset patch # User chunhan # Date 1389265749 -28800 # Node ID dfde07c7cd6b730782cee738432f12549dd21a96 # Parent d9dc04c3ea9066fb7155c44ef358e31d24f01b1d new_childf generalized diff -r d9dc04c3ea90 -r dfde07c7cd6b no_shm_selinux/Enrich.thy --- a/no_shm_selinux/Enrich.thy Thu Jan 09 14:39:00 2014 +0800 +++ b/no_shm_selinux/Enrich.thy Thu Jan 09 19:09:09 2014 +0800 @@ -6,16 +6,8 @@ (* enriched objects, closely related to static objects, so are only 3 kinds *) datatype t_enrich_obj = E_proc "t_process" "t_msgq" "t_msgq" -| E_file_link "t_file" | E_file "t_file" "nat" - (* -| E_fd "t_process" "t_fd" -| E_inum "nat" *) | E_msgq "t_msgq" -(* -| E_msg "t_msgq" "t_msg" -*) - (* objects that need dynamic indexing, all nature-numbers *) datatype t_index_obj = @@ -45,6 +37,7 @@ *) | "no_del_event (_ # \) = no_del_event \" +(* fun all_inums :: "t_state \ t_inode_num set" where "all_inums [] = current_inode_nums []" @@ -71,6 +64,7 @@ "all_msgqs [] = {}" | "all_msgqs (CreateMsgq p q # s) = all_msgqs s \ {q}" | "all_msgqs (e # s) = all_msgqs s" +*) fun all_msgs:: "t_state \ t_msgq \ t_msg set" where @@ -128,6 +122,7 @@ (f \ current_files s \ (\ pf. parent f = Some pf \ is_dir s pf))" | "notin_cur s (E_msgq q) = (q \ current_msgqs s)" +(* lemma not_all_procs_cons: "p \ all_procs (e # s) \ p \ all_procs s" by (case_tac e, auto) @@ -148,10 +143,11 @@ "p' \ all_procs s \ p' \ current_procs s" apply (induct s, simp) by (case_tac a, auto) - +*) +(* lemma not_all_msgqs_cons: "p \ all_msgqs (e # s) \ p \ all_msgqs s" -by (case_tac e, auto) +apply (case_tac e, auto) lemma not_all_msgqs_prop: "\p' \ all_msgqs s; p \ current_msgqs s; valid s\ \ p' \ p" @@ -164,7 +160,8 @@ "p' \ all_msgqs s \ p' \ current_msgqs s" apply (induct s, simp) by (case_tac a, auto) - + *) + fun enrich_not_alive :: "t_state \ t_enrich_obj \ t_index_obj \ bool" where "enrich_not_alive s obj (I_file f) = @@ -427,7 +424,7 @@ by (auto simp:Clone) have p'_not_in: "p' \ current_procs s'" using alive' notin_cur os Clone apply (erule_tac x = "I_proc p'" in allE) - apply (auto dest:not_all_procs_prop3 simp del:nums_of_recvmsg.simps) + apply (auto simp del:nums_of_recvmsg.simps) done have fd_in: "fds \ proc_file_fds s' p" using os alive ffd_remain by (auto simp:Clone proc_file_fds_def) @@ -537,8 +534,7 @@ have fd_not_in: "fd \ current_proc_fds s' p" using os alive' p_in notin_cur Open None apply (erule_tac x = "I_fd p fd" in allE) - apply (case_tac obj') - apply (auto dest:not_all_procs_prop3) + apply (case_tac obj', auto) done have "os_grant s' e" using p_in f_in fd_not_in os by (simp add:Open None) @@ -589,7 +585,7 @@ have fd_not_in: "fd \ current_proc_fds s' p" using os alive' p_in Open Some notin_cur apply (erule_tac x = "I_fd p fd" in allE) - apply (case_tac obj', auto dest:not_all_procs_prop3) + apply (case_tac obj', auto) done have inum_not_in: "inum \ current_inode_nums s'" using os alive' Open Some notin_cur nodel @@ -1137,7 +1133,7 @@ have m_not_in: "m \ set (msgs_of_queue s' q)" using os alive' notin_cur SendMsg q_in nodel vd apply (erule_tac x = "I_msg q m" in allE) - apply (case_tac obj', auto dest:not_all_msgqs_prop3) + apply (case_tac obj', auto) apply (drule nums_of_recv_0, simp+) apply (drule new_msgq_1, simp+) done @@ -1305,6 +1301,1034 @@ apply (auto simp:oflags_check_def perms_of_flags_def perm_of_oflag_def split:bool.splits) done +fun enrich_msgq :: "t_state \ t_msgq \ t_msgq \ t_state" +where + "enrich_msgq [] tq dq = []" +| "enrich_msgq (CreateMsgq p q # s) tq dq = + (if (tq = q) + then (CreateMsgq p dq # CreateMsgq p q # s) + else CreateMsgq p q # (enrich_msgq s tq dq))" +| "enrich_msgq (SendMsg p q m # s) tq dq = + (if (tq = q) + then (SendMsg p dq m # SendMsg p q m # (enrich_msgq s tq dq)) + else SendMsg p q m # (enrich_msgq s tq dq))" +| "enrich_msgq (RecvMsg p q m # s) tq dq = + (if (tq = q) + then (RecvMsg p dq m # RecvMsg p q m # (enrich_msgq s tq dq)) + else RecvMsg p q m # (enrich_msgq s tq dq))" +| "enrich_msgq (e # s) tq dq = e # (enrich_msgq s tq dq)" + +lemma enrich_msgq_duq_in: + "\q' \ current_msgqs s; q \ current_msgqs s; no_del_event s; valid s\ + \ q' \ current_msgqs (enrich_msgq s q q')" +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons, case_tac a) +apply auto +done + +lemma enrich_msgq_duq_sms: + "\q' \ current_msgqs s; q \ current_msgqs s; no_del_event s; valid s\ + \ msgs_of_queue (enrich_msgq s q q') q' = msgs_of_queue s q" +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons, case_tac a) +apply auto +done + +lemma enrich_msgq_cur_inof: + "\no_del_event s; valid s\ + \ inum_of_file (enrich_msgq s q q') f = inum_of_file s f" +apply (induct s arbitrary:f, simp) +apply (frule vt_grant_os, frule vd_cons, case_tac a) +apply (auto split:option.splits) +done + +lemma enrich_msgq_cur_inos: + "\no_del_event s; valid s\ + \ inum_of_socket (enrich_msgq s q q') = inum_of_socket s" +apply (rule ext) +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons, case_tac a) +apply (auto split:option.splits) +done + +lemma enrich_msgq_cur_inos': + "\no_del_event s; valid s\ + \ inum_of_socket (enrich_msgq s q q') sock = inum_of_socket s sock" +apply (simp add:enrich_msgq_cur_inos) +done + +lemma enrich_msgq_cur_inums: + "\no_del_event s; valid s\ + \ current_inode_nums (enrich_msgq s q q') = current_inode_nums s" +apply (auto simp:current_inode_nums_def current_file_inums_def + current_sock_inums_def enrich_msgq_cur_inof enrich_msgq_cur_inos) +done + +lemma enrich_msgq_cur_itag: + "\no_del_event s; valid s\ + \ itag_of_inum (enrich_msgq s q q') = itag_of_inum s" +apply (rule ext) +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons, case_tac a) +apply (auto split:option.splits t_socket_type.splits) +done + +lemma enrich_msgq_cur_tcps: + "\no_del_event s; valid s\ + \ is_tcp_sock (enrich_msgq s q q') = is_tcp_sock s" +apply (rule ext) +apply (auto simp:is_tcp_sock_def enrich_msgq_cur_itag enrich_msgq_cur_inos + split:option.splits t_inode_tag.splits) +done + +lemma enrich_msgq_cur_udps: + "\no_del_event s; valid s\ + \ is_udp_sock (enrich_msgq s q q') = is_udp_sock s" +apply (rule ext) +apply (auto simp:is_udp_sock_def enrich_msgq_cur_itag enrich_msgq_cur_inos + split:option.splits t_inode_tag.splits) +done + +lemma enrich_msgq_cur_msgqs: + "\q \ current_msgqs s; no_del_event s; valid s\ + \ current_msgqs (enrich_msgq s q q') = current_msgqs s \ {q'}" +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons) +apply (case_tac a, auto) +done + +lemma enrich_msgq_cur_msgs: + "\q' \ current_msgqs s; q \ current_msgqs s; no_del_event s; valid s\ + \ msgs_of_queue (enrich_msgq s q q') = (msgs_of_queue s) (q' := msgs_of_queue s q)" +apply (rule ext, simp, rule conjI, rule impI) +apply (simp add:enrich_msgq_duq_sms) +apply (rule impI) +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons) +apply (case_tac a, auto) +done + +lemma enrich_msgq_cur_procs: + "\no_del_event s; valid s\ + \ current_procs (enrich_msgq s q q') = current_procs s" +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons) +apply (case_tac a, auto) +done + +lemma enrich_msgq_cur_files: + "\no_del_event s; valid s\ + \ current_files (enrich_msgq s q q') = current_files s" +apply (auto simp:current_files_def) +apply (simp add:enrich_msgq_cur_inof)+ +done + +lemma enrich_msgq_cur_fds: + "\no_del_event s; valid s\ + \ current_proc_fds (enrich_msgq s q q') = current_proc_fds s" +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons, case_tac a) +apply auto +done + +lemma enrich_msgq_filefd: + "\no_del_event s; valid s\ + \ file_of_proc_fd (enrich_msgq s q q') = file_of_proc_fd s" +apply (rule ext, rule ext) +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons, case_tac a) +apply auto +done + +lemma enrich_msgq_flagfd: + "\no_del_event s; valid s\ + \ flags_of_proc_fd (enrich_msgq s q q') = flags_of_proc_fd s" +apply (rule ext, rule ext) +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons, case_tac a) +apply auto +done + +lemma enrich_msgq_proc_fds: + "\no_del_event s; valid s\ + \ proc_file_fds (enrich_msgq s q q') = proc_file_fds s" +apply (auto simp:proc_file_fds_def enrich_msgq_filefd) +done + +lemma enrich_msgq_hungs: + "\no_del_event s; valid s\ + \ files_hung_by_del (enrich_msgq s q q') = files_hung_by_del s" +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons, case_tac a) +apply (auto simp:files_hung_by_del.simps) +done + +lemma enrich_msgq_is_file: + "\no_del_event s; valid s\ + \ is_file (enrich_msgq s q q') = is_file s" +apply (rule ext) +apply (auto simp add:is_file_def enrich_msgq_cur_itag enrich_msgq_cur_inof + split:option.splits t_inode_tag.splits) +done + +lemma enrich_msgq_is_dir: + "\no_del_event s; valid s\ + \ is_dir (enrich_msgq s q q') = is_dir s" +apply (rule ext) +apply (auto simp add:is_dir_def enrich_msgq_cur_itag enrich_msgq_cur_inof + split:option.splits t_inode_tag.splits) +done + +lemma enrich_msgq_sameinode: + "\no_del_event s; valid s\ + \ (f \ same_inode_files (enrich_msgq s q q') f') = (f \ same_inode_files s f')" +apply (induct s, simp) +apply (simp add:same_inode_files_def) +apply (auto simp:enrich_msgq_is_file enrich_msgq_cur_inof) +done + +lemma enrich_msgq_tainted_aux1: + "\no_del_event s; q \ current_msgqs s; q' \ current_msgqs s; valid s\ + \ (tainted s \ {O_msg q' m| m. O_msg q m \ tainted s}) \ tainted (enrich_msgq s q q')" +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons) +apply (case_tac a) +apply (auto split:option.splits if_splits dest:tainted_in_current + simp:enrich_msgq_filefd enrich_msgq_sameinode) +done + +lemma enrich_msgq_tainted_aux2: + "\no_del_event s; q \ current_msgqs s; q' \ current_msgqs s; valid s; valid (enrich_msgq s q q')\ + \ tainted (enrich_msgq s q q') \ (tainted s \ {O_msg q' m| m. O_msg q m \ tainted s})" +apply (induct s, simp) +apply (frule vt_grant_os, frule vd_cons) +apply (case_tac a) +apply (auto split:option.splits if_splits simp:enrich_msgq_filefd enrich_msgq_sameinode + dest:tainted_in_current vd_cons) +apply (drule_tac vd_cons)+ +apply (simp) +apply (drule set_mp) +apply simp +apply simp +apply (drule_tac s = s in tainted_in_current) +apply simp+ +done + +lemma enrich_msgq_alive: + "\alive s obj; valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ + \ alive (enrich_msgq s q q') obj" +apply (case_tac obj) +apply (simp_all add:enrich_msgq_is_file enrich_msgq_is_dir enrich_msgq_cur_msgqs + enrich_msgq_cur_msgs enrich_msgq_cur_procs enrich_msgq_cur_fds + enrich_msgq_cur_tcps enrich_msgq_cur_udps) +apply (rule impI, simp) +done + +lemma enrich_msgq_alive': + "\alive (enrich_msgq s q q') obj; valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ + \ alive s obj \ obj = O_msgq q' \ (\ m. obj = O_msg q' m \ alive s (O_msg q m))" +apply (case_tac obj) +apply (simp_all add:enrich_msgq_is_file enrich_msgq_is_dir enrich_msgq_cur_msgqs + enrich_msgq_cur_msgs enrich_msgq_cur_procs enrich_msgq_cur_fds + enrich_msgq_cur_tcps enrich_msgq_cur_udps) +apply (auto split:if_splits) +done + +lemma enrich_msgq_not_alive: + "\enrich_not_alive s (E_msgq q') obj; q' \ current_msgqs s; q \ current_msgqs s; no_del_event s; valid s\ + \ enrich_not_alive (enrich_msgq s q q') (E_msgq q') obj" +apply (case_tac obj) +apply (auto simp:enrich_msgq_cur_fds enrich_msgq_cur_files + enrich_msgq_cur_procs enrich_msgq_cur_inums enrich_msgq_cur_msgqs enrich_msgq_cur_msgs) +done + +lemma enrich_msgq_nodel: + "no_del_event (enrich_msgq s q q') = no_del_event s" +apply (induct s, simp) +by (case_tac a, auto) + +lemma enrich_msgq_died_proc: + "died (O_proc p) (enrich_msgq s q q') = died (O_proc p) s" +apply (induct s, simp) +by (case_tac a, auto) + +lemma cf2sfile_execve: + "\ff \ current_files (Execve p f fds # s); valid (Execve p f fds # s)\ + \ cf2sfile (Execve p f fds # s) ff= cf2sfile s ff" +by (auto dest:cf2sfile_other' simp:current_files_simps) +lemma cf2sfile_clone: + "\ff \ current_files (Clone p p' fds # s); valid (Clone p p' fds # s)\ + \ cf2sfile (Clone p p' fds # s) ff= cf2sfile s ff" +by (auto dest:cf2sfile_other' simp:current_files_simps) +lemma cf2sfile_ptrace: + "\ff \ current_files (Ptrace p p' # s); valid (Ptrace p p' # s)\ + \ cf2sfile (Ptrace p p' # s) ff= cf2sfile s ff" +by (auto dest:cf2sfile_other' simp:current_files_simps) +lemma cf2sfile_readfile: + "\ff \ current_files (ReadFile p fd # s); valid (ReadFile p fd # s)\ + \ cf2sfile (ReadFile p fd # s) ff= cf2sfile s ff" +by (auto dest:cf2sfile_other' simp:current_files_simps) +lemma cf2sfile_writefile: + "\ff \ current_files (WriteFile p fd # s); valid (WriteFile p fd # s)\ + \ cf2sfile (WriteFile p fd # s) ff= cf2sfile s ff" +by (auto dest:cf2sfile_other' simp:current_files_simps) +lemma cf2sfile_truncate: + "\ff \ current_files (Truncate p f len # s); valid (Truncate p f len # s)\ + \ cf2sfile (Truncate p f len # s) ff= cf2sfile s ff" +by (auto dest:cf2sfile_other' simp:current_files_simps) +lemma cf2sfile_createmsgq: + "\ff \ current_files (CreateMsgq p q # s); valid (CreateMsgq p q # s)\ + \ cf2sfile (CreateMsgq p q # s) ff= cf2sfile s ff" +by (auto dest:cf2sfile_other' simp:current_files_simps) +lemma cf2sfile_sendmsg: + "\ff \ current_files (SendMsg p q m # s); valid (SendMsg p q m # s)\ + \ cf2sfile (SendMsg p q m # s) ff = cf2sfile s ff" +by (auto dest:cf2sfile_other' simp:current_files_simps) +lemma cf2sfile_recvmsg: + "\ff \ current_files (RecvMsg p q m # s); valid (RecvMsg p q m # s)\ + \ cf2sfile (RecvMsg p q m # s) ff = cf2sfile s ff" +by (auto dest:cf2sfile_other' simp:current_files_simps) +lemmas cf2sfile_other'' = cf2sfile_recvmsg cf2sfile_sendmsg cf2sfile_createmsgq cf2sfile_truncate + cf2sfile_writefile cf2sfile_readfile cf2sfile_ptrace cf2sfile_clone cf2sfile_execve + +lemma enrich_msgq_prop: + "\valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ + \ valid (enrich_msgq s q q') \ + (\ tp. tp \ current_procs s \ cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp) \ + (\ f. f \ current_files s \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f) \ + (\ tq. tq \ current_msgqs s \ cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq) \ + (\ tp fd. fd \ proc_file_fds s tp \ cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd) \ + (cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q) \ + (tainted (enrich_msgq s q q') = (tainted s \ {O_msg q' m| m. O_msg q m \ tainted s}))" +proof (induct s) + case Nil + thus ?case by (auto) +next + case (Cons e s) + hence vd_cons': "valid (e # s)" and curq_cons: "q \ current_msgqs (e # s)" + and curq'_cons: "q' \ current_msgqs (e # s)" and nodel_cons: "no_del_event (e # s)" + and os: "os_grant s e" and grant: "grant s e" and vd: "valid s" + by (auto dest:vd_cons vt_grant_os vt_grant) + from curq'_cons nodel_cons have curq': "q' \ current_msgqs s" by (case_tac e, auto) + from nodel_cons have nodel: "no_del_event s" by (case_tac e, auto) + from nodel curq' vd Cons + have pre: "q \ current_msgqs s \ valid (enrich_msgq s q q') \ + (\tp. tp \ current_procs s \ cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp) \ + (\f. f \ current_files s \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f) \ + (\tq. tq \ current_msgqs s \ cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq) \ + (\tp fd. fd \ proc_file_fds s tp \ cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd) \ + (cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q) \ + (tainted (enrich_msgq s q q') = (tainted s \ {O_msg q' m| m. O_msg q m \ tainted s}))" + by auto + + from pre have pre_vd: "q \ current_msgqs s \ valid (enrich_msgq s q q')" by simp + from pre have pre_sp: "\ tp. \tp \ current_procs s; q \ current_msgqs s\ + \ cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp" by auto + from pre have pre_sf: "\ f. \f \ current_files s; q \ current_msgqs s\ + \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f" by auto + from pre have pre_sq: "\ tq. \tq \ current_msgqs s; q \ current_msgqs s\ + \ cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq" by auto + from pre have pre_sfd: "\ tp fd. \fd \ proc_file_fds s tp; q \ current_msgqs s\ + \ cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" by auto + hence pre_sfd': "\ tp fd f. \file_of_proc_fd s tp fd = Some f; q \ current_msgqs s\ + \ cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" by (auto simp:proc_file_fds_def) + from pre have pre_duq: "q \ current_msgqs s \ cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q" + by auto + have vd_enrich:"q \ current_msgqs s \ valid (e # enrich_msgq s q q')" + apply (frule pre_vd) + apply (erule_tac s = s and obj' = "E_msgq q'" in enrich_valid_intro_cons) + apply (simp_all add:pre nodel nodel_cons curq_cons vd_cons' vd enrich_msgq_hungs) + apply (clarsimp simp:nodel vd curq' enrich_msgq_alive) + apply (rule allI, rule impI, erule enrich_msgq_not_alive) + apply (simp_all add:curq' curq'_cons nodel vd enrich_msgq_cur_msgs enrich_msgq_filefd enrich_msgq_flagfd) + done + + have q_neq_q': "q' \ q" using curq'_cons curq_cons by auto + + have vd_enrich_cons: "valid (enrich_msgq (e # s) q q')" + proof- + have "\ p q''. e = CreateMsgq p q'' \ valid (enrich_msgq (e # s) q q')" + proof- + fix p q'' assume ev: "e = CreateMsgq p q''" + show "valid (enrich_msgq (e # s) q q')" + proof (cases "q'' = q") + case False with ev vd_enrich curq_cons + show ?thesis by simp + next + case True + have "os_grant (CreateMsgq p q # s) (CreateMsgq p q')" + using os ev + by (simp add:q_neq_q' curq') + moreover have "grant (CreateMsgq p q # s) (CreateMsgq p q')" + using grant ev + by (auto simp add:sectxt_of_obj_def split:option.splits) + ultimately + show ?thesis using ev vd_cons' True + by (auto intro: valid.intros(2)) + qed + qed + moreover have "\ p q'' m. \e = SendMsg p q'' m; q \ current_msgqs s\ + \ valid (enrich_msgq (e # s) q q')" + proof- + fix p q'' m assume ev: "e = SendMsg p q'' m" and q_in: "q \ current_msgqs s" + show "valid (enrich_msgq (e # s) q q')" + proof (cases "q'' = q") + case False with ev vd_enrich q_in + show ?thesis by simp + next + case True + from grant os ev True obtain psec qsec + where psec: "sectxt_of_obj s (O_proc p) = Some psec" + and qsec: "sectxt_of_obj s (O_msgq q) = Some qsec" + by (auto split:option.splits) + from psec q_in os ev + have psec':"sectxt_of_obj (enrich_msgq s q q') (O_proc p) = Some psec" + by (auto dest!:pre_sp simp:cp2sproc_def split:option.splits) + from qsec q_in pre_duq vd + have qsec': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = Some qsec" + by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms') + show ?thesis using ev True vd_cons' q_in vd_enrich nodel vd + curq' psec psec' qsec qsec' grant os q_neq_q' + apply (simp, erule_tac valid.intros(2)) + apply (auto simp:q_neq_q' enrich_msgq_cur_msgqs enrich_msgq_cur_procs + enrich_msgq_cur_msgs sectxt_of_obj_simps) + done + qed + qed + moreover have "\ p q'' m. \e = RecvMsg p q'' m; q \ current_msgqs s\ + \ valid (enrich_msgq (e # s) q q')" + proof- + fix p q'' m assume ev: "e = RecvMsg p q'' m" and q_in: "q \ current_msgqs s" + show "valid (enrich_msgq (e # s) q q')" + proof (cases "q'' = q") + case False with ev vd_enrich q_in + show ?thesis by simp + next + case True + from grant os ev True obtain psec qsec msec + where psec: "sectxt_of_obj s (O_proc p) = Some psec" + and qsec: "sectxt_of_obj s (O_msgq q) = Some qsec" + and msec: "sectxt_of_obj s (O_msg q (hd (msgs_of_queue s q))) = Some msec" + by (auto split:option.splits) + from psec q_in os ev + have psec':"sectxt_of_obj (enrich_msgq s q q') (O_proc p) = Some psec" + by (auto dest!:pre_sp simp:cp2sproc_def split:option.splits) + from qsec q_in pre_duq vd + have qsec': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = Some qsec" + by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms') + from qsec q_in vd + have qsec'': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q) = Some qsec" + apply (frule_tac pre_sq, simp_all) + by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms') + from msec q_in pre_duq vd qsec qsec' qsec'' curq' nodel + have msec': "sectxt_of_obj (enrich_msgq s q q') (O_msg q' (hd (msgs_of_queue s q))) = Some msec" + apply (auto simp:cq2smsgq_def enrich_msgq_cur_msgs + split:option.splits dest!:current_has_sms') + apply (case_tac "msgs_of_queue s q") + using os ev True apply simp + apply (simp add:cqm2sms.simps split:option.splits) + apply (auto simp:cm2smsg_def split:option.splits) + done + show ?thesis using ev True vd_cons' q_in vd_enrich nodel vd + curq' grant os q_neq_q' psec psec' msec msec' qsec qsec' + apply (simp, erule_tac valid.intros(2)) + apply (auto simp:enrich_msgq_cur_msgqs enrich_msgq_cur_procs + enrich_msgq_cur_msgs sectxt_of_obj_simps) + done + qed + qed + ultimately + show ?thesis using vd_enrich curq_cons vd_cons' + apply (case_tac e) + apply (auto simp del:enrich_msgq.simps) + apply (auto split:if_splits ) + done + qed + + have curpsec: "\ tp. \tp \ current_procs s; q \ current_msgqs s\ \ + sectxt_of_obj (enrich_msgq s q q') (O_proc tp) = sectxt_of_obj s (O_proc tp)" + using pre_vd vd + apply (frule_tac pre_sp, simp) + by (auto simp:cp2sproc_def split:option.splits if_splits dest!:current_has_sec') + have curfsec: "\ f. \is_file s f; q \ current_msgqs s\ \ + sectxt_of_obj (enrich_msgq s q q') (O_file f) = sectxt_of_obj s (O_file f)" + proof- + fix f + assume a1: "is_file s f" and a2: "q \ current_msgqs s" + from a2 pre_sf pre_vd + have pre': "\f. f \ current_files s \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f" + and vd_enrich: "valid (enrich_msgq s q q')" + by auto + hence csf: "cf2sfile (enrich_msgq s q q') f = cf2sfile s f" + using a1 by (auto simp:is_file_in_current) + from a1 obtain sf where csf_some: "cf2sfile s f = Some sf" + apply (case_tac "cf2sfile s f") + apply (drule current_file_has_sfile') + apply (simp add:vd, simp add:is_file_in_current, simp) + done + from a1 have a1': "is_file (enrich_msgq s q q') f" + using vd nodel by (simp add:enrich_msgq_is_file) + show "sectxt_of_obj (enrich_msgq s q q') (O_file f) = sectxt_of_obj s (O_file f)" + using csf csf_some vd_enrich vd a1 a1' + apply (auto simp:cf2sfile_def split:option.splits if_splits) + apply (case_tac f, simp_all) + apply (drule root_is_dir', simp+) + done + qed + have curdsec: "\ tf. \is_dir s tf; q \ current_msgqs s\ + \ sectxt_of_obj (enrich_msgq s q q') (O_dir tf) = sectxt_of_obj s (O_dir tf)" + proof- + fix tf + assume a1: "is_dir s tf" and a2: "q \ current_msgqs s" + from a2 pre_sf pre_vd + have pre': "\f. f \ current_files s \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f" + and vd_enrich: "valid (enrich_msgq s q q')" + by auto + hence csf: "cf2sfile (enrich_msgq s q q') tf = cf2sfile s tf" + using a1 by (auto simp:is_dir_in_current) + from a1 obtain sf where csf_some: "cf2sfile s tf = Some sf" + apply (case_tac "cf2sfile s tf") + apply (drule current_file_has_sfile') + apply (simp add:vd, simp add:is_dir_in_current, simp) + done + from a1 have a1': "is_dir (enrich_msgq s q q') tf" + using enrich_msgq_is_dir vd nodel by simp + from a1 have a3: "\ is_file s tf" using vd by (simp add:is_dir_not_file) + from a1' vd have a3': "\ is_file (enrich_msgq s q q') tf" by (simp add:is_dir_not_file) + show "sectxt_of_obj (enrich_msgq s q q') (O_dir tf) = sectxt_of_obj s (O_dir tf)" + using csf csf_some a3 a3' vd_enrich vd + apply (auto simp:cf2sfile_def split:option.splits) + apply (case_tac tf) + apply (simp add:root_sec_remains, simp) + done + qed + have curpsecs: "\ tf ctxts'. \is_dir s tf; q \ current_msgqs s; get_parentfs_ctxts s tf = Some ctxts'\ + \ \ ctxts. get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts \ set ctxts = set ctxts'" + proof- + fix tf ctxts' + assume a1: "is_dir s tf" and a2: "q \ current_msgqs s" + and a4: "get_parentfs_ctxts s tf = Some ctxts'" + from a2 pre + have pre': "\f. f \ current_files s \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f" + and vd_enrich': "valid (enrich_msgq s q q')" + by auto + hence csf: "cf2sfile (enrich_msgq s q q') tf = cf2sfile s tf" + using a1 by (auto simp:is_dir_in_current) + from a1 obtain sf where csf_some: "cf2sfile s tf = Some sf" + apply (case_tac "cf2sfile s tf") + apply (drule current_file_has_sfile') + apply (simp add:vd, simp add:is_dir_in_current, simp) + done + from a1 have a1': "is_dir (enrich_msgq s q q') tf" + using enrich_msgq_is_dir vd nodel by simp + from a1 have a5: "\ is_file s tf" using vd by (simp add:is_dir_not_file) + from a1' vd have a5': "\ is_file (enrich_msgq s q q') tf" by (simp add:is_dir_not_file) + + from a1' pre_vd a2 obtain ctxts + where a3: "get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts" + apply (case_tac "get_parentfs_ctxts (enrich_msgq s q q') tf") + apply (drule get_pfs_secs_prop', simp+) + done + moreover have "set ctxts = set ctxts'" + proof (cases tf) + case Nil + with a3 a4 vd vd_enrich' + show ?thesis + by (simp add:get_parentfs_ctxts.simps root_sec_remains split:option.splits) + next + case (Cons a ff) + with csf csf_some a5 a5' vd_enrich' vd a3 a4 + show ?thesis + apply (auto simp:cf2sfile_def split:option.splits if_splits) + done + qed + ultimately + show "\ ctxts. get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts \ set ctxts = set ctxts'" + by auto + qed + + have psec_cons: "\ tp. tp \ current_procs (e # s) \ + sectxt_of_obj (enrich_msgq (e # s) q q') (O_proc tp) = sectxt_of_obj (e # s) (O_proc tp)" + using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd + apply (case_tac e) + apply (auto intro:curpsec simp:sectxt_of_obj_simps) + apply (frule curpsec, simp, frule curfsec, simp) + apply (auto split:option.splits)[1] + apply (frule vd_cons) defer apply (frule vd_cons) + apply (auto intro:curpsec simp:sectxt_of_obj_simps) + done + + + have sf_cons: "\ f. f \ current_files (e # s) \ cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" + proof- + fix f + assume a1: "f \ current_files (e # s)" + hence a1': "f \ current_files (enrich_msgq (e # s) q q')" + using nodel_cons os vd vd_cons' vd_enrich_cons + apply (case_tac e) + apply (auto simp:current_files_simps enrich_msgq_cur_files dest:is_file_in_current split:option.splits) + done + have b1: "\ p f' flags fd opt. e = Open p f' flags fd opt \ + cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" + proof- + fix p f' flags fd opt + assume ev: "e = Open p f' flags fd opt" + show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" + proof (cases opt) + case None + with a1 a1' ev vd_cons' vd_enrich_cons curq_cons + show ?thesis + apply (simp add:cf2sfile_open_none) + apply (simp add:pre_sf current_files_simps) + done + next + case (Some inum) + show ?thesis + proof (cases "f = f'") + case False + with a1 a1' ev vd_cons' vd_enrich_cons curq_cons Some + show ?thesis + apply (simp add:cf2sfile_open) + apply (simp add:pre_sf current_files_simps) + done + next + case True + with vd_cons' ev os Some + obtain pf where pf: "parent f = Some pf" by auto + then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts" + using os vd ev Some True + apply (case_tac "get_parentfs_ctxts s pf") + apply (drule get_pfs_secs_prop', simp, simp) + apply auto + done + + have "sectxt_of_obj (Open p f' flags fd (Some inum) # enrich_msgq s q q') (O_file f') = + sectxt_of_obj (Open p f' flags fd (Some inum) # s) (O_file f')" + using Some vd_enrich_cons vd_cons' ev pf True os curq_cons + by (simp add:sectxt_of_obj_simps curpsec curdsec) + moreover + have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)" + using curq_cons ev pf Some True os + by (simp add:curdsec) + moreover + have "\ ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \ set ctxts' = set ctxts" + using curq_cons ev pf Some True os vd psecs + apply (case_tac "get_parentfs_ctxts s pf") + apply (drule get_pfs_secs_prop', simp+) + apply (rule curpsecs, simp+) + done + then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'" + and psecs_eq: "set ctxts' = set ctxts" by auto + ultimately show ?thesis + using a1 a1' ev vd_cons' vd_enrich_cons Some True pf psecs + by (simp add:cf2sfile_open split:option.splits) + qed + qed + qed + have b2: "\ p f' inum. e = Mkdir p f' inum \ + cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" + proof- + fix p f' inum + assume ev: "e = Mkdir p f' inum" + show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" + proof (cases "f' = f") + case False + with a1 a1' ev vd_cons' vd_enrich_cons curq_cons + show ?thesis + apply (simp add:cf2sfile_mkdir) + apply (simp add:pre_sf current_files_simps) + done + next + case True + with vd_cons' ev os + obtain pf where pf: "parent f = Some pf" by auto + then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts" + using os vd ev True + apply (case_tac "get_parentfs_ctxts s pf") + apply (drule get_pfs_secs_prop', simp, simp) + apply auto + done + + have "sectxt_of_obj (Mkdir p f' inum # enrich_msgq s q q') (O_dir f') = + sectxt_of_obj (Mkdir p f' inum # s) (O_dir f')" + using vd_enrich_cons vd_cons' ev pf True os curq_cons + by (simp add:sectxt_of_obj_simps curpsec curdsec) + moreover + have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)" + using curq_cons ev pf True os + by (simp add:curdsec) + moreover + have "\ ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \ set ctxts' = set ctxts" + using curq_cons ev pf True os vd psecs + apply (case_tac "get_parentfs_ctxts s pf") + apply (drule get_pfs_secs_prop', simp+) + apply (rule curpsecs, simp+) + done + then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'" + and psecs_eq: "set ctxts' = set ctxts" by auto + ultimately show ?thesis + using a1 a1' ev vd_cons' vd_enrich_cons True pf psecs + apply (simp add:cf2sfile_mkdir split:option.splits) + done + qed + qed + have b3: "\ p f' f''. e = LinkHard p f' f'' \ + cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" + proof- + fix p f' f'' + assume ev: "e = LinkHard p f' f''" + show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" + proof (cases "f'' = f") + case False + with a1 a1' ev vd_cons' vd_enrich_cons curq_cons + show ?thesis + apply (simp add:cf2sfile_linkhard) + apply (simp add:pre_sf current_files_simps) + done + next + case True + with vd_cons' ev os + obtain pf where pf: "parent f = Some pf" by auto + then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts" + using os vd ev True + apply (case_tac "get_parentfs_ctxts s pf") + apply (drule get_pfs_secs_prop', simp, simp) + apply auto + done + + have "sectxt_of_obj (LinkHard p f' f'' # enrich_msgq s q q') (O_file f) = + sectxt_of_obj (LinkHard p f' f'' # s) (O_file f)" + using vd_enrich_cons vd_cons' ev pf True os curq_cons + by (simp add:sectxt_of_obj_simps curpsec curdsec) + moreover + have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)" + using curq_cons ev pf True os + by (simp add:curdsec) + moreover + have "\ ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \ set ctxts' = set ctxts" + using curq_cons ev pf True os vd psecs + apply (case_tac "get_parentfs_ctxts s pf") + apply (drule get_pfs_secs_prop', simp+) + apply (rule curpsecs, simp+) + done + then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'" + and psecs_eq: "set ctxts' = set ctxts" by auto + ultimately show ?thesis + using a1 a1' ev vd_cons' vd_enrich_cons True pf psecs + apply (simp add:cf2sfile_linkhard split:option.splits) + done + qed + qed + show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" + apply (case_tac e) + prefer 6 apply (erule b1) + prefer 11 apply (erule b2) + prefer 11 apply (erule b3) + apply (simp_all only:b1 b2 b3) + using a1' a1 vd_enrich_cons vd_cons' curq_cons nodel_cons + apply (simp_all add:cf2sfile_other'' cf2sfile_simps enrich_msgq.simps no_del_event.simps split:if_splits) + apply (simp_all add:pre_sf cf2sfile_other' current_files_simps split:if_splits) + apply (drule vd_cons, simp add:cf2sfile_other', drule pre_sf, simp+)+ + done + qed + + have sfd_cons:"\ tp fd f. file_of_proc_fd (e # s) tp fd = Some f \ + cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" + proof- + fix tp fd f + assume a1: "file_of_proc_fd (e # s) tp fd = Some f" + hence a1': "file_of_proc_fd (enrich_msgq (e # s) q q') tp fd = Some f" + using nodel_cons vd_enrich os vd_cons' + apply (case_tac e, auto simp:enrich_msgq_filefd simp del:enrich_msgq.simps) + done + have b1: "\ p f' flags fd' opt. e = Open p f' flags fd' opt \ + cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" + proof- + fix p f' flags fd' opt + assume ev: "e = Open p f' flags fd' opt" + have c1': "file_of_proc_fd (Open p f' flags fd' opt # s) tp fd = Some f" + using a1' ev a1 by (simp split:if_splits) + show "cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" thm cfd2sfd_open + proof (cases "tp = p \ fd = fd'") + case False + show ?thesis using ev vd_enrich_cons vd_cons' a1' a1 False curq_cons + apply (simp add:cfd2sfd_open split:if_splits del:file_of_proc_fd.simps) + apply (rule conjI, rule impI, simp) + apply (rule conjI, rule impI, simp) + apply (auto simp: False intro!:pre_sfd' split:if_splits) + done + next + case True + have "f' \ current_files (Open p f' flags fd' opt # s)" using ev vd_cons' os + by (auto simp:current_files_simps is_file_in_current split:option.splits) + hence "cf2sfile (Open p f' flags fd' opt # enrich_msgq s q q') f' + = cf2sfile (Open p f' flags fd' opt # s) f'" + using sf_cons ev by auto + moreover have "sectxt_of_obj (enrich_msgq s q q') (O_proc p) = sectxt_of_obj s (O_proc p)" + apply (rule curpsec) + using os ev curq_cons + by (auto split:option.splits) + ultimately show ?thesis + using ev True a1 a1' vd_enrich_cons vd_cons' + apply (simp add:cfd2sfd_open sectxt_of_obj_simps del:file_of_proc_fd.simps) + done + qed + qed + show "cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" + apply (case_tac e) + prefer 6 apply (erule b1) + using a1' a1 vd_enrich_cons vd_cons' curq_cons + apply (simp_all only:cfd2sfd_simps enrich_msgq.simps) + apply (auto simp:cfd2sfd_simps pre_sfd' dest:vd_cons cfd2sfd_other split:if_splits) + done + qed + + have pfds_cons: "\ tp. tp \ current_procs (e # s) \ + cpfd2sfds (enrich_msgq (e # s) q q') tp = cpfd2sfds (e # s) tp" + apply (auto simp add:cpfd2sfds_def proc_file_fds_def) + apply (rule_tac x = fd in exI, rule conjI, rule_tac x = f in exI) + prefer 3 + apply (rule_tac x = fd in exI, rule conjI, rule_tac x = f in exI) + apply (auto simp:sfd_cons enrich_msgq_filefd nodel_cons vd_cons') + done + + have tainted_cons: "tainted (enrich_msgq (e # s) q q') = + (tainted (e # s) \ {O_msg q' m | m. O_msg q m \ tainted (e # s)})" + apply (rule equalityI) + using nodel_cons curq_cons curq'_cons vd_cons' vd_enrich_cons + apply (rule enrich_msgq_tainted_aux2) + using nodel_cons curq_cons curq'_cons vd_cons' + apply (rule enrich_msgq_tainted_aux1) + done + have pre_tainted: "q \ current_msgqs s \ tainted (enrich_msgq s q q') = + (tainted s \ {O_msg q' m| m. O_msg q m \ tainted s})" by (simp add:pre) + + have "\tp fd. fd \ proc_file_fds (e # s) tp \ cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" + by (auto simp:proc_file_fds_def elim!:sfd_cons) + moreover + have "\tp. tp \ current_procs (e # s) \ cp2sproc (enrich_msgq (e # s) q q') tp = cp2sproc (e # s) tp" + by (auto simp:cp2sproc_def pfds_cons psec_cons enrich_msgq_died_proc split:option.splits) + moreover + have "\tq. tq \ current_msgqs (e # s) \ cq2smsgq (enrich_msgq (e # s) q q') tq = cq2smsgq (e # s) tq" + proof clarify + fix tq assume a1: "tq \ current_msgqs (e # s)" + + have curqsec: "\ tq. \tq \ current_msgqs s; q \ current_msgqs s\ \ + sectxt_of_obj (enrich_msgq s q q') (O_msgq tq) = sectxt_of_obj s (O_msgq tq)" + using pre_vd vd + apply (frule_tac pre_sq, simp) + by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') + have cursms: "\ q''. \q'' \ current_msgqs s; q \ current_msgqs s\ \ + cqm2sms (enrich_msgq s q q') q'' (msgs_of_queue (enrich_msgq s q q') q'') = + cqm2sms s q'' (msgs_of_queue s q'')" + using pre_vd vd + apply (frule_tac pre_sq, simp) + by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') + have qsec_cons: "sectxt_of_obj (enrich_msgq (e # s) q q') (O_msgq tq) = sectxt_of_obj (e # s) (O_msgq tq)" + using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons curq_cons a1 + apply (case_tac e) + apply (auto intro:curqsec simp:sectxt_of_obj_simps curpsec split:option.splits dest!:current_has_sec') + apply (frule vd_cons) defer apply (frule vd_cons) + apply (auto intro:curqsec simp:sectxt_of_obj_simps) + done + have sms_cons: "cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = + cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" + proof- + have b1: "\ p q'' m. e = SendMsg p q'' m \ + cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = + cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" + apply (case_tac e) + using a1 curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons + apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms pre_tainted + split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) + apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*}) + apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms pre_tainted + split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) + done + have b2: "\ p q''. e = CreateMsgq p q'' \ + cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = + cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" + using a1 curq_cons curq'_cons vd_enrich_cons vd_cons' + apply (auto simp:cqm2sms_simps intro:cursms) + apply (auto simp:cqm2sms.simps) + done + have b3: "\ p q'' m. e = RecvMsg p q'' m \ + cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = + cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" + using a1 curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons + apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms + split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) + apply (frule vd_cons) defer apply (frule vd_cons) + apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms + split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) + done + show "cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = + cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" + apply (case_tac e) + prefer 15 apply (erule b2) + prefer 15 apply (erule b1) + prefer 15 apply (erule b3) + using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons a1 + apply (auto intro:cursms simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec + split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) + done + qed + + show "cq2smsgq (enrich_msgq (e # s) q q') tq = cq2smsgq (e # s) tq" + using a1 curq_cons + apply (simp add:cq2smsgq_def qsec_cons sms_cons) + done + qed + moreover + have "cq2smsgq (enrich_msgq (e # s) q q') q' = cq2smsgq (e # s) q" + proof- + have duqsec: "q \ current_msgqs s \ + sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = sectxt_of_obj s (O_msgq q)" + apply (frule pre_duq) using vd + by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') + have duqsms: "q \ current_msgqs s \ + cqm2sms (enrich_msgq s q q') q' (msgs_of_queue (enrich_msgq s q q') q') = + cqm2sms s q (msgs_of_queue s q)" + apply (frule pre_duq) using vd + by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') + have qsec_cons: "sectxt_of_obj (enrich_msgq (e # s) q q') (O_msgq q') = sectxt_of_obj (e # s) (O_msgq q)" + using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons curq_cons + apply (case_tac e) + apply (auto simp:duqsec sectxt_of_obj_simps curpsec split:option.splits dest!:current_has_sec') + apply (frule vd_cons) defer apply (frule vd_cons) + apply (auto intro:duqsec simp:sectxt_of_obj_simps) + done + have sms_cons: "cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = + cqm2sms (e # s) q (msgs_of_queue (e # s) q)" + proof- + have b1: "\ p q'' m. e = SendMsg p q'' m \ + cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = + cqm2sms (e # s) q (msgs_of_queue (e # s) q)" + apply (case_tac e) + using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons + apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted + enrich_msgq_cur_procs enrich_msgq_cur_msgqs + split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) + apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*}) + apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted + enrich_msgq_cur_procs enrich_msgq_cur_msgqs dest:tainted_in_current + split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) + done + have b2: "\ p q''. e = CreateMsgq p q'' \ + cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = + cqm2sms (e # s) q (msgs_of_queue (e # s) q)" + using curq_cons curq'_cons vd_enrich_cons vd_cons' + apply (auto simp:cqm2sms_simps intro:duqsms) + apply (auto simp:cqm2sms.simps) + done + have b3: "\ p q'' m. e = RecvMsg p q'' m \ + cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = + cqm2sms (e # s) q (msgs_of_queue (e # s) q)" + using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons + apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted + enrich_msgq_cur_procs enrich_msgq_cur_msgqs + split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) + apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*}) + apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted + enrich_msgq_cur_procs enrich_msgq_cur_msgqs dest:tainted_in_current + split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) + done + show ?thesis + apply (case_tac e) + prefer 15 apply (erule b2) + prefer 15 apply (erule b1) + prefer 15 apply (erule b3) + using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons + apply (auto intro:duqsms simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsms + split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) + done + qed + show ?thesis + using curq_cons + apply (simp add:cq2smsgq_def qsec_cons sms_cons) + done + qed + ultimately show ?case using vd_enrich_cons sf_cons tainted_cons + by auto +qed + +lemma enrich_msgq_vd: + "\q \ current_msgqs s; q' \ current_msgqs s; no_del_event s; valid s\ \ + valid (enrich_msgq s q q')" +by (auto dest:enrich_msgq_prop) + +lemma enrich_msgq_sp: + "\tp \ current_procs s; valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ + \ cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp" +by (auto dest:enrich_msgq_prop) + +lemma enrich_msgq_sf: + "\f \ current_files s; valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ + \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f" +by (auto dest:enrich_msgq_prop) + +lemma enrich_msgq_sfs: + "\is_file s f; valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ + \ cf2sfiles (enrich_msgq s q q') f = cf2sfiles s f" +apply (auto simp add:cf2sfiles_def) +apply (rule_tac x = f' in bexI) defer +apply (simp add:enrich_msgq_sameinode) +apply (rule_tac x = f' in bexI) defer +apply (simp add:enrich_msgq_sameinode)+ +apply (drule same_inode_files_prop11, drule_tac f = f' in is_file_in_current) +apply (simp add:enrich_msgq_sf) +apply (drule same_inode_files_prop11, drule_tac f = f' in is_file_in_current) +apply (simp add:enrich_msgq_sf) +done + +lemma enrich_msgq_sq: + "\tq \ current_msgqs s; valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ + \ cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq" +by (auto dest:enrich_msgq_prop) + +lemma enrich_msgq_sfd': + "\fd \ proc_file_fds s tp; valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ + \ cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" +by (auto dest:enrich_msgq_prop) + +lemma enrich_msgq_sfd: + "\file_of_proc_fd s tp fd = Some f; valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ + \ cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" +by (auto intro:enrich_msgq_sfd' simp:proc_file_fds_def) + +lemma enrich_msgq_duq: + "\valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ + \ cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q" +by (auto dest:enrich_msgq_prop) + +lemma enrich_msgq_tainted: + "\valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ + \ tainted (enrich_msgq s q q') = (tainted s \ {O_msg q' m| m. O_msg q m \ tainted s})" +by (auto dest:enrich_msgq_prop) + +lemma enrich_msgq_dalive: + "\q \ current_msgqs s; q' \ current_msgqs s; no_del_event s; valid s\ + \ dalive (enrich_msgq s q q') obj = (dalive s obj \ obj = D_msgq q')" +apply (case_tac obj) +apply (auto simp:enrich_msgq_is_file enrich_msgq_is_dir enrich_msgq_cur_msgqs enrich_msgq_cur_procs) +done + +lemma enrich_msgq_s2ss: + "\q \ current_msgqs s; q' \ current_msgqs s; no_del_event s; valid s\ \ + s2ss (enrich_msgq s q q') = s2ss s" +apply (auto simp add:s2ss_def) +apply (simp add:enrich_msgq_dalive) +apply (erule disjE) +apply (rule_tac x = obj in exI) defer +apply (rule_tac x = "D_msgq q" in exI) defer +apply (rule_tac x = obj in exI) +apply (case_tac[!] obj) +apply (auto simp:enrich_msgq_duq enrich_msgq_tainted enrich_msgq_sq enrich_msgq_sf + enrich_msgq_sp co2sobj.simps enrich_msgq_is_file enrich_msgq_is_dir + enrich_msgq_cur_procs enrich_msgq_cur_msgqs enrich_msgq_sfs + split:option.splits dest:is_dir_in_current) +done + end end \ No newline at end of file diff -r d9dc04c3ea90 -r dfde07c7cd6b no_shm_selinux/Enrich2.thy --- a/no_shm_selinux/Enrich2.thy Thu Jan 09 14:39:00 2014 +0800 +++ b/no_shm_selinux/Enrich2.thy Thu Jan 09 19:09:09 2014 +0800 @@ -5,965 +5,11 @@ context tainting_s begin -fun enrich_msgq :: "t_state \ t_msgq \ t_msgq \ t_state" +fun enrich_file :: "t_state \ t_file \ t_state \ (t_file set \ t_state)" where - "enrich_msgq [] tq dq = []" -| "enrich_msgq (CreateMsgq p q # s) tq dq = - (if (tq = q) - then (CreateMsgq p dq # CreateMsgq p q # s) - else CreateMsgq p q # (enrich_msgq s tq dq))" -| "enrich_msgq (SendMsg p q m # s) tq dq = - (if (tq = q) - then (SendMsg p dq m # SendMsg p q m # (enrich_msgq s tq dq)) - else SendMsg p q m # (enrich_msgq s tq dq))" -| "enrich_msgq (RecvMsg p q m # s) tq dq = - (if (tq = q) - then (RecvMsg p dq m # RecvMsg p q m # (enrich_msgq s tq dq)) - else RecvMsg p q m # (enrich_msgq s tq dq))" -| "enrich_msgq (e # s) tq dq = e # (enrich_msgq s tq dq)" - -lemma enrich_msgq_duq_in: - "\q' \ current_msgqs s; q \ current_msgqs s; no_del_event s; valid s\ - \ q' \ current_msgqs (enrich_msgq s q q')" -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons, case_tac a) -apply auto -done - -lemma enrich_msgq_duq_sms: - "\q' \ current_msgqs s; q \ current_msgqs s; no_del_event s; valid s\ - \ msgs_of_queue (enrich_msgq s q q') q' = msgs_of_queue s q" -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons, case_tac a) -apply auto -done - -lemma enrich_msgq_cur_inof: - "\no_del_event s; valid s\ - \ inum_of_file (enrich_msgq s q q') f = inum_of_file s f" -apply (induct s arbitrary:f, simp) -apply (frule vt_grant_os, frule vd_cons, case_tac a) -apply (auto split:option.splits) -done - -lemma enrich_msgq_cur_inos: - "\no_del_event s; valid s\ - \ inum_of_socket (enrich_msgq s q q') = inum_of_socket s" -apply (rule ext) -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons, case_tac a) -apply (auto split:option.splits) -done - -lemma enrich_msgq_cur_inos': - "\no_del_event s; valid s\ - \ inum_of_socket (enrich_msgq s q q') sock = inum_of_socket s sock" -apply (simp add:enrich_msgq_cur_inos) -done - -lemma enrich_msgq_cur_inums: - "\no_del_event s; valid s\ - \ current_inode_nums (enrich_msgq s q q') = current_inode_nums s" -apply (auto simp:current_inode_nums_def current_file_inums_def - current_sock_inums_def enrich_msgq_cur_inof enrich_msgq_cur_inos) -done - -lemma enrich_msgq_cur_itag: - "\no_del_event s; valid s\ - \ itag_of_inum (enrich_msgq s q q') = itag_of_inum s" -apply (rule ext) -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons, case_tac a) -apply (auto split:option.splits t_socket_type.splits) -done - -lemma enrich_msgq_cur_tcps: - "\no_del_event s; valid s\ - \ is_tcp_sock (enrich_msgq s q q') = is_tcp_sock s" -apply (rule ext) -apply (auto simp:is_tcp_sock_def enrich_msgq_cur_itag enrich_msgq_cur_inos - split:option.splits t_inode_tag.splits) -done - -lemma enrich_msgq_cur_udps: - "\no_del_event s; valid s\ - \ is_udp_sock (enrich_msgq s q q') = is_udp_sock s" -apply (rule ext) -apply (auto simp:is_udp_sock_def enrich_msgq_cur_itag enrich_msgq_cur_inos - split:option.splits t_inode_tag.splits) -done - -lemma enrich_msgq_cur_msgqs: - "\q \ current_msgqs s; no_del_event s; valid s\ - \ current_msgqs (enrich_msgq s q q') = current_msgqs s \ {q'}" -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons) -apply (case_tac a, auto) -done - -lemma enrich_msgq_cur_msgs: - "\q' \ current_msgqs s; q \ current_msgqs s; no_del_event s; valid s\ - \ msgs_of_queue (enrich_msgq s q q') = (msgs_of_queue s) (q' := msgs_of_queue s q)" -apply (rule ext, simp, rule conjI, rule impI) -apply (simp add:enrich_msgq_duq_sms) -apply (rule impI) -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons) -apply (case_tac a, auto) -done - -lemma enrich_msgq_cur_procs: - "\no_del_event s; valid s\ - \ current_procs (enrich_msgq s q q') = current_procs s" -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons) -apply (case_tac a, auto) -done - -lemma enrich_msgq_cur_files: - "\no_del_event s; valid s\ - \ current_files (enrich_msgq s q q') = current_files s" -apply (auto simp:current_files_def) -apply (simp add:enrich_msgq_cur_inof)+ -done - -lemma enrich_msgq_cur_fds: - "\no_del_event s; valid s\ - \ current_proc_fds (enrich_msgq s q q') = current_proc_fds s" -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons, case_tac a) -apply auto -done - -lemma enrich_msgq_filefd: - "\no_del_event s; valid s\ - \ file_of_proc_fd (enrich_msgq s q q') = file_of_proc_fd s" -apply (rule ext, rule ext) -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons, case_tac a) -apply auto -done - -lemma enrich_msgq_flagfd: - "\no_del_event s; valid s\ - \ flags_of_proc_fd (enrich_msgq s q q') = flags_of_proc_fd s" -apply (rule ext, rule ext) -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons, case_tac a) -apply auto -done - -lemma enrich_msgq_proc_fds: - "\no_del_event s; valid s\ - \ proc_file_fds (enrich_msgq s q q') = proc_file_fds s" -apply (auto simp:proc_file_fds_def enrich_msgq_filefd) -done - -lemma enrich_msgq_hungs: - "\no_del_event s; valid s\ - \ files_hung_by_del (enrich_msgq s q q') = files_hung_by_del s" -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons, case_tac a) -apply (auto simp:files_hung_by_del.simps) -done - -lemma enrich_msgq_is_file: - "\no_del_event s; valid s\ - \ is_file (enrich_msgq s q q') = is_file s" -apply (rule ext) -apply (auto simp add:is_file_def enrich_msgq_cur_itag enrich_msgq_cur_inof - split:option.splits t_inode_tag.splits) -done - -lemma enrich_msgq_is_dir: - "\no_del_event s; valid s\ - \ is_dir (enrich_msgq s q q') = is_dir s" -apply (rule ext) -apply (auto simp add:is_dir_def enrich_msgq_cur_itag enrich_msgq_cur_inof - split:option.splits t_inode_tag.splits) -done - -lemma enrich_msgq_sameinode: - "\no_del_event s; valid s\ - \ (f \ same_inode_files (enrich_msgq s q q') f') = (f \ same_inode_files s f')" -apply (induct s, simp) -apply (simp add:same_inode_files_def) -apply (auto simp:enrich_msgq_is_file enrich_msgq_cur_inof) -done - -lemma enrich_msgq_tainted_aux1: - "\no_del_event s; q \ current_msgqs s; q' \ current_msgqs s; valid s\ - \ (tainted s \ {O_msg q' m| m. O_msg q m \ tainted s}) \ tainted (enrich_msgq s q q')" -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons) -apply (case_tac a) -apply (auto split:option.splits if_splits dest:tainted_in_current - simp:enrich_msgq_filefd enrich_msgq_sameinode) -done - -lemma enrich_msgq_tainted_aux2: - "\no_del_event s; q \ current_msgqs s; q' \ current_msgqs s; valid s; valid (enrich_msgq s q q')\ - \ tainted (enrich_msgq s q q') \ (tainted s \ {O_msg q' m| m. O_msg q m \ tainted s})" -apply (induct s, simp) -apply (frule vt_grant_os, frule vd_cons) -apply (case_tac a) -apply (auto split:option.splits if_splits simp:enrich_msgq_filefd enrich_msgq_sameinode - dest:tainted_in_current vd_cons) -apply (drule_tac vd_cons)+ -apply (simp) -apply (drule set_mp) -apply simp -apply simp -apply (drule_tac s = s in tainted_in_current) -apply simp+ -done - -lemma enrich_msgq_alive: - "\alive s obj; valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ - \ alive (enrich_msgq s q q') obj" -apply (case_tac obj) -apply (simp_all add:enrich_msgq_is_file enrich_msgq_is_dir enrich_msgq_cur_msgqs - enrich_msgq_cur_msgs enrich_msgq_cur_procs enrich_msgq_cur_fds - enrich_msgq_cur_tcps enrich_msgq_cur_udps) -apply (rule impI, simp) -done - -lemma enrich_msgq_alive': - "\alive (enrich_msgq s q q') obj; valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ - \ alive s obj \ obj = O_msgq q' \ (\ m. obj = O_msg q' m \ alive s (O_msg q m))" -apply (case_tac obj) -apply (simp_all add:enrich_msgq_is_file enrich_msgq_is_dir enrich_msgq_cur_msgqs - enrich_msgq_cur_msgs enrich_msgq_cur_procs enrich_msgq_cur_fds - enrich_msgq_cur_tcps enrich_msgq_cur_udps) -apply (auto split:if_splits) -done - -lemma enrich_msgq_not_alive: - "\enrich_not_alive s (E_msgq q') obj; q' \ current_msgqs s; q \ current_msgqs s; no_del_event s; valid s\ - \ enrich_not_alive (enrich_msgq s q q') (E_msgq q') obj" -apply (case_tac obj) -apply (auto simp:enrich_msgq_cur_fds enrich_msgq_cur_files - enrich_msgq_cur_procs enrich_msgq_cur_inums enrich_msgq_cur_msgqs enrich_msgq_cur_msgs) -done - -lemma enrich_msgq_nodel: - "no_del_event (enrich_msgq s q q') = no_del_event s" -apply (induct s, simp) -by (case_tac a, auto) - -lemma enrich_msgq_died_proc: - "died (O_proc p) (enrich_msgq s q q') = died (O_proc p) s" -apply (induct s, simp) -by (case_tac a, auto) - -lemma cf2sfile_execve: - "\ff \ current_files (Execve p f fds # s); valid (Execve p f fds # s)\ - \ cf2sfile (Execve p f fds # s) ff= cf2sfile s ff" -by (auto dest:cf2sfile_other' simp:current_files_simps) -lemma cf2sfile_clone: - "\ff \ current_files (Clone p p' fds # s); valid (Clone p p' fds # s)\ - \ cf2sfile (Clone p p' fds # s) ff= cf2sfile s ff" -by (auto dest:cf2sfile_other' simp:current_files_simps) -lemma cf2sfile_ptrace: - "\ff \ current_files (Ptrace p p' # s); valid (Ptrace p p' # s)\ - \ cf2sfile (Ptrace p p' # s) ff= cf2sfile s ff" -by (auto dest:cf2sfile_other' simp:current_files_simps) -lemma cf2sfile_readfile: - "\ff \ current_files (ReadFile p fd # s); valid (ReadFile p fd # s)\ - \ cf2sfile (ReadFile p fd # s) ff= cf2sfile s ff" -by (auto dest:cf2sfile_other' simp:current_files_simps) -lemma cf2sfile_writefile: - "\ff \ current_files (WriteFile p fd # s); valid (WriteFile p fd # s)\ - \ cf2sfile (WriteFile p fd # s) ff= cf2sfile s ff" -by (auto dest:cf2sfile_other' simp:current_files_simps) -lemma cf2sfile_truncate: - "\ff \ current_files (Truncate p f len # s); valid (Truncate p f len # s)\ - \ cf2sfile (Truncate p f len # s) ff= cf2sfile s ff" -by (auto dest:cf2sfile_other' simp:current_files_simps) -lemma cf2sfile_createmsgq: - "\ff \ current_files (CreateMsgq p q # s); valid (CreateMsgq p q # s)\ - \ cf2sfile (CreateMsgq p q # s) ff= cf2sfile s ff" -by (auto dest:cf2sfile_other' simp:current_files_simps) -lemma cf2sfile_sendmsg: - "\ff \ current_files (SendMsg p q m # s); valid (SendMsg p q m # s)\ - \ cf2sfile (SendMsg p q m # s) ff = cf2sfile s ff" -by (auto dest:cf2sfile_other' simp:current_files_simps) -lemma cf2sfile_recvmsg: - "\ff \ current_files (RecvMsg p q m # s); valid (RecvMsg p q m # s)\ - \ cf2sfile (RecvMsg p q m # s) ff = cf2sfile s ff" -by (auto dest:cf2sfile_other' simp:current_files_simps) -lemmas cf2sfile_other'' = cf2sfile_recvmsg cf2sfile_sendmsg cf2sfile_createmsgq cf2sfile_truncate - cf2sfile_writefile cf2sfile_readfile cf2sfile_ptrace cf2sfile_clone cf2sfile_execve - -lemma enrich_msgq_prop: - "\valid s; q \ current_msgqs s; q' \ current_msgqs s; no_del_event s\ - \ valid (enrich_msgq s q q') \ - (\ tp. tp \ current_procs s \ cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp) \ - (\ f. f \ current_files s \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f) \ - (\ tq. tq \ current_msgqs s \ cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq) \ - (\ tp fd. fd \ proc_file_fds s tp \ cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd) \ - (cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q) \ - (tainted (enrich_msgq s q q') = (tainted s \ {O_msg q' m| m. O_msg q m \ tainted s}))" -proof (induct s) - case Nil - thus ?case by (auto) -next - case (Cons e s) - hence vd_cons': "valid (e # s)" and curq_cons: "q \ current_msgqs (e # s)" - and curq'_cons: "q' \ current_msgqs (e # s)" and nodel_cons: "no_del_event (e # s)" - and os: "os_grant s e" and grant: "grant s e" and vd: "valid s" - by (auto dest:vd_cons vt_grant_os vt_grant) - from curq'_cons nodel_cons have curq': "q' \ current_msgqs s" by (case_tac e, auto) - from nodel_cons have nodel: "no_del_event s" by (case_tac e, auto) - from nodel curq' vd Cons - have pre: "q \ current_msgqs s \ valid (enrich_msgq s q q') \ - (\tp. tp \ current_procs s \ cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp) \ - (\f. f \ current_files s \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f) \ - (\tq. tq \ current_msgqs s \ cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq) \ - (\tp fd. fd \ proc_file_fds s tp \ cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd) \ - (cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q) \ - (tainted (enrich_msgq s q q') = (tainted s \ {O_msg q' m| m. O_msg q m \ tainted s}))" - by auto - - from pre have pre_vd: "q \ current_msgqs s \ valid (enrich_msgq s q q')" by simp - from pre have pre_sp: "\ tp. \tp \ current_procs s; q \ current_msgqs s\ - \ cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp" by auto - from pre have pre_sf: "\ f. \f \ current_files s; q \ current_msgqs s\ - \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f" by auto - from pre have pre_sq: "\ tq. \tq \ current_msgqs s; q \ current_msgqs s\ - \ cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq" by auto - from pre have pre_sfd: "\ tp fd. \fd \ proc_file_fds s tp; q \ current_msgqs s\ - \ cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" by auto - hence pre_sfd': "\ tp fd f. \file_of_proc_fd s tp fd = Some f; q \ current_msgqs s\ - \ cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" by (auto simp:proc_file_fds_def) - from pre have pre_duq: "q \ current_msgqs s \ cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q" - by auto - have vd_enrich:"q \ current_msgqs s \ valid (e # enrich_msgq s q q')" - apply (frule pre_vd) - apply (erule_tac s = s and obj' = "E_msgq q'" in enrich_valid_intro_cons) - apply (simp_all add:pre nodel nodel_cons curq_cons vd_cons' vd enrich_msgq_hungs) - apply (clarsimp simp:nodel vd curq' enrich_msgq_alive) - apply (rule allI, rule impI, erule enrich_msgq_not_alive) - apply (simp_all add:curq' curq'_cons nodel vd enrich_msgq_cur_msgs enrich_msgq_filefd enrich_msgq_flagfd) - done - - have q_neq_q': "q' \ q" using curq'_cons curq_cons by auto - - have vd_enrich_cons: "valid (enrich_msgq (e # s) q q')" - proof- - have "\ p q''. e = CreateMsgq p q'' \ valid (enrich_msgq (e # s) q q')" - proof- - fix p q'' assume ev: "e = CreateMsgq p q''" - show "valid (enrich_msgq (e # s) q q')" - proof (cases "q'' = q") - case False with ev vd_enrich curq_cons - show ?thesis by simp - next - case True - have "os_grant (CreateMsgq p q # s) (CreateMsgq p q')" - using os ev - by (simp add:q_neq_q' curq') - moreover have "grant (CreateMsgq p q # s) (CreateMsgq p q')" - using grant ev - by (auto simp add:sectxt_of_obj_def split:option.splits) - ultimately - show ?thesis using ev vd_cons' True - by (auto intro: valid.intros(2)) - qed - qed - moreover have "\ p q'' m. \e = SendMsg p q'' m; q \ current_msgqs s\ - \ valid (enrich_msgq (e # s) q q')" - proof- - fix p q'' m assume ev: "e = SendMsg p q'' m" and q_in: "q \ current_msgqs s" - show "valid (enrich_msgq (e # s) q q')" - proof (cases "q'' = q") - case False with ev vd_enrich q_in - show ?thesis by simp - next - case True - from grant os ev True obtain psec qsec - where psec: "sectxt_of_obj s (O_proc p) = Some psec" - and qsec: "sectxt_of_obj s (O_msgq q) = Some qsec" - by (auto split:option.splits) - from psec q_in os ev - have psec':"sectxt_of_obj (enrich_msgq s q q') (O_proc p) = Some psec" - by (auto dest!:pre_sp simp:cp2sproc_def split:option.splits) - from qsec q_in pre_duq vd - have qsec': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = Some qsec" - by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms') - show ?thesis using ev True vd_cons' q_in vd_enrich nodel vd - curq' psec psec' qsec qsec' grant os q_neq_q' - apply (simp, erule_tac valid.intros(2)) - apply (auto simp:q_neq_q' enrich_msgq_cur_msgqs enrich_msgq_cur_procs - enrich_msgq_cur_msgs sectxt_of_obj_simps) - done - qed - qed - moreover have "\ p q'' m. \e = RecvMsg p q'' m; q \ current_msgqs s\ - \ valid (enrich_msgq (e # s) q q')" - proof- - fix p q'' m assume ev: "e = RecvMsg p q'' m" and q_in: "q \ current_msgqs s" - show "valid (enrich_msgq (e # s) q q')" - proof (cases "q'' = q") - case False with ev vd_enrich q_in - show ?thesis by simp - next - case True - from grant os ev True obtain psec qsec msec - where psec: "sectxt_of_obj s (O_proc p) = Some psec" - and qsec: "sectxt_of_obj s (O_msgq q) = Some qsec" - and msec: "sectxt_of_obj s (O_msg q (hd (msgs_of_queue s q))) = Some msec" - by (auto split:option.splits) - from psec q_in os ev - have psec':"sectxt_of_obj (enrich_msgq s q q') (O_proc p) = Some psec" - by (auto dest!:pre_sp simp:cp2sproc_def split:option.splits) - from qsec q_in pre_duq vd - have qsec': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = Some qsec" - by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms') - from qsec q_in vd - have qsec'': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q) = Some qsec" - apply (frule_tac pre_sq, simp_all) - by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms') - from msec q_in pre_duq vd qsec qsec' qsec'' curq' nodel - have msec': "sectxt_of_obj (enrich_msgq s q q') (O_msg q' (hd (msgs_of_queue s q))) = Some msec" - apply (auto simp:cq2smsgq_def enrich_msgq_cur_msgs - split:option.splits dest!:current_has_sms') - apply (case_tac "msgs_of_queue s q") - using os ev True apply simp - apply (simp add:cqm2sms.simps split:option.splits) - apply (auto simp:cm2smsg_def split:option.splits) - done - show ?thesis using ev True vd_cons' q_in vd_enrich nodel vd - curq' grant os q_neq_q' psec psec' msec msec' qsec qsec' - apply (simp, erule_tac valid.intros(2)) - apply (auto simp:enrich_msgq_cur_msgqs enrich_msgq_cur_procs - enrich_msgq_cur_msgs sectxt_of_obj_simps) - done - qed - qed - ultimately - show ?thesis using vd_enrich curq_cons vd_cons' - apply (case_tac e) - apply (auto simp del:enrich_msgq.simps) - apply (auto split:if_splits ) - done - qed - - have curpsec: "\ tp. \tp \ current_procs s; q \ current_msgqs s\ \ - sectxt_of_obj (enrich_msgq s q q') (O_proc tp) = sectxt_of_obj s (O_proc tp)" - using pre_vd vd - apply (frule_tac pre_sp, simp) - by (auto simp:cp2sproc_def split:option.splits if_splits dest!:current_has_sec') - have curfsec: "\ f. \is_file s f; q \ current_msgqs s\ \ - sectxt_of_obj (enrich_msgq s q q') (O_file f) = sectxt_of_obj s (O_file f)" - proof- - fix f - assume a1: "is_file s f" and a2: "q \ current_msgqs s" - from a2 pre_sf pre_vd - have pre': "\f. f \ current_files s \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f" - and vd_enrich: "valid (enrich_msgq s q q')" - by auto - hence csf: "cf2sfile (enrich_msgq s q q') f = cf2sfile s f" - using a1 by (auto simp:is_file_in_current) - from a1 obtain sf where csf_some: "cf2sfile s f = Some sf" - apply (case_tac "cf2sfile s f") - apply (drule current_file_has_sfile') - apply (simp add:vd, simp add:is_file_in_current, simp) - done - from a1 have a1': "is_file (enrich_msgq s q q') f" - using vd nodel by (simp add:enrich_msgq_is_file) - show "sectxt_of_obj (enrich_msgq s q q') (O_file f) = sectxt_of_obj s (O_file f)" - using csf csf_some vd_enrich vd a1 a1' - apply (auto simp:cf2sfile_def split:option.splits if_splits) - apply (case_tac f, simp_all) - apply (drule root_is_dir', simp+) - done - qed - have curdsec: "\ tf. \is_dir s tf; q \ current_msgqs s\ - \ sectxt_of_obj (enrich_msgq s q q') (O_dir tf) = sectxt_of_obj s (O_dir tf)" - proof- - fix tf - assume a1: "is_dir s tf" and a2: "q \ current_msgqs s" - from a2 pre_sf pre_vd - have pre': "\f. f \ current_files s \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f" - and vd_enrich: "valid (enrich_msgq s q q')" - by auto - hence csf: "cf2sfile (enrich_msgq s q q') tf = cf2sfile s tf" - using a1 by (auto simp:is_dir_in_current) - from a1 obtain sf where csf_some: "cf2sfile s tf = Some sf" - apply (case_tac "cf2sfile s tf") - apply (drule current_file_has_sfile') - apply (simp add:vd, simp add:is_dir_in_current, simp) - done - from a1 have a1': "is_dir (enrich_msgq s q q') tf" - using enrich_msgq_is_dir vd nodel by simp - from a1 have a3: "\ is_file s tf" using vd by (simp add:is_dir_not_file) - from a1' vd have a3': "\ is_file (enrich_msgq s q q') tf" by (simp add:is_dir_not_file) - show "sectxt_of_obj (enrich_msgq s q q') (O_dir tf) = sectxt_of_obj s (O_dir tf)" - using csf csf_some a3 a3' vd_enrich vd - apply (auto simp:cf2sfile_def split:option.splits) - apply (case_tac tf) - apply (simp add:root_sec_remains, simp) - done - qed - have curpsecs: "\ tf ctxts'. \is_dir s tf; q \ current_msgqs s; get_parentfs_ctxts s tf = Some ctxts'\ - \ \ ctxts. get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts \ set ctxts = set ctxts'" - proof- - fix tf ctxts' - assume a1: "is_dir s tf" and a2: "q \ current_msgqs s" - and a4: "get_parentfs_ctxts s tf = Some ctxts'" - from a2 pre - have pre': "\f. f \ current_files s \ cf2sfile (enrich_msgq s q q') f = cf2sfile s f" - and vd_enrich': "valid (enrich_msgq s q q')" - by auto - hence csf: "cf2sfile (enrich_msgq s q q') tf = cf2sfile s tf" - using a1 by (auto simp:is_dir_in_current) - from a1 obtain sf where csf_some: "cf2sfile s tf = Some sf" - apply (case_tac "cf2sfile s tf") - apply (drule current_file_has_sfile') - apply (simp add:vd, simp add:is_dir_in_current, simp) - done - from a1 have a1': "is_dir (enrich_msgq s q q') tf" - using enrich_msgq_is_dir vd nodel by simp - from a1 have a5: "\ is_file s tf" using vd by (simp add:is_dir_not_file) - from a1' vd have a5': "\ is_file (enrich_msgq s q q') tf" by (simp add:is_dir_not_file) - - from a1' pre_vd a2 obtain ctxts - where a3: "get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts" - apply (case_tac "get_parentfs_ctxts (enrich_msgq s q q') tf") - apply (drule get_pfs_secs_prop', simp+) - done - moreover have "set ctxts = set ctxts'" - proof (cases tf) - case Nil - with a3 a4 vd vd_enrich' - show ?thesis - by (simp add:get_parentfs_ctxts.simps root_sec_remains split:option.splits) - next - case (Cons a ff) - with csf csf_some a5 a5' vd_enrich' vd a3 a4 - show ?thesis - apply (auto simp:cf2sfile_def split:option.splits if_splits) - done - qed - ultimately - show "\ ctxts. get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts \ set ctxts = set ctxts'" - by auto - qed - - have psec_cons: "\ tp. tp \ current_procs (e # s) \ - sectxt_of_obj (enrich_msgq (e # s) q q') (O_proc tp) = sectxt_of_obj (e # s) (O_proc tp)" - using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd - apply (case_tac e) - apply (auto intro:curpsec simp:sectxt_of_obj_simps) - apply (frule curpsec, simp, frule curfsec, simp) - apply (auto split:option.splits)[1] - apply (frule vd_cons) defer apply (frule vd_cons) - apply (auto intro:curpsec simp:sectxt_of_obj_simps) - done - - - have sf_cons: "\ f. f \ current_files (e # s) \ cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" - proof- - fix f - assume a1: "f \ current_files (e # s)" - hence a1': "f \ current_files (enrich_msgq (e # s) q q')" - using nodel_cons os vd vd_cons' vd_enrich_cons - apply (case_tac e) - apply (auto simp:current_files_simps enrich_msgq_cur_files dest:is_file_in_current split:option.splits) - done - have b1: "\ p f' flags fd opt. e = Open p f' flags fd opt \ - cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" - proof- - fix p f' flags fd opt - assume ev: "e = Open p f' flags fd opt" - show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" - proof (cases opt) - case None - with a1 a1' ev vd_cons' vd_enrich_cons curq_cons - show ?thesis - apply (simp add:cf2sfile_open_none) - apply (simp add:pre_sf current_files_simps) - done - next - case (Some inum) - show ?thesis - proof (cases "f = f'") - case False - with a1 a1' ev vd_cons' vd_enrich_cons curq_cons Some - show ?thesis - apply (simp add:cf2sfile_open) - apply (simp add:pre_sf current_files_simps) - done - next - case True - with vd_cons' ev os Some - obtain pf where pf: "parent f = Some pf" by auto - then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts" - using os vd ev Some True - apply (case_tac "get_parentfs_ctxts s pf") - apply (drule get_pfs_secs_prop', simp, simp) - apply auto - done - - have "sectxt_of_obj (Open p f' flags fd (Some inum) # enrich_msgq s q q') (O_file f') = - sectxt_of_obj (Open p f' flags fd (Some inum) # s) (O_file f')" - using Some vd_enrich_cons vd_cons' ev pf True os curq_cons - by (simp add:sectxt_of_obj_simps curpsec curdsec) - moreover - have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)" - using curq_cons ev pf Some True os - by (simp add:curdsec) - moreover - have "\ ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \ set ctxts' = set ctxts" - using curq_cons ev pf Some True os vd psecs - apply (case_tac "get_parentfs_ctxts s pf") - apply (drule get_pfs_secs_prop', simp+) - apply (rule curpsecs, simp+) - done - then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'" - and psecs_eq: "set ctxts' = set ctxts" by auto - ultimately show ?thesis - using a1 a1' ev vd_cons' vd_enrich_cons Some True pf psecs - by (simp add:cf2sfile_open split:option.splits) - qed - qed - qed - have b2: "\ p f' inum. e = Mkdir p f' inum \ - cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" - proof- - fix p f' inum - assume ev: "e = Mkdir p f' inum" - show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" - proof (cases "f' = f") - case False - with a1 a1' ev vd_cons' vd_enrich_cons curq_cons - show ?thesis - apply (simp add:cf2sfile_mkdir) - apply (simp add:pre_sf current_files_simps) - done - next - case True - with vd_cons' ev os - obtain pf where pf: "parent f = Some pf" by auto - then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts" - using os vd ev True - apply (case_tac "get_parentfs_ctxts s pf") - apply (drule get_pfs_secs_prop', simp, simp) - apply auto - done - - have "sectxt_of_obj (Mkdir p f' inum # enrich_msgq s q q') (O_dir f') = - sectxt_of_obj (Mkdir p f' inum # s) (O_dir f')" - using vd_enrich_cons vd_cons' ev pf True os curq_cons - by (simp add:sectxt_of_obj_simps curpsec curdsec) - moreover - have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)" - using curq_cons ev pf True os - by (simp add:curdsec) - moreover - have "\ ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \ set ctxts' = set ctxts" - using curq_cons ev pf True os vd psecs - apply (case_tac "get_parentfs_ctxts s pf") - apply (drule get_pfs_secs_prop', simp+) - apply (rule curpsecs, simp+) - done - then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'" - and psecs_eq: "set ctxts' = set ctxts" by auto - ultimately show ?thesis - using a1 a1' ev vd_cons' vd_enrich_cons True pf psecs - apply (simp add:cf2sfile_mkdir split:option.splits) - done - qed - qed - have b3: "\ p f' f''. e = LinkHard p f' f'' \ - cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" - proof- - fix p f' f'' - assume ev: "e = LinkHard p f' f''" - show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" - proof (cases "f'' = f") - case False - with a1 a1' ev vd_cons' vd_enrich_cons curq_cons - show ?thesis - apply (simp add:cf2sfile_linkhard) - apply (simp add:pre_sf current_files_simps) - done - next - case True - with vd_cons' ev os - obtain pf where pf: "parent f = Some pf" by auto - then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts" - using os vd ev True - apply (case_tac "get_parentfs_ctxts s pf") - apply (drule get_pfs_secs_prop', simp, simp) - apply auto - done - - have "sectxt_of_obj (LinkHard p f' f'' # enrich_msgq s q q') (O_file f) = - sectxt_of_obj (LinkHard p f' f'' # s) (O_file f)" - using vd_enrich_cons vd_cons' ev pf True os curq_cons - by (simp add:sectxt_of_obj_simps curpsec curdsec) - moreover - have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)" - using curq_cons ev pf True os - by (simp add:curdsec) - moreover - have "\ ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \ set ctxts' = set ctxts" - using curq_cons ev pf True os vd psecs - apply (case_tac "get_parentfs_ctxts s pf") - apply (drule get_pfs_secs_prop', simp+) - apply (rule curpsecs, simp+) - done - then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'" - and psecs_eq: "set ctxts' = set ctxts" by auto - ultimately show ?thesis - using a1 a1' ev vd_cons' vd_enrich_cons True pf psecs - apply (simp add:cf2sfile_linkhard split:option.splits) - done - qed - qed - show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" - apply (case_tac e) - prefer 6 apply (erule b1) - prefer 11 apply (erule b2) - prefer 11 apply (erule b3) - apply (simp_all only:b1 b2 b3) - using a1' a1 vd_enrich_cons vd_cons' curq_cons nodel_cons - apply (simp_all add:cf2sfile_other'' cf2sfile_simps enrich_msgq.simps no_del_event.simps split:if_splits) - apply (simp_all add:pre_sf cf2sfile_other' current_files_simps split:if_splits) - apply (drule vd_cons, simp add:cf2sfile_other', drule pre_sf, simp+)+ - done - qed - - have sfd_cons:"\ tp fd f. file_of_proc_fd (e # s) tp fd = Some f \ - cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" - proof- - fix tp fd f - assume a1: "file_of_proc_fd (e # s) tp fd = Some f" - hence a1': "file_of_proc_fd (enrich_msgq (e # s) q q') tp fd = Some f" - using nodel_cons vd_enrich os vd_cons' - apply (case_tac e, auto simp:enrich_msgq_filefd simp del:enrich_msgq.simps) - done - have b1: "\ p f' flags fd' opt. e = Open p f' flags fd' opt \ - cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" - proof- - fix p f' flags fd' opt - assume ev: "e = Open p f' flags fd' opt" - have c1': "file_of_proc_fd (Open p f' flags fd' opt # s) tp fd = Some f" - using a1' ev a1 by (simp split:if_splits) - show "cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" thm cfd2sfd_open - proof (cases "tp = p \ fd = fd'") - case False - show ?thesis using ev vd_enrich_cons vd_cons' a1' a1 False curq_cons - apply (simp add:cfd2sfd_open split:if_splits del:file_of_proc_fd.simps) - apply (rule conjI, rule impI, simp) - apply (rule conjI, rule impI, simp) - apply (auto simp: False intro!:pre_sfd' split:if_splits) - done - next - case True - have "f' \ current_files (Open p f' flags fd' opt # s)" using ev vd_cons' os - by (auto simp:current_files_simps is_file_in_current split:option.splits) - hence "cf2sfile (Open p f' flags fd' opt # enrich_msgq s q q') f' - = cf2sfile (Open p f' flags fd' opt # s) f'" - using sf_cons ev by auto - moreover have "sectxt_of_obj (enrich_msgq s q q') (O_proc p) = sectxt_of_obj s (O_proc p)" - apply (rule curpsec) - using os ev curq_cons - by (auto split:option.splits) - ultimately show ?thesis - using ev True a1 a1' vd_enrich_cons vd_cons' - apply (simp add:cfd2sfd_open sectxt_of_obj_simps del:file_of_proc_fd.simps) - done - qed - qed - show "cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" - apply (case_tac e) - prefer 6 apply (erule b1) - using a1' a1 vd_enrich_cons vd_cons' curq_cons - apply (simp_all only:cfd2sfd_simps enrich_msgq.simps) - apply (auto simp:cfd2sfd_simps pre_sfd' dest:vd_cons cfd2sfd_other split:if_splits) - done - qed - - have pfds_cons: "\ tp. tp \ current_procs (e # s) \ - cpfd2sfds (enrich_msgq (e # s) q q') tp = cpfd2sfds (e # s) tp" - apply (auto simp add:cpfd2sfds_def proc_file_fds_def) - apply (rule_tac x = fd in exI, rule conjI, rule_tac x = f in exI) - prefer 3 - apply (rule_tac x = fd in exI, rule conjI, rule_tac x = f in exI) - apply (auto simp:sfd_cons enrich_msgq_filefd nodel_cons vd_cons') - done - - have tainted_cons: "tainted (enrich_msgq (e # s) q q') = - (tainted (e # s) \ {O_msg q' m | m. O_msg q m \ tainted (e # s)})" - apply (rule equalityI) - using nodel_cons curq_cons curq'_cons vd_cons' vd_enrich_cons - apply (rule enrich_msgq_tainted_aux2) - using nodel_cons curq_cons curq'_cons vd_cons' - apply (rule enrich_msgq_tainted_aux1) - done - have pre_tainted: "q \ current_msgqs s \ tainted (enrich_msgq s q q') = - (tainted s \ {O_msg q' m| m. O_msg q m \ tainted s})" by (simp add:pre) - - have "\tp fd. fd \ proc_file_fds (e # s) tp \ cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" - by (auto simp:proc_file_fds_def elim!:sfd_cons) - moreover - have "\tp. tp \ current_procs (e # s) \ cp2sproc (enrich_msgq (e # s) q q') tp = cp2sproc (e # s) tp" - by (auto simp:cp2sproc_def pfds_cons psec_cons enrich_msgq_died_proc split:option.splits) - moreover - have "\tq. tq \ current_msgqs (e # s) \ cq2smsgq (enrich_msgq (e # s) q q') tq = cq2smsgq (e # s) tq" - proof clarify - fix tq assume a1: "tq \ current_msgqs (e # s)" - - have curqsec: "\ tq. \tq \ current_msgqs s; q \ current_msgqs s\ \ - sectxt_of_obj (enrich_msgq s q q') (O_msgq tq) = sectxt_of_obj s (O_msgq tq)" - using pre_vd vd - apply (frule_tac pre_sq, simp) - by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') - have cursms: "\ q''. \q'' \ current_msgqs s; q \ current_msgqs s\ \ - cqm2sms (enrich_msgq s q q') q'' (msgs_of_queue (enrich_msgq s q q') q'') = - cqm2sms s q'' (msgs_of_queue s q'')" - using pre_vd vd - apply (frule_tac pre_sq, simp) - by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') - have qsec_cons: "sectxt_of_obj (enrich_msgq (e # s) q q') (O_msgq tq) = sectxt_of_obj (e # s) (O_msgq tq)" - using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons curq_cons a1 - apply (case_tac e) - apply (auto intro:curqsec simp:sectxt_of_obj_simps curpsec split:option.splits dest!:current_has_sec') - apply (frule vd_cons) defer apply (frule vd_cons) - apply (auto intro:curqsec simp:sectxt_of_obj_simps) - done - have sms_cons: "cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = - cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" - proof- - have b1: "\ p q'' m. e = SendMsg p q'' m \ - cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = - cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" - apply (case_tac e) - using a1 curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons - apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms pre_tainted - split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) - apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*}) - apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms pre_tainted - split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) - done - have b2: "\ p q''. e = CreateMsgq p q'' \ - cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = - cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" - using a1 curq_cons curq'_cons vd_enrich_cons vd_cons' - apply (auto simp:cqm2sms_simps intro:cursms) - apply (auto simp:cqm2sms.simps) - done - have b3: "\ p q'' m. e = RecvMsg p q'' m \ - cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = - cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" - using a1 curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons - apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms - split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) - apply (frule vd_cons) defer apply (frule vd_cons) - apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms - split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) - done - show "cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = - cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" - apply (case_tac e) - prefer 15 apply (erule b2) - prefer 15 apply (erule b1) - prefer 15 apply (erule b3) - using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons a1 - apply (auto intro:cursms simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec - split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) - done - qed - - show "cq2smsgq (enrich_msgq (e # s) q q') tq = cq2smsgq (e # s) tq" - using a1 curq_cons - apply (simp add:cq2smsgq_def qsec_cons sms_cons) - done - qed - moreover - have "cq2smsgq (enrich_msgq (e # s) q q') q' = cq2smsgq (e # s) q" - proof- - have duqsec: "q \ current_msgqs s \ - sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = sectxt_of_obj s (O_msgq q)" - apply (frule pre_duq) using vd - by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') - have duqsms: "q \ current_msgqs s \ - cqm2sms (enrich_msgq s q q') q' (msgs_of_queue (enrich_msgq s q q') q') = - cqm2sms s q (msgs_of_queue s q)" - apply (frule pre_duq) using vd - by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') - have qsec_cons: "sectxt_of_obj (enrich_msgq (e # s) q q') (O_msgq q') = sectxt_of_obj (e # s) (O_msgq q)" - using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons curq_cons - apply (case_tac e) - apply (auto simp:duqsec sectxt_of_obj_simps curpsec split:option.splits dest!:current_has_sec') - apply (frule vd_cons) defer apply (frule vd_cons) - apply (auto intro:duqsec simp:sectxt_of_obj_simps) - done - have sms_cons: "cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = - cqm2sms (e # s) q (msgs_of_queue (e # s) q)" - proof- - have b1: "\ p q'' m. e = SendMsg p q'' m \ - cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = - cqm2sms (e # s) q (msgs_of_queue (e # s) q)" - apply (case_tac e) - using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons - apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted - enrich_msgq_cur_procs enrich_msgq_cur_msgqs - split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) - apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*}) - apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted - enrich_msgq_cur_procs enrich_msgq_cur_msgqs dest:tainted_in_current - split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) - done - have b2: "\ p q''. e = CreateMsgq p q'' \ - cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = - cqm2sms (e # s) q (msgs_of_queue (e # s) q)" - using curq_cons curq'_cons vd_enrich_cons vd_cons' - apply (auto simp:cqm2sms_simps intro:duqsms) - apply (auto simp:cqm2sms.simps) - done - have b3: "\ p q'' m. e = RecvMsg p q'' m \ - cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = - cqm2sms (e # s) q (msgs_of_queue (e # s) q)" - using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons - apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted - enrich_msgq_cur_procs enrich_msgq_cur_msgqs - split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) - apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*}) - apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted - enrich_msgq_cur_procs enrich_msgq_cur_msgqs dest:tainted_in_current - split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) - done - show ?thesis - apply (case_tac e) - prefer 15 apply (erule b2) - prefer 15 apply (erule b1) - prefer 15 apply (erule b3) - using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons - apply (auto intro:duqsms simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsms - split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) - done - qed - show ?thesis - using curq_cons - apply (simp add:cq2smsgq_def qsec_cons sms_cons) - done - qed - ultimately show ?case using vd_enrich_cons sf_cons tainted_cons - by auto -qed - -fun enrich_file_link :: "t_state \ t_file \ t_state \ (t_file \ t_state)" -where - "enrich_file_link [] tf df = []" - -fun enrich_file_set :: "t_state \ t_file \ t_file \ t_state" -where - "enrich_file_set [] tf df = []" -| "enrich_file_set [] (Open p f flags fd (Some inum) # s) = + "enrich_file [] tf s = ({}, [])" +| "enrich_file (Open p f flags fd (Some inum) # s) tf s = + " diff -r d9dc04c3ea90 -r dfde07c7cd6b no_shm_selinux/New_obj_prop.thy --- a/no_shm_selinux/New_obj_prop.thy Thu Jan 09 14:39:00 2014 +0800 +++ b/no_shm_selinux/New_obj_prop.thy Thu Jan 09 19:09:09 2014 +0800 @@ -63,31 +63,57 @@ 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) +lemma finite_fs_fnames: + "finite (fs::t_file set) \ finite {fn. fn # f \ fs}" +thm 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) +apply (rule_tac B = "(list_case [] (\fn pf'. if pf' = f then fn else []) ` fs)" in finite_subset) unfolding image_def apply (clarsimp split:if_splits) -apply (rule_tac x = "x # f" in bexI, simp+) +apply (rule_tac x = "x # f" in bexI, simp+) +done -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) +lemma ncf_notin_curf_aux: + "finite (fs::t_file set) \ Suc (Max (fname_length_set {fn. fn # f \ fs})) \ fname_length_set {fn. fn # f \ fs}" +apply (drule_tac f = f in finite_fs_fnames) +apply (drule_tac h = length in finite_imageI) +apply (simp add:fname_length_set_def) +apply (drule nn_notin_aux) +apply auto done +lemma ncf_notin_curf_general: + assumes vd: "valid \" and fin_fs: "finite fs" + shows "new_childf_general f \ fs \ (current_files \ \ fs)" +proof- + from vd fin_fs have "finite (current_files \ \ fs)" + by (auto dest:finite_cf) + hence a1: "Suc (Max (fname_length_set {fn. fn # f \ (current_files \ \ fs)})) \ + fname_length_set {fn. fn # f \ (current_files \ \ fs)}" + by (erule_tac ncf_notin_curf_aux) + have a2: "\ f pf fs. f # pf \ fs \ f \ {fn. fn # pf \ fs}" by auto + have a3: "\ f pf fs. f \ {fn. fn # pf \ fs} \ length f \ fname_length_set {fn. fn # pf \ fs}" + by (auto simp:fname_length_set_def) + show ?thesis + apply (simp only:new_childf_general_def next_fname_def all_fname_under_dir_def) + apply (rule notI) + apply (drule a2, drule a3) + apply (simp only:len_fname_all) + using a1 by auto +qed + +lemma ncf_notin_curf: + "valid \ \ new_childf f \ \ (current_files \)" +apply (drule_tac fs = "{}" in ncf_notin_curf_general) +apply (simp) +apply (simp add:new_childf_def) +done + +lemma ncf_parent_general: "valid \ \ parent (new_childf_general f \ fs) = Some f" +by (simp add:new_childf_general_def) + lemma ncf_parent: "valid \ \ parent (new_childf f \) = Some f" -by (simp add:new_childf_def) +by (simp add:new_childf_def ncf_parent_general) end diff -r d9dc04c3ea90 -r dfde07c7cd6b no_shm_selinux/ROOT --- a/no_shm_selinux/ROOT Thu Jan 09 14:39:00 2014 +0800 +++ b/no_shm_selinux/ROOT Thu Jan 09 19:09:09 2014 +0800 @@ -40,5 +40,4 @@ session "enrich" = "s2ss" + options [document = false] theories - Enrich - Enrich1 \ No newline at end of file + Enrich \ No newline at end of file diff -r d9dc04c3ea90 -r dfde07c7cd6b no_shm_selinux/Static.thy --- a/no_shm_selinux/Static.thy Thu Jan 09 14:39:00 2014 +0800 +++ b/no_shm_selinux/Static.thy Thu Jan 09 19:09:09 2014 +0800 @@ -646,7 +646,6 @@ "init_obj_related (S_proc (pi, sec, fds) tag) (O_proc p') = (pi = Init 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 (qi, sec, sms)) (O_msgq q') = (qi = Init q')" (* | "init_obj_related (S_shm (hi, sec)) (O_shm h') = (hi = Init h')" @@ -658,11 +657,13 @@ (***************** for backward proof when Instancing static objects ******************) +(* fun all_procs :: "t_state \ t_process set" where "all_procs [] = init_procs" | "all_procs (Clone p p' fds # s) = insert p' (all_procs s)" | "all_procs (e # s) = all_procs s" +*) definition next_nat :: "nat set \ nat" where @@ -672,9 +673,12 @@ where "new_proc \ = next_nat (current_procs \)" +(* definition brandnew_proc :: "t_state \ t_process" where "brandnew_proc s \ next_nat (all_procs s)" +*) + definition new_inode_num :: "t_state \ nat" where @@ -698,9 +702,9 @@ where "new_proc_fd \ p = next_nat (current_proc_fds \ p)" -definition all_fname_under_dir:: "t_file \ t_state \ t_fname set" +definition all_fname_under_dir:: "t_file \ t_state \ t_file set \ t_fname set" where - "all_fname_under_dir d \ = {fn. \ f. fn # d = f \ f \ current_files \}" + "all_fname_under_dir d \ fs = {fn. \ f. fn # d = f \ f \ (current_files \ \ fs)}" fun fname_all_a:: "nat \ t_fname" where @@ -711,13 +715,17 @@ where "fname_length_set fns = length`fns" -definition next_fname:: "t_file \ t_state \ t_fname" +definition next_fname:: "t_file \ t_state \ t_file set \ t_fname" where - "next_fname pf \ = fname_all_a ((Max (fname_length_set (all_fname_under_dir pf \))) + 1)" + "next_fname pf \ fs = fname_all_a ((Max (fname_length_set (all_fname_under_dir pf \ fs))) + 1)" + +definition new_childf_general:: "t_file \ t_state \ t_file set \ t_file" +where + "new_childf_general pf \ fs = next_fname pf \ fs # pf" definition new_childf:: "t_file \ t_state \ t_file" where - "new_childf pf \ = next_fname pf \ # pf" + "new_childf pf \ = new_childf_general pf \ {}" end diff -r d9dc04c3ea90 -r dfde07c7cd6b no_shm_selinux/Temp.thy --- a/no_shm_selinux/Temp.thy Thu Jan 09 14:39:00 2014 +0800 +++ b/no_shm_selinux/Temp.thy Thu Jan 09 19:09:09 2014 +0800 @@ -264,11 +264,18 @@ "ss \ sss \ \ ss' \ sss. ss \ ss'" lemma s2ss_included_sobj: - "\alive s obj; co2sobj s obj= Some sobj\ \ sobj \ (s2ss s)" + "\dalive s obj; co2sobj s obj= Some sobj\ \ sobj \ (s2ss s)" by (simp add:s2ss_def, rule_tac x = obj in exI, simp) +fun init_dobj_related :: "t_sobject \ t_dobject \ bool" +where + "init_dobj_related (S_proc (pi, sec, fds) tag) (D_proc p') = (pi = Init p')" +| "init_dobj_related (S_file sfs tag) (D_file f) = (\ sf \ sfs. sfile_related sf f)" +| "init_dobj_related (S_dir sf) (D_dir f) = (sfile_related sf f)" +| "init_dobj_related _ _ = False" + lemma init_ss_in_prop: - "\s2ss s \ static; co2sobj s obj = Some sobj; alive s obj; init_obj_related sobj obj\ + "\s2ss s \ static; co2sobj s obj = Some sobj; dalive s obj; init_dobj_related sobj obj\ \ \ ss \ static. sobj \ ss" apply (simp add:init_ss_in_def init_ss_eq_def) apply (erule bexE, erule conjE)