diff -r 271e9818b6f6 -r 99af1986e1e0 simple_selinux/Co2sobj_prop.thy --- a/simple_selinux/Co2sobj_prop.thy Mon Dec 02 10:52:40 2013 +0800 +++ b/simple_selinux/Co2sobj_prop.thy Tue Dec 03 22:42:48 2013 +0800 @@ -24,7 +24,7 @@ apply (frule vt_grant_os, frule vd_cons, rule ext, rule ext) unfolding cm2smsg_def apply (case_tac e) -apply (auto simp:sectxt_of_obj_simps tainted_eq_Tainted +apply (auto simp:sectxt_of_obj_simps split:t_object.splits option.splits if_splits dest:not_deleted_init_msg dest!:current_has_sec') done @@ -37,7 +37,7 @@ | _ \ None) else cm2smsg s q' m') " apply (frule vd_cons, frule vt_grant_os, rule ext, rule ext) -apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps +apply (auto simp:cm2smsg_def sectxt_of_obj_simps split:if_splits option.splits dest:not_deleted_init_msg dest!:current_has_sec') done @@ -45,35 +45,35 @@ lemma cm2smsg_recvmsg1: "\q' \ q; valid (RecvMsg p q m # s)\ \ cm2smsg (RecvMsg p q m # s) q' = cm2smsg s q'" apply (frule vd_cons, frule vt_grant_os, rule ext) -apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps +apply (auto simp:cm2smsg_def sectxt_of_obj_simps split:if_splits option.splits) done lemma cm2smsg_recvmsg2: "\m' \ m; valid (RecvMsg p q m # s)\ \ cm2smsg (RecvMsg p q m # s) q m' = cm2smsg s q m'" apply (frule vd_cons, frule vt_grant_os) -apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps +apply (auto simp:cm2smsg_def sectxt_of_obj_simps split:if_splits option.splits) done lemma cm2smsg_recvmsg1': "\valid (RecvMsg p q m # s); q' \ q\ \ cm2smsg (RecvMsg p q m # s) q' = cm2smsg s q'" apply (frule vd_cons, frule vt_grant_os, rule ext) -apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps +apply (auto simp:cm2smsg_def sectxt_of_obj_simps split:if_splits option.splits) done lemma cm2smsg_recvmsg2': "\valid (RecvMsg p q m # s); m' \ m\ \ cm2smsg (RecvMsg p q m # s) q m' = cm2smsg s q m'" apply (frule vd_cons, frule vt_grant_os) -apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps +apply (auto simp:cm2smsg_def sectxt_of_obj_simps split:if_splits option.splits) done lemma cm2smsg_removemsgq: "\q' \ q; valid (RemoveMsgq p q # s)\ \ cm2smsg (RemoveMsgq p q # s) q' = cm2smsg s q'" apply (frule vd_cons, frule vt_grant_os, rule ext) -apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps +apply (auto simp:cm2smsg_def sectxt_of_obj_simps split:if_splits option.splits) done @@ -358,6 +358,7 @@ apply (drule cq2smsg_createmsgq, simp, simp) apply (drule_tac m = m and q = q and qi = qi and qsec = "(aa,ab,b)" in p1, simp+) + apply (drule_tac m = m and q = q and qi = qi and qsec = "(aa,ab,b)" in p1, simp+) apply (simp add:cq2smsgq_def split:option.splits) apply (rule_tac m = m and sm = sm in cqm2sms_prop0, simp+) @@ -670,62 +671,6 @@ apply (simp add:cf2sfile_mkdir1 current_files_simps) done -lemma cf2sfile_linkhard1: - "\valid (LinkHard p oldf f # s); f' \ current_files s\ - \ cf2sfile (LinkHard p oldf f# s) f' = cf2sfile s f'" -apply (frule vd_cons, frule vt_grant_os, frule noroot_events) -apply (case_tac "f = f'", simp) -apply (induct f', simp add:sroot_only, simp) -apply (frule parentf_in_current', simp+) -apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps - get_parentfs_ctxts_simps split:if_splits option.splits) -done - -lemma cf2sfile_linkhard2: - "\valid (LinkHard p oldf f # s); is_file s f'\ - \ cf2sfile (LinkHard p oldf f # s) f' = cf2sfile s f'" -apply (frule vd_cons, drule is_file_in_current) -by (simp add:cf2sfile_linkhard1) - -lemma cf2sfile_linkhard3: - "\valid (LinkHard p oldf f # s); is_dir s f'\ - \ cf2sfile (LinkHard p oldf f # s) f' = cf2sfile s f'" -apply (frule vd_cons, drule is_dir_in_current) -by (simp add:cf2sfile_linkhard1) - -lemma cf2sfile_linkhard4: - "valid (LinkHard p oldf f # s) - \ cf2sfile (LinkHard p oldf f # s) f = (case (parent f) of - Some pf \ (case (sectxt_of_obj (LinkHard p oldf f # s) (O_file f), sectxt_of_obj s (O_dir pf), - get_parentfs_ctxts s pf) of - (Some sec, Some psec, Some asecs) \ Some (Created, sec, Some psec, set asecs) - | _ \ None) - | None \ None)" -apply (frule vd_cons, frule vt_grant_os, frule noroot_events) -apply (case_tac f, simp) -apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps - get_parentfs_ctxts_simps) -apply (rule impI, (erule conjE)+) -apply (drule not_deleted_init_file, simp+) -apply (simp add:is_file_in_current) -done - -lemma cf2sfile_linkhard: - "\valid (LinkHard p oldf f # s); f' \ current_files (LinkHard p oldf f # s)\ - \ cf2sfile (LinkHard p oldf f # s) f' = ( - if (f' = f) - then (case (parent f) of - Some pf \ (case (sectxt_of_obj (LinkHard p oldf f # s) (O_file f), sectxt_of_obj s (O_dir pf), - get_parentfs_ctxts s pf) of - (Some sec, Some psec, Some asecs) \ Some (Created, sec, Some psec, set asecs) - | _ \ None) - | None \ None) - else cf2sfile s f')" -apply (case_tac "f = f'") -apply (simp add:cf2sfile_linkhard4 split:option.splits) -apply (simp add:cf2sfile_linkhard1 current_files_simps) -done - lemma cf2sfile_other: "\ff \ current_files s; \ p f flag fd opt. e \ Open p f flag fd opt; @@ -733,7 +678,6 @@ \ p f. e \ UnLink p f; \ p f. e \ Rmdir p f; \ p f i. e \ Mkdir p f i; - \ p f f'. e \ LinkHard p f f'; valid (e # s)\ \ cf2sfile (e # s) ff = cf2sfile s ff" apply (frule vd_cons, frule vt_grant_os) apply (induct ff, simp add:sroot_only) @@ -749,7 +693,6 @@ \ p f. e \ UnLink p f; \ p f. e \ Rmdir p f; \ p f i. e \ Mkdir p f i; - \ p f f'. e \ LinkHard p f f'; ff \ current_files s\ \ cf2sfile (e # s) ff = cf2sfile s ff" by (auto intro!:cf2sfile_other) @@ -785,15 +728,17 @@ apply (frule vd_cons, frule vt_grant_os) apply (simp add:current_files_simps split:if_splits option.splits) (* costs too much time, but solved *) - +(* apply (auto simp:cf2sfile_def sectxt_of_obj_simps get_parentfs_ctxts_simps is_file_simps is_dir_simps split:if_splits option.splits dest:init_file_dir_conflict pfdof_simp5 pfdof_simp6 file_of_pfd_is_file not_deleted_init_file not_deleted_init_dir is_file_not_dir is_dir_not_file dest!:current_has_sec') done +*) +sorry -lemmas cf2sfile_simps = cf2sfile_open cf2sfile_mkdir cf2sfile_linkhard cf2sfile_other +lemmas cf2sfile_simps = cf2sfile_open cf2sfile_mkdir cf2sfile_other cf2sfile_unlink cf2sfile_rmdir cf2sfile_closefd (*********** cfd2sfd simpset *********) @@ -871,8 +816,8 @@ done lemma cfd2sfd_clone: - "\valid (Clone p p' fds shms # s); file_of_proc_fd (Clone p p' fds shms # s) p'' fd' = Some f\ - \ cfd2sfd (Clone p p' fds shms # s) p'' fd' = ( + "\valid (Clone p p' fds # s); file_of_proc_fd (Clone p p' fds # s) p'' fd' = Some f\ + \ cfd2sfd (Clone p p' fds # s) p'' fd' = ( if (p'' = p') then cfd2sfd s p fd' else cfd2sfd s p'' fd')" apply (frule vd_cons, frule vt_grant_os) @@ -927,7 +872,7 @@ lemma cfd2sfd_other: "\valid (e # s); file_of_proc_fd (e # s) p' fd' = Some f'; \ p f flags fd opt. e \ Open p f flags fd opt; - \ p p'' fds shms. e \ Clone p p'' fds shms\ + \ p p'' fds. e \ Clone p p'' fds\ \ cfd2sfd (e # s) p' fd' = cfd2sfd s p' fd'" apply (frule vd_cons, frule vt_grant_os) apply (frule proc_fd_in_fds, simp) @@ -1057,8 +1002,8 @@ done lemma cpfd2sfds_clone: - "valid (Clone p p' fds shms # s) - \ cpfd2sfds (Clone p p' fds shms # s) = (cpfd2sfds s) (p' := {sfd. \ fd \ fds. \ f. file_of_proc_fd s p fd = Some f \ cfd2sfd s p fd = Some sfd})" + "valid (Clone p p' fds # s) + \ cpfd2sfds (Clone p p' fds # s) = (cpfd2sfds s) (p' := {sfd. \ fd \ fds. \ f. file_of_proc_fd s p fd = Some f \ cfd2sfd s p fd = Some sfd})" apply (frule vd_cons, frule vt_grant_os) apply (rule ext) apply (rule set_eqI, rule iffI) @@ -1087,7 +1032,7 @@ \ p p'. e \ Kill p p'; \ p. e \ Exit p; \ p fd. e \ CloseFd p fd; - \ p p' fds shms. e \ Clone p p' fds shms\ \ cpfd2sfds (e # s) = cpfd2sfds s" + \ p p' fds. e \ Clone p p' fds\ \ cpfd2sfds (e # s) = cpfd2sfds s" apply (frule vd_cons, frule vt_grant_os) apply (rule ext) unfolding cpfd2sfds_def proc_file_fds_def @@ -1185,265 +1130,13 @@ lemmas cpfd2sfds_simps = cpfd2sfds_open cpfd2sfds_execve cpfd2sfds_clone cpfd2sfds_kill cpfd2sfds_exit cpfd2sfds_closefd cpfd2sfds_other -(********* ch2sshm simpset ********) - -lemma ch2sshm_createshm: - "valid (CreateShM p h # s) - \ ch2sshm (CreateShM p h # s) = (ch2sshm s) (h := - (case (sectxt_of_obj (CreateShM p h # s) (O_shm h)) of - Some sec \ - Some (if (\ deleted (O_shm h) s \ h \ init_shms) then Init h else Created, sec) - | _ \ None))" -apply (frule vd_cons, frule vt_grant_os) -apply (auto simp:ch2sshm_def sectxt_of_obj_simps dest!:current_has_sec' split:option.splits if_splits) -done - -lemma ch2sshm_other: - "\valid (e # s); - \ p h. e \ CreateShM p h; - h' \ current_shms (e # s)\ \ ch2sshm (e # s) h' = ch2sshm s h'" -apply (frule vd_cons, frule vt_grant_os) -apply (case_tac e) -apply (auto simp:ch2sshm_def sectxt_of_obj_simps dest!:current_has_sec' split:option.splits if_splits) -done - -lemmas ch2sshm_simps = ch2sshm_createshm ch2sshm_other - -lemma current_shm_has_sh: - "\h \ current_shms s; valid s\ \ \ sh. ch2sshm s h = Some sh" -by (auto simp:ch2sshm_def split:option.splits dest!:current_has_sec') - -lemma current_shm_has_sh': - "\ch2sshm s h = None; valid s\ \ h \ current_shms s" -by (auto dest:current_shm_has_sh) - -(********** cph2spshs simpset **********) - -lemma cph2spshs_attach: - "valid (Attach p h flag # s) \ - cph2spshs (Attach p h flag # s) = (cph2spshs s) (p := - (case (ch2sshm s h) of - Some sh \ cph2spshs s p \ {(sh, flag)} - | _ \ cph2spshs s p) )" -apply (frule vd_cons, frule vt_grant_os, rule ext) -using ch2sshm_other[where e = "Attach p h flag" and s = s] -apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits -dest!:current_shm_has_sh' dest: procs_of_shm_prop1 simp:cph2spshs_def) -apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) -apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) -apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) -apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) -apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) -apply (case_tac "ha = h", simp, frule procs_of_shm_prop1, simp) -apply (rule_tac x = ha in exI, simp) -apply (rule_tac x = ha in exI, simp, drule procs_of_shm_prop1, simp, simp) -apply (rule_tac x = ha in exI, simp) -apply (frule procs_of_shm_prop1, simp, simp) -apply (rule impI, simp) -done - -lemma cph2spshs_detach: "valid (Detach p h # s) \ - cph2spshs (Detach p h # s) = (cph2spshs s) (p := - (case (ch2sshm s h, flag_of_proc_shm s p h) of - (Some sh, Some flag) \ if (\ h'. h' \ h \ (p,flag) \ procs_of_shm s h' \ ch2sshm s h' = Some sh) - then cph2spshs s p else cph2spshs s p - {(sh, flag)} - | _ \ cph2spshs s p) )" -apply (frule vd_cons, frule vt_grant_os, rule ext) -apply (case_tac "x = p") defer -using ch2sshm_other[where e = "Detach p h" and s = s] -apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits -dest!:current_shm_has_sh' procs_of_shm_prop4' dest:procs_of_shm_prop1 procs_of_shm_prop3 simp:cph2spshs_def) [1] -apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp, simp) -apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp, simp) -apply (rule impI, simp) - -apply (clarsimp) -apply (frule current_shm_has_sh, simp, erule exE) -apply (frule procs_of_shm_prop4, simp, simp) -apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+ - -apply (simp only:cph2spshs_def, rule set_eqI, rule iffI) -apply (erule CollectE | erule exE| erule conjE| rule CollectI)+ -apply (case_tac "ha = h", simp) -apply (rule_tac x = sha in exI, rule_tac x = flaga in exI, rule_tac x = ha in exI, simp) -using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1) - -apply (erule CollectE | erule exE| erule conjE| rule CollectI)+ -apply (case_tac "ha = h", simp) -apply (rule_tac x = h' in exI, simp) -apply (frule_tac flag = flag and flag' = flaga in procs_of_shm_prop3, simp+) -apply (frule_tac flag = flaga in procs_of_shm_prop4, simp+) -using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1) -apply (frule_tac h = h' in procs_of_shm_prop1, simp, simp) -apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+) -using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1) - -apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+ -apply (simp only:cph2spshs_def, rule set_eqI, rule iffI) -apply (erule CollectE | erule exE| erule conjE| rule DiffI |rule CollectI)+ -apply (simp split:if_splits) -apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+) -using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1) -apply (rule notI, simp split:if_splits) -apply (erule_tac x = ha in allE, simp, frule_tac h = ha in procs_of_shm_prop1, simp+) -using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1) -apply (erule CollectE | erule exE| erule conjE| erule DiffE |rule CollectI)+ -apply (simp split:if_splits) -apply (erule_tac x = ha in allE, simp, rule_tac x = ha in exI, simp) -apply (case_tac "ha = h", simp add:procs_of_shm_prop3, frule_tac h = ha in procs_of_shm_prop1, simp+) -using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1) -done - -lemma cph2spshs_deleteshm: - "valid (DeleteShM p h # s) \ - cph2spshs (DeleteShM p h # s) = (\ p'. - (case (ch2sshm s h, flag_of_proc_shm s p' h) of - (Some sh, Some flag) \ if (\ h'. h' \ h \ (p', flag) \ procs_of_shm s h' \ ch2sshm s h' = Some sh) - then cph2spshs s p' else cph2spshs s p' - {(sh, flag)} - | _ \ cph2spshs s p') )" -apply (frule vd_cons, frule vt_grant_os, rule ext) - -apply (clarsimp) -apply (frule current_shm_has_sh, simp, erule exE, simp, simp split:option.splits) -apply (rule conjI, rule impI) -using ch2sshm_other[where e = "DeleteShM p h" and s = s] -apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits -dest!:current_shm_has_sh' procs_of_shm_prop4' dest:procs_of_shm_prop1 procs_of_shm_prop3 simp:cph2spshs_def) [1] -apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp+) -apply (rule_tac x = ha in exI, simp) -apply (case_tac "ha = h", simp+, frule_tac h = ha in procs_of_shm_prop1, simp+) - -apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+ - -apply (simp only:cph2spshs_def, rule set_eqI, rule iffI) -apply (erule CollectE | erule exE| erule conjE| rule CollectI)+ -apply (case_tac "ha = h", simp) -apply (rule_tac x = sha in exI, rule_tac x = flag in exI, rule_tac x = ha in exI, simp) -using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1) - -apply (erule CollectE | erule exE| erule conjE| rule CollectI)+ -apply (case_tac "ha = h", simp) -apply (rule_tac x = h' in exI, simp) -apply (frule_tac flag = flag in procs_of_shm_prop4, simp+) -using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1) -apply (frule_tac h = h' in procs_of_shm_prop1, simp, simp) -apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+) -using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1) - -apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+ -apply (simp only:cph2spshs_def, rule set_eqI, rule iffI) -apply (erule CollectE | erule exE| erule conjE| rule DiffI |rule CollectI)+ -apply (simp split:if_splits) -apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+) -using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1) -apply (rule notI, simp split:if_splits) -apply (erule_tac x = ha in allE, simp, frule_tac h = ha in procs_of_shm_prop1, simp+) -using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1) -apply (erule CollectE | erule exE| erule conjE| erule DiffE |rule CollectI)+ -apply (simp split:if_splits) -apply (erule_tac x = ha in allE, simp, rule_tac x = ha in exI, simp) -apply (case_tac "ha = h", simp add:procs_of_shm_prop4) -apply (frule_tac h = ha in procs_of_shm_prop1, simp+) -using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1) -done - -lemma flag_of_proc_shm_prop1: - "\flag_of_proc_shm s p h = Some flag; valid s\ \ (p, flag) \ procs_of_shm s h" -apply (induct s arbitrary:p) -apply (simp add:init_shmflag_has_proc) -apply (frule vt_grant_os, frule vd_cons, case_tac a, auto split:if_splits option.splits) -done - -lemma cph2spshs_clone: - "valid (Clone p p' fds shms # s) \ - cph2spshs (Clone p p' fds shms # s) = (cph2spshs s) (p' := - {(sh, flag) | h sh flag. h \ shms \ ch2sshm s h = Some sh \ flag_of_proc_shm s p h = Some flag} )" -apply (frule vd_cons, frule vt_grant_os, rule ext) -using ch2sshm_other[where e = "Clone p p' fds shms" and s = s] -apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits -dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def) -apply (erule_tac x = h in allE, frule procs_of_shm_prop1, simp, simp add:procs_of_shm_prop4) -apply (rule_tac x = h in exI, simp, frule flag_of_proc_shm_prop1, simp+, simp add:procs_of_shm_prop1) -apply (rule_tac x = h in exI, simp, frule procs_of_shm_prop1, simp+)+ -done - -lemma cph2spshs_execve: - "valid (Execve p f fds # s) \ - cph2spshs (Execve p f fds # s) = (cph2spshs s) (p := {})" -apply (frule vd_cons, frule vt_grant_os, rule ext) -using ch2sshm_other[where e = "Execve p f fds" and s = s] -apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits -dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def) -apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+ -done - -lemma cph2spshs_kill: - "valid (Kill p p' # s) \ cph2spshs (Kill p p' # s) = (cph2spshs s) (p' := {})" -apply (frule vd_cons, frule vt_grant_os, rule ext) -using ch2sshm_other[where e = "Kill p p'" and s = s] -apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits -dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def) -apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+ -done - -lemma cph2spshs_exit: - "valid (Exit p # s) \ cph2spshs (Exit p # s) = (cph2spshs s) (p := {})" -apply (frule vd_cons, frule vt_grant_os, rule ext) -using ch2sshm_other[where e = "Exit p" and s = s] -apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits -dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def) -apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+ -done - -lemma cph2spshs_createshm: - "valid (CreateShM p h # s) \ cph2spshs (CreateShM p h # s) = cph2spshs s" -apply (frule vt_grant_os, frule vd_cons, rule ext) -apply (auto simp:cph2spshs_def) -using ch2sshm_createshm -apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits -dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def) -apply (rule_tac x = ha in exI, auto simp:procs_of_shm_prop1) -done - -lemma cph2spshs_other: - "\valid (e # s); - \ p h flag. e \ Attach p h flag; - \ p h. e \ Detach p h; - \ p h. e \ DeleteShM p h; - \ p p' fds shms. e \ Clone p p' fds shms; - \ p f fds. e \ Execve p f fds; - \ p p'. e \ Kill p p'; - \ p. e \ Exit p\ \ cph2spshs (e # s) = cph2spshs s" -apply (frule vt_grant_os, frule vd_cons, rule ext) -unfolding cph2spshs_def -apply (rule set_eqI, rule iffI) -apply (erule CollectE | erule conjE| erule exE| rule conjI| rule impI| rule CollectI)+ -apply (frule procs_of_shm_prop1, simp, rule_tac x= sh in exI, rule_tac x = flag in exI, rule_tac x = h in exI) -apply (case_tac e) -using ch2sshm_other[where e = e and s = s] -apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits -dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 - simp:ch2sshm_createshm) -apply (rule_tac x = h in exI, case_tac e) -using ch2sshm_other[where e = e and s = s] -apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits -dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 - simp:ch2sshm_createshm) -done - -lemmas cph2spshs_simps = cph2spshs_attach cph2spshs_detach cph2spshs_deleteshm cph2spshs_clone cph2spshs_execve - cph2spshs_exit cph2spshs_kill cph2spshs_other - -(******** cp2sproc simpset *********) - lemma cp2sproc_clone: - "valid (Clone p p' fds shms # s) \ cp2sproc (Clone p p' fds shms # s) = (cp2sproc s) (p' := + "valid (Clone p p' fds # s) \ cp2sproc (Clone p p' fds # s) = (cp2sproc s) (p' := case (sectxt_of_obj s (O_proc p)) of Some sec \ Some (Created, sec, - {sfd. \ fd \ fds. \ f. file_of_proc_fd s p fd = Some f \ cfd2sfd s p fd = Some sfd}, - {(sh, flag) | h sh flag. h \ shms \ ch2sshm s h = Some sh \ flag_of_proc_shm s p h = Some flag}) + {sfd. \ fd \ fds. \ f. file_of_proc_fd s p fd = Some f \ cfd2sfd s p fd = Some sfd}) | _ \ None)" -apply (frule cph2spshs_clone, frule cpfd2sfds_clone) +apply (frule cpfd2sfds_clone) apply (frule vd_cons, frule vt_grant_os, simp only:os_grant.simps) apply ((erule exE| erule conjE)+, frule not_init_intro_proc, simp, rule ext, case_tac "x = p'", simp) apply (auto simp:cp2sproc_def sectxt_of_obj_simps dest!:current_has_sec' @@ -1454,14 +1147,11 @@ "\valid (e # s); \ p f flags fd opt. e \ Open p f flags fd opt; \ p fd. e \ CloseFd p fd; - \ p h flag. e \ Attach p h flag; - \ p h. e \ Detach p h; - \ p h. e \ DeleteShM p h; - \ p p' fds shms. e \ Clone p p' fds shms; + \ p p' fds. e \ Clone p p' fds; \ p f fds. e \ Execve p f fds; \ p p'. e \ Kill p p'; \ p. e \ Exit p\ \ cp2sproc (e # s) = cp2sproc s" -apply (frule cph2spshs_other, simp+, frule cpfd2sfds_other, simp+) +apply (frule cpfd2sfds_other, simp+) apply (frule vt_grant_os, frule vd_cons, rule ext, case_tac e, simp_all) apply (auto simp:cp2sproc_def sectxt_of_obj_simps dest!:current_has_sec' dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits) @@ -1474,9 +1164,8 @@ cf2sfile (Open p f flags fd opt # s) f) of (Some sec, Some fdsec, Some sf) \ Some (if (\ deleted (O_proc p) s \ p \ init_procs) then Init p else Created, sec, - (cpfd2sfds s p) \ {(fdsec, flags, sf)}, cph2spshs s p) + (cpfd2sfds s p) \ {(fdsec, flags, sf)}) | _ \ None)" -apply (frule cph2spshs_other, simp, simp, simp, simp, simp, simp, simp) apply (frule cpfd2sfds_open, frule vt_grant_os, frule vd_cons, rule ext) apply (case_tac "x = p") apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps @@ -1494,29 +1183,11 @@ else (case (sectxt_of_obj s (O_proc p)) of Some sec \ Some (if (\ deleted (O_proc p) s \ p \ init_procs) then Init p else Created, - sec, cpfd2sfds s p - {sfd}, cph2spshs s p) + sec, cpfd2sfds s p - {sfd}) | _ \ None)) | _ \ cp2sproc s p) else cp2sproc s p)" -apply (frule cpfd2sfds_closefd, frule cph2spshs_other, simp, simp, simp, simp, simp, simp, simp) -apply (frule vt_grant_os, frule vd_cons, rule ext) -apply (case_tac "x = p") -apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps - dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current - dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits) -done - -lemma cp2sproc_attach: - "valid (Attach p h flag # s) \ - cp2sproc (Attach p h flag # s) = (cp2sproc s) (p := - (case (ch2sshm s h) of - Some sh \ (case (sectxt_of_obj s (O_proc p)) of - Some sec \ Some (if (\ deleted (O_proc p) s \ p \ init_procs) - then Init p else Created, - sec, cpfd2sfds s p, cph2spshs s p \ {(sh, flag)}) - | _ \ None) - | _ \ cp2sproc s p) )" -apply (frule cph2spshs_attach, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp) +apply (frule cpfd2sfds_closefd) apply (frule vt_grant_os, frule vd_cons, rule ext) apply (case_tac "x = p") apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps @@ -1524,53 +1195,14 @@ dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits) done -lemma cp2sproc_detach: - "valid (Detach p h # s) \ - cp2sproc (Detach p h # s) = (cp2sproc s) (p := - (case (ch2sshm s h, flag_of_proc_shm s p h) of - (Some sh, Some flag) \ if (\ h'. h' \ h \ (p,flag) \ procs_of_shm s h' \ ch2sshm s h' = Some sh) - then cp2sproc s p - else (case (sectxt_of_obj s (O_proc p)) of - Some sec \ Some (if (\ deleted (O_proc p) s \ p \ init_procs) - then Init p else Created, sec, - cpfd2sfds s p, cph2spshs s p - {(sh, flag)}) - | None \ None) - | _ \ cp2sproc s p) )" -apply (frule cph2spshs_detach, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp) -apply (frule vt_grant_os, frule vd_cons, rule ext) -apply (case_tac "x = p") -apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps - dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current - dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits) -done - -lemma cp2sproc_deleteshm: - "valid (DeleteShM p h # s) \ - cp2sproc (DeleteShM p h # s) = (\ p'. - (case (ch2sshm s h, flag_of_proc_shm s p' h) of - (Some sh, Some flag) \ if (\ h'. h' \ h \ (p', flag) \ procs_of_shm s h' \ ch2sshm s h' = Some sh) - then cp2sproc s p' - else (case (sectxt_of_obj s (O_proc p')) of - Some sec \ Some (if (\ deleted (O_proc p') s \ p' \ init_procs) - then Init p' else Created, sec, - cpfd2sfds s p', cph2spshs s p' - {(sh, flag)}) - | None \ None) - | _ \ cp2sproc s p') )" -apply (frule cph2spshs_deleteshm, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp) -apply (frule vt_grant_os, frule vd_cons, rule ext) -apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps - dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current - dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits) -done - lemma cp2sproc_execve: "valid (Execve p f fds # s) \ cp2sproc (Execve p f fds # s) = (cp2sproc s) (p := (case (sectxt_of_obj (Execve p f fds # s) (O_proc p)) of Some sec \ Some (if (\ deleted (O_proc p) s \ p \ init_procs) then Init p else Created, sec, - {sfd. \ fd \ fds. \ f. file_of_proc_fd s p fd = Some f \ cfd2sfd s p fd = Some sfd}, {}) + {sfd. \ fd \ fds. \ f. file_of_proc_fd s p fd = Some f \ cfd2sfd s p fd = Some sfd}) | _ \ None) )" -apply (frule cph2spshs_execve, frule cpfd2sfds_execve) +apply (frule cpfd2sfds_execve) apply (frule vt_grant_os, frule vd_cons, rule ext) apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current @@ -1580,7 +1212,7 @@ lemma cp2sproc_kill: "\valid (Kill p p' # s); p'' \ p'\ \ cp2sproc (Kill p p' # s) p'' = (cp2sproc s) p''" -apply (frule cph2spshs_kill, frule cpfd2sfds_kill) +apply (frule cpfd2sfds_kill) apply (frule vt_grant_os, frule vd_cons) apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current @@ -1595,7 +1227,7 @@ lemma cp2sproc_exit: "\valid (Exit p # s); p' \ p\ \ cp2sproc (Exit p # s) p' = (cp2sproc s) p'" -apply (frule cph2spshs_exit, frule cpfd2sfds_exit) +apply (frule cpfd2sfds_exit) apply (frule vt_grant_os, frule vd_cons) apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current @@ -1607,8 +1239,8 @@ cp2sproc (Exit p # s) p' = (cp2sproc s) p'" by (drule_tac p'= p' in cp2sproc_exit, simp+) -lemmas cp2sproc_simps = cp2sproc_open cp2sproc_closefd cp2sproc_attach cp2sproc_detach cp2sproc_deleteshm - cp2sproc_clone cp2sproc_execve cp2sproc_kill cp2sproc_exit cp2sproc_other +lemmas cp2sproc_simps = cp2sproc_open cp2sproc_closefd cp2sproc_clone cp2sproc_execve + cp2sproc_kill cp2sproc_exit cp2sproc_other lemma current_proc_has_sp: "\p \ current_procs s; valid s\ \ \ sp. cp2sproc s p = Some sp" @@ -1618,230 +1250,23 @@ "\cp2sproc s p = None; valid s\ \ p \ current_procs s" by (auto dest:current_proc_has_sp) - - -(* simpset for cf2sfiles *) - -lemma cf2sfiles_open: - "\valid (Open p f flag fd opt # s); f' \ current_files (Open p f flag fd opt # s)\ - \ cf2sfiles (Open p f flag fd opt # s) f' = ( - if (f' = f \ opt \ None) - then (case cf2sfile (Open p f flag fd opt # s) f of - Some sf \ {sf} - | _ \ {} ) - else cf2sfiles s f')" -apply (frule vt_grant_os, frule vd_cons) -apply (auto simp:cf2sfiles_def cf2sfile_open_none cf2sfile_simps same_inode_files_open - split:if_splits option.splits dest!:current_file_has_sfile' dest:cf2sfile_open) -apply (rule_tac x = "f'a" in bexI, drule same_inode_files_prop1, simp add:cf2sfile_open_some1, simp)+ -done - -lemma cf2sfiles_other: - "\valid (e # s); - \ p f flag fd opt. e \ Open p f flag fd opt; - \ p fd. e \ CloseFd p fd; - \ p f. e \ UnLink p f; - \ p f f'. e \ LinkHard p f f'\ \ cf2sfiles (e # s) = cf2sfiles s" -apply (frule vt_grant_os, frule vd_cons, rule ext) -apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI) -apply (drule Set.CollectD, erule bexE, rule CollectI) -apply (rule_tac x = f' in bexI, case_tac e) -apply (auto simp:cf2sfiles_def cf2sfile_simps same_inode_files_simps current_files_simps - split:if_splits option.splits dest!:current_file_has_sfile' dest:same_inode_files_prop1 cf2sfile_other') -apply (drule_tac f' = f' in cf2sfile_rmdir) -apply (simp add:current_files_simps same_inode_files_prop1 same_inode_files_prop3 dir_is_empty_def)+ - -apply (rule_tac x = f' in bexI, case_tac e) -apply (auto simp:cf2sfiles_def cf2sfile_simps same_inode_files_simps current_files_simps - split:if_splits option.splits dest!:current_file_has_sfile' dest:same_inode_files_prop1 cf2sfile_other') -apply (drule_tac f' = f' in cf2sfile_rmdir) -apply (simp add:current_files_simps same_inode_files_prop1 same_inode_files_prop3 dir_is_empty_def)+ -done - -lemma cf2sfile_linkhard1': - "\valid (LinkHard p oldf f # s); f' \ same_inode_files s f''\ - \ cf2sfile (LinkHard p oldf f# s) f' = cf2sfile s f'" -apply (drule same_inode_files_prop1) -by (simp add:cf2sfile_linkhard1) - -lemma cf2sfiles_linkhard: - "valid (LinkHard p oldf f # s) \ cf2sfiles (LinkHard p oldf f # s) = (\ f'. - if (f' = f \ f' \ same_inode_files s oldf) - then (case (cf2sfile (LinkHard p oldf f # s) f) of - Some sf \ cf2sfiles s oldf \ {sf} - | _ \ {}) - else cf2sfiles s f')" -apply (frule vt_grant_os, frule vd_cons, rule ext) -apply (auto simp:cf2sfiles_def cf2sfile_linkhard1' same_inode_files_linkhard current_files_linkhard - split:if_splits option.splits dest!:current_file_has_sfile' current_has_sec' dest:same_inode_files_prop1) -done - -lemma cf2sfile_unlink': - "\valid (UnLink p f # s); f' \ same_inode_files (UnLink p f # s) f''\ - \ cf2sfile (UnLink p f # s) f' = cf2sfile s f'" -apply (drule same_inode_files_prop1) -by (simp add:cf2sfile_unlink) - -lemma cf2sfiles_unlink: - "\valid (UnLink p f # s); f' \ current_files (UnLink p f # s)\ \ cf2sfiles (UnLink p f # s) f' = ( - if (f' \ same_inode_files s f \ proc_fd_of_file s f = {} \ - (\ f'' \ same_inode_files s f. f'' \ f \ cf2sfile s f'' \ cf2sfile s f)) then - (case (cf2sfile s f) of - Some sf \ cf2sfiles s f' - {sf} - | _ \ {}) - else cf2sfiles s f')" -apply (frule vt_grant_os, frule vd_cons, simp add:current_files_simps split:if_splits) -apply (rule conjI, clarify, frule is_file_has_sfile', simp, erule exE, simp) -apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) -apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink) -apply (simp add:current_files_unlink, simp, erule conjE) -apply (erule_tac x = f'a in ballE, frule_tac f' = "f'a" in cf2sfile_unlink) -apply (simp add:current_files_unlink same_inode_files_prop1, simp) -apply (rule_tac x = f'a in bexI, simp, simp) -apply (drule_tac f = f and f' = f' and f'' = f'a in same_inode_files_prop4, simp+) -apply (erule conjE|erule exE|erule bexE)+ -apply (case_tac "f'a = f", simp) -apply (frule_tac f' = f' in same_inode_files_unlink, simp add:current_files_unlink) -apply (frule_tac f' = f'a in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1) -apply (rule_tac x = f'a in bexI, simp, simp) - -apply (rule impI)+ -apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) -apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink) -apply (simp add:current_files_unlink, simp, (erule conjE)+) -apply (rule_tac x = f'a in bexI, frule_tac f' = "f'a" in cf2sfile_unlink) -apply (simp add:current_files_unlink same_inode_files_prop1, simp, simp) -apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_unlink) -apply (simp add:current_files_unlink, simp) -apply (case_tac "f'a = f", simp) -apply (frule_tac f = f' and f' = f in same_inode_files_prop5, simp) -apply (erule bexE, erule conjE) -apply (rule_tac x = f'' in bexI) -apply (drule_tac f' = f'' in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1) -apply (simp, simp, erule same_inode_files_prop4, simp) -apply (rule_tac x = f'a in bexI) -apply (drule_tac f' = f'a in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1) -apply (simp, simp) - - -apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) -apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink) -apply (simp add:current_files_unlink, simp) -apply (rule_tac x = f'a in bexI) -apply (frule_tac f' = f'a in cf2sfile_unlink) -apply (simp add:same_inode_files_prop1 current_files_unlink, simp, simp) -apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_unlink) -apply (simp add:current_files_unlink, simp) -apply (rule_tac x = f'a in bexI) -apply (frule_tac f' = f'a in cf2sfile_unlink) -apply (simp add:same_inode_files_prop1 current_files_unlink, simp, simp) -done - -lemma cf2sfiles_closefd: - "\valid (CloseFd p fd # s); f' \ current_files (CloseFd p fd # s)\ \ cf2sfiles (CloseFd p fd # s) f' = ( - case (file_of_proc_fd s p fd) of - Some f \ if (f' \ same_inode_files s f \ proc_fd_of_file s f = {(p, fd)} \ f \ files_hung_by_del s \ - (\ f'' \ same_inode_files s f. f'' \ f \ cf2sfile s f'' \ cf2sfile s f)) - then (case (cf2sfile s f) of - Some sf \ cf2sfiles s f' - {sf} - | _ \ {}) - else cf2sfiles s f' - | _ \ cf2sfiles s f')" - -apply (frule vt_grant_os, frule vd_cons, case_tac "file_of_proc_fd s p fd") -apply (simp_all add:current_files_simps split:if_splits) - -apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) -apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd) -apply (simp add:current_files_closefd, simp) -apply (rule_tac x = f'a in bexI) -apply (frule_tac f = f'a in cf2sfile_closefd) -apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) -apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd) -apply (simp add:current_files_closefd, simp) -apply (rule_tac x = f'a in bexI) -apply (frule_tac f = f'a in cf2sfile_closefd) -apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) - -apply (rule conjI, clarify, frule file_of_pfd_is_file, simp) -apply (frule is_file_has_sfile', simp, erule exE, simp) -apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) -apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd) -apply (simp add:current_files_closefd, simp, erule conjE) -apply (erule_tac x = f'a in ballE, frule_tac f = "f'a" in cf2sfile_closefd) -apply (simp add:current_files_closefd same_inode_files_prop1, simp) -apply (rule_tac x = f'a in bexI, simp, simp) -apply (drule_tac f = a and f' = f' and f'' = f'a in same_inode_files_prop4, simp+) -apply (erule conjE|erule exE|erule bexE)+ -apply (case_tac "f'a = a", simp) -apply (frule_tac f' = f' in same_inode_files_closefd, simp add:current_files_closefd, simp) -apply (frule_tac f = f'a in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1) -apply (rule_tac x = f'a in bexI, simp, simp) - -apply (rule impI)+ -apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) -apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd) -apply (simp add:current_files_closefd, simp, (erule conjE)+) -apply (rule_tac x = f'a in bexI, frule_tac f = f'a in cf2sfile_closefd) -apply (simp add:current_files_closefd same_inode_files_prop1, simp, simp) -apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd) -apply (simp add:current_files_closefd, simp) -apply (case_tac "f'a = a", simp) -apply (frule_tac f = f' and f' = a in same_inode_files_prop5, simp) -apply (erule bexE, erule conjE) -apply (rule_tac x = f'' in bexI) -apply (drule_tac f = f'' in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1) -apply (simp, simp, erule same_inode_files_prop4, simp) -apply (rule_tac x = f'a in bexI) -apply (drule_tac f = f'a in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1) -apply (simp, simp) - -apply (rule conjI, clarify) - -apply (rule impI) -apply (case_tac "a \ files_hung_by_del s", simp_all) -apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) -apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd) -apply (simp add:current_files_closefd, simp) -apply (rule_tac x = f'a in bexI) -apply (frule_tac f = f'a in cf2sfile_closefd) -apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) -apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd) -apply (simp add:current_files_closefd, simp) -apply (rule_tac x = f'a in bexI) -apply (frule_tac f = f'a in cf2sfile_closefd) -apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) -apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) -apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd) -apply (simp add:current_files_closefd, simp) -apply (rule_tac x = f'a in bexI) -apply (frule_tac f = f'a in cf2sfile_closefd) -apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) -apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd) -apply (simp add:current_files_closefd, simp) -apply (rule_tac x = f'a in bexI) -apply (frule_tac f = f'a in cf2sfile_closefd) -apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) -done - -lemmas cf2sfiles_simps = cf2sfiles_open cf2sfiles_linkhard cf2sfiles_other - cf2sfiles_unlink cf2sfiles_closefd - - (* simpset for co2sobj *) lemma co2sobj_execve: "\valid (Execve p f fds # s); alive (Execve p f fds # s) obj\ \ co2sobj (Execve p f fds # s) obj = ( if (obj = O_proc p) then (case (cp2sproc (Execve p f fds # s) p) of - Some sp \ Some (S_proc sp (O_proc p \ Tainted s \ O_file f \ Tainted s)) + Some sp \ Some (S_proc sp (O_proc p \ tainted s \ O_file f \ tainted s)) | _ \ None) else co2sobj s obj )" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps cf2sfiles_other ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps cq2smsgq_other cf2sfile_other) apply (case_tac "cp2sproc (Execve p f fds # s) p") apply (drule current_proc_has_sp', simp, simp) -apply (simp (no_asm_simp) add:cp2sproc_execve tainted_eq_Tainted split:option.splits) +apply (simp (no_asm_simp) add:cp2sproc_execve split:option.splits) +apply (simp add:is_file_simps, frule_tac s = s in is_file_has_sfile', simp, erule exE, simp) +apply (frule_tac ff = list in cf2sfile_other', simp_all) +apply (simp add:is_file_in_current) apply (simp add:is_dir_simps, frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp) apply (frule_tac ff = list in cf2sfile_other', simp_all) apply (simp add:is_dir_in_current) @@ -1851,60 +1276,66 @@ "\valid (Execve p f fds # s); alive s obj\ \ co2sobj (Execve p f fds # s) obj = ( if (obj = O_proc p) then (case (cp2sproc (Execve p f fds # s) p) of - Some sp \ Some (S_proc sp (O_proc p \ Tainted s \ O_file f \ Tainted s)) + Some sp \ Some (S_proc sp (O_proc p \ tainted s \ O_file f \ tainted s)) | _ \ None) else co2sobj s obj )" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps cf2sfiles_other ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps cf2sfile_other cq2smsgq_other) apply (case_tac "cp2sproc (Execve p f fds # s) p") apply (drule current_proc_has_sp', simp, simp) -apply (simp (no_asm_simp) add:cp2sproc_execve tainted_eq_Tainted split:option.splits) +apply (simp (no_asm_simp) add:cp2sproc_execve split:option.splits) +apply (frule_tac s = s in is_file_has_sfile', simp, erule exE, simp) +apply (frule_tac ff = list in cf2sfile_other', simp_all) +apply (simp add:is_file_in_current) apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp) apply (frule_tac ff = list in cf2sfile_other', simp_all) apply (simp add:is_dir_in_current) done lemma co2sobj_clone': - "\valid (Clone p p' fds shms # s); alive s obj\ \ co2sobj (Clone p p' fds shms # s) obj = ( + "\valid (Clone p p' fds # s); alive s obj\ \ co2sobj (Clone p p' fds # s) obj = ( if (obj = O_proc p') - then (case (cp2sproc (Clone p p' fds shms # s) p') of - Some sp \ Some (S_proc sp (O_proc p \ Tainted s)) + then (case (cp2sproc (Clone p p' fds # s) p') of + Some sp \ Some (S_proc sp (O_proc p \ tainted s)) | _ \ None) else co2sobj s obj )" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps cf2sfiles_other ch2sshm_other cq2smsgq_other tainted_eq_Tainted) -apply (case_tac "cp2sproc (Clone p p' fds shms # s) p'") +apply (simp_all add:current_files_simps cf2sfile_other cq2smsgq_other ) +apply (case_tac "cp2sproc (Clone p p' fds # s) p'") apply (drule current_proc_has_sp', simp, simp) apply ((erule conjE)+, frule_tac p = p in current_proc_has_sec, simp, erule exE, simp) -apply (rule conjI, rule impI, simp) -apply (simp (no_asm_simp) add:cp2sproc_clone tainted_eq_Tainted split:option.splits) - +apply (simp (no_asm_simp) add:cp2sproc_clone split:option.splits) +apply (case_tac "nat = p'", simp, simp) +apply (frule_tac s = s in is_file_has_sfile', simp, erule exE, simp) +apply (frule_tac ff = list in cf2sfile_other', simp_all) +apply (simp add:is_file_in_current) apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp) apply (frule_tac ff = list in cf2sfile_other', simp_all) apply (simp add:is_dir_in_current) done lemma co2sobj_clone: - "\valid (Clone p p' fds shms # s); alive (Clone p p' fds shms # s) obj\ - \ co2sobj (Clone p p' fds shms # s) obj = ( + "\valid (Clone p p' fds # s); alive (Clone p p' fds # s) obj\ + \ co2sobj (Clone p p' fds # s) obj = ( if (obj = O_proc p') - then (case (cp2sproc (Clone p p' fds shms # s) p') of - Some sp \ Some (S_proc sp (O_proc p \ Tainted s)) + then (case (cp2sproc (Clone p p' fds # s) p') of + Some sp \ Some (S_proc sp (O_proc p \ tainted s)) | _ \ None) else co2sobj s obj )" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps cf2sfiles_other ch2sshm_other - cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps cf2sfile_other cq2smsgq_other) apply (rule conjI, rule impI, simp) -apply (case_tac "cp2sproc (Clone p p' fds shms # s) p'") +apply (case_tac "cp2sproc (Clone p p' fds # s) p'") apply (drule current_proc_has_sp', simp, simp) apply ((erule conjE)+, frule_tac p = p in current_proc_has_sec, simp, erule exE, simp) -apply (simp add:tainted_eq_Tainted, rule impI, rule notI) -apply (drule Tainted_in_current, simp+) +apply (rule impI,rule notI, drule tainted_in_current, simp+) apply (rule impI, simp) apply (drule current_proc_has_sp, simp, (erule exE|erule conjE)+) -apply (simp (no_asm_simp) add:cp2sproc_clone tainted_eq_Tainted Tainted_in_current split:option.splits) +apply (simp (no_asm_simp) add:cp2sproc_clone tainted_in_current split:option.splits) +apply (simp add:is_file_simps, frule_tac s = s in is_file_has_sfile', simp, erule exE, simp) +apply (frule_tac ff = list in cf2sfile_other', simp_all) +apply (simp add:is_file_in_current) apply (simp add:is_dir_simps, frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp) apply (frule_tac ff = list in cf2sfile_other', simp_all) apply (simp add:is_dir_in_current) @@ -1913,22 +1344,25 @@ lemma co2sobj_ptrace: "\valid (Ptrace p p' # s); alive s obj\\ co2sobj (Ptrace p p' # s) obj = ( case obj of - O_proc p'' \ if (info_flow_shm s p' p'') + O_proc p'' \ if (p'' = p') then (case (cp2sproc s p'') of - Some sp \ Some (S_proc sp (O_proc p'' \ Tainted s \ O_proc p \ Tainted s)) + Some sp \ Some (S_proc sp (O_proc p'' \ tainted s \ O_proc p \ tainted s)) | _ \ None) - else if (info_flow_shm s p p'') + else if (p'' = p) then (case (cp2sproc s p'') of - Some sp \ Some (S_proc sp (O_proc p'' \ Tainted s \ O_proc p' \ Tainted s)) + Some sp \ Some (S_proc sp (O_proc p'' \ tainted s \ O_proc p' \ tainted s)) | _ \ None) else co2sobj s (O_proc p'') | _ \ co2sobj s obj)" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other cf2sfiles_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps cq2smsgq_other cf2sfile_other) -apply (auto simp:cp2sproc_other tainted_eq_Tainted split:option.splits - dest!:current_proc_has_sec' current_proc_has_sp' intro:info_flow_shm_Tainted)[1] +apply (auto simp:cp2sproc_other split:option.splits + dest!:current_proc_has_sec' current_proc_has_sp')[1] +apply (frule_tac s = s in is_file_has_sfile', simp, erule exE, simp) +apply (frule_tac ff = list in cf2sfile_other', simp_all) +apply (simp add:is_file_in_current) apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp) apply (frule_tac ff = list in cf2sfile_other', simp_all) apply (simp add:is_dir_in_current) @@ -1939,29 +1373,27 @@ \ co2sobj (Open p f flag fd opt # s) obj = (case obj of O_file f' \ if (f' = f \ opt \ None) then (case (cf2sfile (Open p f flag fd opt # s) f) of - Some sf \ Some (S_file {sf} (O_proc p \ Tainted s)) + Some sf \ Some (S_file sf (O_proc p \ tainted s)) | _ \ None) else co2sobj s (O_file f') | O_proc p' \ if (p' = p) then (case (cp2sproc (Open p f flag fd opt # s) p) of - Some sp \ Some (S_proc sp (O_proc p \ Tainted s)) + Some sp \ Some (S_proc sp (O_proc p \ tainted s)) | _ \ None) else co2sobj s (O_proc p') | _ \ co2sobj s obj )" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits - dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1] +apply (auto simp:cp2sproc_simps split:option.splits dest!:current_proc_has_sp')[1] apply (simp split:if_splits t_object.splits) apply (rule conjI, rule impI, erule conjE, erule exE, simp, (erule exE|erule conjE)+) apply (case_tac "cf2sfile (Open p f flag fd (Some y) # s) f") apply (drule current_file_has_sfile', simp, simp add:current_files_simps, simp) -apply (frule_tac f' = f in cf2sfiles_open, simp add:current_files_simps) -apply (simp add:tainted_eq_Tainted) -apply (rule impI, rule notI, drule Tainted_in_current, simp, simp add:is_file_in_current) -apply (rule impI, simp add:tainted_eq_Tainted cf2sfiles_open is_file_in_current split:option.splits) +apply (frule_tac f' = f in cf2sfile_open, simp add:current_files_simps) +apply (rule impI, rule notI, drule tainted_in_current, simp, simp add:is_file_in_current) +apply (rule impI, simp add:cf2sfile_open is_file_in_current split:option.splits) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other) apply (frule is_dir_in_current) apply (frule_tac f' = list in cf2sfile_open) @@ -1973,7 +1405,7 @@ "\valid (ReadFile p fd # s); alive s obj\ \ co2sobj (ReadFile p fd # s) obj = ( case obj of O_proc p' \ (case (file_of_proc_fd s p fd) of - Some f \ (if (info_flow_shm s p p' \ O_file f \ Tainted s) + Some f \ (if (p' = p \ O_file f \ tainted s) then (case (cp2sproc s p') of Some sp \ Some (S_proc sp True) | _ \ None) @@ -1981,12 +1413,11 @@ | _ \ None) | _ \ co2sobj s obj)" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits - dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1] -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (auto simp:cp2sproc_simps split:option.splits dest!:current_proc_has_sp')[1] +apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other) apply (auto split:if_splits option.splits dest!:current_file_has_sfile' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps + simp:current_files_simps cf2sfile_simps dest:is_file_in_current is_dir_in_current) done @@ -1994,54 +1425,43 @@ "\valid (WriteFile p fd # s); alive s obj\ \ co2sobj (WriteFile p fd # s) obj = ( case obj of O_file f' \ (case (file_of_proc_fd s p fd) of - Some f \ (if (f \ same_inode_files s f') - then Some (S_file (cf2sfiles s f') - (O_file f' \ Tainted s \ O_proc p \ Tainted s)) + Some f \ (if (f' = f) + then (case cf2sfile s f of + Some sf \ Some (S_file sf (O_file f \ tainted s \ O_proc p \ tainted s)) + | _ \ None) else co2sobj s obj) | _ \ None) | _ \ co2sobj s obj)" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other) apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted - same_inode_files_prop6 + simp:current_files_simps cf2sfile_simps cp2sproc_simps dest:is_file_in_current is_dir_in_current) - -(* should delete has_same_inode !?!?*) -by (auto simp:same_inode_files_def is_file_def has_same_inode_def split:if_splits option.splits) +done lemma co2sobj_closefd: "\valid (CloseFd p fd # s); alive (CloseFd p fd # s) obj\ \ co2sobj (CloseFd p fd # s) obj = ( case obj of - O_file f' \ (case (file_of_proc_fd s p fd) of - Some f \ (if (f' \ same_inode_files s f \ proc_fd_of_file s f = {(p, fd)} \ - f \ files_hung_by_del s \ (\ f'' \ same_inode_files s f. - f'' \ f \ cf2sfile s f'' \ cf2sfile s f)) - then (case (cf2sfile s f, co2sobj s (O_file f')) of - (Some sf, Some (S_file sfs b)) \ Some (S_file (sfs - {sf}) b) - | _ \ None) - else co2sobj s obj) - | _ \ co2sobj s obj) - | O_proc p' \ (if (p = p') + O_proc p' \ (if (p = p') then (case (cp2sproc (CloseFd p fd # s) p) of - Some sp \ Some (S_proc sp (O_proc p \ Tainted s)) + Some sp \ Some (S_proc sp (O_proc p \ tainted s)) | _ \ None) else co2sobj s obj) | _ \ co2sobj s obj) " apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) -apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits if_splits - dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1] +apply (simp_all add:current_files_simps cq2smsgq_other) +apply (auto simp:cp2sproc_simps split:option.splits if_splits dest!:current_proc_has_sp')[1] apply (frule is_file_in_current) apply (case_tac "file_of_proc_fd s p fd") -apply (simp add:tainted_eq_Tainted) -apply (drule_tac f' = list in cf2sfiles_closefd, simp add:current_files_closefd, simp) -apply (frule_tac f' = list in cf2sfiles_closefd, simp) -apply (simp add:is_file_simps current_files_simps) -apply (auto simp add:tainted_eq_Tainted cf2sfile_closefd split:if_splits option.splits - dest!:current_file_has_sfile' dest:hung_file_in_current)[1] +apply (simp add:current_files_simps is_file_simps tainted.simps) +apply (drule_tac f = list in cf2sfile_closefd, simp add:current_files_simps) +apply (simp split:option.splits) +apply (frule_tac f = list in cf2sfile_closefd, simp) +apply (simp split:option.splits) +apply (simp add:is_file_simps current_files_simps split:if_splits) +apply (rule impI, simp) apply (simp add:is_dir_simps, frule is_dir_in_current) apply (frule_tac f = list in cf2sfile_closefd) @@ -2050,25 +1470,14 @@ done lemma co2sobj_unlink: - "\valid (UnLink p f # s); alive (UnLink p f # s) obj\ \ co2sobj (UnLink p f # s) obj = ( - case obj of - O_file f' \ if (f' \ same_inode_files s f \ proc_fd_of_file s f = {} \ - (\ f'' \ same_inode_files s f. f'' \ f \ cf2sfile s f'' \ cf2sfile s f)) - then (case (cf2sfile s f, co2sobj s (O_file f')) of - (Some sf, Some (S_file sfs b)) \ Some (S_file (sfs - {sf}) b) - | _ \ None) - else co2sobj s obj - | _ \ co2sobj s obj)" + "\valid (UnLink p f # s); alive (UnLink p f # s) obj\ + \ co2sobj (UnLink p f # s) obj = co2sobj s obj" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) -apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits if_splits - dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1] +apply (simp_all add:current_files_simps cq2smsgq_other) +apply (auto simp:cp2sproc_simps split:option.splits if_splits dest!:current_proc_has_sp')[1] apply (frule is_file_in_current) apply (frule_tac f' = list in cf2sfile_unlink, simp) -apply (frule_tac f' = list in cf2sfiles_unlink, simp) -apply (simp add:is_file_simps current_files_simps) -apply (auto simp add:tainted_eq_Tainted is_file_in_current split:if_splits option.splits - dest!:current_file_has_sfile')[1] +apply (simp add:is_file_simps current_files_simps split:option.splits if_splits) apply (simp add:is_dir_simps, frule is_dir_in_current) apply (frule_tac f' = list in cf2sfile_unlink) @@ -2079,13 +1488,13 @@ lemma co2sobj_rmdir: "\valid (Rmdir p f # s); alive (Rmdir p f # s) obj\ \ co2sobj (Rmdir p f # s) obj = co2sobj s obj" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) -apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits if_splits - dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1] +apply (simp_all add:current_files_simps cq2smsgq_other) +apply (auto simp:cp2sproc_simps split:option.splits if_splits dest!:current_proc_has_sp')[1] apply (simp add:is_file_simps dir_is_empty_def) apply (case_tac "f = list", drule file_dir_conflict, simp, simp) -apply (simp add:cf2sfiles_other) -apply (auto simp:cf2sfile_simps dest:is_dir_in_current) +apply (frule_tac f' = list in cf2sfile_rmdir, simp add:is_file_in_current current_files_simps) +apply (simp split:option.splits) +apply (auto simp:cf2sfile_simps cf2sfile_other dest:is_dir_in_current) done lemma co2sobj_mkdir: @@ -2096,75 +1505,50 @@ | _ \ None) else co2sobj s obj)" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) -apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits if_splits - dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1] -apply (frule_tac cf2sfiles_other, simp+) +apply (simp_all add:current_files_simps cq2smsgq_other) +apply (auto simp:cp2sproc_simps split:option.splits if_splits dest!:current_proc_has_sp')[1] +apply (case_tac "f = list", simp add:is_file_simps is_file_in_current) +apply (frule_tac f' = list in cf2sfile_mkdir) +apply (simp add:is_file_in_current) +apply (simp split:option.splits) + apply (frule is_dir_in_current, rule impI, simp add:current_files_mkdir) apply (frule current_file_has_sfile, simp, erule exE, simp) apply (drule cf2sfile_mkdir1, simp+) done - -lemma co2sobj_linkhard: - "\valid (LinkHard p oldf f # s); alive (LinkHard p oldf f # s) obj\ - \ co2sobj (LinkHard p oldf f # s) obj = ( - case obj of - O_file f' \ if (f' = f \ f' \ same_inode_files s oldf) - then (case (cf2sfile (LinkHard p oldf f # s) f) of - Some sf \ Some (S_file (cf2sfiles s oldf \ {sf}) (O_file oldf \ Tainted s)) - | _ \ None) - else co2sobj s obj - | _ \ co2sobj s obj)" -apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) -apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits if_splits - dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1] -apply (frule_tac cf2sfiles_linkhard, simp+) -apply (frule_tac f' = f in cf2sfile_linkhard, simp add:current_files_linkhard) -apply (auto simp:is_file_simps sectxt_of_obj_simps current_files_simps is_file_in_current same_inodes_Tainted - split:option.splits if_splits dest:Tainted_in_current - dest!:current_has_sec' current_file_has_sfile')[1] - -apply (frule is_dir_in_current, simp add:current_files_linkhard is_dir_simps is_dir_in_current) -apply (frule is_dir_in_current) -apply (frule current_file_has_sfile, simp) -apply (drule cf2sfile_linkhard1, simp+) -done - lemma co2sobj_truncate: "\valid (Truncate p f len # s); alive s obj\ \ co2sobj (Truncate p f len # s) obj = ( case obj of - O_file f' \ if (f' \ same_inode_files s f \ len > 0) - then Some (S_file (cf2sfiles s f') (O_file f' \ Tainted s \ O_proc p \ Tainted s)) + O_file f' \ if (f' = f \ len > 0) + then (case cf2sfile s f of + Some sf \ Some (S_file sf (O_file f' \ tainted s \ O_proc p \ tainted s)) + | _ \ None) else co2sobj s obj | _ \ co2sobj s obj)" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other) apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted - same_inode_files_prop6 + simp:current_files_simps cf2sfile_simps cp2sproc_simps dest:is_file_in_current is_dir_in_current) done lemma co2sobj_kill: "\valid (Kill p p' # s); alive (Kill p p' # s) obj\ \ co2sobj (Kill p p' # s) obj = co2sobj s obj" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other) apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted - same_inode_files_prop6 + simp:current_files_simps cf2sfile_simps cp2sproc_simps dest:is_file_in_current is_dir_in_current) done lemma co2sobj_exit: "\valid (Exit p # s); alive (Exit p # s) obj\ \ co2sobj (Exit p # s) obj = co2sobj s obj" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other ) apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted - same_inode_files_prop6 + simp:current_files_simps cf2sfile_simps cp2sproc_simps dest:is_file_in_current is_dir_in_current) done @@ -2177,10 +1561,9 @@ else co2sobj s obj | _ \ co2sobj s obj)" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other ) apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted - same_inode_files_prop6 + simp:current_files_simps cf2sfile_simps cp2sproc_simps dest!:current_has_smsgq' dest:is_file_in_current is_dir_in_current cq2smsg_createmsgq) done @@ -2194,10 +1577,9 @@ else co2sobj s obj | _ \ co2sobj s obj)" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other ) apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted - same_inode_files_prop6 + simp:current_files_simps cf2sfile_simps cp2sproc_simps dest!:current_has_smsgq' dest:is_file_in_current is_dir_in_current cq2smsg_sendmsg) done @@ -2209,17 +1591,16 @@ Some sq \ Some (S_msgq sq) | _ \ None) else co2sobj s obj - | O_proc p' \ if (info_flow_shm s p p' \ O_msg q m \ Tainted s) + | O_proc p' \ if (p' = p \ O_msg q m \ tainted s) then (case (cp2sproc s p') of Some sp \ Some (S_proc sp True) | _ \ None) else co2sobj s obj | _ \ co2sobj s obj)" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other) apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted - same_inode_files_prop6 + simp:current_files_simps cf2sfile_simps cp2sproc_simps dest!:current_has_smsgq' dest:is_file_in_current is_dir_in_current cq2smsg_recvmsg) done @@ -2228,118 +1609,20 @@ "\valid (RemoveMsgq p q # s); alive (RemoveMsgq p q # s) obj\ \ co2sobj (RemoveMsgq p q # s) obj = co2sobj s obj" apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) +apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other) apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted - same_inode_files_prop6 + simp:current_files_simps cf2sfile_simps cp2sproc_simps dest!:current_has_smsgq' dest:is_file_in_current is_dir_in_current cq2smsg_removemsgq) done -lemma co2sobj_createshm: - "\valid (CreateShM p h # s); alive (CreateShM p h # s) obj\ \ co2sobj (CreateShM p h # s) obj = ( - case obj of - O_shm h' \ if (h' = h) then (case (ch2sshm (CreateShM p h # s) h) of - Some sh \ Some (S_shm sh) - | _ \ None) - else co2sobj s obj - | _ \ co2sobj s obj)" -apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) -apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted - same_inode_files_prop6 ch2sshm_simps - dest!:current_shm_has_sh' - dest:is_file_in_current is_dir_in_current) -done - -lemma co2sobj_detach: - "\valid (Detach p h # s); alive s obj\ \ co2sobj (Detach p h # s) obj = ( - case obj of - O_proc p' \ if (p' = p) then (case (cp2sproc (Detach p h # s) p) of - Some sp \ Some (S_proc sp (O_proc p \ Tainted s)) - | _ \ None) - else co2sobj s obj - | _ \ co2sobj s obj)" -apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) - -apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted - same_inode_files_prop6 ch2sshm_simps - dest!:current_shm_has_sh' - dest:is_file_in_current is_dir_in_current) -done - -lemma co2sobj_deleteshm: - "\valid (DeleteShM p h # s); alive (DeleteShM p h # s) obj\ \ co2sobj (DeleteShM p h # s) obj = ( - case obj of - O_proc p' \ (case (cp2sproc (DeleteShM p h # s) p') of - Some sp \ Some (S_proc sp (O_proc p' \ Tainted s)) - | _ \ None) - | _ \ co2sobj s obj)" -apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) - -apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps tainted_eq_Tainted - same_inode_files_prop6 ch2sshm_simps - dest!:current_shm_has_sh' - dest:is_file_in_current is_dir_in_current) -done declare Product_Type.split_paired_Ex Product_Type.split_paired_All [simp del] -lemma info_flow_shm_prop1: - "p \ current_procs s \ info_flow_shm s p p" -by (simp add:info_flow_shm_def) - -lemma co2sobj_attach: - "\valid (Attach p h flag # s); alive s obj\ \ co2sobj (Attach p h flag # s) obj = ( - case obj of - O_proc p' \ if (info_flow_shm s p p') - then (case (cp2sproc (Attach p h flag # s) p') of - Some sp \ Some (S_proc sp (O_proc p' \ Tainted s \ - (\ p''. O_proc p'' \ Tainted s \ (p'', SHM_RDWR) \ procs_of_shm s h))) - | _ \ None) - else if (\ p'' flag'. (p'', flag') \ procs_of_shm s h \ flag = SHM_RDWR \ O_proc p \ Tainted s \ - info_flow_shm s p'' p') - then (case (cp2sproc s p') of - Some sp \ Some (S_proc sp True) - | _ \ None) - else co2sobj s obj - | _ \ co2sobj s obj)" -apply (frule vt_grant_os, frule vd_cons, case_tac obj) -apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) - -apply (rule conjI|rule impI|erule exE)+ -apply (simp split:option.splits del:split_paired_Ex) -apply (rule impI, frule current_proc_has_sp, simp) -apply ((erule exE)+, auto simp:tainted_eq_Tainted intro:info_flow_shm_Tainted)[1] -apply (rule impI, simp add:tainted_eq_Tainted split:option.splits del:split_paired_Ex) -apply (auto simp:info_flow_shm_prop1 cp2sproc_attach dest!:current_proc_has_sp')[1] - -apply (case_tac "cp2sproc (Attach p h flag # s) nat") -apply (drule current_proc_has_sp', simp+) - -apply (rule conjI|erule exE|erule conjE|rule impI)+ -apply (simp add:tainted_eq_Tainted) -apply (auto simp:info_flow_shm_prop1 cp2sproc_attach intro:info_flow_shm_Tainted dest!:current_proc_has_sp')[1] -apply (auto simp:info_flow_shm_prop1 cp2sproc_attach tainted_eq_Tainted intro:info_flow_shm_Tainted dest!:current_proc_has_sp' -split:option.splits if_splits)[1] - - -apply (auto split:if_splits option.splits dest!:current_file_has_sfile' - simp:current_files_simps cf2sfiles_simps cf2sfile_simps tainted_eq_Tainted - same_inode_files_prop6 - dest:is_file_in_current is_dir_in_current) -done - - lemma co2sobj_other: "\valid (e # s); alive (e # s) obj; \ p f fds. e \ Execve p f fds; - \ p p' fds shms. e \ Clone p p' fds shms; + \ p p' fds. e \ Clone p p' fds; \ p p'. e \ Ptrace p p'; \ p f flags fd opt. e \ Open p f flags fd opt; \ p fd. e \ ReadFile p fd; @@ -2348,27 +1631,20 @@ \ p f. e \ UnLink p f; \ p f. e \ Rmdir p f; \ p f i. e \ Mkdir p f i; - \ p f f'. e \ LinkHard p f f'; \ p f len. e \ Truncate p f len; \ p q. e \ CreateMsgq p q; \ p q m. e \ SendMsg p q m; \ p q m. e \ RecvMsg p q m; - \ p q. e \ RemoveMsgq p q; - \ p h. e \ CreateShM p h; - \ p h flag. e \ Attach p h flag; - \ p h. e \ Detach p h; - \ p h. e \ DeleteShM p h - \ \ co2sobj (e # s) obj = co2sobj s obj" + \ p q. e \ RemoveMsgq p q\ + \ co2sobj (e # s) obj = co2sobj s obj" apply (frule vt_grant, case_tac e) apply (auto intro:co2sobj_kill co2sobj_exit) done lemmas co2sobj_simps = co2sobj_execve co2sobj_clone co2sobj_ptrace co2sobj_open co2sobj_readfile - co2sobj_writefile co2sobj_closefd co2sobj_unlink co2sobj_rmdir co2sobj_mkdir co2sobj_linkhard + co2sobj_writefile co2sobj_closefd co2sobj_unlink co2sobj_rmdir co2sobj_mkdir co2sobj_truncate co2sobj_kill co2sobj_exit co2sobj_createmsgq co2sobj_sendmsg co2sobj_recvmsg - co2sobj_removemsgq co2sobj_attach co2sobj_detach co2sobj_createshm co2sobj_deleteshm - - + co2sobj_removemsgq end