# HG changeset patch # User chunhan # Date 1387799237 -28800 # Node ID 0a68c605ae79481122771a5f2000ee73bcb0d489 # Parent 1ac0c3031ed2e8f6ebf6ccb2ab7b99c13b80dfc8 update diff -r 1ac0c3031ed2 -r 0a68c605ae79 no_shm_selinux/Enrich.thy --- a/no_shm_selinux/Enrich.thy Fri Dec 20 13:48:25 2013 +0800 +++ b/no_shm_selinux/Enrich.thy Mon Dec 23 19:47:17 2013 +0800 @@ -1201,245 +1201,23 @@ apply (erule enrich_proc_dup_ffd, simp+) done -lemma enrich_proc_prop: - "\valid s; is_created_proc s p; p' \ all_procs s\ - \ valid (enrich_proc s p p') \ - (\ obj. alive s obj \ alive (enrich_proc s p p') obj) \ - (\ obj. enrich_not_alive s obj \ enrich_not_alive (enrich_proc s p p') obj) \ - (files_hung_by_del (enrich_proc s p p') = files_hung_by_del s) \ - (\ tp. tp \ current_procs s \ cp2sproc (enrich_proc s p p') tp = cp2sproc s tp) \ - (\ f. f \ current_files s \ cf2sfile (enrich_proc s p p') f = cf2sfile s f) \ - (\ q. q \ current_msgqs s \ cq2smsgq (enrich_proc s p p') q = cq2smsgq s q) \ - (\ tp fd f. file_of_proc_fd s tp fd = Some f \ file_of_proc_fd (enrich_proc s p p') tp fd = Some f) \ - (\ tp fd flags. flags_of_proc_fd s tp fd = Some flags \ - flags_of_proc_fd (enrich_proc s p p') tp fd = Some flags) \ - (\ q. msgs_of_queue (enrich_proc s p p') q = msgs_of_queue s q) \ - (\ tp fd. fd \ proc_file_fds s tp \ cfd2sfd (enrich_proc s p p') tp fd = cfd2sfd s tp fd) \ - (cp2sproc (enrich_proc s p p') p' = cp2sproc s p) \ - (\ fd. fd \ proc_file_fds s p \ cfd2sfd (enrich_proc s p p') p' fd = cfd2sfd s p fd)" -proof (induct s) - case Nil - thus ?case by (auto simp:is_created_proc_def) -next - case (Cons e s) - hence vd_cons: "valid (e # s)" and created_cons: "is_created_proc (e # s) p" - and all_procs_cons: "p' \ all_procs (e # s)" and vd: "valid s" - and os: "os_grant s e" and grant: "grant s e" - by (auto dest:vd_cons vt_grant_os vt_grant) - from all_procs_cons have all_procs: "p' \ all_procs s" by (case_tac e, auto) - from Cons have pre: "is_created_proc s p \ valid (enrich_proc s p p') \ - (\obj. alive s obj \ alive (enrich_proc s p p') obj) \ - (\obj. enrich_not_alive s obj \ enrich_not_alive (enrich_proc s p p') obj) \ - files_hung_by_del (enrich_proc s p p') = files_hung_by_del s \ - (\tp. tp \ current_procs s \ cp2sproc (enrich_proc s p p') tp = cp2sproc s tp) \ - (\f. f \ current_files s \ cf2sfile (enrich_proc s p p') f = cf2sfile s f) \ - (\q. q \ current_msgqs s \ cq2smsgq (enrich_proc s p p') q = cq2smsgq s q) \ - (\tp fd f. file_of_proc_fd s tp fd = Some f \ file_of_proc_fd (enrich_proc s p p') tp fd = Some f) \ - (\tp fd flags. - flags_of_proc_fd s tp fd = Some flags \ flags_of_proc_fd (enrich_proc s p p') tp fd = Some flags) \ - (\q. msgs_of_queue (enrich_proc s p p') q = msgs_of_queue s q) \ - (\tp fd. fd \ proc_file_fds s tp \ cfd2sfd (enrich_proc s p p') tp fd = cfd2sfd s tp fd) \ - (cp2sproc (enrich_proc s p p') p' = cp2sproc s p) \ - (\ fd. fd \ proc_file_fds s p \ cfd2sfd (enrich_proc s p p') p' fd = cfd2sfd s p fd)" - using vd all_procs by auto - have alive_pre: "is_created_proc s p \ (\obj. alive s obj \ alive (enrich_proc s p p') obj)" - using pre by simp - hence curf_pre: "is_created_proc s p \ (\f. f \ current_files s \ f \ current_files (enrich_proc s p p'))" - using vd apply auto - apply (drule is_file_or_dir, simp) - apply (erule disjE) - apply (erule_tac x = "O_file f" in allE, simp add:is_file_in_current) - apply (erule_tac x = "O_dir f" in allE, simp add:is_dir_in_current) - done - have "valid (enrich_proc (e # s) p p')" - proof- - from pre have pre': "is_created_proc s p \ valid (enrich_proc s p p')" by simp - have "is_created_proc s p \ valid (e # enrich_proc s p p')" - apply (frule pre') - apply (erule_tac s = s in enrich_valid_intro_cons) - apply (simp_all add:os grant vd pre) - done - moreover have "\f fds. \valid (Execve p f fds # enrich_proc s p p'); is_created_proc s p; - valid (Execve p f fds # s); p' \ all_procs s\ - \ valid (Execve p' f (fds \ proc_file_fds s p) # Execve p f fds # enrich_proc s p p')" - proof- - fix f fds - assume a1: "valid (Execve p f fds # enrich_proc s p p')" and a2: "is_created_proc s p" - and a3: "valid (Execve p f fds # s)" and a0: "p' \ all_procs s" - have cp2sp: "\ tp. tp \ current_procs s \ cp2sproc (enrich_proc s p p') tp = cp2sproc s tp" - and cf2sf: "\ tf. tf \ current_files s \ cf2sfile (enrich_proc s p p') tf = cf2sfile s tf" - and cfd2sfd: "\ tp tfd. tfd \ proc_file_fds s tp \ cfd2sfd (enrich_proc s p p') tp tfd = cfd2sfd s tp tfd" - and ffd_remain: "\tp fd f. file_of_proc_fd s tp fd = Some f \ - file_of_proc_fd (enrich_proc s p p') tp fd = Some f" - and dup_sp: "cp2sproc (enrich_proc s p p') p' = cp2sproc s p" - and dup_sfd: "\ fd. fd \ proc_file_fds s p \ cfd2sfd (enrich_proc s p p') p' fd = cfd2sfd s p fd" - using pre a2 by auto - show "valid (Execve p' f (fds \ proc_file_fds s p) # Execve p f fds # enrich_proc s p p')" - proof- - from a0 a3 have a0': "p' \ p" by (auto dest!:vt_grant_os not_all_procs_prop3) - from a3 have grant: "grant s (Execve p f fds)" and os: "os_grant s (Execve p f fds)" - by (auto dest:vt_grant_os vt_grant simp del:os_grant.simps) - have f_in: "is_file (enrich_proc s p p') f" - proof- - from pre a2 - have a4: "\ obj. alive s obj \ alive (enrich_proc s p p') obj" - by (auto) - show ?thesis using a3 a4 - apply (erule_tac x = "O_file f" in allE) - by (auto dest:vt_grant_os) - qed - moreover have a5: "proc_file_fds s p \ proc_file_fds (Execve p f fds # enrich_proc s p p') p'" - using a3 a0' - apply (frule_tac vt_grant_os) - apply (auto simp:proc_file_fds_def) - apply (rule_tac x = fa in exI) - apply (erule enrich_proc_dup_ffd) - apply (simp_all add:vd all_procs a2) - done - ultimately have "os_grant (Execve p f fds # enrich_proc s p p') (Execve p' f (fds \ proc_file_fds s p))" - apply (auto simp:is_file_simps enrich_proc_dup_in a2 vd all_procs a1 enrich_proc_dup_ffds) - done - moreover have "grant (Execve p f fds # enrich_proc s p p') (Execve p' f (fds \ proc_file_fds s p))" - proof- - from grant obtain up rp tp uf rf tf - where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" - and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" - by (simp split:option.splits, blast) - with grant obtain pu nr nt where p3: "npctxt_execve (up, rp, tp) (uf, rf, tf) = Some (pu, nr, nt)" - by (simp split:option.splits del:npctxt_execve.simps, blast) - have p1': "sectxt_of_obj (Execve p f fds # enrich_proc s p p') (O_proc p') = Some (up, rp, tp)" - using p1 dup_sp a1 a0' - apply (simp add:sectxt_of_obj_simps) - by (simp add:cp2sproc_def split:option.splits) - from os have f_in': "is_file s f" by simp - from vd os have "\ sf. cf2sfile s f = Some sf" - by (auto dest!:is_file_in_current current_file_has_sfile) - hence p2': "sectxt_of_obj (Execve p f fds # enrich_proc s p p') (O_file f) = Some (uf, rf, tf)" - using f_in p2 cf2sf os a1 - apply (erule_tac x = f in allE) - apply (auto dest:is_file_in_current simp:cf2sfile_def sectxt_of_obj_simps split:option.splits) - apply (case_tac f, simp) - apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) - done - from dup_sfd a5 have "\fd. fd \ proc_file_fds s p \ - cfd2sfd (Execve p f fds # enrich_proc s p p') p' fd = cfd2sfd s p fd" - apply (rule_tac allI) - apply (erule_tac x = fd in allE, clarsimp) - apply (drule set_mp, simp) - apply (auto simp:cfd2sfd_execve proc_file_fds_def a1) - done - hence "inherit_fds_check (Execve p f fds # enrich_proc s p p') (up, nr, nt) p' (fds \ proc_file_fds s p)" - using grant os p1 p2 p3 vd - apply (clarsimp) - apply (rule_tac s = s and p = p and fds = fds in enrich_inherit_fds_check_dup) - apply simp_all - done - moreover have "search_check (Execve p f fds # enrich_proc s p p') (up, rp, tp) f" - using p1 p2 p2' vd cf2sf f_in f_in' grant p3 f_in a1 - apply (rule_tac s = s in enrich_search_check) - apply (simp_all add:is_file_simps) - apply (rule allI, rule impI, erule_tac x = fa in allE, simp) - apply (drule_tac ff = fa in cf2sfile_other') - by (auto simp:a2 curf_pre) - ultimately show ?thesis - using p1' p2' p3 - apply (simp split:option.splits) - using grant p1 p2 - apply simp - done - qed - ultimately show ?thesis using a1 - by (erule_tac valid.intros(2), auto) - qed - qed - moreover have "\tp fds. \valid (Clone tp p fds # s); p' \ p; p' \ all_procs s\ \ - valid (Clone tp p' (fds \ proc_file_fds s tp) # Clone tp p fds # s)" - apply (frule vt_grant_os, frule vt_grant, drule not_all_procs_prop3) - apply (rule valid.intros(2)) - apply (simp_all split:option.splits add:sectxt_of_obj_simps) - apply (auto simp add:proc_file_fds_def)[1] - apply (auto simp:inherit_fds_check_def sectxt_of_obj_simps sectxts_of_fds_def inherit_fds_check_ctxt_def) - done - moreover have "\f flags fd opt. \valid (Open p f flags fd opt # enrich_proc s p p'); - is_created_proc s p; valid (Open p f flags fd opt # s); p' \ all_procs s\ - \ valid (Open p' f (remove_create_flag flags) fd None # Open p f flags fd opt # enrich_proc s p p')" - proof- - fix f flags fd inum - assume a1: "valid (Open p f flags fd inum # enrich_proc s p p')" and a2: "is_created_proc s p" - and a3: "valid (Open p f flags fd inum # s)" and a4: "p' \ all_procs s" - show "valid (Open p' f (remove_create_flag flags) fd inum # Open p f flags fd inum # enrich_proc s p p')" - proof (cases " - apply (rule_tac valid.intros(2)) - apply (simp_all add:a1) +lemma enrich_proc_dup_ffds_eq_fds: + "\is_created_proc s p; p' \ all_procs s; valid s\ + \ current_proc_fds (enrich_proc s p p') p' = proc_file_fds s p" +apply (induct s arbitrary:p) +apply (simp add: is_created_proc_def) +apply (frule not_all_procs_prop3) +apply (frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto split:if_splits option.splits dest:proc_fd_in_fds set_mp not_all_procs_prop3 + simp:proc_file_fds_def is_created_proc_def) +done - - - +lemma oflags_check_remove_create: + "oflags_check flags sp sf \ oflags_check (remove_create_flag flags) sp sf" +apply (case_tac flags) +apply (auto simp:oflags_check_def perms_of_flags_def perm_of_oflag_def split:bool.splits) +done - sorry - moreover have "\fd. \valid (CloseFd p fd # enrich_proc s p p'); is_created_proc s p; - valid (CloseFd p fd # s); p' \ all_procs s; fd \ proc_file_fds s p\ - \ valid (CloseFd p' fd # CloseFd p fd # enrich_proc s p p')" - proof- - fix fd - assume a1: "valid (CloseFd p fd # enrich_proc s p p')" and a2: "is_created_proc s p" - and a3: "p' \ all_procs s" and a4: "valid (CloseFd p fd # s)" - and a5: "fd \ proc_file_fds s p" - from a4 a3 have a0: "p' \ p" by (auto dest!:vt_grant_os not_all_procs_prop3) - have "p' \ current_procs (enrich_proc s p p')" - using a2 a3 vd - by (auto intro:enrich_proc_dup_in) - moreover have "fd \ current_proc_fds (enrich_proc s p p') p'" - using a5 a2 a3 vd pre' - apply (simp) - apply (drule_tac s = "enrich_proc s p p'" and p = p' in file_fds_subset_pfds) - apply (erule set_mp) - apply (simp add:enrich_proc_dup_ffds) - done - ultimately show "valid (CloseFd p' fd # CloseFd p fd # enrich_proc s p p')" - apply (rule_tac valid.intros(2)) - apply (simp_all add:a1 a0 a2 pre') - done - qed - ultimately show ?thesis using created_cons vd_cons all_procs_cons - apply (case_tac e) - apply (auto simp:is_created_proc_simps split:if_splits) - done - qed - moreover have "\obj. alive (e # s) obj \ alive (enrich_proc (e # s) p p') obj" - sorry - moreover have "\obj. enrich_not_alive (e # s) obj \ enrich_not_alive (enrich_proc (e # s) p p') obj" - sorry - moreover have "files_hung_by_del (enrich_proc (e # s) p p') = files_hung_by_del (e # s)" - sorry - moreover have "\p. p \ current_procs (e # s) \ cp2sproc (enrich_proc (e # s) p p') p = cp2sproc (e # s) p" - sorry - moreover have "\f. f \ current_files (e # s) \ cf2sfile (enrich_proc (e # s) p p') f = cf2sfile (e # s) f" - sorry - moreover have "\q. q \ current_msgqs (e # s) \ cq2smsgq (enrich_proc (e # s) p p') q = cq2smsgq (e # s) q" - sorry - moreover have "\p fd f. file_of_proc_fd (e # s) p fd = Some f \ - file_of_proc_fd (enrich_proc (e # s) p p') p fd = Some f" - sorry - moreover have "\p fd flags. flags_of_proc_fd (e # s) p fd = Some flags \ - flags_of_proc_fd (enrich_proc (e # s) p p') p fd = Some flags" - sorry - moreover have "\q. msgs_of_queue (enrich_proc (e # s) p p') q = msgs_of_queue (e # s) q" - sorry - moreover have "\p fd. fd \ proc_file_fds (e # s) p \ - cfd2sfd (enrich_proc (e # s) p p') p fd = cfd2sfd (e # s) p fd" - sorry - ultimately show ?case - by auto -qed - +end -lemma enrich_proc_valid: - "\valid s; is_created_proc s p; p' \ all_procs s\ \ valid (enrich_proc s p p')" -by (auto dest:enrich_proc_prop) - -lemma enrich_proc_valid: - "\valid s; is_created_proc s p; p' \ all_procs s\ \ " - - - \ No newline at end of file +end \ No newline at end of file diff -r 1ac0c3031ed2 -r 0a68c605ae79 no_shm_selinux/Enrich2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/no_shm_selinux/Enrich2.thy Mon Dec 23 19:47:17 2013 +0800 @@ -0,0 +1,624 @@ +theory Enrich +imports Main Flask Static Init_prop Valid_prop Tainted_prop Delete_prop Co2sobj_prop S2ss_prop S2ss_prop2 + Temp Enrich +begin + +context tainting_s begin + +lemma get_parentfs_ctxts_prop: + "\get_parentfs_ctxts s (a # f) = Some ctxts; sectxt_of_obj s (O_dir f) = Some ctxt; valid s\ + \ ctxt \ set (ctxts)" +apply (induct f) +apply (auto split:option.splits) +done + +lemma search_check_allp_intro: + "\search_check s sp pf; get_parentfs_ctxts s pf = Some ctxts; valid s; is_dir s pf\ + \ search_check_allp sp (set ctxts)" +apply (case_tac pf) +apply (simp split:option.splits if_splits add:search_check_allp_def) +apply (rule ballI) +apply (auto simp:search_check_ctxt_def search_check_dir_def split:if_splits option.splits) +apply (auto simp:search_check_allp_def search_check_file_def) +apply (frule is_dir_not_file, simp) +done + +lemma search_check_leveling_f: + "\search_check s sp pf; parent f = Some pf; is_file s f; valid s; + sectxt_of_obj s (O_file f) = Some fctxt; search_check_file sp fctxt\ + \ search_check s sp f" +apply (case_tac f, simp+) +apply (auto split:option.splits simp:search_check_ctxt_def) +apply (frule parentf_is_dir_prop2, simp) +apply (erule get_pfs_secs_prop, simp) +apply (erule_tac search_check_allp_intro, simp_all) +apply (simp add:parentf_is_dir_prop2) +done + +lemma enrich_proc_prop: + "\valid s; is_created_proc s p; p' \ all_procs s\ + \ valid (enrich_proc s p p') \ + (\ obj. alive s obj \ alive (enrich_proc s p p') obj) \ + (\ obj. enrich_not_alive s obj \ enrich_not_alive (enrich_proc s p p') obj) \ + (files_hung_by_del (enrich_proc s p p') = files_hung_by_del s) \ + (\ tp. tp \ current_procs s \ cp2sproc (enrich_proc s p p') tp = cp2sproc s tp) \ + (\ f. f \ current_files s \ cf2sfile (enrich_proc s p p') f = cf2sfile s f) \ + (\ q. q \ current_msgqs s \ cq2smsgq (enrich_proc s p p') q = cq2smsgq s q) \ + (\ tp fd f. file_of_proc_fd s tp fd = Some f \ file_of_proc_fd (enrich_proc s p p') tp fd = Some f) \ + (\ tp fd flags. flags_of_proc_fd s tp fd = Some flags \ + flags_of_proc_fd (enrich_proc s p p') tp fd = Some flags) \ + (\ q. msgs_of_queue (enrich_proc s p p') q = msgs_of_queue s q) \ + (\ tp fd. fd \ proc_file_fds s tp \ cfd2sfd (enrich_proc s p p') tp fd = cfd2sfd s tp fd) \ + (cp2sproc (enrich_proc s p p') p' = cp2sproc s p) \ + (\ fd. fd \ proc_file_fds s p \ cfd2sfd (enrich_proc s p p') p' fd = cfd2sfd s p fd)" +proof (induct s) + case Nil + thus ?case by (auto simp:is_created_proc_def) +next + case (Cons e s) + hence vd_cons: "valid (e # s)" and created_cons: "is_created_proc (e # s) p" + and all_procs_cons: "p' \ all_procs (e # s)" and vd: "valid s" + and os: "os_grant s e" and grant: "grant s e" + by (auto dest:vd_cons vt_grant_os vt_grant) + from all_procs_cons have all_procs: "p' \ all_procs s" by (case_tac e, auto) + from Cons have pre: "is_created_proc s p \ valid (enrich_proc s p p') \ + (\obj. alive s obj \ alive (enrich_proc s p p') obj) \ + (\obj. enrich_not_alive s obj \ enrich_not_alive (enrich_proc s p p') obj) \ + files_hung_by_del (enrich_proc s p p') = files_hung_by_del s \ + (\tp. tp \ current_procs s \ cp2sproc (enrich_proc s p p') tp = cp2sproc s tp) \ + (\f. f \ current_files s \ cf2sfile (enrich_proc s p p') f = cf2sfile s f) \ + (\q. q \ current_msgqs s \ cq2smsgq (enrich_proc s p p') q = cq2smsgq s q) \ + (\tp fd f. file_of_proc_fd s tp fd = Some f \ file_of_proc_fd (enrich_proc s p p') tp fd = Some f) \ + (\tp fd flags. + flags_of_proc_fd s tp fd = Some flags \ flags_of_proc_fd (enrich_proc s p p') tp fd = Some flags) \ + (\q. msgs_of_queue (enrich_proc s p p') q = msgs_of_queue s q) \ + (\tp fd. fd \ proc_file_fds s tp \ cfd2sfd (enrich_proc s p p') tp fd = cfd2sfd s tp fd) \ + (cp2sproc (enrich_proc s p p') p' = cp2sproc s p) \ + (\ fd. fd \ proc_file_fds s p \ cfd2sfd (enrich_proc s p p') p' fd = cfd2sfd s p fd)" + using vd all_procs by auto + have alive_pre: "is_created_proc s p \ (\obj. alive s obj \ alive (enrich_proc s p p') obj)" + using pre by simp + hence curf_pre: "is_created_proc s p \ (\f. f \ current_files s \ f \ current_files (enrich_proc s p p'))" + using vd apply auto + apply (drule is_file_or_dir, simp) + apply (erule disjE) + apply (erule_tac x = "O_file f" in allE, simp add:is_file_in_current) + apply (erule_tac x = "O_dir f" in allE, simp add:is_dir_in_current) + done + have vd_enrich_cons: "valid (enrich_proc (e # s) p p')" + proof- + from pre have pre': "is_created_proc s p \ valid (enrich_proc s p p')" by simp + have "is_created_proc s p \ valid (e # enrich_proc s p p')" + apply (frule pre') + apply (erule_tac s = s in enrich_valid_intro_cons) + apply (simp_all add:os grant vd pre) + done + moreover have "\f fds. \valid (Execve p f fds # enrich_proc s p p'); is_created_proc s p; + valid (Execve p f fds # s); p' \ all_procs s\ + \ valid (Execve p' f (fds \ proc_file_fds s p) # Execve p f fds # enrich_proc s p p')" + proof- + fix f fds + assume a1: "valid (Execve p f fds # enrich_proc s p p')" and a2: "is_created_proc s p" + and a3: "valid (Execve p f fds # s)" and a0: "p' \ all_procs s" + have cp2sp: "\ tp. tp \ current_procs s \ cp2sproc (enrich_proc s p p') tp = cp2sproc s tp" + and cf2sf: "\ tf. tf \ current_files s \ cf2sfile (enrich_proc s p p') tf = cf2sfile s tf" + and cfd2sfd: "\ tp tfd. tfd \ proc_file_fds s tp \ cfd2sfd (enrich_proc s p p') tp tfd = cfd2sfd s tp tfd" + and ffd_remain: "\tp fd f. file_of_proc_fd s tp fd = Some f \ + file_of_proc_fd (enrich_proc s p p') tp fd = Some f" + and dup_sp: "cp2sproc (enrich_proc s p p') p' = cp2sproc s p" + and dup_sfd: "\ fd. fd \ proc_file_fds s p \ cfd2sfd (enrich_proc s p p') p' fd = cfd2sfd s p fd" + using pre a2 by auto + show "valid (Execve p' f (fds \ proc_file_fds s p) # Execve p f fds # enrich_proc s p p')" + proof- + from a0 a3 have a0': "p' \ p" by (auto dest!:vt_grant_os not_all_procs_prop3) + from a3 have grant: "grant s (Execve p f fds)" and os: "os_grant s (Execve p f fds)" + by (auto dest:vt_grant_os vt_grant simp del:os_grant.simps) + have f_in: "is_file (enrich_proc s p p') f" + proof- + from pre a2 + have a4: "\ obj. alive s obj \ alive (enrich_proc s p p') obj" + by (auto) + show ?thesis using a3 a4 + apply (erule_tac x = "O_file f" in allE) + by (auto dest:vt_grant_os) + qed + moreover have a5: "proc_file_fds s p \ proc_file_fds (Execve p f fds # enrich_proc s p p') p'" + using a3 a0' + apply (frule_tac vt_grant_os) + apply (auto simp:proc_file_fds_def) + apply (rule_tac x = fa in exI) + apply (erule enrich_proc_dup_ffd) + apply (simp_all add:vd all_procs a2) + done + ultimately have "os_grant (Execve p f fds # enrich_proc s p p') (Execve p' f (fds \ proc_file_fds s p))" + apply (auto simp:is_file_simps enrich_proc_dup_in a2 vd all_procs a1 enrich_proc_dup_ffds) + done + moreover have "grant (Execve p f fds # enrich_proc s p p') (Execve p' f (fds \ proc_file_fds s p))" + proof- + from grant obtain up rp tp uf rf tf + where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" + and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" + by (simp split:option.splits, blast) + with grant obtain pu nr nt where p3: "npctxt_execve (up, rp, tp) (uf, rf, tf) = Some (pu, nr, nt)" + by (simp split:option.splits del:npctxt_execve.simps, blast) + have p1': "sectxt_of_obj (Execve p f fds # enrich_proc s p p') (O_proc p') = Some (up, rp, tp)" + using p1 dup_sp a1 a0' + apply (simp add:sectxt_of_obj_simps) + by (simp add:cp2sproc_def split:option.splits) + from os have f_in': "is_file s f" by simp + from vd os have "\ sf. cf2sfile s f = Some sf" + by (auto dest!:is_file_in_current current_file_has_sfile) + hence p2': "sectxt_of_obj (Execve p f fds # enrich_proc s p p') (O_file f) = Some (uf, rf, tf)" + using f_in p2 cf2sf os a1 + apply (erule_tac x = f in allE) + apply (auto dest:is_file_in_current simp:cf2sfile_def sectxt_of_obj_simps split:option.splits) + apply (case_tac f, simp) + apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) + done + from dup_sfd a5 have "\fd. fd \ proc_file_fds s p \ + cfd2sfd (Execve p f fds # enrich_proc s p p') p' fd = cfd2sfd s p fd" + apply (rule_tac allI) + apply (erule_tac x = fd in allE, clarsimp) + apply (drule set_mp, simp) + apply (auto simp:cfd2sfd_execve proc_file_fds_def a1) + done + hence "inherit_fds_check (Execve p f fds # enrich_proc s p p') (up, nr, nt) p' (fds \ proc_file_fds s p)" + using grant os p1 p2 p3 vd + apply (clarsimp) + apply (rule_tac s = s and p = p and fds = fds in enrich_inherit_fds_check_dup) + apply simp_all + done + moreover have "search_check (Execve p f fds # enrich_proc s p p') (up, rp, tp) f" + using p1 p2 p2' vd cf2sf f_in f_in' grant p3 f_in a1 + apply (rule_tac s = s in enrich_search_check) + apply (simp_all add:is_file_simps) + apply (rule allI, rule impI, erule_tac x = fa in allE, simp) + apply (drule_tac ff = fa in cf2sfile_other') + by (auto simp:a2 curf_pre) + ultimately show ?thesis + using p1' p2' p3 + apply (simp split:option.splits) + using grant p1 p2 + apply simp + done + qed + ultimately show ?thesis using a1 + by (erule_tac valid.intros(2), auto) + qed + qed + moreover have "\tp fds. \valid (Clone tp p fds # s); p' \ p; p' \ all_procs s\ \ + valid (Clone tp p' (fds \ proc_file_fds s tp) # Clone tp p fds # s)" + apply (frule vt_grant_os, frule vt_grant, drule not_all_procs_prop3) + apply (rule valid.intros(2)) + apply (simp_all split:option.splits add:sectxt_of_obj_simps) + apply (auto simp add:proc_file_fds_def)[1] + apply (auto simp:inherit_fds_check_def sectxt_of_obj_simps sectxts_of_fds_def inherit_fds_check_ctxt_def) + done + moreover have "\f flags fd opt. \valid (Open p f flags fd opt # enrich_proc s p p'); + is_created_proc s p; valid (Open p f flags fd opt # s); p' \ all_procs s\ + \ valid (Open p' f (remove_create_flag flags) fd None # Open p f flags fd opt # enrich_proc s p p')" + proof- + fix f flags fd opt + assume a1: "valid (Open p f flags fd opt # enrich_proc s p p')" and a2: "is_created_proc s p" + and a3: "valid (Open p f flags fd opt # s)" and a4: "p' \ all_procs s" + have cp2sp: "\ tp. tp \ current_procs s \ cp2sproc (enrich_proc s p p') tp = cp2sproc s tp" + and cf2sf: "\ tf. tf \ current_files s \ cf2sfile (enrich_proc s p p') tf = cf2sfile s tf" + and cfd2sfd: "\ tp tfd. tfd \ proc_file_fds s tp \ cfd2sfd (enrich_proc s p p') tp tfd = cfd2sfd s tp tfd" + and ffd_remain: "\tp fd f. file_of_proc_fd s tp fd = Some f \ + file_of_proc_fd (enrich_proc s p p') tp fd = Some f" + and dup_sp: "cp2sproc (enrich_proc s p p') p' = cp2sproc s p" + and dup_sfd: "\ fd. fd \ proc_file_fds s p \ cfd2sfd (enrich_proc s p p') p' fd = cfd2sfd s p fd" + using pre a2 by auto + from a4 a3 have a0: "p' \ p" by (auto dest!:vt_grant_os not_all_procs_prop3 split:option.splits) + have a5: "p' \ current_procs (enrich_proc s p p')" + using a2 a3 vd + apply (erule_tac enrich_proc_dup_in) + by (simp_all add:vd a4) + have a6: "is_file (Open p f flags fd opt # enrich_proc s p p') f" + using a1 a3 + by (auto simp:is_file_open dest:vt_grant_os) + have a7: "fd \ current_proc_fds (enrich_proc s p p') p'" + using a2 a4 vd + apply (simp add:enrich_proc_dup_ffds_eq_fds) + apply (rule notI) + apply (drule_tac p = p in file_fds_subset_pfds) + apply (drule set_mp, simp) + using a3 + apply (drule_tac vt_grant_os) + apply (auto split:option.splits) + done + from a1 have a8: "valid (enrich_proc s p p')" by (erule_tac valid.cases, auto) + from a3 have grant: "grant s (Open p f flags fd opt)" and os: "os_grant s (Open p f flags fd opt)" + by (auto dest:vt_grant_os vt_grant) + show "valid (Open p' f (remove_create_flag flags) fd None # Open p f flags fd opt # enrich_proc s p p')" + proof (cases opt) + case None + have f_in: "is_file (enrich_proc s p p') f" + proof- + from pre a2 + have a4: "\ obj. alive s obj \ alive (enrich_proc s p p') obj" + by (auto) + show ?thesis using a3 a4 None + apply (erule_tac x = "O_file f" in allE) + by (auto dest:vt_grant_os) + qed + from grant None obtain up rp tp uf rf tf + where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" + and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" + apply (simp split:option.splits) + by (case_tac a, case_tac aa, blast) + have p1': "sectxt_of_obj (Open p f flags fd opt # enrich_proc s p p') (O_proc p') = Some (up, rp, tp)" + using p1 dup_sp a1 + apply (simp add:sectxt_of_obj_simps) + by (simp add:cp2sproc_def split:option.splits) + from os None have f_in': "is_file s f" by simp + from vd os None have "\ sf. cf2sfile s f = Some sf" + by (auto dest!:is_file_in_current current_file_has_sfile) + hence p2': "sectxt_of_obj (Open p f flags fd opt # enrich_proc s p p') (O_file f) = Some (uf, rf, tf)" + using p2 cf2sf os a1 None f_in' vd f_in + apply (erule_tac x = f in allE) + apply (auto dest:is_file_in_current simp:cf2sfile_def sectxt_of_obj_simps split:option.splits) + apply (case_tac f, simp) + apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) + done + have "search_check (Open p f flags fd opt # enrich_proc s p p') (up, rp, tp) f" + using p1 p2 p2' vd cf2sf f_in f_in' grant f_in a1 None + apply (rule_tac s = s in enrich_search_check) + apply (simp_all add:is_file_simps) + apply (rule allI, rule impI, erule_tac x = fa in allE, simp) + apply (simp add:cf2sfile_open_none) + done + thus ?thesis using a0 a5 a6 a7 None a1 + apply (rule_tac valid.intros(2)) + apply (simp_all add:a1) + apply (case_tac flags, simp add:is_creat_excl_flag_def) + using p1' p2' + apply simp + using grant p1 p2 + apply (simp add:oflags_check_remove_create) + done + next + case (Some inum) + with os obtain pf where parent: "parent f = Some pf" by auto + with grant Some obtain pu rp pt pfu pfr pft where + p1: "sectxt_of_obj s (O_proc p) = Some (pu, rp, pt)" + and p2: "sectxt_of_obj s (O_dir pf) = Some (pfu, pfr, pft)" + apply (simp split:option.splits) + apply (case_tac a, case_tac aa, blast) + done + from p1 have p1': "sectxt_of_obj (enrich_proc s p p') (O_proc p) = Some (pu, rp, pt)" + using cp2sp os + apply (erule_tac x = p in allE) + apply (auto split:option.splits simp:cp2sproc_def) + done + from os parent Some + have pf_in: "is_dir s pf" by auto + hence "\ sf. cf2sfile s pf = Some sf" using vd + by (auto dest!:is_dir_in_current current_file_has_sfile) + from a1 parent Some have pf_in': "is_dir (enrich_proc s p p') pf" + apply (frule_tac vt_grant_os) + by (clarsimp) + have p2': "sectxt_of_obj (enrich_proc s p p') (O_dir pf) = Some (pfu, pfr, pft)" + proof- + have "cf2sfile (enrich_proc s p p') pf = cf2sfile s pf" + using cf2sf pf_in + apply (drule_tac is_dir_in_current) + apply (erule_tac x = pf in allE) + by simp + with pf_in pf_in' p2 vd Some a8 + show ?thesis + apply (frule_tac is_dir_not_file) + apply (frule_tac s = "enrich_proc s p p'" in is_dir_not_file) + apply (simp add:cf2sfile_def) + apply (case_tac pf, simp) + apply (simp add:sroot_def root_sec_remains) + by (auto split:option.splits dest!:current_has_sec' get_pfs_secs_prop' dest:parentf_is_dir_prop1) + qed + from p1 have dup: "sectxt_of_obj (Open p f flags fd (Some inum) # enrich_proc s p p') (O_proc p') + = Some (pu, rp, pt)" + using a1 Some + apply (simp add:sec_open) + using dup_sp + apply (auto split:option.splits if_splits simp:cp2sproc_def) + done + have nsf: "sectxt_of_obj (Open p f flags fd (Some inum) # enrich_proc s p p') (O_file f) + = Some (pu, R_object, type_transition pt pft C_file True)" + using a1 Some p1 p2 parent p2' p1' + by (simp add:sec_open) + have "search_check (Open p f flags fd (Some inum) # enrich_proc s p p') (pu, rp, pt) f" + proof- + have "search_check (Open p f flags fd (Some inum) # enrich_proc s p p') (pu, rp, pt) pf" + using grant Some parent p1 p2 vd a1 pf_in pf_in' p2' + apply simp + apply (rule_tac s = s in enrich_search_check') + apply (simp_all add:is_dir_simps sectxt_of_obj_simps) + apply (rule allI, rule impI) + apply (case_tac "fa = f") + using os Some + apply simp + apply (drule_tac f' = fa in cf2sfile_open) + apply (simp add:current_files_simps) + using curf_pre a2 + apply simp + apply simp + using cf2sf + apply simp + done + moreover have "is_file (Open p f flags fd (Some inum) # enrich_proc s p p') f" + using a1 Some + by (simp add:is_file_open) + ultimately + show ?thesis + using parent a1 Some nsf + apply (erule_tac search_check_leveling_f) + apply (simp_all) + apply (simp add:search_check_file_def) + apply (simp add:permission_check.simps) + sorry + qed + thus ?thesis using a0 a5 a6 a7 a1 Some + apply (rule_tac valid.intros(2)) + apply (simp add:a1) + apply simp + apply (case_tac flags, simp add:is_creat_excl_flag_def) + using grant p1 p2 parent dup nsf + apply (simp add:oflags_check_remove_create) + done + qed + qed + moreover have "\fd. \valid (CloseFd p fd # enrich_proc s p p'); is_created_proc s p; + valid (CloseFd p fd # s); p' \ all_procs s; fd \ proc_file_fds s p\ + \ valid (CloseFd p' fd # CloseFd p fd # enrich_proc s p p')" + proof- + fix fd + assume a1: "valid (CloseFd p fd # enrich_proc s p p')" and a2: "is_created_proc s p" + and a3: "p' \ all_procs s" and a4: "valid (CloseFd p fd # s)" + and a5: "fd \ proc_file_fds s p" + from a4 a3 have a0: "p' \ p" by (auto dest!:vt_grant_os not_all_procs_prop3) + have "p' \ current_procs (enrich_proc s p p')" + using a2 a3 vd + by (auto intro:enrich_proc_dup_in) + moreover have "fd \ current_proc_fds (enrich_proc s p p') p'" + using a5 a2 a3 vd pre' + apply (simp) + apply (drule_tac s = "enrich_proc s p p'" and p = p' in file_fds_subset_pfds) + apply (erule set_mp) + apply (simp add:enrich_proc_dup_ffds) + done + ultimately show "valid (CloseFd p' fd # CloseFd p fd # enrich_proc s p p')" + apply (rule_tac valid.intros(2)) + apply (simp_all add:a1 a0 a2 pre') + done + qed + moreover have "\ fd. \valid (ReadFile p fd # enrich_proc s p p'); + is_created_proc s p; valid (ReadFile p fd # s); p' \ all_procs s\ + \ valid (ReadFile p' fd # ReadFile p fd # enrich_proc s p p')" + proof- + fix fd + assume a1: "valid (ReadFile p fd # enrich_proc s p p')" and a2: "is_created_proc s p" + and a3: "valid (ReadFile p fd # s)" and a4: "p' \ all_procs s" + from a3 have os: "os_grant s (ReadFile p fd)" and grant: "grant s (ReadFile p fd)" + by (auto dest:vt_grant_os vt_grant) + have cp2sp: "\ tp. tp \ current_procs s \ cp2sproc (enrich_proc s p p') tp = cp2sproc s tp" + and cf2sf: "\ tf. tf \ current_files s \ cf2sfile (enrich_proc s p p') tf = cf2sfile s tf" + and cfd2sfd: "\ tp tfd. tfd \ proc_file_fds s tp \ cfd2sfd (enrich_proc s p p') tp tfd = cfd2sfd s tp tfd" + and ffd_remain: "\tp fd f. file_of_proc_fd s tp fd = Some f \ + file_of_proc_fd (enrich_proc s p p') tp fd = Some f" + and dup_sp: "cp2sproc (enrich_proc s p p') p' = cp2sproc s p" + and dup_sfd: "\ fd. fd \ proc_file_fds s p \ cfd2sfd (enrich_proc s p p') p' fd = cfd2sfd s p fd" + using pre a2 by auto + have vd_enrich: "valid (enrich_proc s p p')" using a1 by (auto dest:valid.cases) + show "valid (ReadFile p' fd # ReadFile p fd # enrich_proc s p p')" + proof- + have "os_grant (ReadFile p fd # enrich_proc s p p') (ReadFile p' fd)" + using a1 a2 a4 vd os + apply (clarsimp simp:current_files_simps enrich_proc_dup_in enrich_proc_dup_ffds_eq_fds) + apply (simp add:proc_file_fds_def) + apply (rule conjI) + apply (rule_tac x = f in exI, simp add:enrich_proc_dup_ffd) + using curf_pre + apply (simp) + apply (drule enrich_proc_dup_fflags) + apply simp_all + apply (erule disjE) + apply auto + apply (simp add:is_read_flag_def) + done + moreover have "grant (ReadFile p fd # enrich_proc s p p') (ReadFile p' fd)" + proof- + from grant obtain f sp sfd sf where p1: "file_of_proc_fd s p fd = Some f" + and p2: "sectxt_of_obj s (O_proc p) = Some sp" + and p3: "sectxt_of_obj s (O_fd p fd) = Some sfd" + and p4: "sectxt_of_obj s (O_file f) = Some sf" + by (auto split:option.splits) + from os obtain flag where p0: "flags_of_proc_fd s p fd = Some flag" + by auto + from os have f_in_s: "f \ current_files s" using p1 by simp + from p1 vd have isfile_s: "is_file s f" by (erule_tac file_of_pfd_is_file, simp) + with alive_pre a2 have isfile_s': "is_file (enrich_proc s p p') f" + apply simp + apply (erule_tac x = "O_file f" in allE, simp) + done + from p0 obtain flag' where p0': "flags_of_proc_fd (enrich_proc s p p') p' fd = Some flag'" + and p0'': "(flag' = flag) \ (flag' = remove_create_flag flag)" + using a2 a4 vd + apply (drule_tac enrich_proc_dup_fflags) + apply auto + apply (case_tac flag, auto) + apply (case_tac flag, auto) + done + from p1 have p1': "file_of_proc_fd (enrich_proc s p p') p' fd = Some f" + using a2 a4 vd + by (simp add:enrich_proc_dup_ffd) + from p2 have p2': "sectxt_of_obj (enrich_proc s p p') (O_proc p') = Some sp" + using dup_sp + by (auto simp:cp2sproc_def split:option.splits) + from p3 p1 p1' p0 p0' os have p3': "sectxt_of_obj (enrich_proc s p p') (O_fd p' fd) = Some sfd" + using dup_sfd + apply (erule_tac x = fd in allE) + apply (auto simp:proc_file_fds_def cfd2sfd_def split:option.splits) + apply (drule current_file_has_sfile') + apply (simp add:vd, simp) + apply (drule current_file_has_sfile') + apply (simp add:vd, simp) + done + from p4 have p4': "sectxt_of_obj (enrich_proc s p p') (O_file f) = Some sf" + using f_in_s cf2sf isfile_s isfile_s' a1 vd_enrich vd + apply (erule_tac x = f in allE) + apply (simp) + apply (auto simp:cf2sfile_def split:option.splits + dest!:current_has_sec' get_pfs_secs_prop' dest:parentf_is_dir is_file_in_current) + apply (case_tac f, simp, drule root_is_dir', simp, simp, simp) + done + show ?thesis using p1' p2' p3' p4' a1 + apply (simp add:sectxt_of_obj_simps) + using grant p1 p2 p3 p4 + apply simp + done + qed + ultimately show ?thesis + using a1 + by (erule_tac valid.intros(2), simp+) + qed + qed + ultimately show ?thesis + using created_cons vd_cons all_procs_cons + apply (case_tac e) + apply (auto simp:is_created_proc_simps split:if_splits) + done + qed + moreover have "\obj. alive (e # s) obj \ alive (enrich_proc (e # s) p p') obj" + sorry + moreover have "\obj. enrich_not_alive (e # s) obj \ enrich_not_alive (enrich_proc (e # s) p p') obj" + sorry + moreover have "files_hung_by_del (enrich_proc (e # s) p p') = files_hung_by_del (e # s)" + sorry + moreover have "\tp. tp \ current_procs (e # s) \ cp2sproc (enrich_proc (e # s) p p') tp = cp2sproc (e # s) tp" + sorry + moreover have "\f. f \ current_files (e # s) \ cf2sfile (enrich_proc (e # s) p p') f = cf2sfile (e # s) f" + sorry + moreover have "\q. q \ current_msgqs (e # s) \ cq2smsgq (enrich_proc (e # s) p p') q = cq2smsgq (e # s) q" + sorry + moreover have "\tp fd f. file_of_proc_fd (e # s) tp fd = Some f \ + file_of_proc_fd (enrich_proc (e # s) p p') tp fd = Some f" + sorry + moreover have "\tp fd flags. flags_of_proc_fd (e # s) tp fd = Some flags \ + flags_of_proc_fd (enrich_proc (e # s) p p') tp fd = Some flags" + sorry + moreover have "\q. msgs_of_queue (enrich_proc (e # s) p p') q = msgs_of_queue (e # s) q" + sorry + moreover have "\tp fd. fd \ proc_file_fds (e # s) tp \ + cfd2sfd (enrich_proc (e # s) p p') tp fd = cfd2sfd (e # s) tp fd" + sorry + moreover have "cp2sproc (enrich_proc (e # s) p p') p' = cp2sproc (e # s) p" + proof- + from pre have b0: "is_created_proc s p \ cp2sproc (enrich_proc s p p') p' = cp2sproc s p" by simp + have b1: "\ tp f fds. \valid (enrich_proc (Execve tp f fds # s) p p'); valid (Execve tp f fds # s); + is_created_proc (Execve tp f fds # s) p; p' \ all_procs (Execve tp f fds # s)\ + \ cp2sproc (enrich_proc (Execve tp f fds # s) p p') p' = cp2sproc (Execve tp f fds # s) p" + proof- + fix tp f fds + assume a1: "valid (enrich_proc (Execve tp f fds # s) p p')" + and a2: "valid (Execve tp f fds # s)" and a3: "is_created_proc (Execve tp f fds # s) p" + and a4: "p' \ all_procs (Execve tp f fds # s)" + show "cp2sproc (enrich_proc (Execve tp f fds # s) p p') p' = cp2sproc (Execve tp f fds # s) p" + proof (cases "tp = p") + case True + show ?thesis using True a1 a2 a3 a4 + thm not_all_procs_prop3 + apply (auto simp add:cp2sproc_execve is_created_proc_def split:option.splits dest!:current_has_sec' + dest:vt_grant_os not_all_procs_prop2 not_all_procs_prop3) + + sorry + next + case False + show ?thesis sorry + qed + qed + have b2: "\ tp fd. cp2sproc (enrich_proc (ReadFile tp fd # s) p p') p' = cp2sproc (ReadFile tp fd # s) p" + sorry + have b3: "\ tp. cp2sproc (enrich_proc (Exit tp # s) p p') p' = cp2sproc (Exit tp # s) p" + sorry + have b4: "\ tp tp'. cp2sproc (enrich_proc (Kill tp tp' # s) p p') p' = cp2sproc (Kill tp tp' # s) p" + sorry + have b5: "\ tp tp' fds. cp2sproc (enrich_proc (Clone tp tp' fds # s) p p') p' = + cp2sproc (Clone tp tp' fds # s) p" + sorry + have b6: "\ tp f flags fd opt. cp2sproc (enrich_proc (Open tp f flags fd opt # s) p p') p' = + cp2sproc (Open tp f flags fd opt # s) p" + sorry + have b7: "\ tp fd. cp2sproc (enrich_proc (CloseFd tp fd # s) p p') p' = cp2sproc (CloseFd tp fd # s) p" + sorry + show ?thesis using vd_enrich_cons + apply (case_tac e) + apply (simp_all only:b1 b2 b3 b4 b5 b6 b7) + using created_cons vd_enrich_cons vd_cons b0 + apply (auto simp:cp2sproc_other is_created_proc_def) + done + qed + moreover have "\ fd. fd \ proc_file_fds (e # s) p \ + cfd2sfd (enrich_proc (e # s) p p') p' fd = cfd2sfd (e # s) p fd" + sorry + ultimately show ?case + by simp +qed + + +lemma enrich_proc_valid: + "\valid s; is_created_proc s p; p' \ all_procs s\ \ valid (enrich_proc s p p')" +by (auto dest:enrich_proc_prop) + +lemma enrich_proc_valid: + "\valid s; is_created_proc s p; p' \ all_procs s\ \ " + + + +lemma enrich_proc_tainted: + "\is_created_proc s p; p' \ all_procs s; valid s\ + \ tainted (enrich_proc s p p') = (if (O_proc p \ tainted s) + then tainted s \ {O_proc p'} else tainted s)" +apply (induct s) +apply (simp add:is_created_proc_def) +apply (frule vt_grant_os, frule vd_cons, simp) +apply (frule enrich_proc_dup_in, simp+) +apply (frule not_all_procs_prop3) +apply (case_tac a) +prefer 3 +apply (simp split:if_splits) +apply (rule impI|rule conjI)+ +apply (simp add:is_created_proc_def) +apply (auto simp:is_created_proc_def split:if_splits dest:tainted_in_current)[1] +apply (simp add:is_created_proc_def) + +prefer 4 +apply (simp split:if_splits) +apply (rule impI|rule conjI)+ +apply (simp add:is_created_proc_def) +apply (auto simp:is_created_proc_def split:if_splits dest:tainted_in_current)[1] +apply (simp add:is_created_proc_def) + +prefer 4 +apply (auto simp:is_created_proc_def split:if_splits option.splits dest:tainted_in_current)[1] + +prefer 4 +apply (auto simp:is_created_proc_def split:if_splits option.splits dest:tainted_in_current enrich_proc_dup_ffd enrich_proc_dup_ffd')[1] + + + +lemma enrich_proc_dup_tainted: + "\is_created_proc s p; p' \ all_procs s; valid s\ + \ (O_proc p' \ tainted (enrich_proc s p p')) = (O_proc p \ tainted s)" +apply (induct s) +apply (simp add:is_created_proc_def) +apply (frule vt_grant_os, frule vd_cons) +apply (case_tac a) +apply (auto simp:is_created_proc_def)[1] + + +lemma enrich_proc_tainted: + "" + + +end + +end \ No newline at end of file diff -r 1ac0c3031ed2 -r 0a68c605ae79 no_shm_selinux/Flask.thy --- a/no_shm_selinux/Flask.thy Fri Dec 20 13:48:25 2013 +0800 +++ b/no_shm_selinux/Flask.thy Mon Dec 23 19:47:17 2013 +0800 @@ -542,7 +542,8 @@ where "flags_of_proc_fd [] p fd = init_oflags_of_proc_fd p fd" | "flags_of_proc_fd (Open p f flags fd iopt # \) p' fd' = - (if (p = p' \ fd = fd') then Some flags else flags_of_proc_fd \ p' fd')" + (if (p = p' \ fd = fd') then Some flags + else flags_of_proc_fd \ p' fd')" | "flags_of_proc_fd (CloseFd p fd # \) p' fd' = (if (p = p' \ fd = fd') then None else flags_of_proc_fd \ p' fd')" | "flags_of_proc_fd (CreateSock p af st fd ino # \) p' fd' = @@ -666,7 +667,7 @@ (******* admissable check by the OS *******) -fun os_grant :: "t_state \ t_event \ bool" +fun os_grant:: "t_state \ t_event \ bool" where "os_grant \ (Open p f flags fd inumopt) = ( case inumopt of