# HG changeset patch # User chunhan # Date 1388026610 -28800 # Node ID e79e3a8a4ceb3087db8f09a418526f71769a439d # Parent 0a68c605ae79481122771a5f2000ee73bcb0d489 enrich diff -r 0a68c605ae79 -r e79e3a8a4ceb no_shm_selinux/Enrich.thy --- a/no_shm_selinux/Enrich.thy Mon Dec 23 19:47:17 2013 +0800 +++ b/no_shm_selinux/Enrich.thy Thu Dec 26 10:56:50 2013 +0800 @@ -3,6 +3,7 @@ Temp begin +(* objects that need dynamic indexing, all nature-numbers *) datatype t_enrich_obj = E_proc "t_process" | E_file "t_file" @@ -13,6 +14,110 @@ context tainting_s begin +fun no_del_event:: "t_event list \ bool" +where + "no_del_event [] = True" +| "no_del_event (Kill p p' # \) = False" +| "no_del_event (Exit p # s) = False" +| "no_del_event (CloseFd p fd # \) = False" +| "no_del_event (UnLink p f # \) = False" +| "no_del_event (Rmdir p f # \) = False" +(* +| "no_del_event (Rename p f f' # \) = False" +*) +| "no_del_event (RemoveMsgq p q # \) = False" +(* +| "no_del_event (RecvMsg p q m # \) = False" +*) +| "no_del_event (_ # \) = no_del_event \" + +fun all_inums :: "t_state \ t_inode_num set" +where + "all_inums [] = current_inode_nums []" +| "all_inums (Open p f flags fd opt # s) = ( + case opt of + None \ all_inums s + | Some i \ all_inums s \ {i} )" +| "all_inums (Mkdir p f i # s) = (all_inums s \ {i})" +| "all_inums (CreateSock p af st fd i # s) = (all_inums s \ {i})" +| "all_inums (Accept p fd addr lport fd' i # s) = (all_inums s \ {i})" +| "all_inums (_ # s) = all_inums s" + +fun all_fds :: "t_state \ t_process \ t_fd set" +where + "all_fds [] = init_fds_of_proc" +| "all_fds (Open p f flags fd ipt # s) = (all_fds s) (p := all_fds s p \ {fd})" +| "all_fds (CreateSock p sf st fd i # s) = (all_fds s) (p := all_fds s p \ {fd})" +| "all_fds (Accept p fd' raddr port fd i # s) = (all_fds s) (p := all_fds s p \ {fd})" +| "all_fds (Clone p p' fds # s) = (all_fds s) (p' := fds)" +| "all_fds (_ # s) = all_fds s" + +fun all_msgqs:: "t_state \ t_msgq set" +where + "all_msgqs [] = init_msgqs" +| "all_msgqs (CreateMsgq p q # s) = all_msgqs s \ {q}" +| "all_msgqs (e # s) = all_msgqs s" + +fun all_msgs:: "t_state \ t_msgq \ t_msg set" +where + "all_msgs [] q = set (init_msgs_of_queue q)" +| "all_msgs (CreateMsgq p q # s) q' = (if q' = q then {} else all_msgs s q')" +| "all_msgs (SendMsg p q m # s) q' = (if q' = q then all_msgs s q \ {m} else all_msgs s q')" +| "all_msgs (_ # s) q' = all_msgs s q'" + +fun all_files:: "t_state \ t_file set" +where + "all_files [] = init_files " +| "all_files (Open p f flags fd opt # s) = (if opt = None then all_files s else (all_files s \ {f}))" +| "all_files (Mkdir p f inum # s) = all_files s \ {f}" +| "all_files (LinkHard p f f' # s) = all_files s \ {f'}" +| "all_files (e # s) = all_files s" + +fun notin_all:: "t_state \ t_enrich_obj \ bool" +where + "notin_all s (E_proc p) = (p \ all_procs s)" +| "notin_all s (E_file f) = (f \ all_files s \ (\ pf. parent f = Some pf \ is_dir s pf))" +| "notin_all s (E_fd p fd) = (fd \ all_fds s p)" +| "notin_all s (E_inum i) = (i \ all_inums s)" +| "notin_all s (E_msgq q) = (q \ all_msgqs s)" +| "notin_all s (E_msg q m) = (m \ all_msgs s q)" + +lemma not_all_procs_cons: + "p \ all_procs (e # s) \ p \ all_procs s" +by (case_tac e, auto) + +lemma not_all_procs_prop: + "\p' \ all_procs s; p \ current_procs s; valid s\ \ p' \ p" +apply (induct s, rule notI, simp) +apply (frule vt_grant_os, frule vd_cons, frule not_all_procs_cons, simp, rule notI) +apply (case_tac a, auto) +done + +lemma not_all_procs_prop2: + "p' \ all_procs s \ p' \ init_procs" +apply (induct s, simp) +by (case_tac a, auto) + +lemma not_all_procs_prop3: + "p' \ all_procs s \ p' \ current_procs s" +apply (induct s, simp) +by (case_tac a, auto) + +fun enrich_not_alive :: "t_state \ t_enrich_obj \ t_enrich_obj \ bool" +where + "enrich_not_alive s obj (E_file f) = (f \ current_files s \ E_file f \ obj)" +| "enrich_not_alive s obj (E_proc p) = (p \ current_procs s \ E_proc p \ obj)" +| "enrich_not_alive s obj (E_fd p fd) = ((p \ current_procs s \ fd \ current_proc_fds s p) \ E_fd p fd \ obj)" +| "enrich_not_alive s obj (E_msgq q) = (q \ current_msgqs s \ E_msgq q \ obj)" +| "enrich_not_alive s obj (E_inum i) = (i \ current_inode_nums s \ E_inum i \ obj)" +| "enrich_not_alive s obj (E_msg q m) = ((q \ current_msgqs s \ m \ set (msgs_of_queue s q)) \ E_msg q m \ obj)" + +lemma file_has_parent: "\is_file s f; valid s\ \ \ pf. is_dir s pf \ parent f = Some pf" +apply (case_tac f) +apply (simp, drule root_is_dir', simp+) +apply (simp add:parentf_is_dir_prop2) +done + (* enrich s target_proc duplicated_pro *) fun enrich_proc :: "t_state \ t_process \ t_process \ t_state" where @@ -33,10 +138,12 @@ if (tp = p) then ReadFile dp fd # ReadFile p fd # (enrich_proc s tp dp) else ReadFile p fd # (enrich_proc s tp dp))" +(* | "enrich_proc (CloseFd p fd # s) tp dp = ( 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))" +*) (* | "enrich_proc (Attach p h flag # s) tp dp = ( if (tp = p) @@ -47,14 +154,41 @@ then Detach dp h # Detach p h # (enrich_proc s tp dp) else Detach p h # (enrich_proc s tp dp))" *) +(* | "enrich_proc (Kill p p' # s) tp dp = ( if (tp = p') then Kill p p' # s else Kill p p' # (enrich_proc s tp dp))" | "enrich_proc (Exit p # s) tp dp = ( if (tp = p) then Exit p # s else Exit p # (enrich_proc s tp dp))" +*) | "enrich_proc (e # s) tp dp = e # (enrich_proc s tp dp)" +definition is_created_proc:: "t_state \ t_process \ bool" +where + "is_created_proc s p \ p \ current_procs s \ (p \ init_procs \ died (O_proc p) s)" + +definition is_created_proc':: "t_state \ t_process \ bool" +where + "is_created_proc' s p \ p \ current_procs s \ p \ init_procs" + +lemma no_del_died: + "\no_del_event s; died obj s\ \ (\ p fd. obj = O_fd p fd \ obj = O_tcp_sock (p, fd) \ obj = O_udp_sock (p, fd)) + \ (\ q m. obj = O_msg q m) " +apply (induct s) +apply simp +apply (case_tac a) +apply (auto split:option.splits) +done + +lemma no_del_created_eq: + "no_del_event s \ is_created_proc s p = is_created_proc' s p" +apply (induct s) +apply (simp add:is_created_proc_def is_created_proc'_def) +apply (case_tac a) +apply (auto simp add:is_created_proc_def is_created_proc'_def dest:no_del_died) +done + lemma enrich_search_check: assumes grant: "search_check s (up, rp, tp) f" and cf2sf: "\ f. f \ current_files s \ cf2sfile s' f = cf2sfile s f" @@ -214,37 +348,11 @@ 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) - -lemma not_all_procs_prop: - "\p' \ all_procs s; p \ current_procs s; valid s\ \ p' \ p" -apply (induct s, rule notI, simp) -apply (frule vt_grant_os, frule vd_cons, frule not_all_procs_cons, simp, rule notI) -apply (case_tac a, auto) -done - -fun enrich_not_alive :: "t_state \ t_enrich_obj \ bool" -where - "enrich_not_alive s (E_file f) = (f \ current_files s)" -| "enrich_not_alive s (E_proc p) = (p \ current_procs s)" -| "enrich_not_alive s (E_fd p fd) = (p \ current_procs s \ fd \ current_proc_fds s p)" -| "enrich_not_alive s (E_msgq q) = (q \ current_msgqs s)" -| "enrich_not_alive s (E_inum i) = (i \ current_inode_nums s)" -| "enrich_not_alive s (E_msg q m) = (q \ current_msgqs s \ m \ set (msgs_of_queue s q))" - -lemma file_has_parent: "\is_file s f; valid s\ \ \ pf. is_dir s pf \ parent f = Some pf" -apply (case_tac f) -apply (simp, drule root_is_dir', simp+) -apply (simp add:parentf_is_dir_prop2) -done lemma enrich_valid_intro_cons: - assumes vs': "valid s'" - and os: "os_grant s e" and grant: "grant s e" and vd: "valid s" + assumes vs': "valid s'" and vd': "valid (e # s)" and alive: "\ obj. alive s obj \ alive s' obj" - and alive': "\ obj. enrich_not_alive s obj \ enrich_not_alive s' obj" + and alive': "\ obj. enrich_not_alive s obj' obj \ enrich_not_alive s' obj' obj" and hungs: "files_hung_by_del s' = files_hung_by_del s" and cp2sp: "\ p. p \ current_procs s \ cp2sproc s' p = cp2sproc s p" and cf2sf: "\ f. f \ current_files s \ cf2sfile s' f = cf2sfile s f" @@ -254,244 +362,569 @@ and sms_remain: "\ q. msgs_of_queue s' q = msgs_of_queue s q" (* and empty_remain: "\ f. dir_is_empty s f \ dir_is_empty s' f" *) and cfd2sfd: "\ p fd. fd \ proc_file_fds s p \ cfd2sfd s' p fd = cfd2sfd s p fd" + and nodel: "no_del_event (e # s)" + and notin_all: "notin_all (e # s) obj'" shows "valid (e # s')" -proof (cases e) - case (Execve p f fds) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:Execve) - have f_in: "is_file s' f" using os alive - apply (erule_tac x = "O_file f" in allE) - by (auto simp:Execve) - have fd_in: "fds \ proc_file_fds s' p" using os alive ffd_remain - by (auto simp:Execve proc_file_fds_def) - have "os_grant s' e" using p_in f_in fd_in by (simp add:Execve) - moreover have "grant s' e" - 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 add:Execve 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 add:Execve split:option.splits del:npctxt_execve.simps, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:Execve co2sobj.simps cp2sproc_def split:option.splits) - from os have f_in': "is_file s f" by (simp add:Execve) - from vd os have "\ sf. cf2sfile s f = Some sf" - by (auto dest!:is_file_in_current current_file_has_sfile simp:Execve) - hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf - apply (erule_tac x = f in allE) - apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) - apply (case_tac f, simp) - apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) - done - have "inherit_fds_check s' (pu, nr, nt) p fds" - proof- - have "fds \ proc_file_fds s' p" using os ffd_remain Execve - by (auto simp:proc_file_fds_def) - thus ?thesis using Execve grant vd cfd2sfd p1 p2 p3 os - apply (rule_tac s = s in enrich_inherit_fds_check) - by (simp_all split:option.splits) - qed - moreover have "search_check s' (pu, rp, tp) f" - using p1 p2 p2' vd cf2sf f_in' grant Execve p3 f_in - apply (rule_tac s = s in enrich_search_check) - by (simp_all split:option.splits) - ultimately show ?thesis using p1' p2' p3 - apply (simp add:Execve split:option.splits) - using grant Execve p1 p2 - by (simp add:Execve grant p1 p2) - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (Clone p p' fds) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:Clone) - have p'_not_in: "p' \ current_procs s'" using os alive' - apply (erule_tac x = "E_proc p'" in allE) - by (auto simp:Clone) - have fd_in: "fds \ proc_file_fds s' p" using os alive ffd_remain - by (auto simp:Clone proc_file_fds_def) - have "os_grant s' e" using p_in p'_not_in fd_in by (simp add:Clone) - moreover have "grant s' e" - proof- - from grant obtain up rp tp - where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" - apply (simp add:Clone split:option.splits) - by (case_tac a, auto) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:Clone co2sobj.simps cp2sproc_def split:option.splits) - have p2: "inherit_fds_check s' (up, rp, tp) p fds" - proof- - have "fds \ proc_file_fds s' p" using os ffd_remain Clone - by (auto simp:proc_file_fds_def) - thus ?thesis using Clone grant vd cfd2sfd p1 os - apply (rule_tac s = s in enrich_inherit_fds_check) - by (simp_all split:option.splits) - qed - show ?thesis using p1 p2 p1' grant - by (simp add:Clone) - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (Kill p p') - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:Kill) - have p'_in: "p' \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p'" in allE) - by (auto simp:Kill) - have "os_grant s' e" using p_in p'_in by (simp add:Kill) - moreover have "grant s' e" - proof- - from grant obtain up rp tp up' rp' tp' - where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" - and p'1: "sectxt_of_obj s (O_proc p') = Some (up', rp', tp')" - apply (simp add:Kill split:option.splits) - by (case_tac a, case_tac aa, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:Kill co2sobj.simps cp2sproc_def split:option.splits) - from p'1 have p'1': "sectxt_of_obj s' (O_proc p') = Some (up', rp', tp')" - using os cp2sp - apply (erule_tac x = p' in allE) - by (auto simp:Kill co2sobj.simps cp2sproc_def split:option.splits) - show ?thesis using p1 p'1 p1' p'1' grant - by (simp add:Kill) - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (Ptrace p p') - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:Ptrace) - have p'_in: "p' \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p'" in allE) - by (auto simp:Ptrace) - have "os_grant s' e" using p_in p'_in by (simp add:Ptrace) - moreover have "grant s' e" - proof- - from grant obtain up rp tp up' rp' tp' - where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" - and p'1: "sectxt_of_obj s (O_proc p') = Some (up', rp', tp')" - apply (simp add:Ptrace split:option.splits) - by (case_tac a, case_tac aa, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:Ptrace co2sobj.simps cp2sproc_def split:option.splits) - from p'1 have p'1': "sectxt_of_obj s' (O_proc p') = Some (up', rp', tp')" - using os cp2sp - apply (erule_tac x = p' in allE) - by (auto simp:Ptrace co2sobj.simps cp2sproc_def split:option.splits) - show ?thesis using p1 p'1 p1' p'1' grant - by (simp add:Ptrace) - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (Exit p) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:Exit) - have "os_grant s' e" using p_in by (simp add:Exit) - moreover have "grant s' e" - by (simp add:Exit) - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (Open p f flags fd opt) +proof- + from vd' have os: "os_grant s e" and grant: "grant s e" and vd: "valid s" + by (auto dest:vt_grant_os vt_grant vd_cons) show ?thesis - proof (cases opt) - case None + proof (cases e) + case (Execve p f fds) have p_in: "p \ current_procs s'" using os alive apply (erule_tac x = "O_proc p" in allE) - by (auto simp:Open None) + by (auto simp:Execve) have f_in: "is_file s' f" using os alive apply (erule_tac x = "O_file f" in allE) - by (auto simp:Open None) - have fd_not_in: "fd \ current_proc_fds s' p" - using os alive' p_in - apply (erule_tac x = "E_fd p fd" in allE) - by (simp add:Open None) - have "os_grant s' e" using p_in f_in fd_not_in os - by (simp add:Open None) - moreover have "grant s' e" + by (auto simp:Execve) + have fd_in: "fds \ proc_file_fds s' p" using os alive ffd_remain + by (auto simp:Execve proc_file_fds_def) + have "os_grant s' e" using p_in f_in fd_in by (simp add:Execve) + moreover have "grant s' e" 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)" - apply (simp add:Open None split:option.splits) - by (case_tac a, case_tac aa, blast) + by (simp add:Execve 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 add:Execve split:option.splits del:npctxt_execve.simps, blast) from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" using os cp2sp apply (erule_tac x = p in allE) - by (auto simp:Open None co2sobj.simps cp2sproc_def split:option.splits) - from os have f_in': "is_file s f" by (simp add:Open None) + by (auto simp:Execve co2sobj.simps cp2sproc_def split:option.splits) + from os have f_in': "is_file s f" by (simp add:Execve) from vd os have "\ sf. cf2sfile s f = Some sf" - by (auto dest!:is_file_in_current current_file_has_sfile simp:Open None) + by (auto dest!:is_file_in_current current_file_has_sfile simp:Execve) hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf apply (erule_tac x = f in allE) apply (auto dest:is_file_in_current simp:cf2sfile_def 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 s' (up, rp, tp) f" - using p1 p2 p2' vd cf2sf f_in' grant Open None f_in + have "inherit_fds_check s' (pu, nr, nt) p fds" + proof- + have "fds \ proc_file_fds s' p" using os ffd_remain Execve + by (auto simp:proc_file_fds_def) + thus ?thesis using Execve grant vd cfd2sfd p1 p2 p3 os + apply (rule_tac s = s in enrich_inherit_fds_check) + by (simp_all split:option.splits) + qed + moreover have "search_check s' (pu, rp, tp) f" + using p1 p2 p2' vd cf2sf f_in' grant Execve p3 f_in apply (rule_tac s = s in enrich_search_check) by (simp_all split:option.splits) - thus ?thesis using p1' p2' - apply (simp add:Open None split:option.splits) - using grant Open None p1 p2 + ultimately show ?thesis using p1' p2' p3 + apply (simp add:Execve split:option.splits) + using grant Execve p1 p2 + by (simp add:Execve grant p1 p2) + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (Clone p p' fds) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:Clone) + have p'_not_in: "p' \ current_procs s'" using alive' notin_all os Clone + apply (erule_tac x = "E_proc p'" in allE) + apply (auto dest:not_all_procs_prop3) + done + have fd_in: "fds \ proc_file_fds s' p" using os alive ffd_remain + by (auto simp:Clone proc_file_fds_def) + have "os_grant s' e" using p_in p'_not_in fd_in by (simp add:Clone) + moreover have "grant s' e" + proof- + from grant obtain up rp tp + where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" + apply (simp add:Clone split:option.splits) + by (case_tac a, auto) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:Clone co2sobj.simps cp2sproc_def split:option.splits) + have p2: "inherit_fds_check s' (up, rp, tp) p fds" + proof- + have "fds \ proc_file_fds s' p" using os ffd_remain Clone + by (auto simp:proc_file_fds_def) + thus ?thesis using Clone grant vd cfd2sfd p1 os + apply (rule_tac s = s in enrich_inherit_fds_check) + by (simp_all split:option.splits) + qed + show ?thesis using p1 p2 p1' grant + by (simp add:Clone) + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (Kill p p') + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:Kill) + have p'_in: "p' \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p'" in allE) + by (auto simp:Kill) + have "os_grant s' e" using p_in p'_in by (simp add:Kill) + moreover have "grant s' e" + proof- + from grant obtain up rp tp up' rp' tp' + where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" + and p'1: "sectxt_of_obj s (O_proc p') = Some (up', rp', tp')" + apply (simp add:Kill split:option.splits) + by (case_tac a, case_tac aa, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:Kill co2sobj.simps cp2sproc_def split:option.splits) + from p'1 have p'1': "sectxt_of_obj s' (O_proc p') = Some (up', rp', tp')" + using os cp2sp + apply (erule_tac x = p' in allE) + by (auto simp:Kill co2sobj.simps cp2sproc_def split:option.splits) + show ?thesis using p1 p'1 p1' p'1' grant + by (simp add:Kill) + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (Ptrace p p') + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:Ptrace) + have p'_in: "p' \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p'" in allE) + by (auto simp:Ptrace) + have "os_grant s' e" using p_in p'_in by (simp add:Ptrace) + moreover have "grant s' e" + proof- + from grant obtain up rp tp up' rp' tp' + where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" + and p'1: "sectxt_of_obj s (O_proc p') = Some (up', rp', tp')" + apply (simp add:Ptrace split:option.splits) + by (case_tac a, case_tac aa, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:Ptrace co2sobj.simps cp2sproc_def split:option.splits) + from p'1 have p'1': "sectxt_of_obj s' (O_proc p') = Some (up', rp', tp')" + using os cp2sp + apply (erule_tac x = p' in allE) + by (auto simp:Ptrace co2sobj.simps cp2sproc_def split:option.splits) + show ?thesis using p1 p'1 p1' p'1' grant + by (simp add:Ptrace) + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (Exit p) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:Exit) + have "os_grant s' e" using p_in by (simp add:Exit) + moreover have "grant s' e" + by (simp add:Exit) + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (Open p f flags fd opt) + show ?thesis + proof (cases opt) + case None + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:Open None) + have f_in: "is_file s' f" using os alive + apply (erule_tac x = "O_file f" in allE) + by (auto simp:Open None) + have fd_not_in: "fd \ current_proc_fds s' p" + using os alive' p_in notin_all Open None + apply (erule_tac x = "E_fd p fd" in allE) + apply (case_tac obj') + by auto + have "os_grant s' e" using p_in f_in fd_not_in os + by (simp add:Open None) + moreover have "grant s' e" + 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)" + apply (simp add:Open None split:option.splits) + by (case_tac a, case_tac aa, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:Open None co2sobj.simps cp2sproc_def split:option.splits) + from os have f_in': "is_file s f" by (simp add:Open None) + from vd os have "\ sf. cf2sfile s f = Some sf" + by (auto dest!:is_file_in_current current_file_has_sfile simp:Open None) + hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf + apply (erule_tac x = f in allE) + apply (auto dest:is_file_in_current simp:cf2sfile_def 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 s' (up, rp, tp) f" + using p1 p2 p2' vd cf2sf f_in' grant Open None f_in + apply (rule_tac s = s in enrich_search_check) + by (simp_all split:option.splits) + thus ?thesis using p1' p2' + apply (simp add:Open None split:option.splits) + using grant Open None p1 p2 + by simp + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (Some inum) + from os obtain pf where pf_in_s: "is_dir s pf" and parent: "parent f = Some pf" + by (auto simp:Open Some) + have pf_in: "is_dir s' pf" using pf_in_s alive + apply (erule_tac x = "O_dir pf" in allE) + by simp + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:Open Some) + have f_not_in: "f \ current_files s'" using os alive' Open Some notin_all + apply (erule_tac x = "E_file f" in allE) + apply (case_tac obj') + by auto + have fd_not_in: "fd \ current_proc_fds s' p" + using os alive' p_in Open Some notin_all + apply (erule_tac x = "E_fd p fd" in allE) + apply (case_tac obj', auto) + done + have inum_not_in: "inum \ current_inode_nums s'" + using os alive' Open Some notin_all + apply (erule_tac x = "E_inum inum" in allE) + by (case_tac obj', auto) + have "os_grant s' e" using p_in pf_in parent f_not_in fd_not_in inum_not_in os + by (simp add:Open Some hungs) + moreover have "grant s' e" + proof- + from grant parent 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_dir pf) = Some (uf, rf, tf)" + apply (simp add:Open Some split:option.splits) + by (case_tac a, case_tac aa, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:Open Some co2sobj.simps cp2sproc_def split:option.splits) + from vd os pf_in_s have "\ sf. cf2sfile s pf = Some sf" + by (auto dest!:is_dir_in_current current_file_has_sfile simp:Open Some) + hence p2': "sectxt_of_obj s' (O_dir pf) = Some (uf, rf, tf)" using p2 cf2sf pf_in pf_in_s + apply (erule_tac x = pf in allE) + apply (erule exE, frule_tac s = s in is_dir_in_current, simp) + apply (drule is_dir_not_file, drule is_dir_not_file) + apply (auto simp:cf2sfile_def split:option.splits) + apply (case_tac pf, simp_all) + by (simp add:sroot_def root_sec_remains vd vs') + have "search_check s' (up, rp, tp) pf" + using p1 p2 p2' vd cf2sf pf_in grant Open Some pf_in_s parent vs' + apply (rule_tac s = s in enrich_search_check') + by (simp_all split:option.splits) + thus ?thesis using p1' p2' parent + apply (simp add:Open Some split:option.splits) + using grant Open Some p1 p2 + by simp + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + qed + next + case (ReadFile p fd) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:ReadFile) + have fd_in: "fd \ current_proc_fds s' p" using os alive + apply (erule_tac x = "O_fd p fd" in allE) + by (auto simp:ReadFile) + obtain f where ffd: "file_of_proc_fd s p fd = Some f" + using os ReadFile by auto + hence f_in_s: "is_file s f" using vd + by (auto intro:file_of_pfd_is_file) + obtain flags where fflag: "flags_of_proc_fd s p fd = Some flags" + using os ReadFile by auto + have ffd_in: "file_of_proc_fd s' p fd = Some f" + using ffd_remain ffd by auto + hence f_in: "is_file s' f" using vs' + by (auto intro:file_of_pfd_is_file) + have flags_in: "flags_of_proc_fd s' p fd = Some flags" + using fflags_remain fflag by auto + have "os_grant s' e" using p_in fd_in ffd_in flags_in fflag os f_in + by (auto simp add:ReadFile is_file_in_current) + moreover have "grant s' e" + proof- + from grant ffd obtain up rp tp uf rf tf ufd rfd tfd + 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)" + and p3: "sectxt_of_obj s (O_fd p fd) = Some (ufd, rfd, tfd)" + apply (simp add:ReadFile split:option.splits) + by (case_tac a, case_tac aa, case_tac ab, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:ReadFile co2sobj.simps cp2sproc_def split:option.splits) + from vd f_in_s have "\ sf. cf2sfile s f = Some sf" + by (auto dest!:is_file_in_current current_file_has_sfile) + hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in_s p2 cf2sf + apply (erule_tac x = f in allE) + apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) + apply (case_tac f, simp) + apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) + done + have p3': "sectxt_of_obj s' (O_fd p fd) = Some (ufd, rfd, tfd)" + using cfd2sfd ffd_in ffd p3 f_in f_in_s vd + apply (erule_tac x = p in allE) + apply (erule_tac x = fd in allE) + apply (simp add:proc_file_fds_def) + apply (auto simp:cfd2sfd_def fflag flags_in p3 split:option.splits + dest!:current_file_has_sfile' simp:is_file_in_current) + done + show ?thesis using p1' p2' p3' ffd_in ffd + apply (simp add:ReadFile split:option.splits) + using grant p1 p2 p3 ReadFile by simp qed ultimately show ?thesis using vs' by (erule_tac valid.intros(2), simp+) next - case (Some inum) + case (WriteFile p fd) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:WriteFile) + have fd_in: "fd \ current_proc_fds s' p" using os alive + apply (erule_tac x = "O_fd p fd" in allE) + by (auto simp:WriteFile) + obtain f where ffd: "file_of_proc_fd s p fd = Some f" + using os WriteFile by auto + hence f_in_s: "is_file s f" using vd + by (auto intro:file_of_pfd_is_file) + obtain flags where fflag: "flags_of_proc_fd s p fd = Some flags" + using os WriteFile by auto + have ffd_in: "file_of_proc_fd s' p fd = Some f" + using ffd_remain ffd by auto + hence f_in: "is_file s' f" using vs' + by (auto intro:file_of_pfd_is_file) + have flags_in: "flags_of_proc_fd s' p fd = Some flags" + using fflags_remain fflag by auto + have "os_grant s' e" using p_in fd_in ffd_in flags_in fflag os f_in + by (auto simp add:WriteFile is_file_in_current) + moreover have "grant s' e" + proof- + from grant ffd obtain up rp tp uf rf tf ufd rfd tfd + 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)" + and p3: "sectxt_of_obj s (O_fd p fd) = Some (ufd, rfd, tfd)" + apply (simp add:WriteFile split:option.splits) + by (case_tac a, case_tac aa, case_tac ab, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:WriteFile co2sobj.simps cp2sproc_def split:option.splits) + from vd f_in_s have "\ sf. cf2sfile s f = Some sf" + by (auto dest!:is_file_in_current current_file_has_sfile) + hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in_s p2 cf2sf + apply (erule_tac x = f in allE) + apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) + apply (case_tac f, simp) + apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) + done + have p3': "sectxt_of_obj s' (O_fd p fd) = Some (ufd, rfd, tfd)" + using cfd2sfd ffd_in ffd p3 f_in f_in_s vd + apply (erule_tac x = p in allE) + apply (erule_tac x = fd in allE) + apply (simp add:proc_file_fds_def) + apply (auto simp:cfd2sfd_def fflag flags_in p3 split:option.splits + dest!:current_file_has_sfile' simp:is_file_in_current) + done + show ?thesis using p1' p2' p3' ffd_in ffd + apply (simp add:WriteFile split:option.splits) + using grant p1 p2 p3 WriteFile + by simp + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (CloseFd p fd) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:CloseFd) + have fd_in: "fd \ current_proc_fds s' p" using os alive + apply (erule_tac x = "O_fd p fd" in allE) + by (auto simp:CloseFd) + have "os_grant s' e" using p_in fd_in + by (auto simp add:CloseFd) + moreover have "grant s' e" + by(simp add:CloseFd) + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (UnLink p f) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:UnLink) + have f_in: "is_file s' f" using os alive + apply (erule_tac x = "O_file f" in allE) + by (auto simp:UnLink) + from os vd obtain pf where pf_in_s: "is_dir s pf" + and parent: "parent f = Some pf" + by (auto simp:UnLink dest!:file_has_parent) + from pf_in_s alive have pf_in: "is_dir s' pf" + apply (erule_tac x = "O_dir pf" in allE) + by (auto simp:UnLink) + have "os_grant s' e" using p_in f_in os + by (simp add:UnLink hungs) + moreover have "grant s' e" + proof- + from grant parent obtain up rp tp uf rf tf upf rpf tpf + 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)" + and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" + apply (simp add:UnLink split:option.splits) + by (case_tac a, case_tac aa, case_tac ab, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:UnLink co2sobj.simps cp2sproc_def split:option.splits) + from vd os pf_in_s have "\ sf. cf2sfile s f = Some sf" + by (auto dest!:is_file_in_current current_file_has_sfile simp:UnLink) + hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" + using p2 cf2sf f_in os parent + apply (erule_tac x = f in allE) + apply (erule exE, clarsimp simp:UnLink) + apply (frule_tac s = s in is_file_in_current, simp) + by (auto simp:cf2sfile_def split:option.splits) + from vd os pf_in_s have "\ sf. cf2sfile s pf = Some sf" + by (auto dest!:is_dir_in_current current_file_has_sfile simp:UnLink) + hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s + apply (erule_tac x = pf in allE) + apply (erule exE, frule_tac s = s in is_dir_in_current, simp) + apply (drule is_dir_not_file, drule is_dir_not_file) + apply (auto simp:cf2sfile_def split:option.splits) + apply (case_tac pf, simp_all) + by (simp add:sroot_def root_sec_remains vd vs') + have "search_check s' (up, rp, tp) f" + using p1 p2 p2' vd cf2sf f_in grant UnLink os parent vs' + apply (rule_tac s = s in enrich_search_check) + by (simp_all split:option.splits) + thus ?thesis using p1' p2' p3' parent + apply (simp add:UnLink split:option.splits) + using grant UnLink p1 p2 p3 + by simp + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (Rmdir p f) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:Rmdir) + have f_in: "is_dir s' f" using os alive + apply (erule_tac x = "O_dir f" in allE) + by (auto simp:Rmdir dir_is_empty_def) + have not_root: "f \ []" using os + by (auto simp:Rmdir) + from os vd obtain pf where pf_in_s: "is_dir s pf" + and parent: "parent f = Some pf" + apply (auto simp:Rmdir dir_is_empty_def) + apply (case_tac f, simp+) + apply (drule parentf_is_dir_prop1, auto) + done + from pf_in_s alive have pf_in: "is_dir s' pf" + apply (erule_tac x = "O_dir pf" in allE) + by (auto simp:Rmdir) + have empty_in: "dir_is_empty s' f" using os Rmdir notin_all + apply (clarsimp simp add:dir_is_empty_def f_in) + using alive' + apply (erule_tac x = "E_file f'" in allE) + apply simp + apply (erule disjE) + apply (erule_tac x = f' in allE, simp) + apply (case_tac obj', simp_all) + apply (clarsimp) + apply (drule_tac f' = f in parent_ancen) + apply (simp, rule notI, simp add:noJ_Anc) + apply (case_tac "f = pf") + using vd' Rmdir + apply (simp_all add:is_dir_rmdir) + apply (erule_tac x = pf in allE) + apply (drule_tac f = pf in is_dir_in_current) + apply (simp add:noJ_Anc) + done + have "os_grant s' e" using p_in f_in os empty_in + by (simp add:Rmdir hungs) + moreover have "grant s' e" + proof- + from grant parent obtain up rp tp uf rf tf upf rpf tpf + where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" + and p2: "sectxt_of_obj s (O_dir f) = Some (uf, rf, tf)" + and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" + apply (simp add:Rmdir split:option.splits) + by (case_tac a, case_tac aa, case_tac ab, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:Rmdir co2sobj.simps cp2sproc_def split:option.splits) + from vd os pf_in_s have "\ sf. cf2sfile s f = Some sf" + by (auto dest!:is_dir_in_current current_file_has_sfile simp:dir_is_empty_def Rmdir) + hence p2': "sectxt_of_obj s' (O_dir f) = Some (uf, rf, tf)" + using p2 cf2sf f_in os parent + apply (erule_tac x = f in allE) + apply (erule exE, clarsimp simp:Rmdir dir_is_empty_def) + apply (frule_tac s = s in is_dir_in_current, simp) + apply (drule is_dir_not_file, drule is_dir_not_file) + by (auto simp:cf2sfile_def split:option.splits) + from vd os pf_in_s have "\ sf. cf2sfile s pf = Some sf" + by (auto dest!:is_dir_in_current current_file_has_sfile simp:Rmdir) + hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s + apply (erule_tac x = pf in allE) + apply (erule exE, frule_tac s = s in is_dir_in_current, simp) + apply (drule is_dir_not_file, drule is_dir_not_file) + apply (auto simp:cf2sfile_def split:option.splits) + apply (case_tac pf, simp_all) + by (simp add:sroot_def root_sec_remains vd vs') + have "search_check s' (up, rp, tp) f" + using p1 p2 p2' vd cf2sf f_in grant Rmdir os parent vs' + apply (rule_tac s = s in enrich_search_check') + by (simp_all add:dir_is_empty_def split:option.splits) + thus ?thesis using p1' p2' p3' parent + apply (simp add:Rmdir split:option.splits) + using grant Rmdir p1 p2 p3 + by simp + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (Mkdir p f inum) from os obtain pf where pf_in_s: "is_dir s pf" and parent: "parent f = Some pf" - by (auto simp:Open Some) + by (auto simp:Mkdir) have pf_in: "is_dir s' pf" using pf_in_s alive apply (erule_tac x = "O_dir pf" in allE) by simp have p_in: "p \ current_procs s'" using os alive apply (erule_tac x = "O_proc p" in allE) - by (auto simp:Open Some) - have f_not_in: "f \ current_files s'" using os alive' + by (auto simp:Mkdir) + have f_not_in: "f \ current_files s'" using os alive' Mkdir notin_all apply (erule_tac x = "E_file f" in allE) - by (auto simp:Open Some) - have fd_not_in: "fd \ current_proc_fds s' p" - using os alive' p_in - apply (erule_tac x = "E_fd p fd" in allE) - by (simp add:Open Some) + by (auto) have inum_not_in: "inum \ current_inode_nums s'" - using os alive' + using os alive' Mkdir notin_all apply (erule_tac x = "E_inum inum" in allE) - by (simp add:Open Some) - have "os_grant s' e" using p_in pf_in parent f_not_in fd_not_in inum_not_in os - by (simp add:Open Some hungs) + by (auto) + have "os_grant s' e" using p_in pf_in parent f_not_in os inum_not_in + by (simp add:Mkdir hungs) moreover have "grant s' e" proof- from grant parent 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_dir pf) = Some (uf, rf, tf)" - apply (simp add:Open Some split:option.splits) + apply (simp add:Mkdir split:option.splits) by (case_tac a, case_tac aa, blast) from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" using os cp2sp apply (erule_tac x = p in allE) - by (auto simp:Open Some co2sobj.simps cp2sproc_def split:option.splits) + by (auto simp:Mkdir co2sobj.simps cp2sproc_def split:option.splits) from vd os pf_in_s have "\ sf. cf2sfile s pf = Some sf" - by (auto dest!:is_dir_in_current current_file_has_sfile simp:Open Some) + by (auto dest!:is_dir_in_current current_file_has_sfile simp:Mkdir) hence p2': "sectxt_of_obj s' (O_dir pf) = Some (uf, rf, tf)" using p2 cf2sf pf_in pf_in_s apply (erule_tac x = pf in allE) apply (erule exE, frule_tac s = s in is_dir_in_current, simp) @@ -500,610 +933,296 @@ apply (case_tac pf, simp_all) by (simp add:sroot_def root_sec_remains vd vs') have "search_check s' (up, rp, tp) pf" - using p1 p2 p2' vd cf2sf pf_in grant Open Some pf_in_s parent vs' + using p1 p2 p2' vd cf2sf pf_in grant Mkdir pf_in_s parent vs' apply (rule_tac s = s in enrich_search_check') + apply (simp_all split:option.splits) + done + thus ?thesis using p1' p2' parent + apply (simp add:Mkdir split:option.splits) + using grant Mkdir p1 p2 + apply simp + done + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (LinkHard p f f') + from os obtain pf where pf_in_s: "is_dir s pf" and parent: "parent f' = Some pf" + by (auto simp:LinkHard) + have pf_in: "is_dir s' pf" using pf_in_s alive + apply (erule_tac x = "O_dir pf" in allE) + by simp + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:LinkHard) + have f'_not_in: "f' \ current_files s'" using os alive' LinkHard notin_all + apply (erule_tac x = "E_file f'" in allE) + by (auto simp:LinkHard) + have f_in: "is_file s' f" using os alive + apply (erule_tac x = "O_file f" in allE) + by (auto simp:LinkHard) + have "os_grant s' e" using p_in pf_in parent os f_in f'_not_in + by (simp add:LinkHard hungs) + moreover have "grant s' e" + proof- + from grant parent obtain up rp tp uf rf tf upf rpf tpf + 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)" + and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" + apply (simp add:LinkHard split:option.splits) + by (case_tac a, case_tac aa, case_tac ab, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:LinkHard co2sobj.simps cp2sproc_def split:option.splits) + from vd os pf_in_s have "\ sf. cf2sfile s f = Some sf" + by (auto dest!:is_file_in_current current_file_has_sfile simp:LinkHard) + hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" + using p2 cf2sf f_in os parent + apply (erule_tac x = f in allE) + apply (erule exE, clarsimp simp:LinkHard) + apply (frule_tac s = s in is_file_in_current, simp) + apply (auto simp:cf2sfile_def split:option.splits) + apply (case_tac f, simp) + by (drule_tac s = s in root_is_dir', simp add:vd, simp+) + from vd os pf_in_s have "\ sf. cf2sfile s pf = Some sf" + by (auto dest!:is_dir_in_current current_file_has_sfile simp:LinkHard) + hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s + apply (erule_tac x = pf in allE) + apply (erule exE, frule_tac s = s in is_dir_in_current, simp) + apply (drule is_dir_not_file, drule is_dir_not_file) + apply (auto simp:cf2sfile_def split:option.splits) + apply (case_tac pf, simp_all) + by (simp add:sroot_def root_sec_remains vd vs') + have "search_check s' (up, rp, tp) f" + using p1 p2 p2' vd cf2sf f_in grant LinkHard os parent vs' + apply (rule_tac s = s in enrich_search_check) by (simp_all split:option.splits) - thus ?thesis using p1' p2' parent - apply (simp add:Open Some split:option.splits) - using grant Open Some p1 p2 + moreover have "search_check s' (up, rp, tp) pf" + using p1 p3 p3' vd cf2sf pf_in grant LinkHard os parent vs' + apply (rule_tac s = s in enrich_search_check') + apply (simp_all split:option.splits) + done + ultimately show ?thesis using p1' p2' p3' parent + apply (simp add:LinkHard split:option.splits) + using grant LinkHard p1 p2 p3 + apply simp + done + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (Truncate p f len) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:Truncate) + have f_in: "is_file s' f" using os alive + apply (erule_tac x = "O_file f" in allE) + by (auto simp:Truncate) + have "os_grant s' e" using p_in f_in by (simp add:Truncate) + moreover have "grant s' e" + 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)" + apply (simp add:Truncate split:option.splits) + by (case_tac a, case_tac aa, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:Truncate co2sobj.simps cp2sproc_def split:option.splits) + from os have f_in': "is_file s f" by (simp add:Truncate) + from vd os have "\ sf. cf2sfile s f = Some sf" + by (auto dest!:is_file_in_current current_file_has_sfile simp:Truncate) + hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf + apply (erule_tac x = f in allE) + apply (auto dest:is_file_in_current simp:cf2sfile_def 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 s' (up, rp, tp) f" + using p1 p2 p2' vd cf2sf f_in' grant Truncate f_in + apply (rule_tac s = s in enrich_search_check) + by (simp_all split:option.splits) + thus ?thesis using p1' p2' + apply (simp add:Truncate split:option.splits) + using grant Truncate p1 p2 + by (simp add:Truncate grant p1 p2) + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (CreateMsgq p q) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:CreateMsgq) + have q_not_in: "q \ current_msgqs s'" using os alive' CreateMsgq notin_all + apply (erule_tac x = "E_msgq q" in allE) + by auto + have "os_grant s' e" using p_in q_not_in by (simp add:CreateMsgq) + moreover have "grant s' e" + proof- + from grant obtain up rp tp + where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" + apply (simp add:CreateMsgq split:option.splits) + by (case_tac a, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:CreateMsgq co2sobj.simps cp2sproc_def split:option.splits) + show ?thesis using p1' + apply (simp add:CreateMsgq split:option.splits) + using grant CreateMsgq p1 by simp qed ultimately show ?thesis using vs' by (erule_tac valid.intros(2), simp+) - qed -next - case (ReadFile p fd) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:ReadFile) - have fd_in: "fd \ current_proc_fds s' p" using os alive - apply (erule_tac x = "O_fd p fd" in allE) - by (auto simp:ReadFile) - obtain f where ffd: "file_of_proc_fd s p fd = Some f" - using os ReadFile by auto - hence f_in_s: "is_file s f" using vd - by (auto intro:file_of_pfd_is_file) - obtain flags where fflag: "flags_of_proc_fd s p fd = Some flags" - using os ReadFile by auto - have ffd_in: "file_of_proc_fd s' p fd = Some f" - using ffd_remain ffd by auto - hence f_in: "is_file s' f" using vs' - by (auto intro:file_of_pfd_is_file) - have flags_in: "flags_of_proc_fd s' p fd = Some flags" - using fflags_remain fflag by auto - have "os_grant s' e" using p_in fd_in ffd_in flags_in fflag os f_in - by (auto simp add:ReadFile is_file_in_current) - moreover have "grant s' e" - proof- - from grant ffd obtain up rp tp uf rf tf ufd rfd tfd - 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)" - and p3: "sectxt_of_obj s (O_fd p fd) = Some (ufd, rfd, tfd)" - apply (simp add:ReadFile split:option.splits) - by (case_tac a, case_tac aa, case_tac ab, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:ReadFile co2sobj.simps cp2sproc_def split:option.splits) - from vd f_in_s have "\ sf. cf2sfile s f = Some sf" - by (auto dest!:is_file_in_current current_file_has_sfile) - hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in_s p2 cf2sf - apply (erule_tac x = f in allE) - apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) - apply (case_tac f, simp) - apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) - done - have p3': "sectxt_of_obj s' (O_fd p fd) = Some (ufd, rfd, tfd)" - using cfd2sfd ffd_in ffd p3 f_in f_in_s vd - apply (erule_tac x = p in allE) - apply (erule_tac x = fd in allE) - apply (simp add:proc_file_fds_def) - apply (auto simp:cfd2sfd_def fflag flags_in p3 split:option.splits - dest!:current_file_has_sfile' simp:is_file_in_current) - done - show ?thesis using p1' p2' p3' ffd_in ffd - apply (simp add:ReadFile split:option.splits) - using grant p1 p2 p3 ReadFile - by simp - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (WriteFile p fd) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:WriteFile) - have fd_in: "fd \ current_proc_fds s' p" using os alive - apply (erule_tac x = "O_fd p fd" in allE) - by (auto simp:WriteFile) - obtain f where ffd: "file_of_proc_fd s p fd = Some f" - using os WriteFile by auto - hence f_in_s: "is_file s f" using vd - by (auto intro:file_of_pfd_is_file) - obtain flags where fflag: "flags_of_proc_fd s p fd = Some flags" - using os WriteFile by auto - have ffd_in: "file_of_proc_fd s' p fd = Some f" - using ffd_remain ffd by auto - hence f_in: "is_file s' f" using vs' - by (auto intro:file_of_pfd_is_file) - have flags_in: "flags_of_proc_fd s' p fd = Some flags" - using fflags_remain fflag by auto - have "os_grant s' e" using p_in fd_in ffd_in flags_in fflag os f_in - by (auto simp add:WriteFile is_file_in_current) - moreover have "grant s' e" - proof- - from grant ffd obtain up rp tp uf rf tf ufd rfd tfd - 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)" - and p3: "sectxt_of_obj s (O_fd p fd) = Some (ufd, rfd, tfd)" - apply (simp add:WriteFile split:option.splits) - by (case_tac a, case_tac aa, case_tac ab, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:WriteFile co2sobj.simps cp2sproc_def split:option.splits) - from vd f_in_s have "\ sf. cf2sfile s f = Some sf" - by (auto dest!:is_file_in_current current_file_has_sfile) - hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in_s p2 cf2sf - apply (erule_tac x = f in allE) - apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) - apply (case_tac f, simp) - apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) - done - have p3': "sectxt_of_obj s' (O_fd p fd) = Some (ufd, rfd, tfd)" - using cfd2sfd ffd_in ffd p3 f_in f_in_s vd - apply (erule_tac x = p in allE) - apply (erule_tac x = fd in allE) - apply (simp add:proc_file_fds_def) - apply (auto simp:cfd2sfd_def fflag flags_in p3 split:option.splits - dest!:current_file_has_sfile' simp:is_file_in_current) - done - show ?thesis using p1' p2' p3' ffd_in ffd - apply (simp add:WriteFile split:option.splits) - using grant p1 p2 p3 WriteFile - by simp - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (CloseFd p fd) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:CloseFd) - have fd_in: "fd \ current_proc_fds s' p" using os alive - apply (erule_tac x = "O_fd p fd" in allE) - by (auto simp:CloseFd) - have "os_grant s' e" using p_in fd_in - by (auto simp add:CloseFd) - moreover have "grant s' e" - by(simp add:CloseFd) - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (UnLink p f) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:UnLink) - have f_in: "is_file s' f" using os alive - apply (erule_tac x = "O_file f" in allE) - by (auto simp:UnLink) - from os vd obtain pf where pf_in_s: "is_dir s pf" - and parent: "parent f = Some pf" - by (auto simp:UnLink dest!:file_has_parent) - from pf_in_s alive have pf_in: "is_dir s' pf" - apply (erule_tac x = "O_dir pf" in allE) - by (auto simp:UnLink) - have "os_grant s' e" using p_in f_in os - by (simp add:UnLink hungs) - moreover have "grant s' e" - proof- - from grant parent obtain up rp tp uf rf tf upf rpf tpf - 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)" - and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" - apply (simp add:UnLink split:option.splits) - by (case_tac a, case_tac aa, case_tac ab, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:UnLink co2sobj.simps cp2sproc_def split:option.splits) - from vd os pf_in_s have "\ sf. cf2sfile s f = Some sf" - by (auto dest!:is_file_in_current current_file_has_sfile simp:UnLink) - hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" - using p2 cf2sf f_in os parent - apply (erule_tac x = f in allE) - apply (erule exE, clarsimp simp:UnLink) - apply (frule_tac s = s in is_file_in_current, simp) - by (auto simp:cf2sfile_def split:option.splits) - from vd os pf_in_s have "\ sf. cf2sfile s pf = Some sf" - by (auto dest!:is_dir_in_current current_file_has_sfile simp:UnLink) - hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s - apply (erule_tac x = pf in allE) - apply (erule exE, frule_tac s = s in is_dir_in_current, simp) - apply (drule is_dir_not_file, drule is_dir_not_file) - apply (auto simp:cf2sfile_def split:option.splits) - apply (case_tac pf, simp_all) - by (simp add:sroot_def root_sec_remains vd vs') - have "search_check s' (up, rp, tp) f" - using p1 p2 p2' vd cf2sf f_in grant UnLink os parent vs' - apply (rule_tac s = s in enrich_search_check) - by (simp_all split:option.splits) - thus ?thesis using p1' p2' p3' parent - apply (simp add:UnLink split:option.splits) - using grant UnLink p1 p2 p3 - by simp - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (Rmdir p f) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:Rmdir) - have f_in: "is_dir s' f" using os alive - apply (erule_tac x = "O_dir f" in allE) - by (auto simp:Rmdir dir_is_empty_def) - have not_root: "f \ []" using os - by (auto simp:Rmdir) - from os vd obtain pf where pf_in_s: "is_dir s pf" - and parent: "parent f = Some pf" - apply (auto simp:Rmdir dir_is_empty_def) - apply (case_tac f, simp+) - apply (drule parentf_is_dir_prop1, auto) - done - from pf_in_s alive have pf_in: "is_dir s' pf" - apply (erule_tac x = "O_dir pf" in allE) - by (auto simp:Rmdir) - have empty_in: "dir_is_empty s' f" using os - apply (simp add:dir_is_empty_def f_in) - apply auto using alive' - apply (erule_tac x = "E_file f'" in allE) - by (simp add:Rmdir dir_is_empty_def) - have "os_grant s' e" using p_in f_in os empty_in - by (simp add:Rmdir hungs) - moreover have "grant s' e" - proof- - from grant parent obtain up rp tp uf rf tf upf rpf tpf - where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" - and p2: "sectxt_of_obj s (O_dir f) = Some (uf, rf, tf)" - and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" - apply (simp add:Rmdir split:option.splits) - by (case_tac a, case_tac aa, case_tac ab, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:Rmdir co2sobj.simps cp2sproc_def split:option.splits) - from vd os pf_in_s have "\ sf. cf2sfile s f = Some sf" - by (auto dest!:is_dir_in_current current_file_has_sfile simp:dir_is_empty_def Rmdir) - hence p2': "sectxt_of_obj s' (O_dir f) = Some (uf, rf, tf)" - using p2 cf2sf f_in os parent - apply (erule_tac x = f in allE) - apply (erule exE, clarsimp simp:Rmdir dir_is_empty_def) - apply (frule_tac s = s in is_dir_in_current, simp) - apply (drule is_dir_not_file, drule is_dir_not_file) - by (auto simp:cf2sfile_def split:option.splits) - from vd os pf_in_s have "\ sf. cf2sfile s pf = Some sf" - by (auto dest!:is_dir_in_current current_file_has_sfile simp:Rmdir) - hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s - apply (erule_tac x = pf in allE) - apply (erule exE, frule_tac s = s in is_dir_in_current, simp) - apply (drule is_dir_not_file, drule is_dir_not_file) - apply (auto simp:cf2sfile_def split:option.splits) - apply (case_tac pf, simp_all) - by (simp add:sroot_def root_sec_remains vd vs') - have "search_check s' (up, rp, tp) f" - using p1 p2 p2' vd cf2sf f_in grant Rmdir os parent vs' - apply (rule_tac s = s in enrich_search_check') - by (simp_all add:dir_is_empty_def split:option.splits) - thus ?thesis using p1' p2' p3' parent - apply (simp add:Rmdir split:option.splits) - using grant Rmdir p1 p2 p3 - by simp - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (Mkdir p f inum) - from os obtain pf where pf_in_s: "is_dir s pf" and parent: "parent f = Some pf" - by (auto simp:Mkdir) - have pf_in: "is_dir s' pf" using pf_in_s alive - apply (erule_tac x = "O_dir pf" in allE) - by simp - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:Mkdir) - have f_not_in: "f \ current_files s'" using os alive' - apply (erule_tac x = "E_file f" in allE) - by (auto simp:Mkdir) - have inum_not_in: "inum \ current_inode_nums s'" - using os alive' - apply (erule_tac x = "E_inum inum" in allE) - by (simp add:Mkdir) - have "os_grant s' e" using p_in pf_in parent f_not_in os inum_not_in - by (simp add:Mkdir hungs) - moreover have "grant s' e" - proof- - from grant parent 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_dir pf) = Some (uf, rf, tf)" - apply (simp add:Mkdir split:option.splits) - by (case_tac a, case_tac aa, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:Mkdir co2sobj.simps cp2sproc_def split:option.splits) - from vd os pf_in_s have "\ sf. cf2sfile s pf = Some sf" - by (auto dest!:is_dir_in_current current_file_has_sfile simp:Mkdir) - hence p2': "sectxt_of_obj s' (O_dir pf) = Some (uf, rf, tf)" using p2 cf2sf pf_in pf_in_s - apply (erule_tac x = pf in allE) - apply (erule exE, frule_tac s = s in is_dir_in_current, simp) - apply (drule is_dir_not_file, drule is_dir_not_file) - apply (auto simp:cf2sfile_def split:option.splits) - apply (case_tac pf, simp_all) - by (simp add:sroot_def root_sec_remains vd vs') - have "search_check s' (up, rp, tp) pf" - using p1 p2 p2' vd cf2sf pf_in grant Mkdir pf_in_s parent vs' - apply (rule_tac s = s in enrich_search_check') - apply (simp_all split:option.splits) - done - thus ?thesis using p1' p2' parent - apply (simp add:Mkdir split:option.splits) - using grant Mkdir p1 p2 - apply simp - done - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (LinkHard p f f') - from os obtain pf where pf_in_s: "is_dir s pf" and parent: "parent f' = Some pf" - by (auto simp:LinkHard) - have pf_in: "is_dir s' pf" using pf_in_s alive - apply (erule_tac x = "O_dir pf" in allE) - by simp - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:LinkHard) - have f'_not_in: "f' \ current_files s'" using os alive' - apply (erule_tac x = "E_file f'" in allE) - by (auto simp:LinkHard) - have f_in: "is_file s' f" using os alive - apply (erule_tac x = "O_file f" in allE) - by (auto simp:LinkHard) - have "os_grant s' e" using p_in pf_in parent os f_in f'_not_in - by (simp add:LinkHard hungs) - moreover have "grant s' e" - proof- - from grant parent obtain up rp tp uf rf tf upf rpf tpf - 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)" - and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" - apply (simp add:LinkHard split:option.splits) - by (case_tac a, case_tac aa, case_tac ab, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:LinkHard co2sobj.simps cp2sproc_def split:option.splits) - from vd os pf_in_s have "\ sf. cf2sfile s f = Some sf" - by (auto dest!:is_file_in_current current_file_has_sfile simp:LinkHard) - hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" - using p2 cf2sf f_in os parent - apply (erule_tac x = f in allE) - apply (erule exE, clarsimp simp:LinkHard) - apply (frule_tac s = s in is_file_in_current, simp) - apply (auto simp:cf2sfile_def split:option.splits) - apply (case_tac f, simp) - by (drule_tac s = s in root_is_dir', simp add:vd, simp+) - from vd os pf_in_s have "\ sf. cf2sfile s pf = Some sf" - by (auto dest!:is_dir_in_current current_file_has_sfile simp:LinkHard) - hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s - apply (erule_tac x = pf in allE) - apply (erule exE, frule_tac s = s in is_dir_in_current, simp) - apply (drule is_dir_not_file, drule is_dir_not_file) - apply (auto simp:cf2sfile_def split:option.splits) - apply (case_tac pf, simp_all) - by (simp add:sroot_def root_sec_remains vd vs') - have "search_check s' (up, rp, tp) f" - using p1 p2 p2' vd cf2sf f_in grant LinkHard os parent vs' - apply (rule_tac s = s in enrich_search_check) - by (simp_all split:option.splits) - moreover have "search_check s' (up, rp, tp) pf" - using p1 p3 p3' vd cf2sf pf_in grant LinkHard os parent vs' - apply (rule_tac s = s in enrich_search_check') - apply (simp_all split:option.splits) - done - ultimately show ?thesis using p1' p2' p3' parent - apply (simp add:LinkHard split:option.splits) - using grant LinkHard p1 p2 p3 - apply simp - done - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (Truncate p f len) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:Truncate) - have f_in: "is_file s' f" using os alive - apply (erule_tac x = "O_file f" in allE) - by (auto simp:Truncate) - have "os_grant s' e" using p_in f_in by (simp add:Truncate) - moreover have "grant s' e" - 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)" - apply (simp add:Truncate split:option.splits) - by (case_tac a, case_tac aa, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:Truncate co2sobj.simps cp2sproc_def split:option.splits) - from os have f_in': "is_file s f" by (simp add:Truncate) - from vd os have "\ sf. cf2sfile s f = Some sf" - by (auto dest!:is_file_in_current current_file_has_sfile simp:Truncate) - hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf - apply (erule_tac x = f in allE) - apply (auto dest:is_file_in_current simp:cf2sfile_def 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 s' (up, rp, tp) f" - using p1 p2 p2' vd cf2sf f_in' grant Truncate f_in - apply (rule_tac s = s in enrich_search_check) - by (simp_all split:option.splits) - thus ?thesis using p1' p2' - apply (simp add:Truncate split:option.splits) - using grant Truncate p1 p2 - by (simp add:Truncate grant p1 p2) - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (CreateMsgq p q) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:CreateMsgq) - have q_not_in: "q \ current_msgqs s'" using os alive' - apply (erule_tac x = "E_msgq q" in allE) - by (simp add:CreateMsgq) - have "os_grant s' e" using p_in q_not_in by (simp add:CreateMsgq) - moreover have "grant s' e" - proof- - from grant obtain up rp tp - where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" - apply (simp add:CreateMsgq split:option.splits) - by (case_tac a, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:CreateMsgq co2sobj.simps cp2sproc_def split:option.splits) - show ?thesis using p1' - apply (simp add:CreateMsgq split:option.splits) - using grant CreateMsgq p1 - by simp - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (RemoveMsgq p q) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:RemoveMsgq) - have q_in: "q \ current_msgqs s'" using os alive - apply (erule_tac x = "O_msgq q" in allE) - by (simp add:RemoveMsgq) - have "os_grant s' e" using p_in q_in by (simp add:RemoveMsgq) - moreover have "grant s' e" - proof- - from grant obtain up rp tp uq rq tq - where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" - and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" - apply (simp add:RemoveMsgq split:option.splits) - by (case_tac a, case_tac aa, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:RemoveMsgq co2sobj.simps cp2sproc_def split:option.splits) - from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" - using os cq2sq vd - apply (erule_tac x = q in allE) - by (auto simp:RemoveMsgq co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) - show ?thesis using p1' p2' grant p1 p2 + next + case (RemoveMsgq p q) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:RemoveMsgq) + have q_in: "q \ current_msgqs s'" using os alive + apply (erule_tac x = "O_msgq q" in allE) by (simp add:RemoveMsgq) - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (SendMsg p q m) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:SendMsg) - have q_in: "q \ current_msgqs s'" using os alive - apply (erule_tac x = "O_msgq q" in allE) - by (simp add:SendMsg) - have m_not_in: "m \ set (msgs_of_queue s' q)" using os alive' - apply (erule_tac x = "E_msg q m" in allE) - by (simp add:SendMsg q_in) - have "os_grant s' e" using p_in q_in m_not_in - by (simp add:SendMsg) - moreover have "grant s' e" - proof- - from grant obtain up rp tp uq rq tq - where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" - and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" - apply (simp add:SendMsg split:option.splits) - by (case_tac a, case_tac aa, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:SendMsg co2sobj.simps cp2sproc_def split:option.splits) - from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" - using os cq2sq vd - apply (erule_tac x = q in allE) - by (auto simp:SendMsg co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) - show ?thesis using p1' p2' grant p1 p2 + have "os_grant s' e" using p_in q_in by (simp add:RemoveMsgq) + moreover have "grant s' e" + proof- + from grant obtain up rp tp uq rq tq + where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" + and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" + apply (simp add:RemoveMsgq split:option.splits) + by (case_tac a, case_tac aa, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:RemoveMsgq co2sobj.simps cp2sproc_def split:option.splits) + from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" + using os cq2sq vd + apply (erule_tac x = q in allE) + by (auto simp:RemoveMsgq co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) + show ?thesis using p1' p2' grant p1 p2 + by (simp add:RemoveMsgq) + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (SendMsg p q m) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:SendMsg) + have q_in: "q \ current_msgqs s'" using os alive + apply (erule_tac x = "O_msgq q" in allE) + by (simp add:SendMsg) + have m_not_in: "m \ set (msgs_of_queue s' q)" using os alive' notin_all SendMsg q_in + apply (erule_tac x = "E_msg q m" in allE) + apply (case_tac obj') + by auto + have "os_grant s' e" using p_in q_in m_not_in by (simp add:SendMsg) - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (RecvMsg p q m) - have p_in: "p \ current_procs s'" using os alive - apply (erule_tac x = "O_proc p" in allE) - by (auto simp:RecvMsg) - have q_in: "q \ current_msgqs s'" using os alive - apply (erule_tac x = "O_msgq q" in allE) - by (simp add:RecvMsg) - have m_in: "m = hd (msgs_of_queue s' q)" - and sms_not_empty: "msgs_of_queue s' q \ []" - using os sms_remain - by (auto simp:RecvMsg) - have "os_grant s' e" using p_in q_in m_in sms_not_empty os - by (simp add:RecvMsg) - moreover have "grant s' e" - proof- - from grant obtain up rp tp uq rq tq um rm tm - where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" - and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" - and p3: "sectxt_of_obj s (O_msg q m) = Some (um, rm, tm)" - apply (simp add:RecvMsg split:option.splits) - by (case_tac a, case_tac aa, case_tac ab, blast) - from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" - using os cp2sp - apply (erule_tac x = p in allE) - by (auto simp:RecvMsg co2sobj.simps cp2sproc_def split:option.splits) - from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" - using os cq2sq vd - apply (erule_tac x = q in allE) - by (auto simp:RecvMsg co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) - from p3 have p3': "sectxt_of_obj s' (O_msg q m) = Some (um, rm, tm)" - using sms_remain cq2sq vd os p2 p2' p3 - apply (erule_tac x = q in allE) - apply (erule_tac x = q in allE) - apply (clarsimp simp:RecvMsg) - apply (simp add:cq2smsgq_def split:option.splits if_splits) - apply (drule current_has_sms', simp, simp) - apply (case_tac "msgs_of_queue s q", simp) - apply (simp add:cqm2sms.simps split:option.splits) - apply (auto simp add:cm2smsg_def split:option.splits if_splits)[1] - apply (case_tac "msgs_of_queue s q", simp) - apply (simp add:cqm2sms.simps split:option.splits) - apply (auto simp add:cm2smsg_def split:option.splits if_splits)[1] - done - show ?thesis using p1' p2' p3' grant p1 p2 p3 + moreover have "grant s' e" + proof- + from grant obtain up rp tp uq rq tq + where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" + and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" + apply (simp add:SendMsg split:option.splits) + by (case_tac a, case_tac aa, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:SendMsg co2sobj.simps cp2sproc_def split:option.splits) + from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" + using os cq2sq vd + apply (erule_tac x = q in allE) + by (auto simp:SendMsg co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) + show ?thesis using p1' p2' grant p1 p2 + by (simp add:SendMsg) + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (RecvMsg p q m) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:RecvMsg) + have q_in: "q \ current_msgqs s'" using os alive + apply (erule_tac x = "O_msgq q" in allE) + by (simp add:RecvMsg) + have m_in: "m = hd (msgs_of_queue s' q)" + and sms_not_empty: "msgs_of_queue s' q \ []" + using os sms_remain + by (auto simp:RecvMsg) + have "os_grant s' e" using p_in q_in m_in sms_not_empty os by (simp add:RecvMsg) - qed - ultimately show ?thesis using vs' - by (erule_tac valid.intros(2), simp+) -next - case (CreateSock p af st fd inum) - show ?thesis using grant - by (simp add:CreateSock) -next - case (Bind p fd addr) - show ?thesis using grant - by (simp add:Bind) -next - case (Connect p fd addr) - show ?thesis using grant - by (simp add:Connect) -next - case (Listen p fd) - show ?thesis using grant - by (simp add:Listen) -next - case (Accept p fd addr port fd' inum) - show ?thesis using grant - by (simp add:Accept) -next - case (SendSock p fd) - show ?thesis using grant - by (simp add:SendSock) -next - case (RecvSock p fd) - show ?thesis using grant - by (simp add:RecvSock) -next - case (Shutdown p fd how) - show ?thesis using grant - by (simp add:Shutdown) -qed - -lemma not_all_procs_prop2: - "p' \ all_procs s \ p' \ init_procs" -apply (induct s, simp) -by (case_tac a, auto) - -lemma not_all_procs_prop3: - "p' \ all_procs s \ p' \ current_procs s" -apply (induct s, simp) -by (case_tac a, auto) - -definition is_created_proc:: "t_state \ t_process \ bool" -where - "is_created_proc s p \ p \ current_procs s \ (p \ init_procs \ died (O_proc p) s)" + moreover have "grant s' e" + proof- + from grant obtain up rp tp uq rq tq um rm tm + where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" + and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" + and p3: "sectxt_of_obj s (O_msg q m) = Some (um, rm, tm)" + apply (simp add:RecvMsg split:option.splits) + by (case_tac a, case_tac aa, case_tac ab, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:RecvMsg co2sobj.simps cp2sproc_def split:option.splits) + from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" + using os cq2sq vd + apply (erule_tac x = q in allE) + by (auto simp:RecvMsg co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) + from p3 have p3': "sectxt_of_obj s' (O_msg q m) = Some (um, rm, tm)" + using sms_remain cq2sq vd os p2 p2' p3 + apply (erule_tac x = q in allE) + apply (erule_tac x = q in allE) + apply (clarsimp simp:RecvMsg) + apply (simp add:cq2smsgq_def split:option.splits if_splits) + apply (drule current_has_sms', simp, simp) + apply (case_tac "msgs_of_queue s q", simp) + apply (simp add:cqm2sms.simps split:option.splits) + apply (auto simp add:cm2smsg_def split:option.splits if_splits)[1] + apply (case_tac "msgs_of_queue s q", simp) + apply (simp add:cqm2sms.simps split:option.splits) + apply (auto simp add:cm2smsg_def split:option.splits if_splits)[1] + done + show ?thesis using p1' p2' p3' grant p1 p2 p3 + by (simp add:RecvMsg) + qed + ultimately show ?thesis using vs' + by (erule_tac valid.intros(2), simp+) + next + case (CreateSock p af st fd inum) + show ?thesis using grant + by (simp add:CreateSock) + next + case (Bind p fd addr) + show ?thesis using grant + by (simp add:Bind) + next + case (Connect p fd addr) + show ?thesis using grant + by (simp add:Connect) + next + case (Listen p fd) + show ?thesis using grant + by (simp add:Listen) + next + case (Accept p fd addr port fd' inum) + show ?thesis using grant + by (simp add:Accept) + next + case (SendSock p fd) + show ?thesis using grant + by (simp add:SendSock) + next + case (RecvSock p fd) + show ?thesis using grant + by (simp add:RecvSock) + next + case (Shutdown p fd how) + show ?thesis using grant + by (simp add:Shutdown) + qed +qed lemma created_proc_clone: "valid (Clone p p' fds # s) \ @@ -1157,7 +1276,8 @@ 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 (enrich_proc s p p') p' fd = Some f; is_created_proc s p; p' \ all_procs s; + no_del_event 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) @@ -1165,6 +1285,14 @@ dest:not_all_procs_prop3 split:if_splits option.splits) done +lemma enrich_proc_dup_ffd_eq: + "\is_created_proc s p; p' \ all_procs s; no_del_event s; valid s\ + \ file_of_proc_fd (enrich_proc s p p') p' fd = file_of_proc_fd s p fd" +apply (case_tac "file_of_proc_fd s p fd") +apply (case_tac[!] "file_of_proc_fd (enrich_proc s p p') p' fd") +apply (auto dest:enrich_proc_dup_ffd enrich_proc_dup_ffd') +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) @@ -1192,7 +1320,7 @@ done lemma enrich_proc_dup_ffds: - "\is_created_proc s p; p' \ all_procs s; valid s\ + "\is_created_proc s p; p' \ all_procs s; no_del_event 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) @@ -1202,7 +1330,7 @@ done lemma enrich_proc_dup_ffds_eq_fds: - "\is_created_proc s p; p' \ all_procs s; valid s\ + "\is_created_proc s p; p' \ all_procs s; no_del_event 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) diff -r 0a68c605ae79 -r e79e3a8a4ceb no_shm_selinux/Enrich2.thy --- a/no_shm_selinux/Enrich2.thy Mon Dec 23 19:47:17 2013 +0800 +++ b/no_shm_selinux/Enrich2.thy Thu Dec 26 10:56:50 2013 +0800 @@ -1,6 +1,6 @@ theory Enrich imports Main Flask Static Init_prop Valid_prop Tainted_prop Delete_prop Co2sobj_prop S2ss_prop S2ss_prop2 - Temp Enrich + Temp Enrich Proc_fd_of_file_prop begin context tainting_s begin @@ -56,10 +56,10 @@ 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" + 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) + 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) \ @@ -482,16 +482,66 @@ qed qed ultimately show ?thesis - using created_cons vd_cons all_procs_cons + 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" + proof clarify + fix obj + assume a0: "alive (e # s) obj" + have a1: "is_created_proc s p \ \ obj. alive s obj \ alive (enrich_proc s p p') obj" + using pre by auto + show "alive (enrich_proc (e # s) p p') obj" (* + proof (cases e) + case (Execve tp f fds) + with created_cons a1 + have b1: "\ obj. alive s obj \ alive (enrich_proc s p p') obj" + by (auto simp:is_created_proc_simps) + show ?thesis + using created_cons all_procs_cons vd_enrich_cons Execve b1 os a0 + apply (simp add:alive_execve split:if_splits) + apply (frule_tac vd_cons) defer + apply (frule_tac vd_cons) + using vd_cons' Execve vd os + apply (auto simp:is_file_simps is_dir_simps is_created_proc_simps alive.simps + split:t_object.splits if_splits + dest:set_mp file_fds_subset_pfds) + apply (erule_tac x = "O_proc nat" in allE, simp) + apply (erule_tac x = "O_file list" in allE, simp) + apply (drule set_mp, simp) + apply (drule_tac s = s and p = tp in file_fds_subset_pfds) + apply (erule_tac x = "O_fd tp nat2" in allE, simp) + apply (auto)[1] + apply (erule_tac x = "O_fd nat1 nat2" in allE, auto dest:set_mp file_fds_subset_pfds)[1] + apply (erule_tac x = "O_dir list" in allE, simp) sorry + show ?thesis *) sorry moreover have "\obj. enrich_not_alive (e # s) obj \ enrich_not_alive (enrich_proc (e # s) p p') obj" + thm enrich_not_alive.simps sorry moreover have "files_hung_by_del (enrich_proc (e # s) p p') = files_hung_by_del (e # s)" + proof- + have "is_created_proc s p \ files_hung_by_del (enrich_proc s p p') = files_hung_by_del s" + and ffd_remain: "is_created_proc s p \ + \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" + using pre by auto + with created_cons all_procs_cons os vd_cons' vd + show ?thesis + apply (frule_tac not_all_procs_prop3) + apply (case_tac e) + apply (auto simp:files_hung_by_del.simps is_created_proc_simps) + apply (auto simp:enrich_proc_dup_ffd_eq proc_file_fds_def procfd_of_file_eq_fpfd'' + dest:procfd_of_file_imp_fpfd procfd_of_file_imp_fpfd' procfd_of_file_non_empty + ) + apply (auto simp:enrich_proc_dup_ffd_eq proc_file_fds_def split:if_splits)[1] + + apply (auto simp:enrich_proc_dup_ffd_eq proc_file_fds_def files_hung_by_del.simps + split:option.splits)[1] + apply (auto split:option.splits)[1] + thm is_created_proc_simps sorry moreover have "\tp. tp \ current_procs (e # s) \ cp2sproc (enrich_proc (e # s) p p') tp = cp2sproc (e # s) tp" sorry @@ -524,10 +574,13 @@ 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 + show ?thesis using True a1 a2 a3 a4 b0 thm not_all_procs_prop3 + apply (frule_tac not_all_procs_prop2) + apply (frule 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) + dest:vt_grant_os) + apply (auto simp:sectxt_of_obj_simps split:option.splits dest:valid.cases) sorry next @@ -552,7 +605,7 @@ 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 + using created_cons vd_enrich_cons vd_cons' b0 apply (auto simp:cp2sproc_other is_created_proc_def) done qed diff -r 0a68c605ae79 -r e79e3a8a4ceb no_shm_selinux/Flask.thy --- a/no_shm_selinux/Flask.thy Mon Dec 23 19:47:17 2013 +0800 +++ b/no_shm_selinux/Flask.thy Thu Dec 26 10:56:50 2013 +0800 @@ -510,7 +510,7 @@ "info_flow_shm s from to \ (self_shm s from to) \ (\ h. one_flow_shm s h from to)" *) -fun current_msgqs :: "t_state \ t_msg set" +fun current_msgqs :: "t_state \ t_msgq set" where "current_msgqs [] = init_msgqs" | "current_msgqs (CreateMsgq p q # \) = insert q (current_msgqs \)"