# HG changeset patch # User chunhan # Date 1387508280 -28800 # Node ID c9cfc416fa2da143dee3c75fc198acc4b63a9dd3 # Parent c09fcbcdb87137a6eb614e691635dcf01fa299dd fixed bug in enrich_proc diff -r c09fcbcdb871 -r c9cfc416fa2d no_shm_selinux/Enrich.thy --- a/no_shm_selinux/Enrich.thy Thu Dec 19 10:05:13 2013 +0800 +++ b/no_shm_selinux/Enrich.thy Fri Dec 20 10:58:00 2013 +0800 @@ -30,7 +30,7 @@ then Open dp f (remove_create_flag flags) fd opt # Open p f flags fd opt # (enrich_proc s tp dp) else Open p f flags fd opt # (enrich_proc s tp dp))" | "enrich_proc (CloseFd p fd # s) tp dp = ( - if (tp = p) + if (tp = p \ fd \ proc_file_fds s p) then CloseFd dp fd # CloseFd p fd # (enrich_proc s tp dp) else CloseFd p fd # (enrich_proc s tp dp))" (* @@ -163,6 +163,53 @@ by (simp add:inherit_fds_check_def) qed + +lemma enrich_inherit_fds_check_dup: + assumes grant: "inherit_fds_check s (up, nr, nt) p fds" and vd: "valid s" + and cfd2sfd: "\ fd. fd \ proc_file_fds s p \ cfd2sfd s' p' fd = cfd2sfd s p fd" + and fd_in: "fds' \ fds \ proc_file_fds s p" + shows "inherit_fds_check s' (up, nr, nt) p' fds'" +proof- + have "sectxts_of_fds s' p' fds' \ sectxts_of_fds s p fds" + proof- + have "\ fd sfd. \fd \ fds'; sectxt_of_obj s' (O_fd p' fd) = Some sfd\ + \ \ fd \ fds. sectxt_of_obj s (O_fd p fd) = Some sfd" + proof- + fix fd sfd + assume fd_in_fds': "fd \ fds'" + and sec: "sectxt_of_obj s' (O_fd p' fd) = Some sfd" + from fd_in_fds' fd_in + have fd_in_fds: "fd \ fds" and fd_in_cfds: "fd \ proc_file_fds s p" + by auto + from fd_in_cfds obtain f where ffd: "file_of_proc_fd s p fd = Some f" + by (auto simp:proc_file_fds_def) + moreover have "flags_of_proc_fd s p fd \ None" + using ffd vd by (auto dest:current_filefd_has_flags) + moreover have "cf2sfile s f \ None" + apply (rule notI) + apply (drule current_file_has_sfile') + using ffd + by (auto simp:vd is_file_in_current dest:file_of_pfd_is_file) + moreover have "sectxt_of_obj s (O_fd p fd) \ None" + using fd_in_cfds vd + apply (rule_tac notI) + by (auto dest!:current_has_sec' file_fds_subset_pfds[where p = p] intro:vd) + ultimately + have "sectxt_of_obj s (O_fd p fd) = Some sfd" + using fd_in_cfds cfd2sfd sec + apply (erule_tac x = fd in allE) + apply (auto simp:cfd2sfd_def split:option.splits) + done + thus "\ fd \ fds. sectxt_of_obj s (O_fd p fd) = Some sfd" + using fd_in_fds + by (rule_tac x = fd in bexI, auto) + qed + thus ?thesis by (auto simp:sectxts_of_fds_def) + qed + thus ?thesis using grant + by (auto simp:inherit_fds_check_def inherit_fds_check_ctxt_def) +qed + lemma not_all_procs_cons: "p \ all_procs (e # s) \ p \ all_procs s" by (case_tac e, auto) @@ -1105,15 +1152,22 @@ dest:not_all_procs_prop3 split:if_splits option.splits) done +lemma enrich_proc_dup_ffd': + "\file_of_proc_fd (enrich_proc s p p') p' fd = Some f; is_created_proc s p; p' \ all_procs s; valid s\ + \ file_of_proc_fd s p fd = Some f" +apply (induct s, simp add:is_created_proc_def) +apply (frule vt_grant_os, frule vd_cons) +apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def + dest:not_all_procs_prop3 split:if_splits option.splits) +done + lemma current_fflag_in_fds: "\flags_of_proc_fd s p fd = Some flag; valid s\ \ fd \ current_proc_fds s p" apply (induct s arbitrary:p) apply (simp add:flags_of_proc_fd.simps file_of_proc_fd.simps init_oflags_prop2) apply (frule vd_cons, frule vt_grant_os, case_tac a) apply (auto split:if_splits option.splits dest:proc_fd_in_fds) -apply (auto simp:proc_file_fds_def) -sorry - +done lemma current_fflag_has_ffd: "\flags_of_proc_fd s p fd = Some flag; valid s\ \ \ f. file_of_proc_fd s p fd = Some f" @@ -1121,17 +1175,27 @@ apply (simp add: file_of_proc_fd.simps init_fileflag_valid) apply (frule vd_cons, frule vt_grant_os, case_tac a) apply (auto split:if_splits option.splits dest:proc_fd_in_fds) -apply (auto simp:proc_file_fds_def) -sorry +done lemma enrich_proc_dup_fflags: - "\flags_of_proc_fd s p fd = Some f; is_created_proc s p; p' \ all_procs s; valid s\ - \ flags_of_proc_fd (enrich_proc s p p') p' fd = Some f" -apply (induct s, simp add:is_created_proc_def) + "\flags_of_proc_fd s p fd = Some flag; is_created_proc s p; p' \ all_procs s; valid s\ + \ flags_of_proc_fd (enrich_proc s p p') p' fd = Some (remove_create_flag flag) \ + flags_of_proc_fd (enrich_proc s p p') p' fd = Some flag" +apply (induct s arbitrary:p, simp add:is_created_proc_def) apply (frule vt_grant_os, frule vd_cons) -apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def +apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def is_creat_flag_def dest:not_all_procs_prop3 split:if_splits option.splits) -sorry +done + +lemma enrich_proc_dup_ffds: + "\is_created_proc s p; p' \ all_procs s; valid s\ + \ proc_file_fds (enrich_proc s p p') p' = proc_file_fds s p" +apply (auto simp:proc_file_fds_def) +apply (rule_tac x = f in exI) +apply (erule enrich_proc_dup_ffd', simp+) +apply (rule_tac x = f in exI) +apply (erule enrich_proc_dup_ffd, simp+) +done lemma enrich_proc_prop: "\valid s; is_created_proc s p; p' \ all_procs s\ @@ -1170,8 +1234,19 @@ (\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)" + (\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 @@ -1180,14 +1255,99 @@ 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\ + 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')" - apply (frule vt_grant_os, frule vt_grant) - apply (rule valid.intros(2)) - apply (simp_all split:option.splits add:sectxt_of_obj_simps is_file_simps) - apply (auto simp add:proc_file_fds_def)[1] - - sorry + 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) @@ -1200,13 +1360,30 @@ \valid (Open p f flags fd inum # enrich_proc s p p'); is_created_proc s p\ \ valid (Open p' f (remove_create_flag flags) fd inum # Open p f flags fd inum # enrich_proc s p p')" sorry - moreover have "\fd. \valid (CloseFd p fd # enrich_proc s p p'); is_created_proc s p\ + 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\ \ 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)" + 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'" + ultimately show "valid (CloseFd p' fd # CloseFd p fd # enrich_proc s p p')" + apply (rule valid.intros(2)) + apply (simp_all add:a1 a0 a2 pre') + + + sorry - thus ?thesis using created_cons vd_cons all_procs_cons +(* + thus ?thesis using created_cons vd_cons all_procs_cons apply (case_tac e) apply (auto simp:is_created_proc_simps split:if_splits) - + *) 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)