# HG changeset patch # User chunhan # Date 1389279225 -28800 # Node ID 042e1e7fd505e8512f2b799882cedfdb8594da0d # Parent dfde07c7cd6b730782cee738432f12549dd21a96 new childf new-version diff -r dfde07c7cd6b -r 042e1e7fd505 no_shm_selinux/Enrich.thy --- a/no_shm_selinux/Enrich.thy Thu Jan 09 19:09:09 2014 +0800 +++ b/no_shm_selinux/Enrich.thy Thu Jan 09 22:53:45 2014 +0800 @@ -7,6 +7,7 @@ datatype t_enrich_obj = E_proc "t_process" "t_msgq" "t_msgq" | E_file "t_file" "nat" +| E_file_link "t_file" | E_msgq "t_msgq" (* objects that need dynamic indexing, all nature-numbers *) diff -r dfde07c7cd6b -r 042e1e7fd505 no_shm_selinux/Enrich2.thy --- a/no_shm_selinux/Enrich2.thy Thu Jan 09 19:09:09 2014 +0800 +++ b/no_shm_selinux/Enrich2.thy Thu Jan 09 22:53:45 2014 +0800 @@ -5,12 +5,37 @@ context tainting_s begin -fun enrich_file :: "t_state \ t_file \ t_state \ (t_file set \ t_state)" +fun new_cf :: "t_state \ t_file set \ t_file \ t_file" +where + "new_cf s fs [] = []" +| "new_cf s fs (f#pf) = new_childf_general pf s fs" + +fun enrich_dufs:: "t_state \ t_file \ t_state \ (t_file \ t_file) list" where - "enrich_file [] tf s = ({}, [])" -| "enrich_file (Open p f flags fd (Some inum) # s) tf s = - " - + "enrich_dufs [] tf s = []" +| "enrich_dufs (Open p f flags fd (Some inum) # s) tf s' = + (if (f \ same_inode_files s' tf) + then [(f, new_cf s' {} f)] + else enrich_dufs s tf s')" +| "enrich_dufs (LinkHard p f f' # s) tf s' = + (if (f' \ same_inode_files s' tf) + then (f', new_cf s' (set (map snd (enrich_dufs s tf s'))) f') # enrich_dufs s tf s' + else enrich_dufs s tf s')" +| "enrich_dufs (_ # s) tf s' = enrich_dufs s tf s'" + +lemma enrich_dufs_sameinodes: + "valid s \ set (map snd (enrich_dufs s f s)) = same_inode_files s f" +thm enrich_dufs.induct +apply (induct rule:enrich_dufs.induct) + +apply (erule enrich_dufs.induct) + +fun enrich_file :: "t_state \ t_file \ t_inode_num \ t_state \ (t_file set \ t_state)" +where + "enrich_file [] tf ninum s = ({}, [])" +| "enrich_file (Open p f flags fd (Some inum) # s) tf ninum s = + if (f \ same_inode_fies s tf) + then Open p f flags fd (Some" diff -r dfde07c7cd6b -r 042e1e7fd505 no_shm_selinux/New_obj_prop.thy --- a/no_shm_selinux/New_obj_prop.thy Thu Jan 09 19:09:09 2014 +0800 +++ b/no_shm_selinux/New_obj_prop.thy Thu Jan 09 22:53:45 2014 +0800 @@ -83,13 +83,12 @@ done lemma ncf_notin_curf_general: - assumes vd: "valid \" and fin_fs: "finite fs" - shows "new_childf_general f \ fs \ (current_files \ \ fs)" + assumes fin_fs: "finite fs" + shows "new_childf_general f fs \ fs" proof- - from vd fin_fs have "finite (current_files \ \ fs)" - by (auto dest:finite_cf) - hence a1: "Suc (Max (fname_length_set {fn. fn # f \ (current_files \ \ fs)})) \ - fname_length_set {fn. fn # f \ (current_files \ \ fs)}" + have a1: "Suc (Max (fname_length_set {fn. fn # f \ fs})) \ + fname_length_set {fn. fn # f \ fs}" + using fin_fs by (erule_tac ncf_notin_curf_aux) have a2: "\ f pf fs. f # pf \ fs \ f \ {fn. fn # pf \ fs}" by auto have a3: "\ f pf fs. f \ {fn. fn # pf \ fs} \ length f \ fname_length_set {fn. fn # pf \ fs}" @@ -104,12 +103,12 @@ lemma ncf_notin_curf: "valid \ \ new_childf f \ \ (current_files \)" -apply (drule_tac fs = "{}" in ncf_notin_curf_general) -apply (simp) +apply (drule finite_cf) +apply (drule_tac fs = "current_files \" in ncf_notin_curf_general) apply (simp add:new_childf_def) done -lemma ncf_parent_general: "valid \ \ parent (new_childf_general f \ fs) = Some f" +lemma ncf_parent_general: "valid \ \ parent (new_childf_general f fs) = Some f" by (simp add:new_childf_general_def) lemma ncf_parent: "valid \ \ parent (new_childf f \) = Some f" diff -r dfde07c7cd6b -r 042e1e7fd505 no_shm_selinux/Static.thy --- a/no_shm_selinux/Static.thy Thu Jan 09 19:09:09 2014 +0800 +++ b/no_shm_selinux/Static.thy Thu Jan 09 22:53:45 2014 +0800 @@ -702,9 +702,9 @@ where "new_proc_fd \ p = next_nat (current_proc_fds \ p)" -definition all_fname_under_dir:: "t_file \ t_state \ t_file set \ t_fname set" +definition all_fname_under_dir:: "t_file \ t_file set \ t_fname set" where - "all_fname_under_dir d \ fs = {fn. \ f. fn # d = f \ f \ (current_files \ \ fs)}" + "all_fname_under_dir d fs = {fn. fn # d \ fs}" fun fname_all_a:: "nat \ t_fname" where @@ -715,17 +715,17 @@ where "fname_length_set fns = length`fns" -definition next_fname:: "t_file \ t_state \ t_file set \ t_fname" +definition next_fname:: "t_file \ t_file set \ t_fname" where - "next_fname pf \ fs = fname_all_a ((Max (fname_length_set (all_fname_under_dir pf \ fs))) + 1)" + "next_fname pf fs = fname_all_a ((Max (fname_length_set (all_fname_under_dir pf fs))) + 1)" -definition new_childf_general:: "t_file \ t_state \ t_file set \ t_file" +definition new_childf_general:: "t_file \ t_file set \ t_file" where - "new_childf_general pf \ fs = next_fname pf \ fs # pf" + "new_childf_general pf fs = next_fname pf fs # pf" definition new_childf:: "t_file \ t_state \ t_file" where - "new_childf pf \ = new_childf_general pf \ {}" + "new_childf pf \ = new_childf_general pf (current_files \)" end