no_shm_selinux/Enrich.thy
changeset 87 8d18cfc845dd
parent 84 cebdef333899
child 88 e832378a2ff2
--- a/no_shm_selinux/Enrich.thy	Mon Dec 30 23:41:58 2013 +0800
+++ b/no_shm_selinux/Enrich.thy	Tue Dec 31 14:57:13 2013 +0800
@@ -141,77 +141,6 @@
 apply (simp add:parentf_is_dir_prop2)
 done
 
-(* enrich s target_proc duplicated_pro *)
-fun enrich_proc :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> t_state"
-where 
-  "enrich_proc [] tp dp = []"
-| "enrich_proc (Execve p f fds # s) tp dp = (
-     if (tp = p) 
-     then Execve dp f (fds \<inter> proc_file_fds s p) # Execve p f fds # (enrich_proc s tp dp)
-     else Execve p f fds # (enrich_proc s tp dp))"
-| "enrich_proc (Clone p p' fds # s) tp dp = (
-     if (tp = p') 
-     then Clone p dp (fds \<inter> proc_file_fds s p) # Clone p p' fds # s
-     else Clone p p' fds # (enrich_proc s tp dp))"
-| "enrich_proc (Open p f flags fd opt # s) tp dp = (
-     if (tp = p)
-     then Open dp f (remove_create_flag flags) fd None # 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 (ReadFile p fd # s) tp dp = (
-     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 \<and> fd \<in> 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)
-     then Attach dp h flag # Attach p h flag # (enrich_proc s tp dp)
-     else Attach p h flag # (enrich_proc s tp dp))"
-| "enrich_proc (Detach p h # s) tp dp = (
-     if (tp = p) 
-     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 \<Rightarrow> t_process \<Rightarrow> bool"
-where
-  "is_created_proc s p \<equiv> p \<in> current_procs s \<and> (p \<in> init_procs \<longrightarrow> died (O_proc p) s)"
-
-definition is_created_proc':: "t_state \<Rightarrow> t_process \<Rightarrow> bool"
-where
-  "is_created_proc' s p \<equiv> p \<in> current_procs s \<and> p \<notin> init_procs"
-
-lemma no_del_died:
-  "\<lbrakk>no_del_event s; died obj s\<rbrakk> \<Longrightarrow> (\<exists> p fd. obj = O_fd p fd \<or> obj = O_tcp_sock (p, fd) \<or> obj = O_udp_sock (p, fd))
-  \<or> (\<exists> 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 \<Longrightarrow> 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: "\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile s' f = cf2sfile s f"
@@ -1248,74 +1177,44 @@
   qed 
 qed
 
-lemma created_proc_clone:
-  "valid (Clone p p' fds # s) \<Longrightarrow> 
-   is_created_proc (Clone p p' fds # s) tp = (if (tp = p') then True else is_created_proc s tp)"
-apply (drule vt_grant_os)
-apply (auto simp:is_created_proc_def dest:not_all_procs_prop2)
-using not_died_init_proc
-by auto
-
-lemma created_proc_exit: 
-  "is_created_proc (Exit p # s) tp = (if (tp = p) then False else is_created_proc s tp)"
-by (simp add:is_created_proc_def)
-
-lemma created_proc_kill:
-  "is_created_proc (Kill p p' # s) tp = (if (tp = p') then False else is_created_proc s tp)"
-by (simp add:is_created_proc_def)
+lemma current_proc_fds_in_curp:
+  "\<lbrakk>fd \<in> current_proc_fds s p; valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s"
+apply (induct s)
+apply (simp add:init_fds_of_proc_prop1)
+apply (frule vt_grant_os, frule vd_cons)
+apply (case_tac a, auto split:if_splits option.splits)
+done
 
-lemma created_proc_other:
-  "\<lbrakk>\<And> p p' fds. e \<noteq> Clone p p' fds;
-    \<And> p. e \<noteq> Exit p;
-    \<And> p p'. e \<noteq> Kill p p'\<rbrakk> \<Longrightarrow> is_created_proc (e # s) tp = is_created_proc s tp"
-by (case_tac e, auto simp:is_created_proc_def)
-
-lemmas is_created_proc_simps = created_proc_clone created_proc_exit created_proc_kill created_proc_other
-(*
-
-
-
-
-       (p \<in> current_procs s \<longrightarrow> co2sobj (enrich_proc s p p') (O_proc p') = co2sobj (enrich_proc s p p') (O_proc p)) \<and>
-       (\<forall> obj. alive s obj \<longrightarrow> alive (enrich_proc s p p')  obj) \<and> 
-       (\<forall> p'. p' \<in> current_procs s \<longrightarrow> cp2sproc (enrich_proc s p p') p' = cp2sproc s p) \<and>
-       (\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile (enrich_proc s p p') f = cf2sfile s f) \<and>
-       (Tainted (enrich_proc s p p') = (Tainted s \<union> (if (O_proc p \<in> Tainted s) then {O_proc p'} else {})))"*)
-
-lemma enrich_proc_dup_in:
-  "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk>
-   \<Longrightarrow> p' \<in> current_procs (enrich_proc s p p')"
-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 dest:not_all_procs_prop3)
+lemma get_parentfs_ctxts_prop:
+  "\<lbrakk>get_parentfs_ctxts s (a # f) = Some ctxts; sectxt_of_obj s (O_dir f) = Some ctxt; valid s\<rbrakk>
+   \<Longrightarrow> ctxt \<in> set (ctxts)"
+apply (induct f)
+apply (auto split:option.splits)
 done
 
-lemma enrich_proc_dup_ffd:
-  "\<lbrakk>file_of_proc_fd s p fd = Some f; is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk>
-   \<Longrightarrow> file_of_proc_fd (enrich_proc s p p') 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 search_check_allp_intro:
+  "\<lbrakk>search_check s sp pf; get_parentfs_ctxts s pf = Some ctxts; valid s; is_dir s pf\<rbrakk>
+   \<Longrightarrow> search_check_allp sp (set ctxts)"
+apply (case_tac pf)
+apply (simp split:option.splits if_splits add:search_check_allp_def)
+apply (rule ballI)
+apply (auto simp:search_check_ctxt_def search_check_dir_def split:if_splits option.splits)
+apply (auto simp:search_check_allp_def search_check_file_def)
+apply (frule is_dir_not_file, simp)
+done
 
-lemma enrich_proc_dup_ffd':
-  "\<lbrakk>file_of_proc_fd (enrich_proc s p p') p' fd = Some f; is_created_proc s p; p' \<notin> all_procs s;
-    no_del_event s; valid s\<rbrakk>
-   \<Longrightarrow> 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 search_check_leveling_f:
+  "\<lbrakk>search_check s sp pf; parent f = Some pf; is_file s f; valid s;
+    sectxt_of_obj s (O_file f) = Some fctxt; search_check_file sp fctxt\<rbrakk>
+   \<Longrightarrow> search_check s sp f"
+apply (case_tac f, simp+)
+apply (auto split:option.splits simp:search_check_ctxt_def)
+apply (frule parentf_is_dir_prop2, simp)
+apply (erule get_pfs_secs_prop, simp)
+apply (erule_tac search_check_allp_intro, simp_all)
+apply (simp add:parentf_is_dir_prop2)
+done
 
-lemma enrich_proc_dup_ffd_eq:
-  "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; no_del_event s; valid s\<rbrakk>
-  \<Longrightarrow> 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:
   "\<lbrakk>flags_of_proc_fd s p fd = Some flag; valid s\<rbrakk> \<Longrightarrow> fd \<in> current_proc_fds s p"
@@ -1333,37 +1232,6 @@
 apply (auto split:if_splits option.splits dest:proc_fd_in_fds)
 done
 
-lemma enrich_proc_dup_fflags:
-  "\<lbrakk>flags_of_proc_fd s p fd = Some flag; is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk>
-   \<Longrightarrow> flags_of_proc_fd (enrich_proc s p p') p' fd = Some (remove_create_flag flag) \<or>
-       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 is_creat_flag_def
-  dest:not_all_procs_prop3 split:if_splits option.splits)
-done
-
-lemma enrich_proc_dup_ffds:
-  "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; no_del_event s; valid s\<rbrakk>
-   \<Longrightarrow> 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_dup_ffds_eq_fds:
-  "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; no_del_event s; valid s\<rbrakk>
-   \<Longrightarrow> current_proc_fds (enrich_proc s p p') p' = proc_file_fds s p"
-apply (induct s arbitrary:p)
-apply (simp add: is_created_proc_def)
-apply (frule not_all_procs_prop3)
-apply (frule vd_cons, frule vt_grant_os, case_tac a)
-apply (auto split:if_splits option.splits dest:proc_fd_in_fds set_mp not_all_procs_prop3 
-  simp:proc_file_fds_def is_created_proc_def)
-done
-
 lemma oflags_check_remove_create: 
   "oflags_check flags sp sf \<Longrightarrow> oflags_check (remove_create_flag flags) sp sf"
 apply (case_tac flags)