--- a/Dynamic_static.thy Wed Nov 20 21:33:56 2013 +0800
+++ b/Dynamic_static.thy Mon Dec 02 10:52:40 2013 +0800
@@ -117,32 +117,40 @@
by (case_tac a, simp)
qed
+lemma proc_filefd_has_sfd: "\<lbrakk>fd \<in> proc_file_fds s p; valid s\<rbrakk> \<Longrightarrow> \<exists> sfd. cfd2sfd s p fd = Some sfd"
+apply (simp add:proc_file_fds_def)
+apply (auto dest: current_filefd_has_sfd)
+done
+
lemma enrich_inherit_fds_check:
assumes grant: "inherit_fds_check s (up, nr, nt) p fds" and vd: "valid s"
and cp2sp: "\<forall> p. p \<in> current_procs s \<longrightarrow> cp2sproc s' p = cp2sproc s p"
and p_in: "p \<in> current_procs s" and p_in': "p \<in> current_procs s'"
- and fd_in: "fds \<subseteq> current_proc_fds s p" and fd_in': "fds \<subseteq> current_proc_fds s' p"
+ and fd_in: "fds \<subseteq> proc_file_fds s p" and fd_in': "fds \<subseteq> proc_file_fds s' p"
shows "inherit_fds_check s' (up, nr, nt) p fds"
proof-
have "\<And> fd. fd \<in> fds \<Longrightarrow> sectxt_of_obj s' (O_fd p fd) = sectxt_of_obj s (O_fd p fd)"
proof-
fix fd
assume fd_in_fds: "fd \<in> fds"
- hence fd_in_cfds: "fd \<in> current_proc_fds s p"
- and fd_in_cfds': "fd \<in> current_proc_fds s' p"
+ hence fd_in_cfds: "fd \<in> proc_file_fds s p"
+ and fd_in_cfds': "fd \<in> proc_file_fds s' p"
using fd_in fd_in' by auto
from p_in vd obtain sp where csp: "cp2sproc s p = Some sp"
by (drule_tac current_proc_has_sp, simp, erule_tac exE, simp)
with cp2sp have "cpfd2sfds s p = cpfd2sfds s' p"
apply (erule_tac x = p in allE)
by (auto simp:cp2sproc_def split:option.splits simp:p_in)
- hence "cfd2sfd s p fd = cfd2sfd s' p fd"
+ hence "cfd2sfd s p fd = cfd2sfd s' p fd" using fd_in_cfds fd_in_cfds'
apply (simp add:cpfd2sfds_def)
+ apply (frule proc_filefd_has_sfd, simp add:vd, erule exE)
+ apply (drule_tac x = sfd in eqset_imp_iff, simp)
+ (*
thm inherit_fds_check_def
thm sectxts_of_fds_def
thm cpfd2sfds_def
apply (
-
+ *)sorry
show "sectxt_of_obj s' (O_fd p fd) = sectxt_of_obj s (O_fd p fd)"
sorry
qed
@@ -160,6 +168,8 @@
and cp2sp: "\<forall> p. p \<in> current_procs s \<longrightarrow> cp2sproc s' p = cp2sproc s p"
and cf2sf: "\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile s' f = cf2sfile s f"
shows "valid (e # s')"
+
+sorry (*
proof (cases e)
case (Execve p f fds)
have p_in: "p \<in> current_procs s'" using os alive
@@ -198,7 +208,7 @@
have
-
+*)
lemma enrich_proc_prop:
"\<lbrakk>valid s; is_created_proc s p; p' \<notin> all_procs s\<rbrakk>
@@ -208,6 +218,7 @@
(\<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 {})))"
+sorry (*
proof (induct s)
case Nil
thus ?case by (auto simp:is_created_proc_def)
@@ -237,17 +248,17 @@
sorry
ultimately show ?case by auto
qed
-
+*)
lemma "alive s obj \<Longrightarrow> alive (enrich_proc s p p') obj"
apply (induct s, simp)
-apply (case_tac a, case_tac[!] obj)
+apply (case_tac a, case_tac[!] obj) sorry (*
apply (auto simp:is_file_def is_dir_def split:option.splits t_inode_tag.splits)
thm is_file_other
-
+*)
lemma enrich_proc_valid:
- "\<lbrakk>p \<in> current_procs s; valid s; p \<in> init_procs \<longrightarrow> deleted (O_proc p) s; p' \<notin> current_procs s\<rbrakk> \<Longrightarrow> valid (enrich_proc s p p')"
+ "\<lbrakk>p \<in> current_procs s; valid s; p \<in> init_procs \<longrightarrow> deleted (O_proc p) s; p' \<notin> current_procs s\<rbrakk> \<Longrightarrow> valid (enrich_proc s p p')" (*
apply (induct s, simp)
apply (frule vd_cons, frule vt_grant, frule vt_grant_os, case_tac a)
apply (auto intro!:valid.intros(2))
@@ -256,6 +267,9 @@
end
+*)
+sorry
+
(* for any created obj, we can enrich trace with events that create new objs with the same static-properties *)
definition enriched:: "t_state \<Rightarrow> t_object set \<Rightarrow> t_state \<Rightarrow> bool"
@@ -387,7 +401,6 @@
apply (auto simp:search_check_s_def search_check.simps cf2sfile_def sroot_def
root_sec_remains init_sectxt_prop sec_of_root_valid
dest!:root_is_dir' current_has_sec' split:option.splits)
-apply (simp add:search_check_allp_def)
done
lemma grant_execve_intro_execve:
@@ -437,9 +450,10 @@
\<Longrightarrow> inherit_fds_check_s sec' sfds'"
apply (frule vt_grant_os, frule vd_cons, frule vt_grant)
apply (auto simp:cp2sproc_def cpfd2sfds_execve inherit_fds_check_def inherit_fds_check_s_def split:option.splits)
+sorry (*
apply (erule_tac x = "(ad, ae, bc)" in ballE, auto simp:sectxts_of_sfds_def sectxts_of_fds_def)
apply (erule_tac x = fd in ballE, auto simp:cfd2sfd_def split:option.splits)
-done
+done *)
lemma d2s_main_execve_grant_aux:
"\<lbrakk>cp2sproc (Execve p f fds # s) p = Some (pi', sec', sfds', shms'); valid (Execve p f fds # s);
@@ -554,11 +568,10 @@
moreover from a4 os p4 have "is_file s' f"
apply (simp add:reserved_def)
by (erule_tac x = "O_file f" in allE, auto)
- moreover from a4 os p4 have "fds \<subseteq> current_proc_fds s' p"
+ moreover from a4 os p4 vd have "fds \<subseteq> proc_file_fds s' p"
apply (rule_tac subsetI, clarsimp simp:reserved_def current_proc_fds.simps)
apply (erule_tac x = "O_fd p x" in allE, erule impE)
- apply (simp, erule set_mp, simp+)
- done
+ sorry
ultimately have "os_grant s' e"
by (simp add:p4)
moreover have "grant s' e"
@@ -571,6 +584,7 @@
apply (simp)
sorry
qed
+qed
lemma s2d_main_execve:
"\<lbrakk>grant_execve pctxt fsec pctxt'; ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; S_file sfs tagf \<in> ss;
@@ -582,7 +596,7 @@
thm update_ss_def
sorry
-
+(*
lemma s2d_main_execve:
"ss \<in> static \<Longrightarrow> \<exists> s. valid s \<and> s2ss s = ss"
apply (erule static.induct)
@@ -592,24 +606,24 @@
apply (erule exE|erule conjE)+
-apply (simp add:update_ss_def)
sorry
-
+*)
(*********************** uppest-level 3 theorems ***********************)
lemma enrichability:
"\<lbrakk>valid s; \<forall> obj \<in> objs. alive s obj \<and> is_created s obj \<and> recorded obj\<rbrakk>
- \<Longrightarrow> enrichable s objs"
+ \<Longrightarrow> enrichable s objs" sorry (*
proof (induct s arbitrary:objs)
case Nil
hence "objs = {}"
- apply (auto simp:is_created_def)
+ apply (auto)
apply (erule_tac x = x in ballE)
+ apply (case_tac x)
apply (auto simp:init_alive_prop)
- done
+ sorry (* done *)
thus ?case using Nil unfolding enrichable_def enriched_def reserved_def
by (rule_tac x = "[]" in exI, auto)
next
@@ -676,7 +690,8 @@
done
-qed
+qed
+*)
lemma d2s_main:
"valid s \<Longrightarrow> s2ss s \<propto> static"
@@ -702,6 +717,17 @@
sorry
+lemma s2d_main':
+ "ss \<in> static \<Longrightarrow> \<exists> s. valid s \<and> s2ss s \<doteq> ss"
+apply (erule static.induct)
+apply (rule_tac x = "[]" in exI, simp add:s2ss_nil_prop valid.intros)
+
+apply (erule exE|erule conjE)+
+
+apply (simp add:update_ss_def)
+
+sorry
+
end
--- a/Final_theorem.thy Wed Nov 20 21:33:56 2013 +0800
+++ b/Final_theorem.thy Mon Dec 02 10:52:40 2013 +0800
@@ -67,6 +67,7 @@
apply (simp add:init_ss_eq_def)
apply (rule_tac x = "ss'" in bexI)
apply (rule_tac x = "sobj" in exI)
+ thm tainted_s_subset_prop
by (auto intro:tainted_s_subset_prop)
qed
@@ -186,6 +187,43 @@
declare co2sobj.simps [simp add]
+lemma subseteq_D:
+ "\<lbrakk> S \<subseteq> {x. P x}; x \<in> S\<rbrakk> \<Longrightarrow> P x"
+by auto
+
+lemma "\<lbrakk>tainted_s ss sobj; ss \<in> static; is_init_sobj sobj\<rbrakk>
+ \<Longrightarrow> \<exists> s. valid s \<and> co2sobj s obj = Some sobj \<and> obj \<in> tainted s"
+apply (drule s2d_main')
+apply (erule exE, erule conjE, simp add:s2ss_def init_ss_eq_def, erule conjE)
+apply (rule_tac x = s in exI, simp)
+apply (case_tac sobj, simp_all only:tainted_s.simps)
+thm set_eq_D
+apply (simp split:option.splits)
+apply (erule conjE, drule_tac S = ss in set_eq_D, simp, (erule exE|erule conjE)+)
+apply (rule_tac x = obj in exI, simp)
+apply (case_tac obj, (simp split:option.splits if_splits)+)
+
+lemma tainted_s_imp_tainted:
+ "\<lbrakk>tainted_s ss sobj; ss \<in> static; init_obj_related sobj obj\<rbrakk>
+ \<Longrightarrow> \<exists> s. valid s \<and> co2sobj s obj = Some sobj \<and> obj \<in> tainted s"
+apply (drule s2d_main')
+apply (erule exE, erule conjE, simp add:s2ss_def init_ss_eq_def, erule conjE)
+apply (rule_tac x = s in exI, simp)
+apply (case_tac sobj, simp_all)
+apply (case_tac[!] obj, simp_all del:co2sobj.simps)
+apply (simp split:option.splits)
+apply (erule conjE, drule_tac S = ss in set_eq_D, simp, (erule exE|erule conjE)+)
+apply (rule_tac x = obj in exI, simp)
+apply (case_tac obj, (simp split:option.splits if_splits)+)
+
+apply (erule conjE, drule_tac S = ss in set_eq_D, simp, (erule exE|erule conjE)+)
+apply (rule_tac x = obj in exI, simp)
+apply (case_tac obj, (simp split:option.splits if_splits)+)
+sorry
+
+
+
+
lemma tainted_s_imp_tainted:
"\<lbrakk>tainted_s ss sobj; ss \<in> static\<rbrakk> \<Longrightarrow> \<exists> s obj. valid s \<and> co2sobj s obj = Some sobj \<and> obj \<in> tainted s"
apply (drule s2d_main)
--- a/Static.thy Wed Nov 20 21:33:56 2013 +0800
+++ b/Static.thy Mon Dec 02 10:52:40 2013 +0800
@@ -359,7 +359,7 @@
(\<forall> sec flag sf'. (sec,flag,sf') \<in> sfds \<and> (sf \<preceq> sf') \<longrightarrow>
(sec, flag, sf') \<in> sfds' \<and> (sec,flag, file_after_rename sf' from to) \<notin> sfds')"
-(* for not many, choose on renamed or not *)
+ (* for not many, choose on renamed or not *)
definition sfds_rename_choices
:: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
where
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Alive_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,335 @@
+theory Alive_prop
+imports Main Flask_type Flask Current_files_prop Current_sockets_prop Init_prop Proc_fd_of_file_prop
+begin
+
+context flask begin
+
+lemma distinct_queue_msgs:
+ "\<lbrakk>q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> distinct (msgs_of_queue s q)"
+apply (induct s)
+apply (simp add:init_msgs_distinct)
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply auto
+apply (case_tac "msgs_of_queue s q", simp+)
+done
+
+lemma received_msg_notin:
+ "\<lbrakk>hd (msgs_of_queue s q) \<in> set (tl (msgs_of_queue s q)); q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> False"
+apply (drule distinct_queue_msgs, simp)
+apply (case_tac "msgs_of_queue s q", auto)
+done
+
+lemma other_msg_remains:
+ "\<lbrakk>m \<noteq> hd (msgs_of_queue s q); q \<in> current_msgqs s; valid s\<rbrakk>
+ \<Longrightarrow> (m \<in> set (tl (msgs_of_queue s q))) = (m \<in> set (msgs_of_queue s q))"
+apply (drule distinct_queue_msgs, simp)
+apply (case_tac "msgs_of_queue s q", auto)
+done
+
+lemma is_tcp_in_current:
+ "is_tcp_sock \<tau> s \<Longrightarrow> s \<in> current_sockets \<tau>"
+by (auto simp:is_tcp_sock_def current_sockets_def split:option.splits)
+
+lemma is_udp_in_current:
+ "is_udp_sock \<tau> s \<Longrightarrow> s \<in> current_sockets \<tau>"
+by (auto simp:is_udp_sock_def current_sockets_def split:option.splits)
+
+(************ alive simpset **************)
+
+lemma alive_open:
+ "valid (Open p f flag fd opt # s) \<Longrightarrow> alive (Open p f flag fd opt # s) = (
+ \<lambda> obj. case obj of
+ O_fd p' fd' \<Rightarrow> if (p' = p \<and> fd' = fd) then True
+ else alive s obj
+ | O_file f' \<Rightarrow> (if (opt = None) then alive s obj
+ else if (f = f') then True
+ else alive s obj)
+ | _ \<Rightarrow> alive s obj)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps
+ dest:is_dir_in_current split:if_splits option.splits)
+done
+
+(* sock is not inherited on Execve, Clone cases, cause I simplified it with os_grant.
+ * chunhan, 2013-11-20
+lemma alive_execve:
+ "valid (Execve p f fds # s) \<Longrightarrow> alive (Execve p f fds # s) = (
+ \<lambda> obj. case obj of
+ O_fd p' fd \<Rightarrow> (if (p = p' \<and> fd \<in> fds) then alive s (O_fd p fd)
+ else if (p = p') then False
+ else alive s (O_fd p' fd))
+ | O_tcp_sock (p', fd) \<Rightarrow> (if (p = p' \<and> fd \<in> fds) then alive s (O_tcp_sock (p, fd))
+ else if (p = p') then False
+ else alive s (O_tcp_sock (p', fd)))
+ | O_udp_sock (p', fd) \<Rightarrow> (if (p = p' \<and> fd \<in> fds) then alive s (O_udp_sock (p, fd))
+ else if (p = p') then False
+ else alive s (O_udp_sock (p', fd)))
+ | _ \<Rightarrow> alive s obj )"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps
+ dest:is_dir_in_current split:if_splits option.splits)
+
+done
+
+lemma alive_clone:
+ "valid (Clone p p' fds shms # s) \<Longrightarrow> alive (Clone p p' fds shms # s) = (
+ \<lambda> obj. case obj of
+ O_proc p'' \<Rightarrow> if (p'' = p') then True else alive s obj
+ | O_fd p'' fd \<Rightarrow> if (p'' = p' \<and> fd \<in> fds) then True
+ else if (p'' = p') then False
+ else alive s obj
+ | O_tcp_sock (p'', fd) \<Rightarrow> (if (p'' = p' \<and> fd \<in> fds) then alive s (O_tcp_sock (p, fd))
+ else if (p'' = p') then False
+ else alive s (O_tcp_sock (p'', fd)))
+ | O_udp_sock (p'', fd) \<Rightarrow> (if (p'' = p' \<and> fd \<in> fds) then alive s (O_udp_sock (p, fd))
+ else if (p'' = p') then False
+ else alive s (O_udp_sock (p'', fd)))
+ | _ \<Rightarrow> alive s obj )"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps
+ intro:is_tcp_in_current is_udp_in_current
+ dest:is_dir_in_current split:if_splits option.splits)
+done
+*)
+
+lemma alive_execve:
+ "valid (Execve p f fds # s) \<Longrightarrow> alive (Execve p f fds # s) = (
+ \<lambda> obj. case obj of
+ O_fd p' fd \<Rightarrow> (if (p = p' \<and> fd \<in> fds) then alive s (O_fd p fd)
+ else if (p = p') then False
+ else alive s (O_fd p' fd))
+ | O_tcp_sock (p', fd) \<Rightarrow> (if (p = p') then False
+ else alive s (O_tcp_sock (p', fd)))
+ | O_udp_sock (p', fd) \<Rightarrow> (if (p = p') then False
+ else alive s (O_udp_sock (p', fd)))
+ | _ \<Rightarrow> alive s obj )"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps
+ dest:is_dir_in_current split:if_splits option.splits
+ dest:file_fds_subset_pfds[where s = s] udp_notin_file_fds tcp_notin_file_fds)
+done
+
+lemma alive_clone:
+ "valid (Clone p p' fds # s) \<Longrightarrow> alive (Clone p p' fds # s) = (
+ \<lambda> obj. case obj of
+ O_proc p'' \<Rightarrow> if (p'' = p') then True else alive s obj
+ | O_fd p'' fd \<Rightarrow> if (p'' = p' \<and> fd \<in> fds) then True
+ else if (p'' = p') then False
+ else alive s obj
+ | O_tcp_sock (p'', fd) \<Rightarrow> (if (p'' = p') then False
+ else alive s (O_tcp_sock (p'', fd)))
+ | O_udp_sock (p'', fd) \<Rightarrow> (if (p'' = p') then False
+ else alive s (O_udp_sock (p'', fd)))
+ | _ \<Rightarrow> alive s obj )"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps
+ intro:is_tcp_in_current is_udp_in_current
+ dest:is_dir_in_current split:if_splits option.splits
+ dest:file_fds_subset_pfds[where s = s] udp_notin_file_fds tcp_notin_file_fds)
+done
+
+lemma alive_kill:
+ "valid (Kill p p' # s) \<Longrightarrow> alive (Kill p p' # s) = (
+ \<lambda> obj. case obj of
+ O_proc p'' \<Rightarrow> if (p'' = p') then False else alive s obj
+ | O_fd p'' fd \<Rightarrow> if (p'' = p') then False else alive s obj
+ | O_tcp_sock (p'', fd) \<Rightarrow> if (p'' = p') then False else alive s obj
+ | O_udp_sock (p'', fd) \<Rightarrow> if (p'' = p') then False else alive s obj
+ | _ \<Rightarrow> alive s obj)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps
+ intro:is_tcp_in_current is_udp_in_current
+ dest:is_dir_in_current split:if_splits option.splits)
+done
+
+lemma alive_exit:
+ "valid (Exit p' # s) \<Longrightarrow> alive (Exit p' # s) = (
+ \<lambda> obj. case obj of
+ O_proc p'' \<Rightarrow> if (p'' = p') then False else alive s obj
+ | O_fd p'' fd \<Rightarrow> if (p'' = p') then False else alive s obj
+ | O_tcp_sock (p'', fd) \<Rightarrow> if (p'' = p') then False else alive s obj
+ | O_udp_sock (p'', fd) \<Rightarrow> if (p'' = p') then False else alive s obj
+ | _ \<Rightarrow> alive s obj)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps
+ intro:is_tcp_in_current is_udp_in_current
+ dest:is_dir_in_current split:if_splits option.splits)
+done
+
+lemma alive_closefd:
+ "valid (CloseFd p fd # s) \<Longrightarrow> alive (CloseFd p fd # s) = (
+ \<lambda> obj. case obj of
+ O_fd p' fd' \<Rightarrow> if (p' = p \<and> fd' = fd) then False else alive s obj
+ | O_tcp_sock (p', fd') \<Rightarrow> if (p' = p \<and> fd' = fd) then False else alive s obj
+ | O_udp_sock (p', fd') \<Rightarrow> if (p' = p \<and> fd' = fd) then False else alive s obj
+ | O_file f \<Rightarrow> (case (file_of_proc_fd s p fd) of
+ Some f' \<Rightarrow> (if (f = f' \<and> proc_fd_of_file s f = {(p, fd)} \<and> f \<in> files_hung_by_del s)
+ then False else alive s obj)
+ | _ \<Rightarrow> alive s obj)
+ | _ \<Rightarrow> alive s obj)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps
+ intro:is_tcp_in_current is_udp_in_current
+ dest:is_dir_in_current file_of_pfd_is_file' split:if_splits option.splits)
+done
+
+lemma alive_unlink:
+ "valid (UnLink p f # s) \<Longrightarrow> alive (UnLink p f # s) = (
+ \<lambda> obj. case obj of
+ O_file f' \<Rightarrow> if (f' = f \<and> proc_fd_of_file s f = {}) then False else alive s obj
+ | _ \<Rightarrow> alive s obj)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps
+ intro:is_tcp_in_current is_udp_in_current
+ dest:is_dir_in_current file_of_pfd_is_file' file_dir_conflict
+ split:if_splits option.splits)
+done
+
+lemma alive_rmdir:
+ "valid (Rmdir p d # s) \<Longrightarrow> alive (Rmdir p d # s) = (alive s) (O_dir d := False)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def
+ intro:is_tcp_in_current is_udp_in_current
+ dest:is_dir_in_current file_of_pfd_is_file' file_dir_conflict
+ split:if_splits option.splits)
+done
+
+lemma alive_mkdir:
+ "valid (Mkdir p d inum # s) \<Longrightarrow> alive (Mkdir p d inum # s) = (alive s) (O_dir d := True)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def
+ intro:is_tcp_in_current is_udp_in_current
+ dest:is_dir_in_current file_of_pfd_is_file' is_file_in_current
+ split:if_splits option.splits)
+done
+
+(*
+lemma alive_linkhard:
+ "valid (LinkHard p f f' # s) \<Longrightarrow> alive (LinkHard p f f' # s) = (alive s) (O_file f' := True)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def
+ intro:is_tcp_in_current is_udp_in_current
+ dest:is_dir_in_current file_of_pfd_is_file' is_file_in_current
+ split:if_splits option.splits)
+done
+*)
+
+lemma alive_createmsgq:
+ "valid (CreateMsgq p q # s) \<Longrightarrow> alive (CreateMsgq p q # s) = (alive s) (O_msgq q := True)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def)
+done
+
+lemma alive_sendmsg:
+ "valid (SendMsg p q m # s) \<Longrightarrow> alive (SendMsg p q m # s) = (alive s) (O_msg q m := True)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def)
+done
+
+lemma alive_recvmsg:
+ "valid (RecvMsg p q m # s) \<Longrightarrow> alive (RecvMsg p q m # s) = (alive s) (O_msg q m := False)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def other_msg_remains
+ dest:received_msg_notin)
+done
+
+lemma alive_removemsgq:
+ "valid (RemoveMsgq p q # s) \<Longrightarrow> alive (RemoveMsgq p q # s) = (
+ \<lambda> obj. case obj of
+ O_msgq q' \<Rightarrow> if (q' = q) then False else alive s obj
+ | O_msg q' m \<Rightarrow> if (q' = q) then False else alive s obj
+ | _ \<Rightarrow> alive s obj)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def)
+done
+
+(*
+lemma alive_createshm:
+ "valid (CreateShM p h # s) \<Longrightarrow> alive (CreateShM p h # s) = (alive s) (O_shm h := True)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def)
+done
+
+lemma alive_deleteshm:
+ "valid (DeleteShM p h # s) \<Longrightarrow> alive (DeleteShM p h # s) = (alive s) (O_shm h := False)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def)
+done
+*)
+
+lemma alive_createsock:
+ "valid (CreateSock p af st fd inum # s) \<Longrightarrow> alive (CreateSock p af st fd inum # s) = (
+ \<lambda> obj. case obj of
+ O_fd p' fd' \<Rightarrow> if (p' = p \<and> fd' = fd) then True else alive s obj
+ | O_tcp_sock (p', fd') \<Rightarrow> if (p' = p \<and> fd' = fd \<and> st = STREAM) then True else alive s obj
+ | O_udp_sock (p', fd') \<Rightarrow> if (p' = p \<and> fd' = fd \<and> st = DGRAM) then True else alive s obj
+ | _ \<Rightarrow> alive s obj)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def
+ intro:is_tcp_in_current is_udp_in_current split:t_socket_type.splits)
+done
+
+lemma alive_accept:
+ "valid (Accept p fd addr port fd' inum # s) \<Longrightarrow> alive (Accept p fd addr port fd' inum # s) = (
+ \<lambda> obj. case obj of
+ O_fd p' fd'' \<Rightarrow> if (p' = p \<and> fd'' = fd') then True else alive s obj
+ | O_tcp_sock (p', fd'') \<Rightarrow> if (p' = p \<and> fd'' = fd') then True else alive s obj
+ | _ \<Rightarrow> alive s obj)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac obj)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def
+ intro:is_tcp_in_current is_udp_in_current split:t_socket_type.splits)
+done
+
+lemma alive_other:
+ "\<lbrakk>valid (e # s);
+ \<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
+ \<forall> p f fds. e \<noteq> Execve p f fds;
+ \<forall> p p' fds. e \<noteq> Clone p p' fds;
+ \<forall> p p'. e \<noteq> Kill p p';
+ \<forall> p. e \<noteq> Exit p;
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p f. e \<noteq> UnLink p f;
+ \<forall> p d. e \<noteq> Rmdir p d;
+ \<forall> p d inum. e \<noteq> Mkdir p d inum;
+ \<forall> p f f'. e \<noteq> LinkHard p f f';
+ \<forall> p q. e \<noteq> CreateMsgq p q;
+ \<forall> p q m. e \<noteq> SendMsg p q m;
+ \<forall> p q m. e \<noteq> RecvMsg p q m;
+ \<forall> p q. e \<noteq> RemoveMsgq p q;
+ \<forall> p af st fd inum. e \<noteq> CreateSock p af st fd inum;
+ \<forall> p fd addr port fd' inum. e \<noteq> Accept p fd addr port fd' inum\<rbrakk>
+ \<Longrightarrow> alive (e # s) = alive s"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x, case_tac [!] e)
+apply (auto simp:current_files_simps current_sockets_simps is_file_simps is_dir_simps
+ is_tcp_sock_simps is_udp_sock_simps dir_is_empty_def
+ intro:is_tcp_in_current is_udp_in_current split:t_socket_type.splits)
+done
+
+lemmas alive_simps = alive_open alive_execve alive_clone alive_kill alive_exit alive_closefd alive_unlink
+ alive_rmdir alive_mkdir alive_createmsgq alive_removemsgq
+ alive_createsock alive_accept alive_other alive_sendmsg alive_recvmsg
+
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Co2sobj_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,2377 @@
+(*<*)
+theory Co2sobj_prop
+imports Main Flask Flask_type Static Static_type Sectxt_prop Init_prop Current_files_prop Current_sockets_prop Delete_prop Proc_fd_of_file_prop Tainted_prop Valid_prop Init_prop Alive_prop
+begin
+(*<*)
+
+ML {*
+fun print_ss ss =
+let
+val {simps, congs, procs, ...} = Raw_Simplifier.dest_ss ss
+in
+simps
+end
+*}
+
+
+context tainting_s begin
+
+(********************* cm2smsg simpset ***********************)
+
+lemma cm2smsg_other:
+ "\<lbrakk>valid (e # s); \<forall> p q m. e \<noteq> SendMsg p q m; \<forall> p q m. e \<noteq> RecvMsg p q m; \<forall> p q. e \<noteq> RemoveMsgq p q\<rbrakk>
+ \<Longrightarrow> cm2smsg (e # s) = cm2smsg s"
+apply (frule vt_grant_os, frule vd_cons, rule ext, rule ext)
+unfolding cm2smsg_def
+apply (case_tac e)
+apply (auto simp:sectxt_of_obj_simps tainted_eq_Tainted
+ split:t_object.splits option.splits if_splits
+ dest:not_deleted_init_msg dest!:current_has_sec')
+done
+
+lemma cm2smsg_sendmsg:
+ "valid (SendMsg p q m # s) \<Longrightarrow> cm2smsg (SendMsg p q m # s) = (\<lambda> q' m'.
+ if (q' = q \<and> m' = m)
+ then (case (sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of
+ Some sec \<Rightarrow> Some (Created, sec, O_msg q m \<in> tainted (SendMsg p q m # s))
+ | _ \<Rightarrow> None)
+ else cm2smsg s q' m') "
+apply (frule vd_cons, frule vt_grant_os, rule ext, rule ext)
+apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps
+ split:if_splits option.splits
+ dest:not_deleted_init_msg dest!:current_has_sec')
+done
+
+lemma cm2smsg_recvmsg1:
+ "\<lbrakk>q' \<noteq> q; valid (RecvMsg p q m # s)\<rbrakk> \<Longrightarrow> cm2smsg (RecvMsg p q m # s) q' = cm2smsg s q'"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps
+ split:if_splits option.splits)
+done
+
+lemma cm2smsg_recvmsg2:
+ "\<lbrakk>m' \<noteq> m; valid (RecvMsg p q m # s)\<rbrakk> \<Longrightarrow> cm2smsg (RecvMsg p q m # s) q m' = cm2smsg s q m'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps
+ split:if_splits option.splits)
+done
+
+lemma cm2smsg_recvmsg1':
+ "\<lbrakk>valid (RecvMsg p q m # s); q' \<noteq> q\<rbrakk> \<Longrightarrow> cm2smsg (RecvMsg p q m # s) q' = cm2smsg s q'"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps
+ split:if_splits option.splits)
+done
+
+lemma cm2smsg_recvmsg2':
+ "\<lbrakk>valid (RecvMsg p q m # s); m' \<noteq> m\<rbrakk> \<Longrightarrow> cm2smsg (RecvMsg p q m # s) q m' = cm2smsg s q m'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps
+ split:if_splits option.splits)
+done
+
+lemma cm2smsg_removemsgq:
+ "\<lbrakk>q' \<noteq> q; valid (RemoveMsgq p q # s)\<rbrakk> \<Longrightarrow> cm2smsg (RemoveMsgq p q # s) q' = cm2smsg s q'"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps
+ split:if_splits option.splits)
+done
+
+lemmas cm2smsg_simps = cm2smsg_nil_prop cm2smsg_sendmsg cm2smsg_recvmsg1' cm2smsg_recvmsg2'
+ cm2smsg_removemsgq cm2smsg_other
+
+lemma init_cm2smsg_has_smsg:
+ "\<lbrakk>m \<in> set (init_msgs_of_queue q); q \<in> init_msgqs\<rbrakk> \<Longrightarrow> \<exists> sm. init_cm2smsg q m = Some sm"
+by (auto simp:init_cm2smsg_def split:option.splits dest:init_msg_has_ctxt)
+
+lemma hd_set_prop1:
+ "hd l \<notin> set l \<Longrightarrow> l = []"
+by (case_tac l, auto)
+
+lemma tl_set_prop1:
+ "a \<in> set (tl l) \<Longrightarrow> a \<in> set l"
+by (case_tac l, auto)
+
+lemma current_has_smsg:
+ "\<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sm. cm2smsg s q m = Some sm"
+apply (induct s)
+apply (simp only:cm2smsg_nil_prop msgs_of_queue.simps current_msgqs.simps init_cm2smsg_has_smsg)
+
+apply (frule vt_grant_os, frule vd_cons, case_tac a)
+apply (auto simp:cm2smsg_def sectxt_of_obj_simps split:if_splits option.splits
+ dest!:current_has_sec' hd_set_prop1 dest:not_deleted_init_msg tl_set_prop1)
+done
+
+lemma current_has_smsg':
+ "\<lbrakk>cm2smsg s q m = None; q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> m \<notin> set (msgs_of_queue s q)"
+by (auto dest:current_has_smsg)
+
+lemma cqm2sms_has_sms_aux:
+ "\<forall> m \<in> set ms. sectxt_of_obj s (O_msg q m) \<noteq> None \<Longrightarrow> (\<exists> sms. cqm2sms s q ms = Some sms)"
+by (induct ms, auto split:option.splits simp:cm2smsg_def)
+
+lemma current_has_sms:
+ "\<lbrakk>q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sms. cqm2sms s q (msgs_of_queue s q) = Some sms"
+apply (rule cqm2sms_has_sms_aux)
+apply (auto dest:current_msg_has_sec)
+done
+
+lemma current_has_sms':
+ "\<lbrakk>cqm2sms s q (msgs_of_queue s q) = None; valid s\<rbrakk> \<Longrightarrow> q \<notin> current_msgqs s"
+by (auto dest:current_has_sms)
+
+(********************* cqm2sms simpset ***********************)
+
+lemma cqm2sms_other:
+ "\<lbrakk>valid (e # s); \<forall> p m. e \<noteq> CreateMsgq p q; \<forall> p q m. e \<noteq> SendMsg p q m;
+ \<forall> p q m. e \<noteq> RecvMsg p q m; \<forall> p q. e \<noteq> RemoveMsgq p q\<rbrakk>
+ \<Longrightarrow> cqm2sms (e # s) = cqm2sms s"
+apply (rule ext, rule ext)
+apply (induct_tac xa, simp)
+apply (frule vt_grant_os, frule vd_cons, case_tac e)
+apply (auto split:option.splits dest:cm2smsg_other)
+done
+
+lemma cqm2sms_createmsgq:
+ "valid (CreateMsgq p q # s) \<Longrightarrow> cqm2sms (CreateMsgq p q # s) = (\<lambda> q' ms'.
+ if (q' = q \<and> ms' = []) then Some []
+ else cqm2sms s q' ms')"
+apply (rule ext, rule ext)
+apply (frule vt_grant_os, frule vd_cons, induct_tac ms')
+apply (auto split:if_splits option.splits dest:cm2smsg_other)
+done
+
+lemma cqm2sms_2:
+ "cqm2sms s q (ms @ [m]) =
+ (case (cqm2sms s q ms, cm2smsg s q m) of
+ (Some sms, Some sm) \<Rightarrow> Some (sms @ [sm])
+ | _ \<Rightarrow> None)"
+apply (induct ms, simp split:option.splits)
+by (auto split:option.splits)
+
+lemmas cqm2sms_simps2 = cqm2sms.simps(1) cqm2sms_2
+
+declare cqm2sms.simps [simp del]
+
+lemma cqm2sms_prop1:
+ "cqm2sms s q ms = None \<Longrightarrow> \<exists> m \<in> set ms. cm2smsg s q m = None"
+by (induct ms, auto simp:cqm2sms.simps split:option.splits)
+
+lemma cqm2sms_sendmsg1:
+ "\<lbrakk>valid (SendMsg p q m # s); m \<notin> set ms\<rbrakk>
+ \<Longrightarrow> cqm2sms (SendMsg p q m # s) q' ms = cqm2sms s q' ms"
+apply (frule vt_grant_os, frule vd_cons)
+apply (frule cm2smsg_sendmsg)
+apply (induct ms rule:rev_induct)
+apply (auto split:if_splits option.splits simp:cqm2sms_simps2)
+done
+
+lemma cqm2sms_sendmsg2:
+ "\<lbrakk>valid (SendMsg p q m # s); q' \<noteq> q\<rbrakk>
+ \<Longrightarrow> cqm2sms (SendMsg p q m # s) q' ms = cqm2sms s q' ms"
+apply (frule vt_grant_os, frule vd_cons)
+apply (frule cm2smsg_sendmsg)
+apply (induct ms rule:rev_induct)
+apply (auto split:if_splits option.splits simp:cqm2sms_simps2)
+done
+
+lemma cqm2sms_sendmsg3:
+ "\<lbrakk>valid (SendMsg p q m # s); ms' = msgs_of_queue (SendMsg p q m # s) q\<rbrakk>
+ \<Longrightarrow> cqm2sms (SendMsg p q m # s) q ms' =
+ (case (cqm2sms s q (msgs_of_queue s q), sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of
+ (Some sms, Some sec) \<Rightarrow> Some (sms @ [(Created, sec, O_msg q m \<in> tainted (SendMsg p q m # s))])
+ | _ \<Rightarrow> None)"
+apply (frule vt_grant_os, frule vd_cons)
+apply (frule cm2smsg_sendmsg)
+apply (auto split:if_splits option.splits simp:cqm2sms_simps2 cqm2sms_sendmsg1)
+done
+
+lemma cqm2sms_sendmsg:
+ "\<lbrakk>valid (SendMsg p q m # s); ms' = msgs_of_queue (SendMsg p q m # s) q'\<rbrakk>
+ \<Longrightarrow> cqm2sms (SendMsg p q m # s) q' ms' = (
+ if (q' = q)
+ then (case (cqm2sms s q (msgs_of_queue s q), sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of
+ (Some sms, Some sec) \<Rightarrow> Some (sms @ [(Created, sec, O_msg q m \<in> tainted (SendMsg p q m # s))])
+ | _ \<Rightarrow> None)
+ else cqm2sms s q' ms' )"
+apply (simp split:if_splits add:cqm2sms_sendmsg2 cqm2sms_sendmsg3)
+done
+
+lemma cqm2sms_recvmsg1:
+ "\<lbrakk>valid (RecvMsg p q m # s); m \<notin> set ms\<rbrakk>
+ \<Longrightarrow> cqm2sms (RecvMsg p q m # s) q ms = cqm2sms s q ms"
+apply (frule vt_grant_os, frule vd_cons)
+apply (induct ms rule:rev_induct)
+apply (auto split:if_splits option.splits simp:cqm2sms_simps2 cm2smsg_recvmsg2')
+done
+
+lemma cqm2sms_recvmsg2:
+ "\<lbrakk>valid (RecvMsg p q m # s); q' \<noteq> q\<rbrakk>
+ \<Longrightarrow> cqm2sms (RecvMsg p q m # s) q' ms = cqm2sms s q' ms"
+apply (frule vt_grant_os, frule vd_cons)
+apply (induct ms rule:rev_induct)
+apply (auto split:if_splits option.splits simp:cqm2sms_simps2 cm2smsg_recvmsg1')
+done
+
+lemma cqm2sms_recvmsg:
+ "\<lbrakk>valid (RecvMsg p q m # s); ms = msgs_of_queue (RecvMsg p q m # s) q'\<rbrakk>
+ \<Longrightarrow> cqm2sms (RecvMsg p q m # s) q' ms = (
+ if (q' = q)
+ then (case (cqm2sms s q (msgs_of_queue s q)) of
+ Some sms \<Rightarrow> Some (tl sms)
+ | _ \<Rightarrow> None)
+ else cqm2sms s q' ms)"
+apply (frule vt_grant_os, frule vd_cons)
+apply (auto split:if_splits option.splits simp:cqm2sms_recvmsg1 cqm2sms_recvmsg2
+ dest!:current_has_sms')
+apply (case_tac "msgs_of_queue s q", simp)
+apply (frule_tac ms = "tl (msgs_of_queue s q)" in cqm2sms_recvmsg1)
+apply (drule_tac q = q in distinct_queue_msgs, simp+)
+apply (case_tac a, auto simp:cqm2sms.simps split:option.splits if_splits)
+done
+
+lemma cqm2sms_removemsgq:
+ "\<lbrakk>valid (RemoveMsgq p q # s); q' \<noteq> q; q' \<in> current_msgqs s\<rbrakk>
+ \<Longrightarrow> cqm2sms (RemoveMsgq p q # s) q' ms = cqm2sms s q' ms"
+apply (frule vt_grant_os, frule vd_cons)
+apply (induct ms rule:rev_induct)
+apply (auto simp:cqm2sms_simps2 cm2smsg_removemsgq split:option.splits if_splits)
+done
+
+lemmas cqm2sms_simps = cqm2sms_other cqm2sms_createmsgq cqm2sms_sendmsg cqm2sms_recvmsg cqm2sms_removemsgq
+
+(********************* cq2smsgq simpset ***********************)
+
+lemma cq2smsgq_other:
+ "\<lbrakk>valid (e # s); \<forall> p q. e \<noteq> CreateMsgq p q; \<forall> p q m. e \<noteq> SendMsg p q m;
+ \<forall> p q m. e \<noteq> RecvMsg p q m; \<forall> p q. e \<noteq> RemoveMsgq p q\<rbrakk>
+ \<Longrightarrow> cq2smsgq (e # s) = cq2smsgq s"
+apply (frule cqm2sms_other, simp+)
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac e)
+apply (auto simp:cq2smsgq_def sectxt_of_obj_simps
+ split:t_object.splits option.splits if_splits
+ dest:not_deleted_init_msg)
+done
+
+lemma cq2smsg_createmsgq:
+ "valid (CreateMsgq p q # s)
+ \<Longrightarrow> cq2smsgq (CreateMsgq p q # s) = (cq2smsgq s) (q :=
+ case (sectxt_of_obj s (O_proc p)) of
+ Some psec \<Rightarrow> Some (Created, (fst psec, R_object, (snd o snd) psec), [])
+ | None \<Rightarrow> None)"
+apply (frule vd_cons, frule vt_grant_os)
+apply (frule cqm2sms_createmsgq)
+apply (rule ext, auto simp:cq2smsgq_def sec_createmsgq
+ split:option.splits if_splits dest:not_deleted_init_msgq)
+done
+
+lemma cq2smsg_sendmsg:
+ "valid (SendMsg p q m # s)
+ \<Longrightarrow> cq2smsgq (SendMsg p q m # s) = (cq2smsgq s) (q :=
+ case (cq2smsgq s q, sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of
+ (Some (qi, sec, sms), Some msec) \<Rightarrow>
+ Some (qi, sec, sms @ [(Created, msec, O_msg q m \<in> tainted (SendMsg p q m # s))])
+ | _ \<Rightarrow> None)"
+apply (frule vd_cons, frule vt_grant_os)
+apply (rule ext)
+apply (frule_tac q' = x in cqm2sms_sendmsg, simp)
+apply (auto simp:cq2smsgq_def sectxt_of_obj_simps split:option.splits if_splits
+ dest!:current_has_sms' current_has_sec')
+done
+
+lemma current_has_smsgq:
+ "\<lbrakk>q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sq. cq2smsgq s q = Some sq"
+by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sec' current_has_sms')
+
+lemma current_has_smsgq':
+ "\<lbrakk>cq2smsgq s q = None; valid s\<rbrakk> \<Longrightarrow> q \<notin> current_msgqs s"
+by (auto dest:current_has_smsgq)
+
+lemma cq2smsg_recvmsg:
+ "valid (RecvMsg p q m # s)
+ \<Longrightarrow> cq2smsgq (RecvMsg p q m # s) = (cq2smsgq s) (q :=
+ case (cq2smsgq s q) of
+ Some (qi, sec, sms) \<Rightarrow> Some (qi, sec, tl sms)
+ | _ \<Rightarrow> None)"
+apply (frule vd_cons, frule vt_grant_os)
+apply (rule ext, frule_tac q' = x in cqm2sms_recvmsg, simp)
+apply (auto simp add:cq2smsgq_def sectxt_of_obj_simps split:option.splits if_splits
+ dest!:current_has_sec' current_has_sms' current_has_smsgq')
+done
+
+lemma cq2smsg_removemsgq:
+ "\<lbrakk>valid (RemoveMsgq p q # s); q' \<noteq> q; q' \<in> current_msgqs s\<rbrakk>
+ \<Longrightarrow> cq2smsgq (RemoveMsgq p q # s) q' = cq2smsgq s q'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (auto simp:cq2smsgq_def sectxt_of_obj_simps cqm2sms_removemsgq split:if_splits option.splits
+ dest!:current_has_sec' current_has_sms' current_has_smsgq')
+done
+
+lemmas cq2smsgq_simps = cq2smsgq_other cq2smsg_sendmsg cq2smsg_recvmsg cq2smsg_removemsgq
+
+lemma sm_in_sqsms_init_aux:
+ "\<lbrakk>m \<in> set ms; init_cm2smsg q m = Some sm; q \<in> init_msgqs;
+ init_cqm2sms q ms = Some sms\<rbrakk> \<Longrightarrow> sm \<in> set sms"
+apply (induct ms arbitrary:m sm sms)
+by (auto split:if_splits option.splits)
+
+lemma sm_in_sqsms_init:
+ "\<lbrakk>m \<in> set (init_msgs_of_queue q); init_cm2smsg q m = Some sm; q \<in> init_msgqs;
+ init_cqm2sms q (init_msgs_of_queue q) = Some sms\<rbrakk> \<Longrightarrow> sm \<in> set sms"
+by (simp add:sm_in_sqsms_init_aux)
+
+lemma cqm2sms_prop0:
+ "\<lbrakk>m \<in> set ms; cm2smsg s q m = Some sm; cqm2sms s q ms = Some sms\<rbrakk> \<Longrightarrow> sm \<in> set sms"
+apply (induct ms arbitrary:m sm sms)
+apply (auto simp:cqm2sms.simps split:option.splits)
+done
+
+lemma sm_in_sqsms:
+ "\<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s; valid s; cq2smsgq s q = Some (qi, qsec, sms);
+ cm2smsg s q m = Some sm\<rbrakk> \<Longrightarrow> sm \<in> set sms"
+proof (induct s arbitrary:m q qi qsec sms sm)
+ case Nil
+ thus ?case
+ by (simp add:cq2smsga_nil_prop cm2smsg_nil_prop init_cq2smsgq_def sm_in_sqsms_init
+ split:option.splits)
+next
+ case (Cons e s)
+ hence p1:"\<And> m q qi qsec sms sm. \<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s; valid s;
+ cq2smsgq s q = Some (qi, qsec, sms); cm2smsg s q m = Some sm\<rbrakk>
+ \<Longrightarrow> sm \<in> set sms" and p2: "m \<in> set (msgs_of_queue (e # s) q)"
+ and p3: "q \<in> current_msgqs (e # s)" and p4: "valid (e # s)"
+ and p5: "cq2smsgq (e # s) q = Some (qi, qsec, sms)"
+ and p6: "cm2smsg (e # s) q m = Some sm" by auto
+ from p4 have os: "os_grant s e" and vd: "valid s" by (auto dest:vd_cons vt_grant_os)
+(*
+ from p1 have p1':
+ "\<And> m q qi qsec sms sm ms. \<lbrakk>m \<in> set ms; set ms \<subseteq> set (msgs_of_queue s q); q \<in> current_msgqs s;
+ valid s; cq2smsgq s q = Some (qi, qsec, sms); cm2smsg s q m = Some sm\<rbrakk>
+ \<Longrightarrow> sm \<in> set "
+*)
+ show ?case using p2 p3 p4 p5 p6 vd os
+ apply (case_tac e)
+ apply (auto simp:cq2smsgq_simps cm2smsg_simps split:if_splits option.splits intro:p1)
+
+ apply (rule_tac m = m and q = q and qi = qi and qsec = qsec in p1, simp+)
+ apply (case_tac "q = nat2", simp)
+ apply (drule cq2smsg_createmsgq, simp, simp)
+
+ apply (drule_tac m = m and q = q and qi = qi and qsec = "(aa,ab,b)" in p1, simp+)
+
+ apply (simp add:cq2smsgq_def split:option.splits)
+ apply (rule_tac m = m and sm = sm in cqm2sms_prop0, simp+)
+ apply (simp add:cqm2sms_recvmsg)
+ done
+qed
+
+(****************** cf2sfile path simpset ***************)
+
+lemma sroot_only:
+ "cf2sfile s [] = Some sroot"
+by (simp add:cf2sfile_def)
+
+lemma not_file_is_dir:
+ "\<lbrakk>\<not> is_file s f; f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> is_dir s f"
+by (auto simp:is_file_def current_files_def is_dir_def
+ dest:finum_has_itag finum_has_ftag' split:t_inode_tag.splits option.splits)
+
+lemma not_dir_is_file:
+ "\<lbrakk>\<not> is_dir s f; f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> is_file s f"
+by (auto simp:is_file_def current_files_def is_dir_def
+ dest:finum_has_itag finum_has_ftag' split:t_inode_tag.splits option.splits)
+
+lemma is_file_or_dir:
+ "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> is_file s f \<or> is_dir s f"
+by (auto dest:not_dir_is_file)
+
+lemma current_file_has_sfile:
+ "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> sf. cf2sfile s f = Some sf"
+apply (induct f)
+apply (rule_tac x = "sroot" in exI, simp add:sroot_only)
+apply (frule parentf_in_current', simp, clarsimp)
+apply (frule parentf_is_dir'', simp)
+apply (frule is_file_or_dir, simp)
+apply (auto dest!:current_has_sec'
+ simp:cf2sfile_def split:option.splits if_splits dest!:get_pfs_secs_prop')
+done
+
+definition sectxt_of_pf :: "t_state \<Rightarrow> t_file \<Rightarrow> security_context_t option"
+where
+ "sectxt_of_pf s f = (case f of [] \<Rightarrow> None | (a # pf) \<Rightarrow> sectxt_of_obj s (O_dir pf))"
+
+definition get_parentfs_ctxts' :: "t_state \<Rightarrow> t_file \<Rightarrow> (security_context_t list) option"
+where
+ "get_parentfs_ctxts' s f = (case f of [] \<Rightarrow> None | (a # pf) \<Rightarrow> get_parentfs_ctxts s pf)"
+
+lemma is_file_has_sfile:
+ "\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> sec psec asecs. cf2sfile s f = Some
+ (if (\<not> deleted (O_file f) s \<and> is_init_file f) then Init f else Created,
+ sec, Some psec, set asecs) \<and> (sectxt_of_obj s (O_file f) = Some sec) \<and>
+ (sectxt_of_pf s f = Some psec) \<and> (get_parentfs_ctxts' s f = Some asecs)"
+apply (case_tac f, simp, drule root_is_dir', simp, simp)
+apply (frule is_file_in_current)
+apply (drule current_file_has_sfile, simp)
+apply (auto simp:cf2sfile_def sectxt_of_pf_def get_parentfs_ctxts'_def split:if_splits option.splits)
+done
+
+lemma is_file_has_sfile':
+ "\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> sf. cf2sfile s f = Some sf"
+by (drule is_file_has_sfile, auto)
+
+lemma is_dir_has_sfile:
+ "\<lbrakk>is_dir s f; valid s\<rbrakk> \<Longrightarrow> (case f of
+ [] \<Rightarrow> cf2sfile s f = Some sroot
+ | a # pf \<Rightarrow> (\<exists> sec psec asecs. cf2sfile s f = Some
+ (if (\<not> deleted (O_dir f) s \<and> is_init_dir f) then Init f else Created,
+ sec, Some psec, set asecs) \<and> (sectxt_of_obj s (O_dir f) = Some sec) \<and>
+ (sectxt_of_obj s (O_dir pf) = Some psec) \<and> (get_parentfs_ctxts s pf = Some asecs)))"
+apply (case_tac f, simp add:sroot_only)
+apply (frule is_dir_in_current, frule is_dir_not_file)
+apply (drule current_file_has_sfile, simp)
+apply (auto simp:cf2sfile_def split:if_splits option.splits)
+done
+
+lemma is_dir_has_sdir':
+ "\<lbrakk>is_dir s f; valid s\<rbrakk> \<Longrightarrow> \<exists> sf. cf2sfile s f = Some sf"
+apply (case_tac f)
+apply (rule_tac x = sroot in exI)
+apply (simp add:sroot_only)
+apply (drule is_dir_has_sfile, auto)
+done
+
+lemma sroot_set:
+ "valid s \<Longrightarrow> \<exists> sec. sroot = (Init [], sec, None, {}) \<and> sectxt_of_obj s (O_dir []) = Some sec"
+apply (frule root_is_dir)
+apply (drule is_dir_has_sec, simp)
+apply (auto simp:sroot_def sec_of_root_def sectxt_of_obj_def type_of_obj.simps
+ root_type_remains root_user_remains
+ dest!:root_has_type' root_has_user' root_has_init_type' root_has_init_user'
+ split:option.splits)
+done
+
+lemma cf2sfile_path_file:
+ "\<lbrakk>is_file s (f # pf); valid s\<rbrakk>
+ \<Longrightarrow> cf2sfile s (f # pf) = (
+ case (cf2sfile s pf) of
+ Some (pfi, pfsec, psec, asecs) \<Rightarrow>
+ (case (sectxt_of_obj s (O_file (f # pf))) of
+ Some fsec \<Rightarrow> Some (if (\<not> deleted (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
+ else Created, fsec, Some pfsec, asecs \<union> {pfsec})
+ | None \<Rightarrow> None)
+ | _ \<Rightarrow> None)"
+apply (frule is_file_in_current, drule parentf_is_dir'', simp)
+apply (frule is_dir_has_sfile, simp, frule is_file_has_sfile, simp)
+apply (frule sroot_set)
+apply (case_tac pf, (clarsimp simp:get_parentfs_ctxts'_def sectxt_of_pf_def)+)
+done
+
+lemma cf2sfile_path_dir:
+ "\<lbrakk>is_dir s (f # pf); valid s\<rbrakk>
+ \<Longrightarrow> cf2sfile s (f # pf) = (
+ case (cf2sfile s pf) of
+ Some (pfi, pfsec, psec, asecs) \<Rightarrow>
+ (case (sectxt_of_obj s (O_dir (f # pf))) of
+ Some fsec \<Rightarrow> Some (if (\<not> deleted (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
+ else Created, fsec, Some pfsec, asecs \<union> {pfsec})
+ | None \<Rightarrow> None)
+ | _ \<Rightarrow> None)"
+apply (frule is_dir_in_current, drule parentf_is_dir'', simp)
+apply (frule_tac f = "f # pf" in is_dir_has_sfile, simp)
+apply (frule_tac f = "pf" in is_dir_has_sfile, simp)
+apply (frule sroot_set)
+apply (case_tac pf, (clarsimp simp:get_parentfs_ctxts'_def sectxt_of_pf_def)+)
+done
+
+lemma cf2sfile_path:
+ "\<lbrakk>f # pf \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> cf2sfile s (f # pf) = (
+ case (cf2sfile s pf) of
+ Some (pfi, pfsec, psec, asecs) \<Rightarrow> (if (is_file s (f # pf))
+ then (case (sectxt_of_obj s (O_file (f # pf))) of
+ Some fsec \<Rightarrow> Some (if (\<not> deleted (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
+ else Created, fsec, Some pfsec, asecs \<union> {pfsec})
+ | None \<Rightarrow> None)
+ else (case (sectxt_of_obj s (O_dir (f # pf))) of
+ Some fsec \<Rightarrow> Some (if (\<not> deleted (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
+ else Created, fsec, Some pfsec, asecs \<union> {pfsec})
+ | None \<Rightarrow> None) )
+ | None \<Rightarrow> None)"
+apply (drule is_file_or_dir, simp)
+apply (erule disjE)
+apply (frule cf2sfile_path_file, simp) defer
+apply (frule cf2sfile_path_dir, simp, drule is_dir_not_file)
+apply (auto split:option.splits)
+done
+
+lemma cf2sfile_path_file_prop1:
+ "\<lbrakk>is_file s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs); valid s\<rbrakk>
+ \<Longrightarrow> \<exists> fsec. cf2sfile s (f # pf) =
+ Some (if (\<not> deleted (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
+ else Created, fsec, Some pfsec, asecs \<union> {pfsec}) \<and>
+ sectxt_of_obj s (O_file (f # pf)) = Some fsec"
+apply (frule is_file_has_sfile, simp)
+by (auto simp:cf2sfile_path_file)
+
+lemma cf2sfile_path_file_prop2:
+ "\<lbrakk>is_file s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs);
+ sectxt_of_obj s (O_file (f # pf)) = Some fsec; valid s\<rbrakk> \<Longrightarrow> cf2sfile s (f # pf) =
+ Some (if (\<not> deleted (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
+ else Created, fsec, Some pfsec, asecs \<union> {pfsec})"
+by (drule cf2sfile_path_file_prop1, auto)
+
+lemma cf2sfile_path_dir_prop1:
+ "\<lbrakk>is_dir s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs); valid s\<rbrakk>
+ \<Longrightarrow> \<exists> fsec. cf2sfile s (f # pf) =
+ Some (if (\<not> deleted (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
+ else Created, fsec, Some pfsec, asecs \<union> {pfsec}) \<and>
+ sectxt_of_obj s (O_dir (f # pf)) = Some fsec"
+apply (frule is_dir_has_sfile, simp)
+by (auto simp:cf2sfile_path_dir)
+
+lemma cf2sfile_path_dir_prop2:
+ "\<lbrakk>is_dir s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs);
+ sectxt_of_obj s (O_dir (f # pf)) = Some fsec; valid s\<rbrakk> \<Longrightarrow> cf2sfile s (f # pf) =
+ Some (if (\<not> deleted (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
+ else Created, fsec, Some pfsec, asecs \<union> {pfsec})"
+by (drule cf2sfile_path_dir_prop1, auto)
+
+(**************** cf2sfile event list simpset ****************)
+
+lemma cf2sfile_open_none':
+ "valid (Open p f flag fd None # s) \<Longrightarrow> cf2sfile (Open p f flag fd None # s) f'= cf2sfile s f'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (induct f', simp add:cf2sfile_def)
+apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ get_parentfs_ctxts_simps)
+done
+
+lemma cf2sfile_open_none:
+ "valid (Open p f flag fd None # s) \<Longrightarrow> cf2sfile (Open p f flag fd None # s) = cf2sfile s"
+apply (rule ext)
+by (simp add:cf2sfile_open_none')
+
+lemma cf2sfile_open_some1:
+ "\<lbrakk>valid (Open p f flag fd (Some inum) # s); f' \<in> current_files s\<rbrakk>
+ \<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'"
+apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+apply (case_tac "f = f'", simp)
+apply (induct f', simp add:sroot_only, simp)
+apply (frule parentf_in_current', simp+)
+apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ get_parentfs_ctxts_simps)
+done
+
+lemma cf2sfile_open_some2:
+ "\<lbrakk>valid (Open p f flag fd (Some inum) # s); is_file s f'\<rbrakk>
+ \<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'"
+apply (frule vd_cons, drule is_file_in_current)
+by (simp add:cf2sfile_open_some1)
+
+lemma cf2sfile_open_some3:
+ "\<lbrakk>valid (Open p f flag fd (Some inum) # s); is_dir s f'\<rbrakk>
+ \<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'"
+apply (frule vd_cons, drule is_dir_in_current)
+by (simp add:cf2sfile_open_some1)
+
+lemma cf2sfile_open_some4:
+ "valid (Open p f flag fd (Some inum) # s) \<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f = (
+ case (parent f) of
+ Some pf \<Rightarrow> (case (sectxt_of_obj (Open p f flag fd (Some inum) # s) (O_file f), sectxt_of_obj s (O_dir pf),
+ get_parentfs_ctxts s pf) of
+ (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ | _ \<Rightarrow> None)
+ | None \<Rightarrow> None)"
+apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+apply (case_tac f, simp)
+apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ get_parentfs_ctxts_simps)
+apply (rule impI, (erule conjE)+)
+apply (drule not_deleted_init_file, simp+)
+apply (simp add:is_file_in_current)
+done
+
+lemma cf2sfile_open:
+ "\<lbrakk>valid (Open p f flag fd opt # s); f' \<in> current_files (Open p f flag fd opt # s)\<rbrakk>
+ \<Longrightarrow> cf2sfile (Open p f flag fd opt # s) f' = (
+ if (opt = None) then cf2sfile s f'
+ else if (f' = f)
+ then (case (parent f) of
+ Some pf \<Rightarrow> (case (sectxt_of_obj (Open p f flag fd opt # s) (O_file f), sectxt_of_obj s (O_dir pf),
+ get_parentfs_ctxts s pf) of
+ (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ | _ \<Rightarrow> None)
+ | None \<Rightarrow> None)
+ else cf2sfile s f')"
+apply (case_tac opt)
+apply (simp add:cf2sfile_open_none)
+apply (case_tac "f = f'")
+apply (simp add:cf2sfile_open_some4 split:option.splits)
+apply (simp add:cf2sfile_open_some1 current_files_simps)
+done
+
+lemma cf2sfile_mkdir1:
+ "\<lbrakk>valid (Mkdir p f i # s); f' \<in> current_files s\<rbrakk>
+ \<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = cf2sfile s f'"
+apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+apply (case_tac "f = f'", simp)
+apply (induct f', simp add:sroot_only, simp)
+apply (frule parentf_in_current', simp+)
+apply (case_tac "f = f'", simp)
+apply (simp add:cf2sfile_path is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ get_parentfs_ctxts_simps split:if_splits option.splits)
+done
+
+lemma cf2sfile_mkdir2:
+ "\<lbrakk>valid (Mkdir p f i # s); is_file s f'\<rbrakk>
+ \<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = cf2sfile s f'"
+apply (frule vd_cons, drule is_file_in_current)
+by (simp add:cf2sfile_mkdir1)
+
+lemma cf2sfile_mkdir3:
+ "\<lbrakk>valid (Mkdir p f i # s); is_dir s f'\<rbrakk>
+ \<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = cf2sfile s f'"
+apply (frule vd_cons, drule is_dir_in_current)
+by (simp add:cf2sfile_mkdir1)
+
+lemma cf2sfile_mkdir4:
+ "valid (Mkdir p f i # s)
+ \<Longrightarrow> cf2sfile (Mkdir p f i # s) f = (case (parent f) of
+ Some pf \<Rightarrow> (case (sectxt_of_obj (Mkdir p f i # s) (O_dir f), sectxt_of_obj s (O_dir pf),
+ get_parentfs_ctxts s pf) of
+ (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ | _ \<Rightarrow> None)
+ | None \<Rightarrow> None)"
+apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+apply (case_tac f, simp)
+apply (clarsimp simp:os_grant.simps)
+apply (simp add:sectxt_of_obj_simps)
+apply (frule current_proc_has_sec, simp)
+apply (frule is_dir_has_sec, simp)
+apply (frule get_pfs_secs_prop, simp)
+apply (frule is_dir_not_file)
+apply (auto simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ get_parentfs_ctxts_simps split:option.splits if_splits
+ dest:not_deleted_init_dir is_dir_in_current not_deleted_init_file is_file_in_current)
+done
+
+lemma cf2sfile_mkdir:
+ "\<lbrakk>valid (Mkdir p f i # s); f' \<in> current_files (Mkdir p f i # s)\<rbrakk>
+ \<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = (
+ if (f' = f)
+ then (case (parent f) of
+ Some pf \<Rightarrow> (case (sectxt_of_obj (Mkdir p f i # s) (O_dir f), sectxt_of_obj s (O_dir pf),
+ get_parentfs_ctxts s pf) of
+ (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ | _ \<Rightarrow> None)
+ | None \<Rightarrow> None)
+ else cf2sfile s f')"
+apply (case_tac "f = f'")
+apply (simp add:cf2sfile_mkdir4 split:option.splits)
+apply (simp add:cf2sfile_mkdir1 current_files_simps)
+done
+
+lemma cf2sfile_linkhard1:
+ "\<lbrakk>valid (LinkHard p oldf f # s); f' \<in> current_files s\<rbrakk>
+ \<Longrightarrow> cf2sfile (LinkHard p oldf f# s) f' = cf2sfile s f'"
+apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+apply (case_tac "f = f'", simp)
+apply (induct f', simp add:sroot_only, simp)
+apply (frule parentf_in_current', simp+)
+apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ get_parentfs_ctxts_simps split:if_splits option.splits)
+done
+
+lemma cf2sfile_linkhard2:
+ "\<lbrakk>valid (LinkHard p oldf f # s); is_file s f'\<rbrakk>
+ \<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f' = cf2sfile s f'"
+apply (frule vd_cons, drule is_file_in_current)
+by (simp add:cf2sfile_linkhard1)
+
+lemma cf2sfile_linkhard3:
+ "\<lbrakk>valid (LinkHard p oldf f # s); is_dir s f'\<rbrakk>
+ \<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f' = cf2sfile s f'"
+apply (frule vd_cons, drule is_dir_in_current)
+by (simp add:cf2sfile_linkhard1)
+
+lemma cf2sfile_linkhard4:
+ "valid (LinkHard p oldf f # s)
+ \<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f = (case (parent f) of
+ Some pf \<Rightarrow> (case (sectxt_of_obj (LinkHard p oldf f # s) (O_file f), sectxt_of_obj s (O_dir pf),
+ get_parentfs_ctxts s pf) of
+ (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ | _ \<Rightarrow> None)
+ | None \<Rightarrow> None)"
+apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+apply (case_tac f, simp)
+apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ get_parentfs_ctxts_simps)
+apply (rule impI, (erule conjE)+)
+apply (drule not_deleted_init_file, simp+)
+apply (simp add:is_file_in_current)
+done
+
+lemma cf2sfile_linkhard:
+ "\<lbrakk>valid (LinkHard p oldf f # s); f' \<in> current_files (LinkHard p oldf f # s)\<rbrakk>
+ \<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f' = (
+ if (f' = f)
+ then (case (parent f) of
+ Some pf \<Rightarrow> (case (sectxt_of_obj (LinkHard p oldf f # s) (O_file f), sectxt_of_obj s (O_dir pf),
+ get_parentfs_ctxts s pf) of
+ (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ | _ \<Rightarrow> None)
+ | None \<Rightarrow> None)
+ else cf2sfile s f')"
+apply (case_tac "f = f'")
+apply (simp add:cf2sfile_linkhard4 split:option.splits)
+apply (simp add:cf2sfile_linkhard1 current_files_simps)
+done
+
+lemma cf2sfile_other:
+ "\<lbrakk>ff \<in> current_files s;
+ \<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p f. e \<noteq> UnLink p f;
+ \<forall> p f. e \<noteq> Rmdir p f;
+ \<forall> p f i. e \<noteq> Mkdir p f i;
+ \<forall> p f f'. e \<noteq> LinkHard p f f';
+ valid (e # s)\<rbrakk> \<Longrightarrow> cf2sfile (e # s) ff = cf2sfile s ff"
+apply (frule vd_cons, frule vt_grant_os)
+apply (induct ff, simp add:sroot_only)
+apply (frule parentf_in_current', simp+, case_tac e)
+apply (auto simp:current_files_simps is_file_simps is_dir_simps sectxt_of_obj_simps cf2sfile_path
+ split:if_splits option.splits)
+done
+
+lemma cf2sfile_other':
+ "\<lbrakk>valid (e # s);
+ \<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p f. e \<noteq> UnLink p f;
+ \<forall> p f. e \<noteq> Rmdir p f;
+ \<forall> p f i. e \<noteq> Mkdir p f i;
+ \<forall> p f f'. e \<noteq> LinkHard p f f';
+ ff \<in> current_files s\<rbrakk> \<Longrightarrow> cf2sfile (e # s) ff = cf2sfile s ff"
+by (auto intro!:cf2sfile_other)
+
+lemma cf2sfile_unlink:
+ "\<lbrakk>valid (UnLink p f # s); f' \<in> current_files (UnLink p f # s)\<rbrakk>
+ \<Longrightarrow> cf2sfile (UnLink p f # s) f' = cf2sfile s f'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (simp add:current_files_simps split:if_splits)
+apply (auto simp:cf2sfile_def sectxt_of_obj_simps get_parentfs_ctxts_simps is_file_simps is_dir_simps
+ split:if_splits option.splits)
+done
+
+lemma cf2sfile_rmdir:
+ "\<lbrakk>valid (Rmdir p f # s); f' \<in> current_files (Rmdir p f # s)\<rbrakk>
+ \<Longrightarrow> cf2sfile (Rmdir p f # s) f' = cf2sfile s f'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (simp add:current_files_simps split:if_splits)
+apply (auto simp:cf2sfile_def sectxt_of_obj_simps get_parentfs_ctxts_simps is_file_simps is_dir_simps
+ split:if_splits option.splits)
+done
+
+lemma pfdof_simp5: "\<lbrakk>proc_fd_of_file s f = {(p, fd)}; file_of_proc_fd s p fd = None\<rbrakk> \<Longrightarrow> False"
+apply (subgoal_tac "(p, fd) \<in> proc_fd_of_file s f")
+by (simp add:pfdof_simp2, simp)
+
+lemma pfdof_simp6: "proc_fd_of_file s f = {(p, fd)} \<Longrightarrow> file_of_proc_fd s p fd = Some f"
+apply (subgoal_tac "(p, fd) \<in> proc_fd_of_file s f")
+by (simp add:pfdof_simp2, simp)
+
+lemma cf2sfile_closefd:
+ "\<lbrakk>valid (CloseFd p fd # s); f \<in> current_files (CloseFd p fd # s)\<rbrakk>
+ \<Longrightarrow> cf2sfile (CloseFd p fd # s) f = cf2sfile s f"
+apply (frule vd_cons, frule vt_grant_os)
+apply (simp add:current_files_simps split:if_splits option.splits)
+(* costs too much time, but solved *)
+
+apply (auto simp:cf2sfile_def sectxt_of_obj_simps get_parentfs_ctxts_simps is_file_simps is_dir_simps
+ split:if_splits option.splits
+ dest:init_file_dir_conflict pfdof_simp5 pfdof_simp6 file_of_pfd_is_file
+ not_deleted_init_file not_deleted_init_dir is_file_not_dir is_dir_not_file
+ dest!:current_has_sec')
+done
+
+lemmas cf2sfile_simps = cf2sfile_open cf2sfile_mkdir cf2sfile_linkhard cf2sfile_other
+ cf2sfile_unlink cf2sfile_rmdir cf2sfile_closefd
+
+(*********** cfd2sfd simpset *********)
+
+lemma cfd2sfd_open1:
+ "valid (Open p f flags fd opt # s)
+ \<Longrightarrow> cfd2sfd (Open p f flags fd opt # s) p fd =
+ (case (sectxt_of_obj (Open p f flags fd opt # s) (O_fd p fd), cf2sfile (Open p f flags fd opt # s) f) of
+ (Some sec, Some sf) \<Rightarrow> Some (sec, flags, sf)
+ | _ \<Rightarrow> None)"
+by (simp add:cfd2sfd_def sectxt_of_obj_simps split:if_splits)
+
+lemma cfd2sfd_open_some2:
+ "\<lbrakk>valid (Open p f flags fd (Some inum) # s); file_of_proc_fd s p' fd' = Some f'\<rbrakk>
+ \<Longrightarrow> cfd2sfd (Open p f flags fd (Some inum) # s) p' fd' = cfd2sfd s p' fd'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (frule proc_fd_in_fds, simp)
+apply (frule file_of_proc_fd_in_curf, simp)
+apply (case_tac "f = f'", simp)
+apply (simp add:cfd2sfd_def sectxt_of_obj_simps cf2sfile_open_some1)
+apply (case_tac "p = p'", simp)
+apply (rule conjI, rule impI, simp)
+apply (drule cf2sfile_open_some1, simp)
+apply (auto split:option.splits)[1]
+apply simp
+apply (drule cf2sfile_open_some1, simp)
+apply (auto split:option.splits)[1]
+done
+
+lemma cfd2sfd_open_none2:
+ "\<lbrakk>valid (Open p f flags fd None # s); file_of_proc_fd s p' fd' = Some f'\<rbrakk>
+ \<Longrightarrow> cfd2sfd (Open p f flags fd None # s) p' fd' = cfd2sfd s p' fd'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (frule proc_fd_in_fds, simp)
+apply (frule file_of_proc_fd_in_curf, simp)
+apply (simp add:cfd2sfd_def sectxt_of_obj_simps cf2sfile_open_none)
+apply (case_tac "p = p'", simp)
+apply (rule conjI, rule impI, simp)
+apply (drule cf2sfile_open_none)
+apply (auto split:option.splits)[1]
+apply simp
+apply (drule cf2sfile_open_none)
+apply (auto split:option.splits)[1]
+done
+
+lemma cfd2sfd_open2:
+ "\<lbrakk>valid (Open p f flags fd opt # s); file_of_proc_fd s p' fd' = Some f'\<rbrakk>
+ \<Longrightarrow> cfd2sfd (Open p f flags fd opt # s) p' fd' = cfd2sfd s p' fd'"
+apply (case_tac opt)
+apply (simp add:cfd2sfd_open_none2)
+apply (simp add:cfd2sfd_open_some2)
+done
+
+lemma cfd2sfd_open:
+ "\<lbrakk>valid (Open p f flags fd opt # s); file_of_proc_fd (Open p f flags fd opt # s) p' fd' = Some f'\<rbrakk>
+ \<Longrightarrow> cfd2sfd (Open p f flags fd opt # s) p' fd' = (if (p' = p \<and> fd' = fd) then
+ (case (sectxt_of_obj (Open p f flags fd opt # s) (O_fd p fd), cf2sfile (Open p f flags fd opt # s) f) of
+ (Some sec, Some sf) \<Rightarrow> Some (sec, flags, sf)
+ | _ \<Rightarrow> None) else cfd2sfd s p' fd')"
+apply (simp split:if_splits)
+apply (simp add:cfd2sfd_open1 split:option.splits)
+apply (simp add:cfd2sfd_open2)
+apply (rule impI, simp)
+done
+
+lemma cfd2sfd_closefd:
+ "\<lbrakk>valid (CloseFd p fd # s); file_of_proc_fd (CloseFd p fd # s) p' fd' = Some f\<rbrakk>
+ \<Longrightarrow> cfd2sfd (CloseFd p fd # s) p' fd' = cfd2sfd s p' fd'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (frule proc_fd_in_fds, simp)
+apply (frule file_of_proc_fd_in_curf, simp)
+apply (frule cf2sfile_closefd, simp)
+apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
+apply (auto split:option.splits if_splits)
+done
+
+lemma cfd2sfd_clone:
+ "\<lbrakk>valid (Clone p p' fds shms # s); file_of_proc_fd (Clone p p' fds shms # s) p'' fd' = Some f\<rbrakk>
+ \<Longrightarrow> cfd2sfd (Clone p p' fds shms # s) p'' fd' = (
+ if (p'' = p') then cfd2sfd s p fd'
+ else cfd2sfd s p'' fd')"
+apply (frule vd_cons, frule vt_grant_os)
+apply (frule proc_fd_in_fds, simp)
+apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
+apply (frule_tac cf2sfile_other', simp+)
+apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
+apply (case_tac "p'' = p'", simp)
+apply (auto split:option.splits if_splits)[1]
+apply (simp)
+apply (auto split:option.splits if_splits)[1]
+done
+
+lemma cfd2sfd_execve:
+ "\<lbrakk>valid (Execve p f fds # s); file_of_proc_fd (Execve p f fds # s) p' fd' = Some f'\<rbrakk>
+ \<Longrightarrow> cfd2sfd (Execve p f fds # s) p' fd' = cfd2sfd s p' fd'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (frule proc_fd_in_fds, simp)
+apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
+apply (frule_tac cf2sfile_other', simp+)
+apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
+apply (case_tac "p' = p", simp)
+apply (auto split:option.splits if_splits)[1]
+apply (simp)
+apply (auto split:option.splits if_splits)[1]
+done
+
+lemma cfd2sfd_kill:
+ "\<lbrakk>valid (Kill p p'' # s); file_of_proc_fd (Kill p p'' # s) p' fd' = Some f'\<rbrakk>
+ \<Longrightarrow> cfd2sfd (Kill p p'' # s) p' fd' = cfd2sfd s p' fd'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (frule proc_fd_in_fds, simp)
+apply (frule proc_fd_in_procs, simp)
+apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
+apply (frule_tac cf2sfile_other', simp+)
+apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
+apply (auto split:option.splits if_splits)
+done
+
+lemma cfd2sfd_exit:
+ "\<lbrakk>valid (Exit p # s); file_of_proc_fd (Exit p # s) p' fd' = Some f'\<rbrakk>
+ \<Longrightarrow> cfd2sfd (Exit p # s) p' fd' = cfd2sfd s p' fd'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (frule proc_fd_in_fds, simp)
+apply (frule proc_fd_in_procs, simp)
+apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
+apply (frule_tac cf2sfile_other', simp+)
+apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
+apply (auto split:option.splits if_splits)
+done
+
+lemma cfd2sfd_other:
+ "\<lbrakk>valid (e # s); file_of_proc_fd (e # s) p' fd' = Some f';
+ \<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ \<forall> p p'' fds shms. e \<noteq> Clone p p'' fds shms\<rbrakk>
+ \<Longrightarrow> cfd2sfd (e # s) p' fd' = cfd2sfd s p' fd'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (frule proc_fd_in_fds, simp)
+apply (frule proc_fd_in_procs, simp)
+apply (frule file_of_proc_fd_in_curf, simp)
+apply (case_tac e)
+apply (auto intro!:cfd2sfd_execve cfd2sfd_closefd cfd2sfd_kill cfd2sfd_exit)
+apply (auto simp:cfd2sfd_def sectxt_of_obj_simps current_files_simps cf2sfile_simps split:option.splits)
+apply (auto dest!:current_has_sec' dest:file_of_proc_fd_in_curf proc_fd_in_fds)
+done
+
+lemmas cfd2sfd_simps = cfd2sfd_open cfd2sfd_clone cfd2sfd_other
+
+(********** cpfd2sfds simpset **********)
+
+lemma current_filefd_has_flags:
+ "\<lbrakk>file_of_proc_fd s p fd = Some f; valid s\<rbrakk> \<Longrightarrow> \<exists> flags. flags_of_proc_fd s p fd = Some flags"
+apply (induct s arbitrary:p)
+apply (simp only:flags_of_proc_fd.simps file_of_proc_fd.simps init_filefd_prop4)
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto split:if_splits option.splits dest:proc_fd_in_fds)
+done
+
+lemma current_filefd_has_flags':
+ "\<lbrakk>flags_of_proc_fd s p fd = None; valid s\<rbrakk> \<Longrightarrow> file_of_proc_fd s p fd = None"
+apply (case_tac "file_of_proc_fd s p fd")
+apply (simp, drule current_filefd_has_flags, simp+)
+done
+
+lemma current_file_has_sfile':
+ "\<lbrakk>cf2sfile s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
+by (rule notI, drule current_file_has_sfile, simp+)
+
+lemma current_filefd_has_sfd:
+ "\<lbrakk>file_of_proc_fd s p fd = Some f; valid s\<rbrakk> \<Longrightarrow> \<exists>sfd. cfd2sfd s p fd = Some sfd"
+by (auto simp:cfd2sfd_def split:option.splits dest!:current_has_sec' current_file_has_sfile'
+ dest:file_of_proc_fd_in_curf proc_fd_in_fds current_filefd_has_flags)
+
+lemma current_filefd_has_sfd':
+ "\<lbrakk>cfd2sfd s p fd = None; valid s\<rbrakk> \<Longrightarrow> file_of_proc_fd s p fd = None"
+by (case_tac "file_of_proc_fd s p fd", auto dest:current_filefd_has_sfd)
+
+lemma cpfd2sfds_open1:
+ "valid (Open p f flags fd opt # s) \<Longrightarrow>
+ cpfd2sfds (Open p f flags fd opt # s) p = (
+ case (cfd2sfd (Open p f flags fd opt # s) p fd) of
+ Some sfd \<Rightarrow> (cpfd2sfds s p) \<union> {sfd}
+ | _ \<Rightarrow> cpfd2sfds s p)"
+apply (frule vd_cons, frule vt_grant_os)
+apply (split option.splits)
+apply (rule conjI, rule impI, drule current_filefd_has_sfd', simp, simp)
+apply (rule allI, rule impI)
+apply (rule set_eqI, rule iffI)
+apply (case_tac "x = a", simp)
+unfolding cpfd2sfds_def
+apply (erule CollectE, (erule conjE|erule bexE)+)
+apply (simp add:proc_file_fds_def split:if_splits)
+apply (erule exE, rule_tac x = fda in exI)
+apply (simp add:cfd2sfd_open2)
+apply (case_tac "x = a", simp add:proc_file_fds_def)
+apply (rule_tac x = fd in exI, simp+)
+apply (erule conjE|erule bexE)+
+apply (rule_tac x = fda in bexI)
+apply (simp add:proc_file_fds_def, erule exE)
+apply (simp add:cfd2sfd_open2)
+apply (simp add:proc_file_fds_def)
+done
+
+lemma cpfd2sfds_open1':
+ "valid (Open p f flags fd opt # s) \<Longrightarrow>
+ cpfd2sfds (Open p f flags fd opt # s) p = (
+ case (sectxt_of_obj (Open p f flags fd opt # s) (O_fd p fd), cf2sfile (Open p f flags fd opt # s) f) of
+ (Some sec, Some sf) \<Rightarrow> (cpfd2sfds s p) \<union> {(sec, flags, sf)}
+ | _ \<Rightarrow> cpfd2sfds s p)"
+apply (frule cfd2sfd_open1)
+apply (auto dest:cpfd2sfds_open1 split:option.splits)
+done
+
+lemma cpfd2sfds_open2:
+ "\<lbrakk>valid (Open p f flags fd opt # s); p' \<noteq> p\<rbrakk> \<Longrightarrow> cpfd2sfds (Open p f flags fd opt # s) p' = cpfd2sfds s p'"
+apply (frule vt_grant_os, frule vd_cons)
+unfolding cpfd2sfds_def
+apply (rule set_eqI, rule iffI)
+apply (simp add:proc_file_fds_def)
+apply (erule exE|erule conjE)+
+apply (simp only:file_of_proc_fd.simps cfd2sfd_open2 split:if_splits)
+apply (rule_tac x = fda in exI, simp)
+apply (simp add:proc_file_fds_def)
+apply (erule exE|erule conjE)+
+apply (rule_tac x = fda in exI, simp add:cfd2sfd_open2)
+done
+
+lemma cpfd2sfds_open:
+ "valid (Open p f flags fd opt # s)
+ \<Longrightarrow> cpfd2sfds (Open p f flags fd opt # s) = (cpfd2sfds s) (p := (
+ case (sectxt_of_obj (Open p f flags fd opt # s) (O_fd p fd), cf2sfile (Open p f flags fd opt # s) f) of
+ (Some sec, Some sf) \<Rightarrow> (cpfd2sfds s p) \<union> {(sec, flags, sf)}
+ | _ \<Rightarrow> cpfd2sfds s p))"
+apply (rule ext)
+apply (case_tac "x \<noteq> p")
+apply (simp add:cpfd2sfds_open2)
+apply (simp add:cpfd2sfds_open1')
+done
+
+lemma cpfd2sfds_execve:
+ "valid (Execve p f fds # s)
+ \<Longrightarrow> cpfd2sfds (Execve p f fds # s) = (cpfd2sfds s) (p := {sfd. \<exists> fd \<in> fds. \<exists> f. file_of_proc_fd s p fd = Some f \<and> cfd2sfd s p fd = Some sfd})"
+apply (frule vd_cons, frule vt_grant_os)
+apply (rule ext)
+apply (rule set_eqI, rule iffI)
+unfolding cpfd2sfds_def proc_file_fds_def
+apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+apply (simp split:if_splits)
+apply (frule_tac p' = p and fd' = fd in cfd2sfd_other, simp+)
+apply (rule_tac x = fd in bexI, simp+)
+apply (simp add:cpfd2sfds_def proc_file_fds_def)
+apply (frule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+apply (rule_tac x = fd in exI, simp)
+apply (simp split:if_splits)
+apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+apply (rule_tac x = fd in exI, simp)
+apply (frule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+apply (simp add:cpfd2sfds_def proc_file_fds_def)
+apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+apply (rule_tac x = fd in exI, simp)
+apply (frule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+done
+
+lemma cpfd2sfds_clone:
+ "valid (Clone p p' fds shms # s)
+ \<Longrightarrow> cpfd2sfds (Clone p p' fds shms # s) = (cpfd2sfds s) (p' := {sfd. \<exists> fd \<in> fds. \<exists> f. file_of_proc_fd s p fd = Some f \<and> cfd2sfd s p fd = Some sfd})"
+apply (frule vd_cons, frule vt_grant_os)
+apply (rule ext)
+apply (rule set_eqI, rule iffI)
+unfolding cpfd2sfds_def proc_file_fds_def
+apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+apply (simp split:if_splits)
+apply (frule_tac p'' = p' and fd' = fd in cfd2sfd_clone, simp+)
+apply (rule_tac x = fd in bexI, simp+)
+apply (simp add:cpfd2sfds_def proc_file_fds_def)
+apply (frule_tac p'' = x and fd' = fd in cfd2sfd_clone, simp+)
+apply (rule_tac x = fd in exI, simp)
+apply (simp split:if_splits)
+apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+apply (rule_tac x = fd in exI, simp)
+apply (frule_tac p'' = p' and fd' = fd in cfd2sfd_clone, simp+)
+apply (simp add:cpfd2sfds_def proc_file_fds_def)
+apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+apply (rule_tac x = fd in exI, simp)
+apply (frule_tac p'' = x and fd' = fd in cfd2sfd_clone, simp+)
+done
+
+lemma cpfd2sfds_other:
+ "\<lbrakk>valid (e # s);
+ \<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ \<forall> p f fds. e \<noteq> Execve p f fds;
+ \<forall> p p'. e \<noteq> Kill p p';
+ \<forall> p. e \<noteq> Exit p;
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms\<rbrakk> \<Longrightarrow> cpfd2sfds (e # s) = cpfd2sfds s"
+apply (frule vd_cons, frule vt_grant_os)
+apply (rule ext)
+unfolding cpfd2sfds_def proc_file_fds_def
+apply (case_tac e)
+using cfd2sfd_other
+by auto
+
+lemma cpfd2sfds_kill:
+ "valid (Kill p p' # s) \<Longrightarrow> cpfd2sfds (Kill p p' # s) = (cpfd2sfds s) (p' := {})"
+apply (frule vd_cons, frule vt_grant_os)
+apply (rule ext, rule set_eqI)
+unfolding cpfd2sfds_def proc_file_fds_def
+apply (rule iffI)
+apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
+apply (rule_tac x = fd in exI, simp)
+apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
+apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+apply (rule_tac x = fd in exI, simp)
+apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+done
+
+lemma cpfd2sfds_exit:
+ "valid (Exit p # s) \<Longrightarrow> cpfd2sfds (Exit p # s) = (cpfd2sfds s) (p := {})"
+apply (frule vd_cons, frule vt_grant_os)
+apply (rule ext, rule set_eqI)
+unfolding cpfd2sfds_def proc_file_fds_def
+apply (rule iffI)
+apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
+apply (rule_tac x = fd in exI, simp)
+apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
+apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+apply (rule_tac x = fd in exI, simp)
+apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+done
+
+lemma cpfd2sfds_closefd:
+ "valid (CloseFd p fd # s) \<Longrightarrow> cpfd2sfds (CloseFd p fd # s) = (cpfd2sfds s) (p :=
+ if (fd \<in> proc_file_fds s p)
+ then (case (cfd2sfd s p fd) of
+ Some sfd \<Rightarrow> (if (\<exists> fd' f'. fd' \<noteq> fd \<and> file_of_proc_fd s p fd' = Some f' \<and> cfd2sfd s p fd' = Some sfd)
+ then cpfd2sfds s p else cpfd2sfds s p - {sfd})
+ | _ \<Rightarrow> cpfd2sfds s p)
+ else cpfd2sfds s p)"
+apply (frule vd_cons)
+apply (rule ext, rule set_eqI, rule iffI)
+unfolding cpfd2sfds_def proc_file_fds_def
+apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+apply (simp split:if_splits)
+apply (rule conjI, rule impI, rule conjI, rule impI, erule exE)
+apply (frule_tac p = p and fd = fd in current_filefd_has_sfd, simp)
+apply (erule exE, simp)
+apply (rule conjI, rule impI, (erule exE|erule conjE)+)
+apply (rule_tac x = fda in exI, simp, simp add:cfd2sfd_closefd)
+
+apply (rule impI, rule conjI)
+apply (rule_tac x = fda in exI, simp, simp add:cfd2sfd_closefd)
+apply (rule notI, simp)
+apply (erule_tac x = fda in allE, simp add:cfd2sfd_closefd)
+
+apply (rule impI, simp add:cpfd2sfds_def proc_file_fds_def)
+apply (erule exE, rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+
+apply (rule impI| rule conjI)+
+apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+
+apply (rule impI, simp add:cpfd2sfds_def proc_file_fds_def)
+apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+
+apply (simp split:if_splits)
+apply (frule_tac p = p and fd = fd in current_filefd_has_sfd, simp)
+apply (erule exE, simp)
+apply (case_tac "\<exists>fd'. fd' \<noteq> fd \<and> (\<exists>f'. file_of_proc_fd s p fd' = Some f') \<and> cfd2sfd s p fd' = Some sfd")
+apply simp
+apply (case_tac "xa = sfd")
+apply (erule exE|erule conjE)+
+apply (rule_tac x = fd' in exI, simp add:cfd2sfd_closefd)
+apply (erule exE|erule conjE)+
+apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+apply (rule notI, simp)
+apply (simp, (erule exE|erule conjE)+)
+apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+apply (rule notI, simp)
+apply (erule exE|erule conjE)+
+apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+apply (rule notI, simp)
+apply (simp add:cpfd2sfds_def proc_file_fds_def)
+apply (erule exE|erule conjE)+
+apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+done
+
+lemmas cpfd2sfds_simps = cpfd2sfds_open cpfd2sfds_execve cpfd2sfds_clone cpfd2sfds_kill cpfd2sfds_exit
+ cpfd2sfds_closefd cpfd2sfds_other
+
+(********* ch2sshm simpset ********)
+
+lemma ch2sshm_createshm:
+ "valid (CreateShM p h # s)
+ \<Longrightarrow> ch2sshm (CreateShM p h # s) = (ch2sshm s) (h :=
+ (case (sectxt_of_obj (CreateShM p h # s) (O_shm h)) of
+ Some sec \<Rightarrow>
+ Some (if (\<not> deleted (O_shm h) s \<and> h \<in> init_shms) then Init h else Created, sec)
+ | _ \<Rightarrow> None))"
+apply (frule vd_cons, frule vt_grant_os)
+apply (auto simp:ch2sshm_def sectxt_of_obj_simps dest!:current_has_sec' split:option.splits if_splits)
+done
+
+lemma ch2sshm_other:
+ "\<lbrakk>valid (e # s);
+ \<forall> p h. e \<noteq> CreateShM p h;
+ h' \<in> current_shms (e # s)\<rbrakk> \<Longrightarrow> ch2sshm (e # s) h' = ch2sshm s h'"
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac e)
+apply (auto simp:ch2sshm_def sectxt_of_obj_simps dest!:current_has_sec' split:option.splits if_splits)
+done
+
+lemmas ch2sshm_simps = ch2sshm_createshm ch2sshm_other
+
+lemma current_shm_has_sh:
+ "\<lbrakk>h \<in> current_shms s; valid s\<rbrakk> \<Longrightarrow> \<exists> sh. ch2sshm s h = Some sh"
+by (auto simp:ch2sshm_def split:option.splits dest!:current_has_sec')
+
+lemma current_shm_has_sh':
+ "\<lbrakk>ch2sshm s h = None; valid s\<rbrakk> \<Longrightarrow> h \<notin> current_shms s"
+by (auto dest:current_shm_has_sh)
+
+(********** cph2spshs simpset **********)
+
+lemma cph2spshs_attach:
+ "valid (Attach p h flag # s) \<Longrightarrow>
+ cph2spshs (Attach p h flag # s) = (cph2spshs s) (p :=
+ (case (ch2sshm s h) of
+ Some sh \<Rightarrow> cph2spshs s p \<union> {(sh, flag)}
+ | _ \<Rightarrow> cph2spshs s p) )"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+using ch2sshm_other[where e = "Attach p h flag" and s = s]
+apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+dest!:current_shm_has_sh' dest: procs_of_shm_prop1 simp:cph2spshs_def)
+apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
+apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
+apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
+apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
+apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
+apply (case_tac "ha = h", simp, frule procs_of_shm_prop1, simp)
+apply (rule_tac x = ha in exI, simp)
+apply (rule_tac x = ha in exI, simp, drule procs_of_shm_prop1, simp, simp)
+apply (rule_tac x = ha in exI, simp)
+apply (frule procs_of_shm_prop1, simp, simp)
+apply (rule impI, simp)
+done
+
+lemma cph2spshs_detach: "valid (Detach p h # s) \<Longrightarrow>
+ cph2spshs (Detach p h # s) = (cph2spshs s) (p :=
+ (case (ch2sshm s h, flag_of_proc_shm s p h) of
+ (Some sh, Some flag) \<Rightarrow> if (\<exists> h'. h' \<noteq> h \<and> (p,flag) \<in> procs_of_shm s h' \<and> ch2sshm s h' = Some sh)
+ then cph2spshs s p else cph2spshs s p - {(sh, flag)}
+ | _ \<Rightarrow> cph2spshs s p) )"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+apply (case_tac "x = p") defer
+using ch2sshm_other[where e = "Detach p h" and s = s]
+apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+dest!:current_shm_has_sh' procs_of_shm_prop4' dest:procs_of_shm_prop1 procs_of_shm_prop3 simp:cph2spshs_def) [1]
+apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp, simp)
+apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp, simp)
+apply (rule impI, simp)
+
+apply (clarsimp)
+apply (frule current_shm_has_sh, simp, erule exE)
+apply (frule procs_of_shm_prop4, simp, simp)
+apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
+
+apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
+apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
+apply (case_tac "ha = h", simp)
+apply (rule_tac x = sha in exI, rule_tac x = flaga in exI, rule_tac x = ha in exI, simp)
+using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+
+apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
+apply (case_tac "ha = h", simp)
+apply (rule_tac x = h' in exI, simp)
+apply (frule_tac flag = flag and flag' = flaga in procs_of_shm_prop3, simp+)
+apply (frule_tac flag = flaga in procs_of_shm_prop4, simp+)
+using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+apply (frule_tac h = h' in procs_of_shm_prop1, simp, simp)
+apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+
+apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
+apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
+apply (erule CollectE | erule exE| erule conjE| rule DiffI |rule CollectI)+
+apply (simp split:if_splits)
+apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+apply (rule notI, simp split:if_splits)
+apply (erule_tac x = ha in allE, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+apply (erule CollectE | erule exE| erule conjE| erule DiffE |rule CollectI)+
+apply (simp split:if_splits)
+apply (erule_tac x = ha in allE, simp, rule_tac x = ha in exI, simp)
+apply (case_tac "ha = h", simp add:procs_of_shm_prop3, frule_tac h = ha in procs_of_shm_prop1, simp+)
+using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+done
+
+lemma cph2spshs_deleteshm:
+ "valid (DeleteShM p h # s) \<Longrightarrow>
+ cph2spshs (DeleteShM p h # s) = (\<lambda> p'.
+ (case (ch2sshm s h, flag_of_proc_shm s p' h) of
+ (Some sh, Some flag) \<Rightarrow> if (\<exists> h'. h' \<noteq> h \<and> (p', flag) \<in> procs_of_shm s h' \<and> ch2sshm s h' = Some sh)
+ then cph2spshs s p' else cph2spshs s p' - {(sh, flag)}
+ | _ \<Rightarrow> cph2spshs s p') )"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+
+apply (clarsimp)
+apply (frule current_shm_has_sh, simp, erule exE, simp, simp split:option.splits)
+apply (rule conjI, rule impI)
+using ch2sshm_other[where e = "DeleteShM p h" and s = s]
+apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+dest!:current_shm_has_sh' procs_of_shm_prop4' dest:procs_of_shm_prop1 procs_of_shm_prop3 simp:cph2spshs_def) [1]
+apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp+)
+apply (rule_tac x = ha in exI, simp)
+apply (case_tac "ha = h", simp+, frule_tac h = ha in procs_of_shm_prop1, simp+)
+
+apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
+
+apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
+apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
+apply (case_tac "ha = h", simp)
+apply (rule_tac x = sha in exI, rule_tac x = flag in exI, rule_tac x = ha in exI, simp)
+using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+
+apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
+apply (case_tac "ha = h", simp)
+apply (rule_tac x = h' in exI, simp)
+apply (frule_tac flag = flag in procs_of_shm_prop4, simp+)
+using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+apply (frule_tac h = h' in procs_of_shm_prop1, simp, simp)
+apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+
+apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
+apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
+apply (erule CollectE | erule exE| erule conjE| rule DiffI |rule CollectI)+
+apply (simp split:if_splits)
+apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+apply (rule notI, simp split:if_splits)
+apply (erule_tac x = ha in allE, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+apply (erule CollectE | erule exE| erule conjE| erule DiffE |rule CollectI)+
+apply (simp split:if_splits)
+apply (erule_tac x = ha in allE, simp, rule_tac x = ha in exI, simp)
+apply (case_tac "ha = h", simp add:procs_of_shm_prop4)
+apply (frule_tac h = ha in procs_of_shm_prop1, simp+)
+using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+done
+
+lemma flag_of_proc_shm_prop1:
+ "\<lbrakk>flag_of_proc_shm s p h = Some flag; valid s\<rbrakk> \<Longrightarrow> (p, flag) \<in> procs_of_shm s h"
+apply (induct s arbitrary:p)
+apply (simp add:init_shmflag_has_proc)
+apply (frule vt_grant_os, frule vd_cons, case_tac a, auto split:if_splits option.splits)
+done
+
+lemma cph2spshs_clone:
+ "valid (Clone p p' fds shms # s) \<Longrightarrow>
+ cph2spshs (Clone p p' fds shms # s) = (cph2spshs s) (p' :=
+ {(sh, flag) | h sh flag. h \<in> shms \<and> ch2sshm s h = Some sh \<and> flag_of_proc_shm s p h = Some flag} )"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+using ch2sshm_other[where e = "Clone p p' fds shms" and s = s]
+apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
+apply (erule_tac x = h in allE, frule procs_of_shm_prop1, simp, simp add:procs_of_shm_prop4)
+apply (rule_tac x = h in exI, simp, frule flag_of_proc_shm_prop1, simp+, simp add:procs_of_shm_prop1)
+apply (rule_tac x = h in exI, simp, frule procs_of_shm_prop1, simp+)+
+done
+
+lemma cph2spshs_execve:
+ "valid (Execve p f fds # s) \<Longrightarrow>
+ cph2spshs (Execve p f fds # s) = (cph2spshs s) (p := {})"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+using ch2sshm_other[where e = "Execve p f fds" and s = s]
+apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
+apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+
+done
+
+lemma cph2spshs_kill:
+ "valid (Kill p p' # s) \<Longrightarrow> cph2spshs (Kill p p' # s) = (cph2spshs s) (p' := {})"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+using ch2sshm_other[where e = "Kill p p'" and s = s]
+apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
+apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+
+done
+
+lemma cph2spshs_exit:
+ "valid (Exit p # s) \<Longrightarrow> cph2spshs (Exit p # s) = (cph2spshs s) (p := {})"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+using ch2sshm_other[where e = "Exit p" and s = s]
+apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
+apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+
+done
+
+lemma cph2spshs_createshm:
+ "valid (CreateShM p h # s) \<Longrightarrow> cph2spshs (CreateShM p h # s) = cph2spshs s"
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+apply (auto simp:cph2spshs_def)
+using ch2sshm_createshm
+apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
+apply (rule_tac x = ha in exI, auto simp:procs_of_shm_prop1)
+done
+
+lemma cph2spshs_other:
+ "\<lbrakk>valid (e # s);
+ \<forall> p h flag. e \<noteq> Attach p h flag;
+ \<forall> p h. e \<noteq> Detach p h;
+ \<forall> p h. e \<noteq> DeleteShM p h;
+ \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms;
+ \<forall> p f fds. e \<noteq> Execve p f fds;
+ \<forall> p p'. e \<noteq> Kill p p';
+ \<forall> p. e \<noteq> Exit p\<rbrakk> \<Longrightarrow> cph2spshs (e # s) = cph2spshs s"
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+unfolding cph2spshs_def
+apply (rule set_eqI, rule iffI)
+apply (erule CollectE | erule conjE| erule exE| rule conjI| rule impI| rule CollectI)+
+apply (frule procs_of_shm_prop1, simp, rule_tac x= sh in exI, rule_tac x = flag in exI, rule_tac x = h in exI)
+apply (case_tac e)
+using ch2sshm_other[where e = e and s = s]
+apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3
+ simp:ch2sshm_createshm)
+apply (rule_tac x = h in exI, case_tac e)
+using ch2sshm_other[where e = e and s = s]
+apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3
+ simp:ch2sshm_createshm)
+done
+
+lemmas cph2spshs_simps = cph2spshs_attach cph2spshs_detach cph2spshs_deleteshm cph2spshs_clone cph2spshs_execve
+ cph2spshs_exit cph2spshs_kill cph2spshs_other
+
+(******** cp2sproc simpset *********)
+
+lemma cp2sproc_clone:
+ "valid (Clone p p' fds shms # s) \<Longrightarrow> cp2sproc (Clone p p' fds shms # s) = (cp2sproc s) (p' :=
+ case (sectxt_of_obj s (O_proc p)) of
+ Some sec \<Rightarrow> Some (Created, sec,
+ {sfd. \<exists> fd \<in> fds. \<exists> f. file_of_proc_fd s p fd = Some f \<and> cfd2sfd s p fd = Some sfd},
+ {(sh, flag) | h sh flag. h \<in> shms \<and> ch2sshm s h = Some sh \<and> flag_of_proc_shm s p h = Some flag})
+ | _ \<Rightarrow> None)"
+apply (frule cph2spshs_clone, frule cpfd2sfds_clone)
+apply (frule vd_cons, frule vt_grant_os, simp only:os_grant.simps)
+apply ((erule exE| erule conjE)+, frule not_init_intro_proc, simp, rule ext, case_tac "x = p'", simp)
+apply (auto simp:cp2sproc_def sectxt_of_obj_simps dest!:current_has_sec'
+ dest:current_proc_has_sec split:option.splits if_splits)
+done
+
+lemma cp2sproc_other:
+ "\<lbrakk>valid (e # s);
+ \<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p h flag. e \<noteq> Attach p h flag;
+ \<forall> p h. e \<noteq> Detach p h;
+ \<forall> p h. e \<noteq> DeleteShM p h;
+ \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms;
+ \<forall> p f fds. e \<noteq> Execve p f fds;
+ \<forall> p p'. e \<noteq> Kill p p';
+ \<forall> p. e \<noteq> Exit p\<rbrakk> \<Longrightarrow> cp2sproc (e # s) = cp2sproc s"
+apply (frule cph2spshs_other, simp+, frule cpfd2sfds_other, simp+)
+apply (frule vt_grant_os, frule vd_cons, rule ext, case_tac e, simp_all)
+apply (auto simp:cp2sproc_def sectxt_of_obj_simps dest!:current_has_sec'
+ dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
+done
+
+lemma cp2sproc_open:
+ "valid (Open p f flags fd opt # s) \<Longrightarrow>
+ cp2sproc (Open p f flags fd opt # s) = (cp2sproc s) (p :=
+ case (sectxt_of_obj s (O_proc p), sectxt_of_obj (Open p f flags fd opt # s) (O_fd p fd),
+ cf2sfile (Open p f flags fd opt # s) f) of
+ (Some sec, Some fdsec, Some sf) \<Rightarrow> Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs)
+ then Init p else Created, sec,
+ (cpfd2sfds s p) \<union> {(fdsec, flags, sf)}, cph2spshs s p)
+ | _ \<Rightarrow> None)"
+apply (frule cph2spshs_other, simp, simp, simp, simp, simp, simp, simp)
+apply (frule cpfd2sfds_open, frule vt_grant_os, frule vd_cons, rule ext)
+apply (case_tac "x = p")
+apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
+done
+
+lemma cp2sproc_closefd:
+ "valid (CloseFd p fd # s) \<Longrightarrow>
+ cp2sproc (CloseFd p fd # s) = (cp2sproc s) (p :=
+ if (fd \<in> proc_file_fds s p)
+ then (case (cfd2sfd s p fd) of
+ Some sfd \<Rightarrow> (if (\<exists> fd' f'. fd' \<noteq> fd \<and> file_of_proc_fd s p fd' = Some f' \<and> cfd2sfd s p fd' = Some sfd)
+ then cp2sproc s p
+ else (case (sectxt_of_obj s (O_proc p)) of
+ Some sec \<Rightarrow> Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs)
+ then Init p else Created,
+ sec, cpfd2sfds s p - {sfd}, cph2spshs s p)
+ | _ \<Rightarrow> None))
+ | _ \<Rightarrow> cp2sproc s p)
+ else cp2sproc s p)"
+apply (frule cpfd2sfds_closefd, frule cph2spshs_other, simp, simp, simp, simp, simp, simp, simp)
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+apply (case_tac "x = p")
+apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
+done
+
+lemma cp2sproc_attach:
+ "valid (Attach p h flag # s) \<Longrightarrow>
+ cp2sproc (Attach p h flag # s) = (cp2sproc s) (p :=
+ (case (ch2sshm s h) of
+ Some sh \<Rightarrow> (case (sectxt_of_obj s (O_proc p)) of
+ Some sec \<Rightarrow> Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs)
+ then Init p else Created,
+ sec, cpfd2sfds s p, cph2spshs s p \<union> {(sh, flag)})
+ | _ \<Rightarrow> None)
+ | _ \<Rightarrow> cp2sproc s p) )"
+apply (frule cph2spshs_attach, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp)
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+apply (case_tac "x = p")
+apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
+done
+
+lemma cp2sproc_detach:
+ "valid (Detach p h # s) \<Longrightarrow>
+ cp2sproc (Detach p h # s) = (cp2sproc s) (p :=
+ (case (ch2sshm s h, flag_of_proc_shm s p h) of
+ (Some sh, Some flag) \<Rightarrow> if (\<exists> h'. h' \<noteq> h \<and> (p,flag) \<in> procs_of_shm s h' \<and> ch2sshm s h' = Some sh)
+ then cp2sproc s p
+ else (case (sectxt_of_obj s (O_proc p)) of
+ Some sec \<Rightarrow> Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs)
+ then Init p else Created, sec,
+ cpfd2sfds s p, cph2spshs s p - {(sh, flag)})
+ | None \<Rightarrow> None)
+ | _ \<Rightarrow> cp2sproc s p) )"
+apply (frule cph2spshs_detach, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp)
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+apply (case_tac "x = p")
+apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
+done
+
+lemma cp2sproc_deleteshm:
+ "valid (DeleteShM p h # s) \<Longrightarrow>
+ cp2sproc (DeleteShM p h # s) = (\<lambda> p'.
+ (case (ch2sshm s h, flag_of_proc_shm s p' h) of
+ (Some sh, Some flag) \<Rightarrow> if (\<exists> h'. h' \<noteq> h \<and> (p', flag) \<in> procs_of_shm s h' \<and> ch2sshm s h' = Some sh)
+ then cp2sproc s p'
+ else (case (sectxt_of_obj s (O_proc p')) of
+ Some sec \<Rightarrow> Some (if (\<not> deleted (O_proc p') s \<and> p' \<in> init_procs)
+ then Init p' else Created, sec,
+ cpfd2sfds s p', cph2spshs s p' - {(sh, flag)})
+ | None \<Rightarrow> None)
+ | _ \<Rightarrow> cp2sproc s p') )"
+apply (frule cph2spshs_deleteshm, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp)
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
+done
+
+lemma cp2sproc_execve:
+ "valid (Execve p f fds # s) \<Longrightarrow>
+ cp2sproc (Execve p f fds # s) = (cp2sproc s) (p :=
+ (case (sectxt_of_obj (Execve p f fds # s) (O_proc p)) of
+ Some sec \<Rightarrow> Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs) then Init p else Created, sec,
+ {sfd. \<exists> fd \<in> fds. \<exists> f. file_of_proc_fd s p fd = Some f \<and> cfd2sfd s p fd = Some sfd}, {})
+ | _ \<Rightarrow> None) )"
+apply (frule cph2spshs_execve, frule cpfd2sfds_execve)
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
+done
+
+lemma cp2sproc_kill:
+ "\<lbrakk>valid (Kill p p' # s); p'' \<noteq> p'\<rbrakk> \<Longrightarrow>
+ cp2sproc (Kill p p' # s) p'' = (cp2sproc s) p''"
+apply (frule cph2spshs_kill, frule cpfd2sfds_kill)
+apply (frule vt_grant_os, frule vd_cons)
+apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
+done
+
+lemma cp2sproc_kill':
+ "\<lbrakk>valid (Kill p p' # s); p'' \<in> current_procs (Kill p p' # s)\<rbrakk> \<Longrightarrow>
+ cp2sproc (Kill p p' # s) p'' = (cp2sproc s) p''"
+by (drule_tac p'' = p'' in cp2sproc_kill, simp+)
+
+lemma cp2sproc_exit:
+ "\<lbrakk>valid (Exit p # s); p' \<noteq> p\<rbrakk> \<Longrightarrow>
+ cp2sproc (Exit p # s) p' = (cp2sproc s) p'"
+apply (frule cph2spshs_exit, frule cpfd2sfds_exit)
+apply (frule vt_grant_os, frule vd_cons)
+apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
+done
+
+lemma cp2sproc_exit':
+ "\<lbrakk>valid (Exit p # s); p' \<in> current_procs (Exit p # s)\<rbrakk> \<Longrightarrow>
+ cp2sproc (Exit p # s) p' = (cp2sproc s) p'"
+by (drule_tac p'= p' in cp2sproc_exit, simp+)
+
+lemmas cp2sproc_simps = cp2sproc_open cp2sproc_closefd cp2sproc_attach cp2sproc_detach cp2sproc_deleteshm
+ cp2sproc_clone cp2sproc_execve cp2sproc_kill cp2sproc_exit cp2sproc_other
+
+lemma current_proc_has_sp:
+ "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sp. cp2sproc s p = Some sp"
+by (auto simp:cp2sproc_def split:option.splits dest!:current_has_sec')
+
+lemma current_proc_has_sp':
+ "\<lbrakk>cp2sproc s p = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
+by (auto dest:current_proc_has_sp)
+
+
+
+(* simpset for cf2sfiles *)
+
+lemma cf2sfiles_open:
+ "\<lbrakk>valid (Open p f flag fd opt # s); f' \<in> current_files (Open p f flag fd opt # s)\<rbrakk>
+ \<Longrightarrow> cf2sfiles (Open p f flag fd opt # s) f' = (
+ if (f' = f \<and> opt \<noteq> None)
+ then (case cf2sfile (Open p f flag fd opt # s) f of
+ Some sf \<Rightarrow> {sf}
+ | _ \<Rightarrow> {} )
+ else cf2sfiles s f')"
+apply (frule vt_grant_os, frule vd_cons)
+apply (auto simp:cf2sfiles_def cf2sfile_open_none cf2sfile_simps same_inode_files_open
+ split:if_splits option.splits dest!:current_file_has_sfile' dest:cf2sfile_open)
+apply (rule_tac x = "f'a" in bexI, drule same_inode_files_prop1, simp add:cf2sfile_open_some1, simp)+
+done
+
+lemma cf2sfiles_other:
+ "\<lbrakk>valid (e # s);
+ \<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p f. e \<noteq> UnLink p f;
+ \<forall> p f f'. e \<noteq> LinkHard p f f'\<rbrakk> \<Longrightarrow> cf2sfiles (e # s) = cf2sfiles s"
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI)
+apply (drule Set.CollectD, erule bexE, rule CollectI)
+apply (rule_tac x = f' in bexI, case_tac e)
+apply (auto simp:cf2sfiles_def cf2sfile_simps same_inode_files_simps current_files_simps
+ split:if_splits option.splits dest!:current_file_has_sfile' dest:same_inode_files_prop1 cf2sfile_other')
+apply (drule_tac f' = f' in cf2sfile_rmdir)
+apply (simp add:current_files_simps same_inode_files_prop1 same_inode_files_prop3 dir_is_empty_def)+
+
+apply (rule_tac x = f' in bexI, case_tac e)
+apply (auto simp:cf2sfiles_def cf2sfile_simps same_inode_files_simps current_files_simps
+ split:if_splits option.splits dest!:current_file_has_sfile' dest:same_inode_files_prop1 cf2sfile_other')
+apply (drule_tac f' = f' in cf2sfile_rmdir)
+apply (simp add:current_files_simps same_inode_files_prop1 same_inode_files_prop3 dir_is_empty_def)+
+done
+
+lemma cf2sfile_linkhard1':
+ "\<lbrakk>valid (LinkHard p oldf f # s); f' \<in> same_inode_files s f''\<rbrakk>
+ \<Longrightarrow> cf2sfile (LinkHard p oldf f# s) f' = cf2sfile s f'"
+apply (drule same_inode_files_prop1)
+by (simp add:cf2sfile_linkhard1)
+
+lemma cf2sfiles_linkhard:
+ "valid (LinkHard p oldf f # s) \<Longrightarrow> cf2sfiles (LinkHard p oldf f # s) = (\<lambda> f'.
+ if (f' = f \<or> f' \<in> same_inode_files s oldf)
+ then (case (cf2sfile (LinkHard p oldf f # s) f) of
+ Some sf \<Rightarrow> cf2sfiles s oldf \<union> {sf}
+ | _ \<Rightarrow> {})
+ else cf2sfiles s f')"
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+apply (auto simp:cf2sfiles_def cf2sfile_linkhard1' same_inode_files_linkhard current_files_linkhard
+ split:if_splits option.splits dest!:current_file_has_sfile' current_has_sec' dest:same_inode_files_prop1)
+done
+
+lemma cf2sfile_unlink':
+ "\<lbrakk>valid (UnLink p f # s); f' \<in> same_inode_files (UnLink p f # s) f''\<rbrakk>
+ \<Longrightarrow> cf2sfile (UnLink p f # s) f' = cf2sfile s f'"
+apply (drule same_inode_files_prop1)
+by (simp add:cf2sfile_unlink)
+
+lemma cf2sfiles_unlink:
+ "\<lbrakk>valid (UnLink p f # s); f' \<in> current_files (UnLink p f # s)\<rbrakk> \<Longrightarrow> cf2sfiles (UnLink p f # s) f' = (
+ if (f' \<in> same_inode_files s f \<and> proc_fd_of_file s f = {} \<and>
+ (\<forall> f'' \<in> same_inode_files s f. f'' \<noteq> f \<longrightarrow> cf2sfile s f'' \<noteq> cf2sfile s f)) then
+ (case (cf2sfile s f) of
+ Some sf \<Rightarrow> cf2sfiles s f' - {sf}
+ | _ \<Rightarrow> {})
+ else cf2sfiles s f')"
+apply (frule vt_grant_os, frule vd_cons, simp add:current_files_simps split:if_splits)
+apply (rule conjI, clarify, frule is_file_has_sfile', simp, erule exE, simp)
+apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink)
+apply (simp add:current_files_unlink, simp, erule conjE)
+apply (erule_tac x = f'a in ballE, frule_tac f' = "f'a" in cf2sfile_unlink)
+apply (simp add:current_files_unlink same_inode_files_prop1, simp)
+apply (rule_tac x = f'a in bexI, simp, simp)
+apply (drule_tac f = f and f' = f' and f'' = f'a in same_inode_files_prop4, simp+)
+apply (erule conjE|erule exE|erule bexE)+
+apply (case_tac "f'a = f", simp)
+apply (frule_tac f' = f' in same_inode_files_unlink, simp add:current_files_unlink)
+apply (frule_tac f' = f'a in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1)
+apply (rule_tac x = f'a in bexI, simp, simp)
+
+apply (rule impI)+
+apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink)
+apply (simp add:current_files_unlink, simp, (erule conjE)+)
+apply (rule_tac x = f'a in bexI, frule_tac f' = "f'a" in cf2sfile_unlink)
+apply (simp add:current_files_unlink same_inode_files_prop1, simp, simp)
+apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_unlink)
+apply (simp add:current_files_unlink, simp)
+apply (case_tac "f'a = f", simp)
+apply (frule_tac f = f' and f' = f in same_inode_files_prop5, simp)
+apply (erule bexE, erule conjE)
+apply (rule_tac x = f'' in bexI)
+apply (drule_tac f' = f'' in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1)
+apply (simp, simp, erule same_inode_files_prop4, simp)
+apply (rule_tac x = f'a in bexI)
+apply (drule_tac f' = f'a in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1)
+apply (simp, simp)
+
+
+apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink)
+apply (simp add:current_files_unlink, simp)
+apply (rule_tac x = f'a in bexI)
+apply (frule_tac f' = f'a in cf2sfile_unlink)
+apply (simp add:same_inode_files_prop1 current_files_unlink, simp, simp)
+apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_unlink)
+apply (simp add:current_files_unlink, simp)
+apply (rule_tac x = f'a in bexI)
+apply (frule_tac f' = f'a in cf2sfile_unlink)
+apply (simp add:same_inode_files_prop1 current_files_unlink, simp, simp)
+done
+
+lemma cf2sfiles_closefd:
+ "\<lbrakk>valid (CloseFd p fd # s); f' \<in> current_files (CloseFd p fd # s)\<rbrakk> \<Longrightarrow> cf2sfiles (CloseFd p fd # s) f' = (
+ case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> if (f' \<in> same_inode_files s f \<and> proc_fd_of_file s f = {(p, fd)} \<and> f \<in> files_hung_by_del s \<and>
+ (\<forall> f'' \<in> same_inode_files s f. f'' \<noteq> f \<longrightarrow> cf2sfile s f'' \<noteq> cf2sfile s f))
+ then (case (cf2sfile s f) of
+ Some sf \<Rightarrow> cf2sfiles s f' - {sf}
+ | _ \<Rightarrow> {})
+ else cf2sfiles s f'
+ | _ \<Rightarrow> cf2sfiles s f')"
+
+apply (frule vt_grant_os, frule vd_cons, case_tac "file_of_proc_fd s p fd")
+apply (simp_all add:current_files_simps split:if_splits)
+
+apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+apply (simp add:current_files_closefd, simp)
+apply (rule_tac x = f'a in bexI)
+apply (frule_tac f = f'a in cf2sfile_closefd)
+apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+apply (simp add:current_files_closefd, simp)
+apply (rule_tac x = f'a in bexI)
+apply (frule_tac f = f'a in cf2sfile_closefd)
+apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+
+apply (rule conjI, clarify, frule file_of_pfd_is_file, simp)
+apply (frule is_file_has_sfile', simp, erule exE, simp)
+apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+apply (simp add:current_files_closefd, simp, erule conjE)
+apply (erule_tac x = f'a in ballE, frule_tac f = "f'a" in cf2sfile_closefd)
+apply (simp add:current_files_closefd same_inode_files_prop1, simp)
+apply (rule_tac x = f'a in bexI, simp, simp)
+apply (drule_tac f = a and f' = f' and f'' = f'a in same_inode_files_prop4, simp+)
+apply (erule conjE|erule exE|erule bexE)+
+apply (case_tac "f'a = a", simp)
+apply (frule_tac f' = f' in same_inode_files_closefd, simp add:current_files_closefd, simp)
+apply (frule_tac f = f'a in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1)
+apply (rule_tac x = f'a in bexI, simp, simp)
+
+apply (rule impI)+
+apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+apply (simp add:current_files_closefd, simp, (erule conjE)+)
+apply (rule_tac x = f'a in bexI, frule_tac f = f'a in cf2sfile_closefd)
+apply (simp add:current_files_closefd same_inode_files_prop1, simp, simp)
+apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+apply (simp add:current_files_closefd, simp)
+apply (case_tac "f'a = a", simp)
+apply (frule_tac f = f' and f' = a in same_inode_files_prop5, simp)
+apply (erule bexE, erule conjE)
+apply (rule_tac x = f'' in bexI)
+apply (drule_tac f = f'' in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1)
+apply (simp, simp, erule same_inode_files_prop4, simp)
+apply (rule_tac x = f'a in bexI)
+apply (drule_tac f = f'a in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1)
+apply (simp, simp)
+
+apply (rule conjI, clarify)
+
+apply (rule impI)
+apply (case_tac "a \<in> files_hung_by_del s", simp_all)
+apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+apply (simp add:current_files_closefd, simp)
+apply (rule_tac x = f'a in bexI)
+apply (frule_tac f = f'a in cf2sfile_closefd)
+apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+apply (simp add:current_files_closefd, simp)
+apply (rule_tac x = f'a in bexI)
+apply (frule_tac f = f'a in cf2sfile_closefd)
+apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+apply (simp add:current_files_closefd, simp)
+apply (rule_tac x = f'a in bexI)
+apply (frule_tac f = f'a in cf2sfile_closefd)
+apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+apply (simp add:current_files_closefd, simp)
+apply (rule_tac x = f'a in bexI)
+apply (frule_tac f = f'a in cf2sfile_closefd)
+apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+done
+
+lemmas cf2sfiles_simps = cf2sfiles_open cf2sfiles_linkhard cf2sfiles_other
+ cf2sfiles_unlink cf2sfiles_closefd
+
+
+(* simpset for co2sobj *)
+
+lemma co2sobj_execve:
+ "\<lbrakk>valid (Execve p f fds # s); alive (Execve p f fds # s) obj\<rbrakk> \<Longrightarrow> co2sobj (Execve p f fds # s) obj = (
+ if (obj = O_proc p)
+ then (case (cp2sproc (Execve p f fds # s) p) of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s \<or> O_file f \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s obj )"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps cf2sfiles_other ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (case_tac "cp2sproc (Execve p f fds # s) p")
+apply (drule current_proc_has_sp', simp, simp)
+apply (simp (no_asm_simp) add:cp2sproc_execve tainted_eq_Tainted split:option.splits)
+apply (simp add:is_dir_simps, frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp)
+apply (frule_tac ff = list in cf2sfile_other', simp_all)
+apply (simp add:is_dir_in_current)
+done
+
+lemma co2sobj_execve':
+ "\<lbrakk>valid (Execve p f fds # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (Execve p f fds # s) obj = (
+ if (obj = O_proc p)
+ then (case (cp2sproc (Execve p f fds # s) p) of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s \<or> O_file f \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s obj )"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps cf2sfiles_other ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (case_tac "cp2sproc (Execve p f fds # s) p")
+apply (drule current_proc_has_sp', simp, simp)
+apply (simp (no_asm_simp) add:cp2sproc_execve tainted_eq_Tainted split:option.splits)
+apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp)
+apply (frule_tac ff = list in cf2sfile_other', simp_all)
+apply (simp add:is_dir_in_current)
+done
+
+lemma co2sobj_clone':
+ "\<lbrakk>valid (Clone p p' fds shms # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (Clone p p' fds shms # s) obj = (
+ if (obj = O_proc p')
+ then (case (cp2sproc (Clone p p' fds shms # s) p') of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s obj )"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps cf2sfiles_other ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (case_tac "cp2sproc (Clone p p' fds shms # s) p'")
+apply (drule current_proc_has_sp', simp, simp)
+apply ((erule conjE)+, frule_tac p = p in current_proc_has_sec, simp, erule exE, simp)
+apply (rule conjI, rule impI, simp)
+apply (simp (no_asm_simp) add:cp2sproc_clone tainted_eq_Tainted split:option.splits)
+
+apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp)
+apply (frule_tac ff = list in cf2sfile_other', simp_all)
+apply (simp add:is_dir_in_current)
+done
+
+lemma co2sobj_clone:
+ "\<lbrakk>valid (Clone p p' fds shms # s); alive (Clone p p' fds shms # s) obj\<rbrakk>
+ \<Longrightarrow> co2sobj (Clone p p' fds shms # s) obj = (
+ if (obj = O_proc p')
+ then (case (cp2sproc (Clone p p' fds shms # s) p') of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s obj )"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps cf2sfiles_other ch2sshm_other
+ cq2smsgq_other tainted_eq_Tainted)
+apply (rule conjI, rule impI, simp)
+apply (case_tac "cp2sproc (Clone p p' fds shms # s) p'")
+apply (drule current_proc_has_sp', simp, simp)
+apply ((erule conjE)+, frule_tac p = p in current_proc_has_sec, simp, erule exE, simp)
+apply (simp add:tainted_eq_Tainted, rule impI, rule notI)
+apply (drule Tainted_in_current, simp+)
+apply (rule impI, simp)
+apply (drule current_proc_has_sp, simp, (erule exE|erule conjE)+)
+apply (simp (no_asm_simp) add:cp2sproc_clone tainted_eq_Tainted Tainted_in_current split:option.splits)
+
+apply (simp add:is_dir_simps, frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp)
+apply (frule_tac ff = list in cf2sfile_other', simp_all)
+apply (simp add:is_dir_in_current)
+done
+
+lemma co2sobj_ptrace:
+ "\<lbrakk>valid (Ptrace p p' # s); alive s obj\<rbrakk>\<Longrightarrow> co2sobj (Ptrace p p' # s) obj = (
+ case obj of
+ O_proc p'' \<Rightarrow> if (info_flow_shm s p' p'')
+ then (case (cp2sproc s p'') of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p'' \<in> Tainted s \<or> O_proc p \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else if (info_flow_shm s p p'')
+ then (case (cp2sproc s p'') of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p'' \<in> Tainted s \<or> O_proc p' \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s (O_proc p'')
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other cf2sfiles_other tainted_eq_Tainted)
+
+apply (auto simp:cp2sproc_other tainted_eq_Tainted split:option.splits
+ dest!:current_proc_has_sec' current_proc_has_sp' intro:info_flow_shm_Tainted)[1]
+
+apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp)
+apply (frule_tac ff = list in cf2sfile_other', simp_all)
+apply (simp add:is_dir_in_current)
+done
+
+lemma co2sobj_open:
+ "\<lbrakk>valid (Open p f flag fd opt # s); alive (Open p f flag fd opt # s) obj\<rbrakk>
+ \<Longrightarrow> co2sobj (Open p f flag fd opt # s) obj = (case obj of
+ O_file f' \<Rightarrow> if (f' = f \<and> opt \<noteq> None)
+ then (case (cf2sfile (Open p f flag fd opt # s) f) of
+ Some sf \<Rightarrow> Some (S_file {sf} (O_proc p \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s (O_file f')
+ | O_proc p' \<Rightarrow> if (p' = p)
+ then (case (cp2sproc (Open p f flag fd opt # s) p) of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s (O_proc p')
+ | _ \<Rightarrow> co2sobj s obj )"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits
+ dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1]
+
+apply (simp split:if_splits t_object.splits)
+apply (rule conjI, rule impI, erule conjE, erule exE, simp, (erule exE|erule conjE)+)
+apply (case_tac "cf2sfile (Open p f flag fd (Some y) # s) f")
+apply (drule current_file_has_sfile', simp, simp add:current_files_simps, simp)
+apply (frule_tac f' = f in cf2sfiles_open, simp add:current_files_simps)
+apply (simp add:tainted_eq_Tainted)
+apply (rule impI, rule notI, drule Tainted_in_current, simp, simp add:is_file_in_current)
+apply (rule impI, simp add:tainted_eq_Tainted cf2sfiles_open is_file_in_current split:option.splits)
+
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+
+apply (frule is_dir_in_current)
+apply (frule_tac f' = list in cf2sfile_open)
+apply (simp add:current_files_simps split:option.splits)
+apply (simp split:if_splits option.splits)
+done
+
+lemma co2sobj_readfile:
+ "\<lbrakk>valid (ReadFile p fd # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (ReadFile p fd # s) obj = (
+ case obj of
+ O_proc p' \<Rightarrow> (case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> (if (info_flow_shm s p p' \<and> O_file f \<in> Tainted s)
+ then (case (cp2sproc s p') of
+ Some sp \<Rightarrow> Some (S_proc sp True)
+ | _ \<Rightarrow> None)
+ else co2sobj s obj)
+ | _ \<Rightarrow> None)
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits
+ dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1]
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps
+ dest:is_file_in_current is_dir_in_current)
+done
+
+lemma co2sobj_writefile:
+ "\<lbrakk>valid (WriteFile p fd # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (WriteFile p fd # s) obj = (
+ case obj of
+ O_file f' \<Rightarrow> (case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> (if (f \<in> same_inode_files s f')
+ then Some (S_file (cf2sfiles s f')
+ (O_file f' \<in> Tainted s \<or> O_proc p \<in> Tainted s))
+ else co2sobj s obj)
+ | _ \<Rightarrow> None)
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted
+ same_inode_files_prop6
+ dest:is_file_in_current is_dir_in_current)
+
+(* should delete has_same_inode !?!?*)
+by (auto simp:same_inode_files_def is_file_def has_same_inode_def split:if_splits option.splits)
+
+lemma co2sobj_closefd:
+ "\<lbrakk>valid (CloseFd p fd # s); alive (CloseFd p fd # s) obj\<rbrakk> \<Longrightarrow> co2sobj (CloseFd p fd # s) obj = (
+ case obj of
+ O_file f' \<Rightarrow> (case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> (if (f' \<in> same_inode_files s f \<and> proc_fd_of_file s f = {(p, fd)} \<and>
+ f \<in> files_hung_by_del s \<and> (\<forall> f'' \<in> same_inode_files s f.
+ f'' \<noteq> f \<longrightarrow> cf2sfile s f'' \<noteq> cf2sfile s f))
+ then (case (cf2sfile s f, co2sobj s (O_file f')) of
+ (Some sf, Some (S_file sfs b)) \<Rightarrow> Some (S_file (sfs - {sf}) b)
+ | _ \<Rightarrow> None)
+ else co2sobj s obj)
+ | _ \<Rightarrow> co2sobj s obj)
+ | O_proc p' \<Rightarrow> (if (p = p')
+ then (case (cp2sproc (CloseFd p fd # s) p) of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s obj)
+ | _ \<Rightarrow> co2sobj s obj) "
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits if_splits
+ dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1]
+
+apply (frule is_file_in_current)
+apply (case_tac "file_of_proc_fd s p fd")
+apply (simp add:tainted_eq_Tainted)
+apply (drule_tac f' = list in cf2sfiles_closefd, simp add:current_files_closefd, simp)
+apply (frule_tac f' = list in cf2sfiles_closefd, simp)
+apply (simp add:is_file_simps current_files_simps)
+apply (auto simp add:tainted_eq_Tainted cf2sfile_closefd split:if_splits option.splits
+ dest!:current_file_has_sfile' dest:hung_file_in_current)[1]
+
+apply (simp add:is_dir_simps, frule is_dir_in_current)
+apply (frule_tac f = list in cf2sfile_closefd)
+apply (clarsimp simp:current_files_closefd split:option.splits)
+apply (drule file_of_pfd_is_file', simp+)
+done
+
+lemma co2sobj_unlink:
+ "\<lbrakk>valid (UnLink p f # s); alive (UnLink p f # s) obj\<rbrakk> \<Longrightarrow> co2sobj (UnLink p f # s) obj = (
+ case obj of
+ O_file f' \<Rightarrow> if (f' \<in> same_inode_files s f \<and> proc_fd_of_file s f = {} \<and>
+ (\<forall> f'' \<in> same_inode_files s f. f'' \<noteq> f \<longrightarrow> cf2sfile s f'' \<noteq> cf2sfile s f))
+ then (case (cf2sfile s f, co2sobj s (O_file f')) of
+ (Some sf, Some (S_file sfs b)) \<Rightarrow> Some (S_file (sfs - {sf}) b)
+ | _ \<Rightarrow> None)
+ else co2sobj s obj
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits if_splits
+ dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1]
+apply (frule is_file_in_current)
+apply (frule_tac f' = list in cf2sfile_unlink, simp)
+apply (frule_tac f' = list in cf2sfiles_unlink, simp)
+apply (simp add:is_file_simps current_files_simps)
+apply (auto simp add:tainted_eq_Tainted is_file_in_current split:if_splits option.splits
+ dest!:current_file_has_sfile')[1]
+
+apply (simp add:is_dir_simps, frule is_dir_in_current)
+apply (frule_tac f' = list in cf2sfile_unlink)
+apply (clarsimp simp:current_files_unlink split:option.splits)
+apply (drule file_dir_conflict, simp+)
+done
+
+lemma co2sobj_rmdir:
+ "\<lbrakk>valid (Rmdir p f # s); alive (Rmdir p f # s) obj\<rbrakk> \<Longrightarrow> co2sobj (Rmdir p f # s) obj = co2sobj s obj"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits if_splits
+ dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1]
+apply (simp add:is_file_simps dir_is_empty_def)
+apply (case_tac "f = list", drule file_dir_conflict, simp, simp)
+apply (simp add:cf2sfiles_other)
+apply (auto simp:cf2sfile_simps dest:is_dir_in_current)
+done
+
+lemma co2sobj_mkdir:
+ "\<lbrakk>valid (Mkdir p f i # s); alive (Mkdir p f i # s) obj\<rbrakk> \<Longrightarrow> co2sobj (Mkdir p f i # s) obj = (
+ if (obj = O_dir f)
+ then (case (cf2sfile (Mkdir p f i # s) f) of
+ Some sf \<Rightarrow> Some (S_dir sf)
+ | _ \<Rightarrow> None)
+ else co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits if_splits
+ dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1]
+apply (frule_tac cf2sfiles_other, simp+)
+apply (frule is_dir_in_current, rule impI, simp add:current_files_mkdir)
+apply (frule current_file_has_sfile, simp, erule exE, simp)
+apply (drule cf2sfile_mkdir1, simp+)
+done
+
+
+lemma co2sobj_linkhard:
+ "\<lbrakk>valid (LinkHard p oldf f # s); alive (LinkHard p oldf f # s) obj\<rbrakk>
+ \<Longrightarrow> co2sobj (LinkHard p oldf f # s) obj = (
+ case obj of
+ O_file f' \<Rightarrow> if (f' = f \<or> f' \<in> same_inode_files s oldf)
+ then (case (cf2sfile (LinkHard p oldf f # s) f) of
+ Some sf \<Rightarrow> Some (S_file (cf2sfiles s oldf \<union> {sf}) (O_file oldf \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s obj
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits if_splits
+ dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1]
+apply (frule_tac cf2sfiles_linkhard, simp+)
+apply (frule_tac f' = f in cf2sfile_linkhard, simp add:current_files_linkhard)
+apply (auto simp:is_file_simps sectxt_of_obj_simps current_files_simps is_file_in_current same_inodes_Tainted
+ split:option.splits if_splits dest:Tainted_in_current
+ dest!:current_has_sec' current_file_has_sfile')[1]
+
+apply (frule is_dir_in_current, simp add:current_files_linkhard is_dir_simps is_dir_in_current)
+apply (frule is_dir_in_current)
+apply (frule current_file_has_sfile, simp)
+apply (drule cf2sfile_linkhard1, simp+)
+done
+
+lemma co2sobj_truncate:
+ "\<lbrakk>valid (Truncate p f len # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (Truncate p f len # s) obj = (
+ case obj of
+ O_file f' \<Rightarrow> if (f' \<in> same_inode_files s f \<and> len > 0)
+ then Some (S_file (cf2sfiles s f') (O_file f' \<in> Tainted s \<or> O_proc p \<in> Tainted s))
+ else co2sobj s obj
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted
+ same_inode_files_prop6
+ dest:is_file_in_current is_dir_in_current)
+done
+
+lemma co2sobj_kill:
+ "\<lbrakk>valid (Kill p p' # s); alive (Kill p p' # s) obj\<rbrakk> \<Longrightarrow> co2sobj (Kill p p' # s) obj = co2sobj s obj"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted
+ same_inode_files_prop6
+ dest:is_file_in_current is_dir_in_current)
+done
+
+lemma co2sobj_exit:
+ "\<lbrakk>valid (Exit p # s); alive (Exit p # s) obj\<rbrakk> \<Longrightarrow> co2sobj (Exit p # s) obj = co2sobj s obj"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted
+ same_inode_files_prop6
+ dest:is_file_in_current is_dir_in_current)
+done
+
+lemma co2sobj_createmsgq:
+ "\<lbrakk>valid (CreateMsgq p q # s); alive (CreateMsgq p q # s) obj\<rbrakk> \<Longrightarrow> co2sobj (CreateMsgq p q # s) obj = (
+ case obj of
+ O_msgq q' \<Rightarrow> if (q' = q) then (case (cq2smsgq (CreateMsgq p q # s) q) of
+ Some sq \<Rightarrow> Some (S_msgq sq)
+ | _ \<Rightarrow> None)
+ else co2sobj s obj
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted
+ same_inode_files_prop6
+ dest!:current_has_smsgq'
+ dest:is_file_in_current is_dir_in_current cq2smsg_createmsgq)
+done
+
+lemma co2sobj_sendmsg:
+ "\<lbrakk>valid (SendMsg p q m # s); alive (SendMsg p q m # s) obj\<rbrakk> \<Longrightarrow> co2sobj (SendMsg p q m # s) obj = (
+ case obj of
+ O_msgq q' \<Rightarrow> if (q' = q) then (case (cq2smsgq (SendMsg p q m # s) q) of
+ Some sq \<Rightarrow> Some (S_msgq sq)
+ | _ \<Rightarrow> None)
+ else co2sobj s obj
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted
+ same_inode_files_prop6
+ dest!:current_has_smsgq'
+ dest:is_file_in_current is_dir_in_current cq2smsg_sendmsg)
+done
+
+lemma co2sobj_recvmsg:
+ "\<lbrakk>valid (RecvMsg p q m # s); alive (RecvMsg p q m # s) obj\<rbrakk> \<Longrightarrow> co2sobj (RecvMsg p q m # s) obj = (
+ case obj of
+ O_msgq q' \<Rightarrow> if (q' = q) then (case (cq2smsgq (RecvMsg p q m # s) q) of
+ Some sq \<Rightarrow> Some (S_msgq sq)
+ | _ \<Rightarrow> None)
+ else co2sobj s obj
+ | O_proc p' \<Rightarrow> if (info_flow_shm s p p' \<and> O_msg q m \<in> Tainted s)
+ then (case (cp2sproc s p') of
+ Some sp \<Rightarrow> Some (S_proc sp True)
+ | _ \<Rightarrow> None)
+ else co2sobj s obj
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted
+ same_inode_files_prop6
+ dest!:current_has_smsgq'
+ dest:is_file_in_current is_dir_in_current cq2smsg_recvmsg)
+done
+
+lemma co2sobj_removemsgq:
+ "\<lbrakk>valid (RemoveMsgq p q # s); alive (RemoveMsgq p q # s) obj\<rbrakk>
+ \<Longrightarrow> co2sobj (RemoveMsgq p q # s) obj = co2sobj s obj"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted
+ same_inode_files_prop6
+ dest!:current_has_smsgq'
+ dest:is_file_in_current is_dir_in_current cq2smsg_removemsgq)
+done
+
+lemma co2sobj_createshm:
+ "\<lbrakk>valid (CreateShM p h # s); alive (CreateShM p h # s) obj\<rbrakk> \<Longrightarrow> co2sobj (CreateShM p h # s) obj = (
+ case obj of
+ O_shm h' \<Rightarrow> if (h' = h) then (case (ch2sshm (CreateShM p h # s) h) of
+ Some sh \<Rightarrow> Some (S_shm sh)
+ | _ \<Rightarrow> None)
+ else co2sobj s obj
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted
+ same_inode_files_prop6 ch2sshm_simps
+ dest!:current_shm_has_sh'
+ dest:is_file_in_current is_dir_in_current)
+done
+
+lemma co2sobj_detach:
+ "\<lbrakk>valid (Detach p h # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (Detach p h # s) obj = (
+ case obj of
+ O_proc p' \<Rightarrow> if (p' = p) then (case (cp2sproc (Detach p h # s) p) of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s obj
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted
+ same_inode_files_prop6 ch2sshm_simps
+ dest!:current_shm_has_sh'
+ dest:is_file_in_current is_dir_in_current)
+done
+
+lemma co2sobj_deleteshm:
+ "\<lbrakk>valid (DeleteShM p h # s); alive (DeleteShM p h # s) obj\<rbrakk> \<Longrightarrow> co2sobj (DeleteShM p h # s) obj = (
+ case obj of
+ O_proc p' \<Rightarrow> (case (cp2sproc (DeleteShM p h # s) p') of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p' \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps tainted_eq_Tainted
+ same_inode_files_prop6 ch2sshm_simps
+ dest!:current_shm_has_sh'
+ dest:is_file_in_current is_dir_in_current)
+done
+
+declare Product_Type.split_paired_Ex Product_Type.split_paired_All [simp del]
+
+lemma info_flow_shm_prop1:
+ "p \<in> current_procs s \<Longrightarrow> info_flow_shm s p p"
+by (simp add:info_flow_shm_def)
+
+lemma co2sobj_attach:
+ "\<lbrakk>valid (Attach p h flag # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (Attach p h flag # s) obj = (
+ case obj of
+ O_proc p' \<Rightarrow> if (info_flow_shm s p p')
+ then (case (cp2sproc (Attach p h flag # s) p') of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p' \<in> Tainted s \<or>
+ (\<exists> p''. O_proc p'' \<in> Tainted s \<and> (p'', SHM_RDWR) \<in> procs_of_shm s h)))
+ | _ \<Rightarrow> None)
+ else if (\<exists> p'' flag'. (p'', flag') \<in> procs_of_shm s h \<and> flag = SHM_RDWR \<and> O_proc p \<in> Tainted s \<and>
+ info_flow_shm s p'' p')
+ then (case (cp2sproc s p') of
+ Some sp \<Rightarrow> Some (S_proc sp True)
+ | _ \<Rightarrow> None)
+ else co2sobj s obj
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted)
+
+apply (rule conjI|rule impI|erule exE)+
+apply (simp split:option.splits del:split_paired_Ex)
+apply (rule impI, frule current_proc_has_sp, simp)
+apply ((erule exE)+, auto simp:tainted_eq_Tainted intro:info_flow_shm_Tainted)[1]
+apply (rule impI, simp add:tainted_eq_Tainted split:option.splits del:split_paired_Ex)
+apply (auto simp:info_flow_shm_prop1 cp2sproc_attach dest!:current_proc_has_sp')[1]
+
+apply (case_tac "cp2sproc (Attach p h flag # s) nat")
+apply (drule current_proc_has_sp', simp+)
+
+apply (rule conjI|erule exE|erule conjE|rule impI)+
+apply (simp add:tainted_eq_Tainted)
+apply (auto simp:info_flow_shm_prop1 cp2sproc_attach intro:info_flow_shm_Tainted dest!:current_proc_has_sp')[1]
+apply (auto simp:info_flow_shm_prop1 cp2sproc_attach tainted_eq_Tainted intro:info_flow_shm_Tainted dest!:current_proc_has_sp'
+split:option.splits if_splits)[1]
+
+
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps tainted_eq_Tainted
+ same_inode_files_prop6
+ dest:is_file_in_current is_dir_in_current)
+done
+
+
+lemma co2sobj_other:
+ "\<lbrakk>valid (e # s); alive (e # s) obj;
+ \<forall> p f fds. e \<noteq> Execve p f fds;
+ \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms;
+ \<forall> p p'. e \<noteq> Ptrace p p';
+ \<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ \<forall> p fd. e \<noteq> ReadFile p fd;
+ \<forall> p fd. e \<noteq> WriteFile p fd;
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p f. e \<noteq> UnLink p f;
+ \<forall> p f. e \<noteq> Rmdir p f;
+ \<forall> p f i. e \<noteq> Mkdir p f i;
+ \<forall> p f f'. e \<noteq> LinkHard p f f';
+ \<forall> p f len. e \<noteq> Truncate p f len;
+ \<forall> p q. e \<noteq> CreateMsgq p q;
+ \<forall> p q m. e \<noteq> SendMsg p q m;
+ \<forall> p q m. e \<noteq> RecvMsg p q m;
+ \<forall> p q. e \<noteq> RemoveMsgq p q;
+ \<forall> p h. e \<noteq> CreateShM p h;
+ \<forall> p h flag. e \<noteq> Attach p h flag;
+ \<forall> p h. e \<noteq> Detach p h;
+ \<forall> p h. e \<noteq> DeleteShM p h
+ \<rbrakk> \<Longrightarrow> co2sobj (e # s) obj = co2sobj s obj"
+apply (frule vt_grant, case_tac e)
+apply (auto intro:co2sobj_kill co2sobj_exit)
+done
+
+lemmas co2sobj_simps = co2sobj_execve co2sobj_clone co2sobj_ptrace co2sobj_open co2sobj_readfile
+ co2sobj_writefile co2sobj_closefd co2sobj_unlink co2sobj_rmdir co2sobj_mkdir co2sobj_linkhard
+ co2sobj_truncate co2sobj_kill co2sobj_exit co2sobj_createmsgq co2sobj_sendmsg co2sobj_recvmsg
+ co2sobj_removemsgq co2sobj_attach co2sobj_detach co2sobj_createshm co2sobj_deleteshm
+
+
+
+end
+
+(*<*)
+end
+(*>*)
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Current_files_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,878 @@
+(*<*)
+theory Current_files_prop
+imports Main Flask_type Flask My_list_prefix Init_prop Valid_prop
+begin
+(*<*)
+
+context init begin
+
+lemma current_files_ndef: "f \<notin> current_files \<tau> \<Longrightarrow> inum_of_file \<tau> f = None"
+by (simp add:current_files_def)
+
+(************** file_of_proc_fd vs proc_fd_of_file *****************)
+lemma pfdof_simp1: "file_of_proc_fd \<tau> p fd = Some f \<Longrightarrow> (p, fd) \<in> proc_fd_of_file \<tau> f"
+by (simp add:proc_fd_of_file_def)
+
+lemma pfdof_simp2: "(p, fd) \<in> proc_fd_of_file \<tau> f \<Longrightarrow> file_of_proc_fd \<tau> p fd = Some f"
+by (simp add:proc_fd_of_file_def)
+
+lemma pfdof_simp3: "proc_fd_of_file \<tau> f = {(p, fd)} \<Longrightarrow> \<forall> p' fd'. (file_of_proc_fd \<tau> p' fd' = Some f \<longrightarrow> p = p' \<and> fd = fd')"
+by (simp add:proc_fd_of_file_def, auto)
+
+lemma pfdof_simp4: "\<lbrakk>file_of_proc_fd \<tau> p' fd' = Some f; proc_fd_of_file \<tau> f = {(p, fd)}\<rbrakk> \<Longrightarrow> p' = p \<and> fd' = fd"
+by (drule pfdof_simp3, auto)
+
+end
+
+context flask begin
+
+(***************** inode number lemmas *************************)
+
+lemma iof's_im_in_cim: "inum_of_file \<tau> f = Some im \<Longrightarrow> im \<in> current_inode_nums \<tau>"
+by (auto simp add:current_inode_nums_def current_file_inums_def)
+
+lemma ios's_im_in_cim: "inum_of_socket \<tau> s = Some im \<Longrightarrow> im \<in> current_inode_nums \<tau>"
+by (case_tac s, auto simp add:current_inode_nums_def current_sock_inums_def)
+
+lemma fim_noninter_sim_aux[rule_format]:
+ "\<forall> f s. inum_of_file \<tau> f = Some im \<and> inum_of_socket \<tau> s = Some im \<and> valid \<tau> \<longrightarrow> False"
+apply (induct \<tau>)
+apply (clarsimp simp:inum_of_file.simps inum_of_socket.simps)
+apply (drule init_inum_sock_file_noninter, simp, simp)
+apply ((rule allI|rule impI|erule conjE)+)
+apply (frule vd_cons, frule vt_grant_os, simp, case_tac a)
+apply (auto simp:inum_of_file.simps inum_of_socket.simps split:if_splits option.splits
+ dest:ios's_im_in_cim iof's_im_in_cim)
+done
+
+lemma fim_noninter_sim':"\<lbrakk>inum_of_file \<tau> f = Some im; inum_of_socket \<tau> s = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+by (auto intro:fim_noninter_sim_aux)
+
+lemma fim_noninter_sim'':"\<lbrakk>inum_of_socket \<tau> s = Some im; inum_of_file \<tau> f = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+by (auto intro:fim_noninter_sim_aux)
+
+lemma fim_noninter_sim: "valid \<tau> \<Longrightarrow> (current_file_inums \<tau>) \<inter> (current_sock_inums \<tau>) = {}"
+by (auto simp:current_file_inums_def current_sock_inums_def intro:fim_noninter_sim_aux[rule_format])
+
+(******************* file inum has inode tag ************************)
+
+lemma finum_has_itag_aux[rule_format]:
+ "\<forall> f im. inum_of_file \<tau> f = Some im \<and> valid \<tau> \<longrightarrow> itag_of_inum \<tau> im \<noteq> None"
+apply (induct \<tau>)
+apply (clarsimp simp:inum_of_file.simps itag_of_inum.simps init_inum_of_file_props)
+apply (clarify, frule vt_grant_os, frule vd_cons, case_tac a)
+apply (auto simp add:inum_of_file.simps itag_of_inum.simps os_grant.simps current_files_def
+ dest:fim_noninter_sim'' split:option.splits if_splits t_socket_type.splits)
+done
+
+lemma finum_has_itag: "\<lbrakk>inum_of_file \<tau> f = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> tag. itag_of_inum \<tau> im = Some tag"
+by (auto dest:conjI[THEN finum_has_itag_aux])
+
+(*********************** file inum is file itag *************************)
+
+lemma finum_has_ftag_aux[rule_format]:
+ "\<forall> f tag. inum_of_file \<tau> f = Some im \<and> itag_of_inum \<tau> im = Some tag \<and> valid \<tau> \<longrightarrow> is_file_dir_itag tag"
+apply (induct \<tau>)
+apply (clarsimp simp:inum_of_file.simps itag_of_inum.simps init_inum_of_file_props)
+apply (clarify, frule vt_grant_os, frule vd_cons, case_tac a)
+apply (auto simp:inum_of_file.simps os_grant.simps current_files_def itag_of_inum.simps
+ split:if_splits option.splits t_socket_type.splits
+ dest:ios's_im_in_cim iof's_im_in_cim)
+done
+
+lemma finum_has_ftag:
+ "\<lbrakk>inum_of_file \<tau> f = Some im; itag_of_inum \<tau> im = Some tag; valid \<tau>\<rbrakk> \<Longrightarrow> is_file_dir_itag tag"
+by (auto intro:finum_has_ftag_aux)
+
+lemma finum_has_ftag':
+ "\<lbrakk>inum_of_file \<tau> f = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> itag_of_inum \<tau> im = Some Tag_FILE \<or> itag_of_inum \<tau> im = Some Tag_DIR"
+apply (frule finum_has_itag, simp, erule exE, drule finum_has_ftag, simp+)
+apply (case_tac tag, auto)
+done
+
+(******************* sock inum has inode tag ************************)
+
+lemma sinum_has_itag_aux[rule_format]:
+ "\<forall> s im. inum_of_socket \<tau> s = Some im \<and> valid \<tau> \<longrightarrow> itag_of_inum \<tau> im \<noteq> None"
+apply (induct \<tau>)
+apply (clarsimp simp:inum_of_socket.simps itag_of_inum.simps)
+apply (drule init_inumos_prop4, clarsimp)
+apply (clarify, frule vt_grant_os, frule vd_cons, case_tac a)
+apply (auto simp add:inum_of_socket.simps itag_of_inum.simps os_grant.simps
+ dest:fim_noninter_sim'' ios's_im_in_cim iof's_im_in_cim
+ split:option.splits if_splits t_socket_type.splits)
+done
+
+lemma sinum_has_itag: "\<lbrakk>inum_of_socket \<tau> s = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> tag. itag_of_inum \<tau> im = Some tag"
+by (auto dest:conjI[THEN sinum_has_itag_aux])
+
+(********************** socket inum is socket itag **********************)
+
+lemma sinum_has_stag_aux[rule_format]:
+ "\<forall> s tag. inum_of_socket \<tau> s = Some im \<and> itag_of_inum \<tau> im = Some tag \<and> valid \<tau> \<longrightarrow> is_sock_itag tag"
+apply (induct \<tau>)
+apply (clarsimp simp:inum_of_socket.simps itag_of_inum.simps)
+apply (drule init_inumos_prop4, clarsimp)
+apply (clarify, frule vt_grant_os, frule vd_cons, case_tac a)
+apply (auto simp add:inum_of_socket.simps itag_of_inum.simps os_grant.simps
+ dest:fim_noninter_sim'' ios's_im_in_cim iof's_im_in_cim
+ split:option.splits if_splits t_socket_type.splits)
+done
+
+lemma sinum_has_stag: "\<lbrakk>inum_of_socket \<tau> s = Some im; itag_of_inum \<tau> im = Some tag; valid \<tau>\<rbrakk> \<Longrightarrow> is_sock_itag tag"
+by (auto dest:conjI[THEN sinum_has_stag_aux])
+
+lemma sinum_has_stag':
+ "\<lbrakk>inum_of_socket \<tau> s = Some im; valid \<tau>\<rbrakk>
+ \<Longrightarrow> itag_of_inum \<tau> im = Some Tag_UDP_SOCK \<or> itag_of_inum \<tau> im = Some Tag_TCP_SOCK"
+apply (frule sinum_has_itag, simp, erule exE)
+apply (drule sinum_has_stag, simp+, case_tac tag, simp+)
+done
+
+(************************************ 4 in 1 *************************************)
+
+lemma file_leveling: "valid \<tau> \<longrightarrow> (
+ (\<forall> f. f \<in> files_hung_by_del \<tau> \<longrightarrow> inum_of_file \<tau> f \<noteq> None) \<and>
+ (\<forall> f pf im. parent f = Some pf \<and> inum_of_file \<tau> f = Some im \<longrightarrow> inum_of_file \<tau> pf \<noteq> None) \<and>
+ (\<forall> f f' im. is_file \<tau> f \<and> parent f' = Some f \<and> inum_of_file \<tau> f' = Some im \<longrightarrow> False) \<and>
+ (\<forall> f f' im. f \<in> files_hung_by_del \<tau> \<and> inum_of_file \<tau> f' = Some im \<and> parent f' = Some f \<longrightarrow> False) )"
+proof (induct \<tau>)
+ case Nil
+ show ?case
+ apply (auto simp:inum_of_file.simps files_hung_by_del.simps is_file_def itag_of_inum.simps parent_file_in_init split:option.splits t_inode_tag.splits)
+ apply (drule init_files_hung_by_del_props, simp add:init_file_has_inum)
+ apply (rule init_parent_file_has_inum, simp+)
+ apply (rule init_file_has_no_son', simp+)
+ apply (rule init_file_hung_has_no_son, simp+)
+ done
+next
+ case (Cons a \<tau>)
+ assume pre: "valid \<tau> \<longrightarrow>
+ (\<forall> f. f \<in> files_hung_by_del \<tau> \<longrightarrow> inum_of_file \<tau> f \<noteq> None) \<and>
+ (\<forall>f pf im. parent f = Some pf \<and> inum_of_file \<tau> f = Some im \<longrightarrow> inum_of_file \<tau> pf \<noteq> None) \<and>
+ (\<forall>f f' im. is_file \<tau> f \<and> parent f' = Some f \<and> inum_of_file \<tau> f' = Some im \<longrightarrow> False) \<and>
+ (\<forall>f f' im. f \<in> files_hung_by_del \<tau> \<and> inum_of_file \<tau> f' = Some im \<and> parent f' = Some f \<longrightarrow> False)"
+ show ?case
+ proof
+ assume cons:"valid (a # \<tau>)"
+ show "(\<forall>f. f \<in> files_hung_by_del (a # \<tau>) \<longrightarrow> inum_of_file (a # \<tau>) f \<noteq> None) \<and>
+ (\<forall>f pf im. parent f = Some pf \<and> inum_of_file (a # \<tau>) f = Some im \<longrightarrow> inum_of_file (a # \<tau>) pf \<noteq> None) \<and>
+ (\<forall>f f' im. is_file (a # \<tau>) f \<and> parent f' = Some f \<and> inum_of_file (a # \<tau>) f' = Some im \<longrightarrow> False) \<and>
+ (\<forall>f f' im. f \<in> files_hung_by_del (a # \<tau>) \<and> inum_of_file (a # \<tau>) f' = Some im \<and> parent f' = Some f \<longrightarrow> False)"
+ proof-
+ have vt: "valid \<tau>" using cons by (auto dest:vd_cons)
+ have os: "os_grant \<tau> a" using cons by (auto dest:vt_grant_os)
+ have fin: "\<forall>f. f \<in> files_hung_by_del \<tau> \<longrightarrow> inum_of_file \<tau> f \<noteq> None" using vt pre by auto
+ have pin: "\<forall>f pf im. parent f = Some pf \<and> inum_of_file \<tau> f = Some im \<longrightarrow> inum_of_file \<tau> pf \<noteq> None"
+ using vt pre apply (erule_tac impE, simp) by ((erule_tac conjE)+, assumption)
+ have fns: "\<forall>f f' im. is_file \<tau> f \<and> parent f' = Some f \<and> inum_of_file \<tau> f' = Some im \<longrightarrow> False"
+ using vt pre apply (erule_tac impE, simp) by ((erule_tac conjE)+, assumption)
+ have hns: "\<forall>f f' im. f \<in> files_hung_by_del \<tau> \<and> inum_of_file \<tau> f' = Some im \<and> parent f' = Some f \<longrightarrow> False"
+ using vt pre apply (erule_tac impE, simp) by ((erule_tac conjE)+, assumption)
+ have ain: "\<forall>f' f im. f \<preceq> f' \<and> inum_of_file \<tau> f' = Some im \<longrightarrow> inum_of_file \<tau> f \<noteq> None"
+ proof
+ fix f'
+ show " \<forall>f im. f \<preceq> f' \<and> inum_of_file \<tau> f' = Some im \<longrightarrow> inum_of_file \<tau> f \<noteq> None"
+ proof (induct f')
+ case Nil show ?case by (auto simp: no_junior_def)
+ next
+ case (Cons a f')
+ assume pre:"\<forall>f im. f \<preceq> f' \<and> inum_of_file \<tau> f' = Some im \<longrightarrow> inum_of_file \<tau> f \<noteq> None"
+ show "\<forall>f im. f \<preceq> a # f' \<and> inum_of_file \<tau> (a # f') = Some im \<longrightarrow> inum_of_file \<tau> f \<noteq> None"
+ proof clarify
+ fix f im
+ assume h1: "f \<preceq> a # f'"
+ and h2: "inum_of_file \<tau> (a # f') = Some im"
+ show "\<exists>y. inum_of_file \<tau> f = Some y"
+ proof-
+ have h3: "\<exists> y. inum_of_file \<tau> f' = Some y"
+ proof-
+ have "parent (a # f') = Some f'" by simp
+ hence "\<exists> y. inum_of_file \<tau> f' = Some y" using pin h2 by blast
+ with h1 show ?thesis by simp
+ qed
+ from h1 have "f = a # f' \<or> f = f' \<or> f \<preceq> f'" by (induct f, auto simp:no_junior_def)
+ moreover have "f = a # f' \<Longrightarrow> \<exists>y. inum_of_file \<tau> f = Some y" using h2 by simp
+ moreover have "f = f' \<Longrightarrow> \<exists>y. inum_of_file \<tau> f = Some y" using h3 by simp
+ moreover have "f \<preceq> f' \<Longrightarrow> \<exists>y. inum_of_file \<tau> f = Some y" using pre h3 by simp
+ ultimately show ?thesis by auto
+ qed
+ qed
+ qed
+ qed
+
+ have fin': "\<And> f. f \<in> files_hung_by_del \<tau> \<Longrightarrow> \<exists> im. inum_of_file \<tau> f = Some im" using fin by auto
+ have pin': "\<And> f pf im. \<lbrakk>parent f = Some pf; inum_of_file \<tau> f = Some im\<rbrakk> \<Longrightarrow> \<exists> im'. inum_of_file \<tau> pf = Some im'"
+ using pin by auto
+ have fns': "\<And> f f' im. \<lbrakk>is_file \<tau> f; parent f' = Some f; inum_of_file \<tau> f' = Some im\<rbrakk> \<Longrightarrow> False" using fns by auto
+ have fns'': "\<And> f f' im im'. \<lbrakk>itag_of_inum \<tau> im = Some Tag_FILE; inum_of_file \<tau> f = Some im; parent f' = Some f; inum_of_file \<tau> f' = Some im'\<rbrakk> \<Longrightarrow> False"
+ by (rule_tac f = f and f' = f' in fns', auto simp:is_file_def)
+ have hns': "\<And> f f' im. \<lbrakk>f \<in> files_hung_by_del \<tau>; inum_of_file \<tau> f' = Some im; parent f' = Some f\<rbrakk> \<Longrightarrow> False" using hns by auto
+ have ain': "\<And> f f' im. \<lbrakk>f \<preceq> f'; inum_of_file \<tau> f' = Some im\<rbrakk> \<Longrightarrow> \<exists> im'. inum_of_file \<tau> f = Some im'" using ain by auto
+ have dns': "\<And> f f' im. \<lbrakk>dir_is_empty \<tau> f; parent f' = Some f; inum_of_file \<tau> f' = Some im\<rbrakk> \<Longrightarrow> False"
+ apply (auto simp:dir_is_empty_def current_files_def is_dir_def split:option.splits)
+ by (erule_tac x = f' in allE, simp add:noJ_Anc parent_is_ancen, drule parent_is_parent, simp+)
+
+ have "\<forall> f. f \<in> files_hung_by_del (a # \<tau>) \<longrightarrow> inum_of_file (a # \<tau>) f \<noteq> None"
+ apply (clarify, case_tac a) using os fin
+ apply (auto simp:files_hung_by_del.simps inum_of_file.simps os_grant.simps current_files_def is_file_def
+ split:if_splits option.splits)
+ done
+ moreover
+ have "\<forall>f pf im. parent f = Some pf \<and> inum_of_file (a # \<tau>) f = Some im \<longrightarrow> inum_of_file (a # \<tau>) pf \<noteq> None"
+ apply (clarify, case_tac a)
+ using vt os pin'
+ apply (auto simp:os_grant.simps current_files_def inum_of_file.simps is_file_def is_dir_def
+ split:if_splits option.splits t_inode_tag.splits)
+ apply (drule_tac f = pf and f' = f in hns', simp, simp, simp)
+ apply (rule_tac f = list and f' = f in fns', simp add:is_file_def, simp, simp)
+ apply (rule_tac f = list and f' = f in dns', simp add:is_dir_def, simp, simp)
+ done
+ moreover have "\<forall>f f' im. is_file (a # \<tau>) f \<and> parent f' = Some f \<and> inum_of_file (a # \<tau>) f' = Some im \<longrightarrow> False"
+ apply (clarify, case_tac a)
+ using vt os fns'' cons
+ apply (auto simp:os_grant.simps current_files_def inum_of_file.simps itag_of_inum.simps
+ is_file_def is_dir_def
+ dest:ios's_im_in_cim iof's_im_in_cim
+ split:if_splits option.splits t_inode_tag.splits t_socket_type.splits)
+ apply (rule_tac im = a and f = f and f' = f' in fns'', simp+)
+ apply (drule_tac f = f' and pf = list in pin', simp, simp)
+ done
+ moreover have "\<forall>f f' im. f \<in> files_hung_by_del (a # \<tau>) \<and> inum_of_file (a # \<tau>) f' = Some im \<and>
+ parent f' = Some f \<longrightarrow> False"
+ apply (clarify, case_tac a)
+ using vt os hns'
+ apply (auto simp:os_grant.simps current_files_def inum_of_file.simps files_hung_by_del.simps
+ split:if_splits option.splits t_sock_addr.splits)
+ apply (drule fns', simp+)
+ done
+ ultimately show ?thesis by blast
+ qed
+ qed
+qed
+
+(**************** hung file in current ***********************)
+
+lemma hung_file_has_inum:"\<lbrakk>f \<in> files_hung_by_del \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> inum_of_file \<tau> f \<noteq> None"
+by (drule file_leveling[rule_format], blast)
+
+lemma hung_file_has_inum': "\<lbrakk>f \<in> files_hung_by_del \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> im. inum_of_file \<tau> f = Some im"
+by (auto dest:hung_file_has_inum)
+
+lemma hung_file_in_current: "\<lbrakk>f \<in> files_hung_by_del \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> f \<in> current_files \<tau>"
+by (clarsimp simp add:current_files_def hung_file_has_inum')
+
+lemma parentf_has_inum: "\<lbrakk>parent f = Some pf; inum_of_file \<tau> f = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> inum_of_file \<tau> pf \<noteq> None"
+by (drule file_leveling[rule_format], blast)
+
+lemma parentf_has_inum': "\<lbrakk>parent f = Some pf; inum_of_file \<tau> f = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> im'. inum_of_file \<tau> pf = Some im'"
+by (auto dest:parentf_has_inum)
+
+lemma parentf_in_current: "\<lbrakk>parent f = Some pf; f \<in> current_files \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> pf \<in> current_files \<tau>"
+by (clarsimp simp add:current_files_def parentf_has_inum')
+
+lemma parentf_in_current': "\<lbrakk>a # pf \<in> current_files \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> pf \<in> current_files \<tau>"
+apply (subgoal_tac "parent (a # pf) = Some pf")
+by (erule parentf_in_current, simp+)
+
+lemma ancenf_has_inum_aux: "\<forall> f im. f \<preceq> f' \<and> inum_of_file \<tau> f' = Some im \<and> valid \<tau> \<longrightarrow> inum_of_file \<tau> f \<noteq> None"
+proof (induct f')
+ case Nil show ?case by (auto simp: no_junior_def)
+next
+ case (Cons a f')
+ assume pre:"\<forall>f im. f \<preceq> f' \<and> inum_of_file \<tau> f' = Some im \<and> valid \<tau> \<longrightarrow> inum_of_file \<tau> f \<noteq> None"
+ show "\<forall>f im. f \<preceq> a # f' \<and> inum_of_file \<tau> (a # f') = Some im \<and> valid \<tau> \<longrightarrow> inum_of_file \<tau> f \<noteq> None"
+ proof clarify
+ fix f im
+ assume h1: "f \<preceq> a # f'"
+ and h2: "inum_of_file \<tau> (a # f') = Some im"
+ and h3: "valid \<tau>"
+ show "\<exists>y. inum_of_file \<tau> f = Some y"
+ proof-
+ have h4: "\<exists> y. inum_of_file \<tau> f' = Some y"
+ proof-
+ have "parent (a # f') = Some f'" by simp
+ hence "\<exists> y. inum_of_file \<tau> f' = Some y" using parentf_has_inum' h2 h3 by blast
+ with h1 show ?thesis by simp
+ qed
+ from h1 have "f = a # f' \<or> f = f' \<or> f \<preceq> f'" by (induct f, auto simp:no_junior_def)
+ moreover have "f = a # f' \<Longrightarrow> \<exists>y. inum_of_file \<tau> f = Some y" using h2 by simp
+ moreover have "f = f' \<Longrightarrow> \<exists>y. inum_of_file \<tau> f = Some y" using h4 by simp
+ moreover have "f \<preceq> f' \<Longrightarrow> \<exists>y. inum_of_file \<tau> f = Some y" using pre h3 h4 by simp
+ ultimately show ?thesis by auto
+ qed
+ qed
+qed
+
+lemma ancenf_has_inum: "\<lbrakk>f \<preceq> f'; inum_of_file \<tau> f' = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> inum_of_file \<tau> f \<noteq> None"
+by (rule ancenf_has_inum_aux[rule_format], auto)
+
+lemma ancenf_has_inum': "\<lbrakk>f \<preceq> f'; inum_of_file \<tau> f' = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> im'. inum_of_file \<tau> f = Some im'"
+by (auto dest:ancenf_has_inum)
+
+lemma ancenf_in_current: "\<lbrakk>f \<preceq> f'; f' \<in> current_files \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> f \<in> current_files \<tau>"
+by (simp add:current_files_def, erule exE, simp add:ancenf_has_inum')
+
+lemma file_has_no_son: "\<lbrakk>is_file \<tau> f; parent f' = Some f; inum_of_file \<tau> f' = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+by (drule file_leveling[rule_format], blast)
+
+lemma file_has_no_son': "\<lbrakk>is_file \<tau> f; parent f' = Some f; f' \<in> current_files \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+by (simp add:current_files_def, erule exE, auto intro:file_has_no_son)
+
+lemma hung_file_no_son: "\<lbrakk>f \<in> files_hung_by_del \<tau>; valid \<tau>; inum_of_file \<tau> f' = Some im; parent f' = Some f\<rbrakk> \<Longrightarrow> False"
+by (drule file_leveling[rule_format], blast)
+
+lemma hung_file_no_son': "\<lbrakk>f \<in> files_hung_by_del \<tau>; valid \<tau>; f' \<in> current_files \<tau>; parent f' = Some f\<rbrakk> \<Longrightarrow> False"
+by (simp add:current_files_def, erule exE, auto intro:hung_file_no_son)
+
+lemma hung_file_no_des_aux: "\<forall> f. f \<in> files_hung_by_del \<tau> \<and> valid \<tau> \<and> f' \<in> current_files \<tau> \<and> f \<preceq> f' \<and> f \<noteq> f' \<longrightarrow> False"
+proof (induct f')
+ case Nil
+ show ?case
+ by (auto simp:files_hung_by_del.simps current_files_def inum_of_file.simps no_junior_def split:if_splits option.splits)
+next
+ case (Cons a pf)
+ assume pre: "\<forall>f. f \<in> files_hung_by_del \<tau> \<and> valid \<tau> \<and> pf \<in> current_files \<tau> \<and> f \<preceq> pf \<and> f \<noteq> pf\<longrightarrow> False"
+ show ?case
+ proof clarify
+ fix f
+ assume h1: "f \<in> files_hung_by_del \<tau>"
+ and h2: "valid \<tau>"
+ and h3: "a # pf \<in> current_files \<tau>"
+ and h4: "f \<preceq> a # pf"
+ and h5: "f \<noteq> a # pf"
+ have h6: "parent (a # pf) = Some pf" by simp
+ with h2 h3 have h7: "pf \<in> current_files \<tau>" by (drule_tac parentf_in_current, auto)
+ from h4 h5 have h8: "f \<preceq> pf" by (erule_tac no_juniorE, case_tac zs, auto simp:no_junior_def)
+ show False
+ proof (cases "f = pf")
+ case True with h6 h2 h3 h1
+ show False by (auto intro!:hung_file_no_son')
+ next
+ case False with pre h1 h2 h7 h8
+ show False by blast
+ qed
+ qed
+qed
+
+lemma hung_file_no_des: "\<lbrakk>f \<in> files_hung_by_del \<tau>; valid \<tau>; f' \<in> current_files \<tau>; f \<preceq> f'; f \<noteq> f'\<rbrakk> \<Longrightarrow> False"
+by (rule hung_file_no_des_aux[rule_format], blast)
+
+(* current version, dir can not be opened, so hung_files are all files
+lemma hung_file_is_leaf: "\<lbrakk>f \<in> files_hung_by_del \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> is_file \<tau> f \<or> dir_is_empty \<tau> f"
+apply (frule hung_file_has_inum', simp, erule exE)
+apply (auto simp add:is_file_def dir_is_empty_def is_dir_def dest:finum_has_itag finum_has_ftag split:option.splits if_splits t_inode_tag.splits)
+by (simp add: noJ_Anc, auto dest:hung_file_no_des)
+*)
+
+lemma hung_file_has_filetag:
+ "\<lbrakk>f \<in> files_hung_by_del s; inum_of_file s f = Some im; valid s\<rbrakk> \<Longrightarrow> itag_of_inum s im = Some Tag_FILE"
+apply (induct s)
+apply (simp add:files_hung_by_del.simps)
+apply (drule init_files_hung_prop2, (erule exE)+)
+apply (drule init_filefd_prop5, clarsimp simp:is_init_file_def split:t_inode_tag.splits option.splits)
+
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:files_hung_by_del.simps is_file_def is_dir_def current_files_def current_inode_nums_def
+ split:if_splits option.splits t_inode_tag.splits t_socket_type.splits
+ dest:hung_file_has_inum iof's_im_in_cim)
+done
+
+lemma hung_file_is_file: "\<lbrakk>f \<in> files_hung_by_del \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> is_file \<tau> f"
+apply (frule hung_file_has_inum', simp, erule exE)
+apply (drule hung_file_has_filetag, auto simp:is_file_def)
+done
+
+(*********************** 2 in 1 *********************)
+
+lemma file_of_pfd_2in1: "valid s \<Longrightarrow> (
+ (\<forall> p fd f. file_of_proc_fd s p fd = Some f \<longrightarrow> inum_of_file s f \<noteq> None) \<and>
+ (\<forall> p fd f. file_of_proc_fd s p fd = Some f \<longrightarrow> is_file s f) )"
+proof (induct s)
+ case Nil
+ show ?case
+ by (auto dest:init_filefd_valid simp:is_file_def)
+next
+ case (Cons e s)
+ hence vd_e: "valid (e # s)" and vd_s: "valid s" and os: "os_grant s e"
+ and pfd: "\<And> p fd f. file_of_proc_fd s p fd = Some f \<Longrightarrow> inum_of_file s f \<noteq> None"
+ and isf: "\<And> p fd f. file_of_proc_fd s p fd = Some f \<Longrightarrow> is_file s f"
+ by (auto dest:vd_cons vt_grant_os)
+ from pfd have pfd': "\<And> p fd f. inum_of_file s f = None \<Longrightarrow> file_of_proc_fd s p fd \<noteq> Some f"
+ by (rule_tac notI, drule_tac pfd, simp)
+
+ have "\<forall>p fd f. file_of_proc_fd (e # s) p fd = Some f \<longrightarrow> inum_of_file (e # s) f \<noteq> None"
+ apply (case_tac e) using os
+ apply (auto simp:inum_of_file.simps file_of_proc_fd.simps os_grant.simps current_files_def
+ dir_is_empty_def is_file_def is_dir_def
+ split:if_splits option.splits dest:pfd)
+ apply (simp add:pfdof_simp3)+
+ apply (simp add:proc_fd_of_file_def)
+ apply (drule isf, simp add:is_file_def split:t_inode_tag.splits)
+ done
+ moreover
+ have "\<forall>p fd f. file_of_proc_fd (e # s) p fd = Some f \<longrightarrow> is_file (e # s) f"
+ apply (case_tac e) using os
+ apply (auto split:option.splits t_inode_tag.splits if_splits t_socket_type.splits
+ dest:pfd isf iof's_im_in_cim
+ simp:is_file_def is_dir_def dir_is_empty_def current_files_def)
+ apply (simp add:pfdof_simp3)+
+ apply (simp add:proc_fd_of_file_def)
+ done
+ ultimately show ?case by auto
+qed
+
+
+(************** file_of_proc_fd in current ********************)
+
+lemma file_of_pfd_imp_inode': "\<lbrakk>file_of_proc_fd \<tau> p fd = Some f; valid \<tau>\<rbrakk> \<Longrightarrow> inum_of_file \<tau> f \<noteq> None"
+by (drule file_of_pfd_2in1, blast)
+
+lemma file_of_pfd_imp_inode: "\<lbrakk>file_of_proc_fd \<tau> p fd = Some f; valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> im. inum_of_file \<tau> f = Some im"
+by (auto dest!:file_of_pfd_imp_inode')
+
+lemma file_of_pfd_in_current: "\<lbrakk>file_of_proc_fd \<tau> p fd = Some f; valid \<tau>\<rbrakk> \<Longrightarrow> f \<in> current_files \<tau>"
+by (auto dest!:file_of_pfd_imp_inode' simp:current_files_def)
+
+
+(*************** file_of_proc_fd is file *********************)
+
+lemma file_of_pfd_is_file:
+ "\<lbrakk>file_of_proc_fd \<tau> p fd = Some f; valid \<tau>\<rbrakk> \<Longrightarrow> is_file \<tau> f"
+by (drule file_of_pfd_2in1, auto simp:is_file_def)
+
+lemma file_of_pfd_is_file':
+ "\<lbrakk>file_of_proc_fd \<tau> p fd = Some f; is_dir \<tau> f; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+by (drule file_of_pfd_is_file, auto simp:is_file_def is_dir_def split:option.splits t_inode_tag.splits)
+
+lemma file_of_pfd_is_file_tag:
+ "\<lbrakk>file_of_proc_fd \<tau> p fd = Some f; valid \<tau>; inum_of_file \<tau> f = Some im\<rbrakk> \<Longrightarrow> itag_of_inum \<tau> im = Some Tag_FILE"
+by (drule file_of_pfd_is_file, auto simp:is_file_def split:option.splits t_inode_tag.splits)
+
+(************** parent file / ancestral file is dir *******************)
+
+lemma parentf_is_dir_aux: "\<forall> f pf. parent f = Some pf \<and> inum_of_file \<tau> f = Some im \<and> inum_of_file \<tau> pf = Some ipm \<and> valid \<tau> \<longrightarrow> itag_of_inum \<tau> ipm = Some Tag_DIR"
+apply (induct \<tau>)
+apply (clarsimp simp:inum_of_file.simps itag_of_inum.simps init_parent_file_is_dir')
+apply (clarify, frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:inum_of_file.simps itag_of_inum.simps os_grant.simps
+ current_files_def is_dir_def is_file_def
+ dest: ios's_im_in_cim iof's_im_in_cim
+ split:if_splits option.splits t_sock_addr.splits t_inode_tag.splits t_socket_type.splits)
+apply (drule parentf_has_inum', simp, simp, simp)+
+done
+
+lemma parentf_has_dirtag:
+ "\<lbrakk>parent f = Some pf; inum_of_file \<tau> f = Some im; inum_of_file \<tau> pf = Some ipm; valid \<tau>\<rbrakk>
+ \<Longrightarrow> itag_of_inum \<tau> ipm = Some Tag_DIR"
+by (auto intro:parentf_is_dir_aux[rule_format])
+
+lemma parentf_is_dir': "\<lbrakk>parent f = Some pf; inum_of_file \<tau> f = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> is_dir \<tau> pf"
+apply (frule parentf_has_inum', simp+, erule exE, simp add:is_dir_def split:t_inode_tag.splits option.splits)
+by (auto dest:parentf_has_dirtag)
+
+lemma parentf_is_dir: "\<lbrakk>parent f = Some pf; f \<in> current_files \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> is_dir \<tau> pf"
+by (clarsimp simp:current_files_def parentf_is_dir')
+
+lemma parentf_is_dir'': "\<lbrakk>f # pf \<in> current_files \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> is_dir \<tau> pf"
+by (auto intro!:parentf_is_dir)
+
+lemma ancenf_is_dir_aux: "\<forall> f. f \<preceq> f' \<and> f \<noteq> f' \<and> f' \<in> current_files \<tau> \<and> valid \<tau> \<longrightarrow> is_dir \<tau> f"
+proof (induct f')
+ case Nil show ?case
+ by (auto simp:current_files_def no_junior_def)
+next
+ case (Cons a pf)
+ assume pre: "\<forall>f. f \<preceq> pf \<and> f \<noteq> pf \<and> pf \<in> current_files \<tau> \<and> valid \<tau> \<longrightarrow> is_dir \<tau> f"
+ show ?case
+ proof clarify
+ fix f
+ assume h1: "f \<preceq> a # pf"
+ and h2: "f \<noteq> a # pf"
+ and h3: "a # pf \<in> current_files \<tau>"
+ and h4: "valid \<tau>"
+ have h5: "parent (a # pf) = Some pf" by simp
+ with h3 h4 have h6: "pf \<in> current_files \<tau>" by (drule_tac parentf_in_current, auto)
+ from h1 h2 have h7: "f \<preceq> pf" by (erule_tac no_juniorE, case_tac zs, auto simp:no_junior_def)
+ show "is_dir \<tau> f"
+ proof (cases "f = pf")
+ case True with h3 h4 h5 show "is_dir \<tau> f" by (drule_tac parentf_is_dir, auto)
+ next
+ case False with pre h6 h7 h4 show "is_dir \<tau> f" by blast
+ qed
+ qed
+qed
+
+lemma ancenf_is_dir: "\<lbrakk>f \<preceq> f'; f \<noteq> f'; f' \<in> current_files \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> is_dir \<tau> f"
+by (auto intro:ancenf_is_dir_aux[rule_format])
+
+(************* rebuild current_files simpset ***********************)
+
+lemma current_files_nil: "current_files [] = init_files"
+apply (simp add:current_files_def inum_of_file.simps)
+by (auto dest:inof_has_file_tag init_file_has_inum)
+
+lemma current_files_open: "current_files (Open p f flags fd (Some i) # \<tau>) = insert f (current_files \<tau>)"
+by (auto simp add:current_files_def inum_of_file.simps split:option.splits)
+
+lemma current_files_open': "current_files (Open p f flags fd None # \<tau>) = current_files \<tau>"
+by (simp add:current_files_def inum_of_file.simps split:option.splits)
+
+lemma current_files_closefd: "current_files (CloseFd p fd # \<tau>) = (
+ case (file_of_proc_fd \<tau> p fd) of
+ Some f \<Rightarrow> ( if ((proc_fd_of_file \<tau> f = {(p, fd)}) \<and> (f \<in> files_hung_by_del \<tau>))
+ then current_files \<tau> - {f}
+ else current_files \<tau>)
+ | _ \<Rightarrow> current_files \<tau> )"
+by (auto simp:current_files_def inum_of_file.simps split:option.splits)
+
+lemma current_files_unlink: "current_files (UnLink p f # \<tau>) = (if (proc_fd_of_file \<tau> f = {}) then (current_files \<tau>) - {f} else current_files \<tau>)"
+by (auto simp:current_files_def inum_of_file.simps split:option.splits)
+
+lemma current_files_rmdir: "current_files (Rmdir p f # \<tau>) = current_files \<tau> - {f}"
+by (auto simp:current_files_def inum_of_file.simps split:option.splits)
+
+lemma current_files_mkdir: "current_files (Mkdir p f ino # \<tau>) = insert f (current_files \<tau>)"
+by (auto simp:current_files_def inum_of_file.simps split:option.splits)
+
+(*
+lemma rename_renaming_decom:
+ "\<lbrakk>f\<^isub>3 \<preceq> file_after_rename f\<^isub>2 f\<^isub>3 f; Rename p f\<^isub>2 f\<^isub>3 # valid \<tau>; f \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> f\<^isub>2 \<preceq> f"
+apply (case_tac "f\<^isub>2 \<preceq> f", simp)
+apply (simp add:file_after_rename_def split:if_splits)
+by (frule vd_cons, frule vt_grant_os, auto simp:os_grant.simps dest!:ancenf_in_current)
+
+lemma rename_renaming_decom':
+ "\<lbrakk>\<not> f\<^isub>3 \<preceq> file_after_rename f\<^isub>2 f\<^isub>3 f; Rename p f\<^isub>2 f\<^isub>3 # valid \<tau>; f \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> \<not> f\<^isub>2 \<preceq> f"
+by (case_tac "f\<^isub>2 \<preceq> f", drule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop1, simp+)
+
+lemma current_files_rename: "Rename p f\<^isub>2 f\<^isub>3 # valid \<tau> \<Longrightarrow> current_files (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) = {file_after_rename f\<^isub>2 f\<^isub>3 f\<^isub>1| f\<^isub>1. f\<^isub>1 \<in> current_files \<tau>}"
+apply (frule vt_grant_os, frule vd_cons)
+apply (auto simp:current_files_def inum_of_file.simps os_grant.simps split:if_splits option.splits)
+apply (rule_tac x = x in exI, simp add:file_after_rename_def)
+apply (frule_tac f\<^isub>2 = f\<^isub>2 in file_renaming_prop1', drule_tac f\<^isub>2 = f\<^isub>2 in file_renaming_prop5')
+apply (erule_tac x = "file_before_rename f\<^isub>2 f\<^isub>3 x" in allE, simp)
+apply (rule_tac x = x in exI, simp add:file_after_rename_def)
+apply (drule_tac a = f\<^isub>3 and b = f\<^isub>2 in no_junior_conf, simp, simp)
+apply (drule_tac f = f\<^isub>3 and f' = f\<^isub>2 in ancenf_has_inum', simp, simp, simp)
+apply (drule_tac f\<^isub>2 = f\<^isub>2 in rename_renaming_decom, simp, simp add:current_files_def, simp add:file_renaming_prop5)
+apply (drule_tac f\<^isub>2 = f\<^isub>2 in rename_renaming_decom', simp, simp add:current_files_def)
+apply (simp add:file_after_rename_def)
+apply (drule_tac f\<^isub>2 = f\<^isub>2 in rename_renaming_decom', simp, simp add:current_files_def)
+apply (simp add:file_after_rename_def)
+done
+*)
+
+lemma current_files_other:
+ "\<lbrakk>\<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p f. e \<noteq> UnLink p f;
+ \<forall> p f. e \<noteq> Rmdir p f;
+ \<forall> p f i. e \<noteq> Mkdir p f i\<rbrakk> \<Longrightarrow> current_files (e # \<tau>) = current_files \<tau>"
+by (case_tac e, auto simp:current_files_def inum_of_file.simps)
+
+lemmas current_files_simps = current_files_nil current_files_open current_files_open'
+ current_files_closefd current_files_unlink current_files_rmdir
+ current_files_mkdir current_files_other
+
+
+(******************** is_file simpset *********************)
+
+lemma is_file_open:
+ "valid (Open p f flags fd opt # s) \<Longrightarrow>
+ is_file (Open p f flags fd opt # s) = (if (opt = None) then is_file s else (is_file s) (f:= True))"
+apply (frule vd_cons, drule vt_grant_os, rule ext)
+apply (auto dest:finum_has_itag iof's_im_in_cim
+ split:if_splits option.splits t_inode_tag.splits
+ simp:is_file_def current_files_def)
+done
+
+lemma is_file_closefd:
+ "valid (CloseFd p fd # s) \<Longrightarrow> is_file (CloseFd p fd # s) = (
+ case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> ( if ((proc_fd_of_file s f = {(p, fd)}) \<and> (f \<in> files_hung_by_del s))
+ then (is_file s) (f := False)
+ else is_file s)
+ | _ \<Rightarrow> is_file s )"
+apply (frule vd_cons, drule vt_grant_os, rule ext)
+apply (auto dest:finum_has_itag iof's_im_in_cim
+ split:if_splits option.splits t_inode_tag.splits
+ simp:is_file_def)
+done
+
+lemma is_file_unlink:
+ "valid (UnLink p f # s) \<Longrightarrow> is_file (UnLink p f # s) = (
+ if (proc_fd_of_file s f = {}) then (is_file s) (f := False) else is_file s)"
+apply (frule vd_cons, drule vt_grant_os, rule ext)
+apply (auto dest:finum_has_itag iof's_im_in_cim
+ split:if_splits option.splits t_inode_tag.splits
+ simp:is_file_def)
+done
+
+lemma is_file_other:
+ "\<lbrakk>valid (e # \<tau>);
+ \<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p f. e \<noteq> UnLink p f\<rbrakk> \<Longrightarrow> is_file (e # \<tau>) = is_file \<tau>"
+apply (frule vd_cons, drule vt_grant_os, rule_tac ext, case_tac e)
+apply (auto dest:finum_has_itag iof's_im_in_cim intro!:ext
+ split:if_splits option.splits t_inode_tag.split t_socket_type.splits
+ simp:is_file_def dir_is_empty_def is_dir_def current_files_def)
+done
+
+lemma file_dir_conflict: "\<lbrakk>is_file s f; is_dir s f\<rbrakk> \<Longrightarrow> False"
+by (auto simp:is_file_def is_dir_def split:option.splits t_inode_tag.splits)
+
+lemma is_file_not_dir: "is_file s f \<Longrightarrow> \<not> is_dir s f"
+by (rule notI, erule file_dir_conflict, simp)
+
+lemma is_dir_not_file: "is_dir s f \<Longrightarrow> \<not> is_file s f"
+by (rule notI, erule file_dir_conflict, simp)
+
+lemmas is_file_simps = is_file_nil is_file_open is_file_closefd is_file_unlink is_file_other
+
+(********* is_dir simpset **********)
+
+lemma is_dir_mkdir: "valid (Mkdir p f i # s) \<Longrightarrow> is_dir (Mkdir p f i # s) = (is_dir s) (f := True)"
+apply (frule vd_cons, drule vt_grant_os, rule_tac ext)
+apply (auto dest:finum_has_itag iof's_im_in_cim intro!:ext
+ split:if_splits option.splits t_inode_tag.split t_socket_type.splits
+ simp:is_dir_def dir_is_empty_def is_dir_def current_files_def)
+done
+
+lemma is_dir_rmdir: "valid (Rmdir p f # s) \<Longrightarrow> is_dir (Rmdir p f # s) = (is_dir s) (f := False)"
+apply (frule vd_cons, drule vt_grant_os, rule_tac ext)
+apply (auto dest:finum_has_itag iof's_im_in_cim intro!:ext
+ split:if_splits option.splits t_inode_tag.split t_socket_type.splits
+ simp:is_dir_def dir_is_empty_def is_dir_def current_files_def)
+done
+
+lemma is_dir_other:
+ "\<lbrakk>valid (e # s);
+ \<forall> p f. e \<noteq> Rmdir p f;
+ \<forall> p f i. e \<noteq> Mkdir p f i\<rbrakk> \<Longrightarrow> is_dir (e # s) = is_dir s"
+apply (frule vd_cons, drule vt_grant_os, rule_tac ext, case_tac e)
+apply (auto dest:finum_has_itag iof's_im_in_cim intro!:ext
+ split:if_splits option.splits t_inode_tag.split t_socket_type.splits
+ simp:is_file_def dir_is_empty_def is_dir_def current_files_def)
+apply (drule file_of_pfd_is_file, simp)
+apply (simp add:is_file_def split:t_inode_tag.splits option.splits)
+done
+
+lemmas is_dir_simps = is_dir_nil is_dir_mkdir is_dir_rmdir is_dir_other
+
+(*********** no root dir involved ***********)
+
+lemma root_is_dir: "valid s \<Longrightarrow> is_dir s []"
+apply (induct s, simp add:is_dir_nil root_is_init_dir)
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:is_dir_simps)
+done
+
+lemma root_is_dir': "\<lbrakk>is_file s []; valid s\<rbrakk> \<Longrightarrow> False"
+apply (drule root_is_dir)
+apply (erule file_dir_conflict, simp)
+done
+
+lemma noroot_execve:
+ "valid (Execve p f fds # s) \<Longrightarrow> f \<noteq> []"
+by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir')
+
+lemma noroot_execve':
+ "valid (Execve p [] fds # s) \<Longrightarrow> False"
+by (drule noroot_execve, simp)
+
+lemma noroot_open:
+ "valid (Open p f flags fd opt # s) \<Longrightarrow> f \<noteq> []"
+by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir' split:option.splits)
+
+lemma noroot_open':
+ "valid (Open p [] flags fd opt # s) \<Longrightarrow> False"
+by (drule noroot_open, simp)
+
+lemma noroot_filefd':
+ "\<lbrakk>file_of_proc_fd s p fd = Some []; valid s\<rbrakk> \<Longrightarrow> False"
+apply (induct s arbitrary:p, simp)
+apply (drule init_filefd_prop5, erule root_is_init_dir')
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto split:if_splits option.splits dest!:root_is_dir')
+done
+
+lemma noroot_filefd:
+ "\<lbrakk>file_of_proc_fd s p fd = Some f; valid s\<rbrakk> \<Longrightarrow> f \<noteq> []"
+by (rule notI, simp, erule noroot_filefd', simp)
+
+lemma noroot_unlink:
+ "valid (UnLink p f # s) \<Longrightarrow> f \<noteq> []"
+by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir')
+
+lemma noroot_unlink':
+ "valid (UnLink p [] # s) \<Longrightarrow> False"
+by (drule noroot_unlink, simp)
+
+lemma noroot_rmdir:
+ "valid (Rmdir p f # s) \<Longrightarrow> f \<noteq> []"
+by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir')
+
+lemma noroot_rmdir':
+ "valid (Rmdir p [] # s) \<Longrightarrow> False"
+by (drule noroot_rmdir, simp)
+
+lemma noroot_mkdir:
+ "valid (Mkdir p f inum # s) \<Longrightarrow> f \<noteq> []"
+by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir')
+
+lemma noroot_mkdir':
+ "valid (Mkdir p [] inum # s) \<Longrightarrow> False"
+by (drule noroot_mkdir, simp)
+
+lemma noroot_truncate:
+ "valid (Truncate p f len # s) \<Longrightarrow> f \<noteq> []"
+by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir')
+
+lemma noroot_truncate':
+ "valid (Truncate p [] len # s) \<Longrightarrow> False"
+by (drule noroot_truncate, simp)
+
+lemmas noroot_events = noroot_execve noroot_open noroot_filefd noroot_unlink noroot_rmdir
+ noroot_mkdir noroot_truncate
+
+lemmas noroot_events' = noroot_execve' noroot_open' noroot_filefd' noroot_unlink' noroot_rmdir'
+ noroot_mkdir' noroot_truncate'
+
+
+lemma is_file_in_current:
+ "is_file s f \<Longrightarrow> f \<in> current_files s"
+by (auto simp:is_file_def current_files_def split:option.splits)
+
+lemma is_dir_in_current:
+ "is_dir s f \<Longrightarrow> f \<in> current_files s"
+by (auto simp:is_dir_def current_files_def split:option.splits)
+
+
+(* simpset for same_inode_files: Current_files_prop.thy *)
+(*
+lemma same_inode_files_nil:
+ "same_inode_files [] = init_same_inode_files"
+by (rule ext, simp add:same_inode_files_def init_same_inode_files_def is_file_nil)
+
+lemma iof's_im_in_cim': "Some im = inum_of_file \<tau> f \<Longrightarrow> im \<in> current_inode_nums \<tau>"
+by (auto simp add:current_inode_nums_def current_file_inums_def)
+
+lemma same_inode_files_open:
+ "valid (Open p f flags fd opt # s) \<Longrightarrow> same_inode_files (Open p f flags fd opt # s) = (\<lambda> f'.
+ if (f' = f \<and> opt \<noteq> None) then {f} else same_inode_files s f')"
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+apply (auto simp:same_inode_files_def is_file_simps split:if_splits option.splits
+ dest:iof's_im_in_cim iof's_im_in_cim')
+apply (drule is_file_in_current)
+apply (simp add:current_files_def)
+done
+
+lemma same_inode_files_linkhard:
+ "valid (LinkHard p oldf f # s) \<Longrightarrow> same_inode_files (LinkHard p oldf f # s) = (\<lambda> f'.
+ if (f' = f \<or> f' \<in> same_inode_files s oldf)
+ then same_inode_files s oldf \<union> {f}
+ else same_inode_files s f')"
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+apply (auto simp:same_inode_files_def is_file_simps split:if_splits option.splits
+ dest:iof's_im_in_cim iof's_im_in_cim')
+apply (drule is_file_in_current)
+apply (simp add:current_files_def is_file_def)
+apply (simp add:is_file_def)
+done
+
+lemma inum_of_file_none_prop:
+ "\<lbrakk>inum_of_file s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
+by (simp add:current_files_def)
+
+lemma same_inode_files_closefd:
+ "\<lbrakk>valid (CloseFd p fd # s); f' \<in> current_files (CloseFd p fd # s)\<rbrakk> \<Longrightarrow>
+ same_inode_files (CloseFd p fd # s) f' = (
+ case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> (if ((proc_fd_of_file s f = {(p, fd)}) \<and> (f \<in> files_hung_by_del s))
+ then same_inode_files s f' - {f}
+ else same_inode_files s f' )
+ | None \<Rightarrow> same_inode_files s f' )"
+apply (frule vt_grant_os, frule vd_cons)
+apply (auto simp:same_inode_files_def is_file_closefd current_files_closefd
+ split:if_splits option.splits
+ dest:iof's_im_in_cim iof's_im_in_cim' inum_of_file_none_prop)
+done
+
+lemma same_inode_files_unlink:
+ "\<lbrakk>valid (UnLink p f # s); f' \<in> current_files (UnLink p f # s)\<rbrakk>
+ \<Longrightarrow> same_inode_files (UnLink p f # s) f' = (
+ if (proc_fd_of_file s f = {})
+ then same_inode_files s f' - {f}
+ else same_inode_files s f')"
+apply (frule vt_grant_os, frule vd_cons)
+apply (auto simp:same_inode_files_def is_file_unlink current_files_unlink
+ split:if_splits option.splits
+ dest:iof's_im_in_cim iof's_im_in_cim' inum_of_file_none_prop)
+done
+
+lemma same_inode_files_mkdir:
+ "valid (Mkdir p f inum # s) \<Longrightarrow> same_inode_files (Mkdir p f inum # s) = (same_inode_files s)"
+apply (frule vt_grant_os, frule vd_cons, rule ext)
+apply (auto simp:same_inode_files_def is_file_simps current_files_simps
+ split:if_splits option.splits
+ dest:iof's_im_in_cim iof's_im_in_cim' inum_of_file_none_prop is_file_in_current)
+apply (simp add:current_files_def is_file_def)
+done
+
+lemma same_inode_files_other:
+ "\<lbrakk>valid (e # s);
+ \<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p f. e \<noteq> UnLink p f;
+ \<forall> p f f'. e \<noteq> LinkHard p f f'\<rbrakk> \<Longrightarrow> same_inode_files (e # s) = same_inode_files s"
+apply (frule vt_grant_os, frule vd_cons, rule ext, case_tac e)
+apply (auto simp:same_inode_files_def is_file_simps current_files_simps dir_is_empty_def
+ split:if_splits option.splits
+ dest:iof's_im_in_cim iof's_im_in_cim' inum_of_file_none_prop is_file_not_dir)
+apply (simp add:is_file_def is_dir_def current_files_def split:option.splits t_inode_tag.splits)+
+done
+
+lemmas same_inode_files_simps = same_inode_files_nil same_inode_files_open same_inode_files_linkhard
+ same_inode_files_closefd same_inode_files_unlink same_inode_files_mkdir same_inode_files_other
+
+lemma same_inode_files_prop1:
+ "f \<in> same_inode_files s f' \<Longrightarrow> f \<in> current_files s"
+by (simp add:same_inode_files_def is_file_def current_files_def split:if_splits option.splits)
+
+lemma same_inode_files_prop2:
+ "\<lbrakk>f \<in> same_inode_files s f'; f'' \<notin> current_files s\<rbrakk> \<Longrightarrow> f \<noteq> f''"
+by (auto dest:same_inode_files_prop1)
+
+lemma same_inode_files_prop3:
+ "\<lbrakk>f \<in> same_inode_files s f'; is_dir s f''\<rbrakk> \<Longrightarrow> f \<noteq> f''"
+apply (rule notI)
+apply (simp add:same_inode_files_def is_file_def is_dir_def
+ split:if_splits option.splits t_inode_tag.splits)
+done
+
+lemma same_inode_files_prop4:
+ "\<lbrakk>f' \<in> same_inode_files s f; f'' \<in> same_inode_files s f'\<rbrakk> \<Longrightarrow> f'' \<in> same_inode_files s f"
+by (auto simp:same_inode_files_def split:if_splits)
+
+lemma same_inode_files_prop5:
+ "f' \<in> same_inode_files s f \<Longrightarrow> f \<in> same_inode_files s f'"
+by (auto simp:same_inode_files_def is_file_def split:if_splits)
+
+lemma same_inode_files_prop6:
+ "f' \<in> same_inode_files s f \<Longrightarrow> same_inode_files s f' = same_inode_files s f"
+by (auto simp:same_inode_files_def is_file_def split:if_splits)
+
+lemma same_inode_files_prop7:
+ "f' \<in> same_inode_files s f \<Longrightarrow> has_same_inode s f f'"
+by (auto simp:same_inode_files_def is_file_def has_same_inode_def split:if_splits option.splits)
+
+lemma same_inode_files_prop8:
+ "f' \<in> same_inode_files s f \<Longrightarrow> has_same_inode s f' f"
+by (auto simp:same_inode_files_def is_file_def has_same_inode_def split:if_splits option.splits)
+*)
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Current_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,105 @@
+(*<*)
+theory Current_prop
+imports Main Flask_type Flask My_list_prefix Init_prop Valid_prop Delete_prop
+begin
+(*>*)
+
+context flask begin
+
+lemma procs_of_shm_prop1: "\<lbrakk> p_flag \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> h \<in> current_shms s"
+apply (induct s arbitrary:p_flag)
+apply (case_tac p_flag, simp, drule init_procs_has_shm, simp)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a, auto split:if_splits option.splits)
+done
+
+lemma procs_of_shm_prop2: "\<lbrakk>(p, flag) \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s"
+apply (induct s arbitrary:p flag)
+apply (simp, drule init_procs_has_shm, simp)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a, auto split:if_splits option.splits)
+done
+
+lemma procs_of_shm_prop2':
+ "\<lbrakk>p \<notin> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<forall> flag h. (p, flag) \<notin> procs_of_shm s h"
+by (auto dest:procs_of_shm_prop2)
+
+lemma procs_of_shm_prop3: "\<lbrakk>(p, flag) \<in> procs_of_shm s h; (p, flag') \<in> procs_of_shm s h; valid s\<rbrakk>
+ \<Longrightarrow> flag = flag'"
+apply (induct s arbitrary:p flag flag')
+apply (simp, drule_tac flag = flag in init_procs_has_shm, drule_tac flag = flag' in init_procs_has_shm, simp)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a, auto split:if_splits option.splits dest:procs_of_shm_prop2)
+done
+
+lemma procs_of_shm_prop4: "\<lbrakk>(p, flag) \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> flag_of_proc_shm s p h = Some flag"
+apply (induct s arbitrary:p flag)
+apply (simp, drule init_procs_has_shm, simp)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a, auto split:if_splits option.splits dest:procs_of_shm_prop2)
+done
+
+lemma procs_of_shm_prop4':
+ "\<lbrakk>flag_of_proc_shm s p h = None; valid s\<rbrakk> \<Longrightarrow> \<forall> flag. (p, flag) \<notin> procs_of_shm s h"
+by (auto dest:procs_of_shm_prop4)
+
+lemma not_init_intro_proc:
+ "\<lbrakk>p \<notin> current_procs s; valid s\<rbrakk> \<Longrightarrow> deleted (O_proc p) s \<or> p \<notin> init_procs"
+using not_deleted_init_proc by auto
+
+lemma not_init_intro_proc':
+ "\<lbrakk>p \<notin> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<not> (\<not> deleted (O_proc p) s \<and> p \<in> init_procs)"
+using not_deleted_init_proc by auto
+
+lemma info_shm_flow_in_procs:
+ "\<lbrakk>info_flow_shm s p p'; valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s \<and> p' \<in> current_procs s"
+by (auto intro:procs_of_shm_prop2 simp:info_flow_shm_def one_flow_shm_def)
+
+lemma flag_of_proc_shm_prop1:
+ "\<lbrakk>flag_of_proc_shm s p h = Some flag; valid s\<rbrakk> \<Longrightarrow> (p, flag) \<in> procs_of_shm s h"
+apply (induct s arbitrary:p flag)
+apply (simp, drule init_shmflag_has_proc, simp)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a, auto split:if_splits option.splits dest:procs_of_shm_prop2)
+done
+
+
+
+lemma has_same_inode_in_current:
+ "\<lbrakk>has_same_inode s f f'; valid s\<rbrakk> \<Longrightarrow> f \<in> current_files s \<and> f' \<in> current_files s"
+by (auto simp add:has_same_inode_def current_files_def same_inode_files_def
+ is_file_def split:if_splits option.splits)
+
+
+lemma has_same_inode_prop1:
+ "\<lbrakk>has_same_inode s f f'; is_file s f; valid s\<rbrakk> \<Longrightarrow> is_file s f'"
+by (auto simp:has_same_inode_def same_inode_files_def is_file_def)
+
+lemma has_same_inode_prop1':
+ "\<lbrakk>has_same_inode s f f'; is_file s f'; valid s\<rbrakk> \<Longrightarrow> is_file s f"
+by (auto simp:has_same_inode_def is_file_def same_inode_files_def
+ split:if_splits option.splits)
+
+lemma has_same_inode_prop2:
+ "\<lbrakk>has_same_inode s f f'; file_of_proc_fd s p fd = Some f; valid s\<rbrakk> \<Longrightarrow> is_file s f'"
+apply (drule has_same_inode_prop1)
+apply (simp add:file_of_pfd_is_file, simp+)
+done
+
+lemma has_same_inode_prop2':
+ "\<lbrakk>has_same_inode s f f'; file_of_proc_fd s p fd = Some f'; valid s\<rbrakk> \<Longrightarrow> is_file s f"
+apply (drule has_same_inode_prop1')
+apply (simp add:file_of_pfd_is_file, simp+)
+done
+
+lemma tobj_in_init_alive:
+ "tobj_in_init obj \<Longrightarrow> init_alive obj"
+by (case_tac obj, auto)
+
+lemma tobj_in_alive:
+ "tobj_in_init obj \<Longrightarrow> alive [] obj"
+by (case_tac obj, auto simp:is_file_nil)
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Current_sockets_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,380 @@
+theory Current_sockets_prop
+imports Main Flask Flask_type Current_files_prop
+begin
+
+context flask begin
+
+lemma cn_in_curp: "\<lbrakk>(p, fd) \<in> current_sockets \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> p \<in> current_procs \<tau>"
+apply (induct \<tau>) defer
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:current_sockets_def inum_of_socket.simps current_procs.simps os_grant.simps split:if_splits option.splits
+ dest:inos_has_sock_tag init_socket_has_inode)
+done
+
+lemma cn_in_curfd_aux: "\<forall> p. (p, fd) \<in> current_sockets \<tau> \<and> valid \<tau> \<longrightarrow> fd \<in> current_proc_fds \<tau> p"
+apply (induct \<tau>) defer
+apply (rule allI, rule impI, erule conjE, frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:current_sockets_def current_proc_fds.simps os_grant.simps inum_of_socket.simps
+ split:if_splits option.splits t_socket_af.splits dest:inos_has_sock_tag init_socket_has_inode)
+done
+
+lemma cn_in_curfd: "\<lbrakk>(p, fd) \<in> current_sockets \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> fd \<in> current_proc_fds \<tau> p"
+by (simp add:cn_in_curfd_aux)
+
+lemma cn_in_curfd': "\<lbrakk>inum_of_socket \<tau> (p, fd) = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> fd \<in> current_proc_fds \<tau> p"
+by (rule cn_in_curfd, simp_all add:current_sockets_def)
+
+lemma cn_in_curp': "\<lbrakk>inum_of_socket \<tau> (p, fd) = Some im; valid \<tau>\<rbrakk> \<Longrightarrow> p \<in> current_procs \<tau>"
+apply (rule cn_in_curp)
+by (auto simp:current_sockets_def)
+
+(***************** current_sockets simpset *****************)
+
+lemma current_sockets_nil: "current_sockets [] = init_sockets"
+apply (simp add:current_sockets_def inum_of_socket.simps)
+using init_socket_has_inode inos_has_sock_tag
+by auto
+
+lemma current_sockets_closefd: "current_sockets (CloseFd p fd # \<tau>) = current_sockets \<tau> - {(p, fd)}"
+by (auto simp add:current_sockets_def inum_of_socket.simps)
+
+lemma current_sockets_createsock: "current_sockets (CreateSock p aft st fd ino # \<tau>) = insert (p, fd) (current_sockets \<tau>)"
+by (auto simp add:current_sockets_def inum_of_socket.simps)
+
+lemma current_sockets_accept: "current_sockets (Accept p fd addr lport' fd' ino # \<tau>) = insert (p, fd') (current_sockets \<tau>)"
+by (auto simp add:current_sockets_def inum_of_socket.simps)
+
+lemma current_sockets_clone: "valid (Clone p p' fds # \<tau>)
+ \<Longrightarrow> current_sockets (Clone p p' fds # \<tau>) =
+ current_sockets \<tau> \<union> {(p', fd)| fd. (p, fd) \<in> current_sockets \<tau> \<and> fd \<in> fds}"
+apply (frule vd_cons, drule vt_grant_os)
+apply (auto simp add:os_grant.simps current_sockets_def inum_of_socket.simps split:if_splits
+ dest:cn_in_curp')
+done
+
+lemma current_sockets_execve: "valid (Execve p f fds # \<tau>)
+ \<Longrightarrow> current_sockets (Execve p f fds # \<tau>) =
+ current_sockets \<tau> - {(p, fd)| fd. (p, fd) \<in> current_sockets \<tau> \<and> fd \<notin> fds}"
+apply (frule vd_cons, drule vt_grant_os)
+apply (auto simp add:os_grant.simps current_sockets_def inum_of_socket.simps split:if_splits
+ dest:cn_in_curp')
+done
+
+lemma current_sockets_kill: "current_sockets (Kill p p' # \<tau>) = current_sockets \<tau> - {(p', fd)| fd. (p', fd) \<in> current_sockets \<tau>}"
+by (auto simp add:current_sockets_def inum_of_socket.simps split:if_splits)
+
+lemma current_sockets_exit: "current_sockets (Exit p # \<tau>) = current_sockets \<tau> - {(p, fd)| fd. (p, fd) \<in> current_sockets \<tau>}"
+by (auto simp add:current_sockets_def inum_of_socket.simps split:if_splits)
+
+lemma current_sockets_other:
+ "\<lbrakk>\<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p aft st fd ino. e \<noteq> CreateSock p aft st fd ino;
+ \<forall> p fd addr lport' fd' ino. e \<noteq> Accept p fd addr lport' fd' ino;
+ \<forall> p p' fds. e \<noteq> Clone p p' fds;
+ \<forall> p f fds. e \<noteq> Execve p f fds;
+ \<forall> p p'. e \<noteq> Kill p p';
+ \<forall> p. e \<noteq> Exit p\<rbrakk> \<Longrightarrow> current_sockets (e # \<tau>) = current_sockets \<tau>"
+apply (case_tac e)
+by (auto simp add:current_sockets_def inum_of_socket.simps split:if_splits)
+
+lemmas current_sockets_simps = current_sockets_nil current_sockets_closefd current_sockets_createsock
+ current_sockets_accept current_sockets_execve current_sockets_clone
+ current_sockets_kill current_sockets_exit current_sockets_other
+
+(********** is_tcp_sock simpset ***************)
+
+lemma is_tcp_sock_nil:
+ "is_tcp_sock [] = is_init_tcp_sock"
+apply (rule ext)
+apply (auto simp add:is_init_tcp_sock_def is_tcp_sock_def init_inum_of_socket_props
+ dest:init_socket_has_inode split:option.splits)
+done
+
+lemma is_tcp_sock_closefd:
+ "is_tcp_sock (CloseFd p fd # s) = (is_tcp_sock s) ((p, fd) := False)"
+apply (rule ext)
+by (auto simp add:is_tcp_sock_def inum_of_socket.simps split:option.splits t_inode_tag.splits)
+
+lemma is_tcp_sock_createsock:
+ "valid (CreateSock p af st fd ino # s) \<Longrightarrow> is_tcp_sock (CreateSock p af st fd ino # s) = (
+ case st of
+ STREAM \<Rightarrow> (is_tcp_sock s) ((p, fd) := True)
+ | _ \<Rightarrow> is_tcp_sock s )"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemma is_tcp_sock_accept:
+ "valid (Accept p fd addr lport fd' ino # s)
+ \<Longrightarrow> is_tcp_sock (Accept p fd addr lport fd' ino # s) = (is_tcp_sock s) ((p,fd') := True)"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemma is_tcp_sock_execve:
+ "valid (Execve p f fds # s) \<Longrightarrow> is_tcp_sock (Execve p f fds # s) = (
+ \<lambda> (p', fd). if (p' = p \<and> fd \<in> fds) then is_tcp_sock s (p, fd)
+ else if (p' = p) then False
+ else is_tcp_sock s (p', fd))"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemma is_tcp_sock_clone:
+ "valid (Clone p p' fds # s) \<Longrightarrow> is_tcp_sock (Clone p p' fds # s) = (
+ \<lambda> (p'', fd). if (p'' = p' \<and> fd \<in> fds) then is_tcp_sock s (p, fd)
+ else if (p'' = p') then False
+ else is_tcp_sock s (p'', fd))"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemma is_tcp_sock_kill:
+ "valid (Kill p p' # s) \<Longrightarrow> is_tcp_sock (Kill p p' # s) = (
+ \<lambda> (p'', fd). if (p'' = p') then False
+ else is_tcp_sock s (p'', fd))"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemma is_tcp_sock_exit:
+ "valid (Exit p # s) \<Longrightarrow> is_tcp_sock (Exit p # s) = (
+ \<lambda> (p', fd). if (p' = p) then False
+ else is_tcp_sock s (p', fd))"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemma is_tcp_sock_other:
+ "\<lbrakk>valid (e # \<tau>);
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p aft st fd ino. e \<noteq> CreateSock p aft st fd ino;
+ \<forall> p fd addr lport' fd' ino. e \<noteq> Accept p fd addr lport' fd' ino;
+ \<forall> p p' fds. e \<noteq> Clone p p' fds;
+ \<forall> p f fds. e \<noteq> Execve p f fds;
+ \<forall> p p'. e \<noteq> Kill p p';
+ \<forall> p. e \<noteq> Exit p\<rbrakk> \<Longrightarrow> is_tcp_sock (e # \<tau>) = is_tcp_sock \<tau>"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x, case_tac e)
+prefer 6 apply (case_tac option)
+by (auto simp add:is_tcp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemmas is_tcp_sock_simps = is_tcp_sock_nil is_tcp_sock_closefd is_tcp_sock_createsock is_tcp_sock_accept
+ is_tcp_sock_clone is_tcp_sock_execve is_tcp_sock_kill is_tcp_sock_exit is_tcp_sock_other
+
+
+(************ is_udp_sock simpset *************)
+
+lemma is_udp_sock_nil:
+ "is_udp_sock [] = is_init_udp_sock"
+apply (rule ext)
+apply (auto simp add:is_init_udp_sock_def is_udp_sock_def init_inum_of_socket_props
+ dest:init_socket_has_inode split:option.splits)
+done
+
+lemma is_udp_sock_closefd:
+ "is_udp_sock (CloseFd p fd # s) = (is_udp_sock s) ((p, fd) := False)"
+apply (rule ext)
+by (auto simp add:is_udp_sock_def inum_of_socket.simps split:option.splits t_inode_tag.splits)
+
+lemma is_udp_sock_createsock:
+ "valid (CreateSock p af st fd ino # s) \<Longrightarrow> is_udp_sock (CreateSock p af st fd ino # s) = (
+ case st of
+ DGRAM \<Rightarrow> (is_udp_sock s) ((p, fd) := True)
+ | _ \<Rightarrow> is_udp_sock s )"
+apply (frule vd_cons, frule vt_grant_os, rule ext)
+by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemma is_udp_sock_execve:
+ "valid (Execve p f fds # s) \<Longrightarrow> is_udp_sock (Execve p f fds # s) = (
+ \<lambda> (p', fd). if (p' = p \<and> fd \<in> fds) then is_udp_sock s (p, fd)
+ else if (p' = p) then False
+ else is_udp_sock s (p', fd))"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemma is_udp_sock_clone:
+ "valid (Clone p p' fds # s) \<Longrightarrow> is_udp_sock (Clone p p' fds # s) = (
+ \<lambda> (p'', fd). if (p'' = p' \<and> fd \<in> fds) then is_udp_sock s (p, fd)
+ else if (p'' = p') then False
+ else is_udp_sock s (p'', fd))"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemma is_udp_sock_kill:
+ "valid (Kill p p' # s) \<Longrightarrow> is_udp_sock (Kill p p' # s) = (
+ \<lambda> (p'', fd). if (p'' = p') then False
+ else is_udp_sock s (p'', fd))"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemma is_udp_sock_exit:
+ "valid (Exit p # s) \<Longrightarrow> is_udp_sock (Exit p # s) = (
+ \<lambda> (p', fd). if (p' = p) then False
+ else is_udp_sock s (p', fd))"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x)
+by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemma is_udp_sock_other:
+ "\<lbrakk>valid (e # \<tau>);
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p aft st fd ino. e \<noteq> CreateSock p aft st fd ino;
+ \<forall> p p' fds. e \<noteq> Clone p p' fds;
+ \<forall> p f fds. e \<noteq> Execve p f fds;
+ \<forall> p p'. e \<noteq> Kill p p';
+ \<forall> p. e \<noteq> Exit p\<rbrakk> \<Longrightarrow> is_udp_sock (e # \<tau>) = is_udp_sock \<tau>"
+apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac x, case_tac e)
+prefer 6 apply (case_tac option)
+by (auto simp add:is_udp_sock_def inum_of_socket.simps dest:ios's_im_in_cim cn_in_curfd'
+ split:option.splits t_inode_tag.splits t_socket_type.splits)
+
+lemmas is_udp_sock_simps = is_udp_sock_nil is_udp_sock_closefd is_udp_sock_createsock
+ is_udp_sock_clone is_udp_sock_execve is_udp_sock_kill is_udp_sock_exit is_udp_sock_other
+
+lemma update_with_shuthow_someI: "socket_in_trans \<tau> s \<Longrightarrow> \<exists> st'. update_with_shuthow (socket_state \<tau> s) how = Some st'"
+apply (case_tac "socket_state \<tau> s", simp add:socket_in_trans_def, case_tac a)
+apply (auto simp:update_with_shuthow.simps socket_in_trans_def split:option.splits)
+apply (case_tac t_sock_trans_state, case_tac [!] how)
+by auto
+
+lemma is_tcp_sock_imp_curernt_proc:
+ "\<lbrakk>is_tcp_sock s (p,fd); valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s"
+apply (induct s)
+apply (simp add:is_tcp_sock_nil is_init_tcp_sock_prop3)
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+by (auto simp:is_tcp_sock_simps split:if_splits option.splits t_socket_type.splits)
+
+lemma is_udp_sock_imp_curernt_proc:
+ "\<lbrakk>is_udp_sock s (p,fd); valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s"
+apply (induct s)
+apply (simp add:is_udp_sock_nil is_init_udp_sock_prop3)
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+by (auto simp:is_udp_sock_simps split:if_splits option.splits t_socket_type.splits)
+
+
+(***** below should be adjusted after starting static analysis of sockets ! ??? ****)
+
+(*
+lemma cn_has_sockstate: "\<lbrakk>(p, fd) \<in> current_sockets \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> socket_state \<tau> (p, fd) \<noteq> None"
+apply (induct \<tau> arbitrary: p) using init_socket_has_state
+apply (simp add:current_sockets_simps socket_state.simps bidirect_in_init_def)
+
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:current_sockets_simps os_grant.simps socket_state.simps intro:update_with_shuthow_someI dest:cn_no_newproc' split:option.splits)
+done
+
+lemma after_bind_update: "\<lbrakk>after_bind (update_with_shuthow (socket_state \<tau> s) t_shutdown_how); valid \<tau>\<rbrakk> \<Longrightarrow> after_bind (socket_state \<tau> s)"
+apply (case_tac "socket_state \<tau> s", simp, case_tac a)
+apply (auto simp:update_with_shuthow.simps split:option.splits)
+done
+
+lemma after_bind_has_laddr_aux: "\<forall> p fd. after_bind (socket_state \<tau> (p, fd)) \<and> valid \<tau> \<longrightarrow> (\<exists> ip port. local_addr \<tau> (p, fd) = Some (INET_ADDR ip port))"
+apply (induct \<tau>)
+apply (clarsimp simp add:socket_state.simps local_addr.simps)
+apply (drule init_socket_has_laddr, simp, erule exE)
+apply (case_tac y, simp)
+
+apply (clarify, frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:socket_state.simps local_addr.simps os_grant.simps dest:after_bind_update split:if_splits option.splits t_sock_addr.splits)
+apply (erule_tac x = nat1 in allE, erule_tac x = nat2 in allE, erule impE, simp+)
+done
+
+lemma after_bind_has_laddr: "\<lbrakk>after_bind (socket_state \<tau> (p, fd)); valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> ip port. local_addr \<tau> (p, fd) = Some (INET_ADDR ip port)"
+by (rule after_bind_has_laddr_aux[rule_format], simp+)
+
+lemma after_bind_has_laddr': "\<lbrakk>after_bind (socket_state \<tau> s); valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> ip port. local_addr \<tau> s = Some (INET_ADDR ip port)"
+by (case_tac s, auto dest:after_bind_has_laddr)
+
+lemma after_bind_has_laddr_D: "\<lbrakk>local_addr \<tau> s = None; valid \<tau>\<rbrakk> \<Longrightarrow> \<not> after_bind (socket_state \<tau> s)"
+by (auto dest:after_bind_has_laddr')
+
+lemma local_addr_valid: "\<lbrakk>local_addr \<tau> (p, fd) = Some (INET_ADDR ip port); (p, fd) \<in> current_sockets \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> ip \<in> local_ipaddrs"
+apply (induct \<tau> arbitrary:p fd port)
+apply (simp add:local_addr.simps current_sockets_simps local_ip_valid)
+
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:current_sockets_simps local_addr.simps os_grant.simps split:if_splits option.splits t_sock_addr.splits dest:cn_no_newproc' after_bind_has_laddr_D)
+done
+
+lemma local_addr_valid': "\<lbrakk>local_addr \<tau> s = Some (INET_ADDR ip port); s \<in> current_sockets \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> ip \<in> local_ipaddrs"
+by (case_tac s, simp add:local_addr_valid)
+
+lemma local_ip_has_dev_D: "\<lbrakk>local_ip2dev ip = None; local_addr \<tau> s = Some (INET_ADDR ip port); s \<in> current_sockets \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+using local_ip2dev_valid
+by (drule_tac local_addr_valid', simp+)
+
+lemma after_bind_has_netdev: "\<lbrakk>after_bind (socket_state \<tau> s); s \<in> current_sockets \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> dev. netdev_of_socket \<tau> s = Some dev"
+apply (drule after_bind_has_laddr', simp, clarsimp)
+apply (frule local_addr_valid', simp+)
+apply (drule local_ip2dev_valid, simp add:netdev_of_socket_def)
+done
+
+lemma after_bind_has_netdev_D: "\<lbrakk>netdev_of_socket \<tau> s = None; s \<in> current_sockets \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> \<not> after_bind (socket_state \<tau> s)"
+by (auto dest:after_bind_has_netdev)
+
+lemma after_conacc_update: "\<lbrakk>after_conacc (update_with_shuthow (socket_state \<tau> s) t_shutdown_how); valid \<tau>\<rbrakk> \<Longrightarrow> after_conacc (socket_state \<tau> s)"
+apply (case_tac "socket_state \<tau> s", simp, case_tac a)
+apply (auto simp:update_with_shuthow.simps split:option.splits)
+done
+
+lemma after_conacc_has_raddr: "\<lbrakk>after_conacc (socket_state \<tau> (p, fd)); valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> addr. remote_addr \<tau> (p, fd) = Some addr"
+apply (induct \<tau> arbitrary:p fd)
+apply (simp add:socket_state.simps local_addr.simps) using init_socket_has_raddr apply blast
+
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:socket_state.simps local_addr.simps os_grant.simps dest:after_conacc_update split:if_splits option.splits t_sock_addr.splits)
+done
+
+lemma after_conacc_has_raddr': "\<lbrakk>after_conacc (socket_state \<tau> s); valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> addr. remote_addr \<tau> s = Some addr"
+by (case_tac s, auto dest:after_conacc_has_raddr)
+
+lemma after_conacc_has_raddr_D: "\<lbrakk>remote_addr \<tau> s = None; valid \<tau>\<rbrakk> \<Longrightarrow> \<not> after_conacc (socket_state \<tau> s)"
+by (auto dest:after_conacc_has_raddr')
+
+
+lemma cn_has_socktype: "\<lbrakk>(p, fd) \<in> current_sockets \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> \<exists> st. socket_type \<tau> (p, fd) = Some st"
+apply (induct \<tau> arbitrary:p fd)
+using init_socket_has_socktype
+apply (simp add:socket_type.simps current_sockets_nil bidirect_in_init_def)
+
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:socket_type.simps current_sockets_simps os_grant.simps dest:cn_no_newproc' split:if_splits option.splits t_sock_addr.splits)
+done
+
+lemma accept_has_temps: "Accept p fd\<^isub>1 addr port fd\<^isub>2 inum # valid \<tau> \<Longrightarrow> \<exists> temp. nettemp_of_socket (Accept p fd\<^isub>1 addr port fd\<^isub>2 inum # \<tau>) (p, fd\<^isub>2) is_remote = Some temp"
+apply (frule vd_cons, frule vt_grant_os)
+apply (auto dest:cn_has_socktype simp:os_grant.simps nettemp_of_socket_def netdev_of_socket_def
+ dest!:after_conacc_has_raddr_D after_bind_has_laddr_D local_ip_has_dev_D
+ simp:socket_state.simps split:if_splits option.splits t_sock_addr.splits)
+done
+
+lemma accept_has_temps':
+ "Accept p fd\<^isub>1 addr port fd\<^isub>2 inum # valid \<tau>
+ \<Longrightarrow> (\<exists> ltemp. nettemp_of_socket (Accept p fd\<^isub>1 addr port fd\<^isub>2 inum # \<tau>) (p, fd\<^isub>2) False = Some ltemp) \<and>
+ (\<exists> rtemp. nettemp_of_socket (Accept p fd\<^isub>1 addr port fd\<^isub>2 inum # \<tau>) (p, fd\<^isub>2) True = Some rtemp)"
+apply (frule accept_has_temps[where is_remote = True])
+apply (frule accept_has_temps[where is_remote = False])
+by auto
+
+lemma bind_has_ltemp: "Bind p fd addr # valid \<tau> \<Longrightarrow> \<exists> temp. nettemp_of_socket (Bind p fd addr # \<tau>) (p, fd) False = Some temp"
+apply (frule vd_cons, frule vt_grant_os)
+apply (auto dest:cn_has_socktype simp:os_grant.simps current_sockets_simps nettemp_of_socket_def dest!:after_bind_has_netdev_D after_bind_has_laddr_D
+ simp:socket_state.simps split:if_splits option.splits t_sock_addr.splits)
+done
+
+lemma connect_has_rtemp: "Connect p fd addr # valid \<tau> \<Longrightarrow> \<exists> temp. nettemp_of_socket (Connect p fd addr # \<tau>) (p, fd) True = Some temp"
+apply (frule vd_cons, frule vt_grant_os)
+apply (auto dest:cn_has_socktype simp:os_grant.simps current_sockets_simps nettemp_of_socket_def dest!:after_bind_has_netdev_D after_bind_has_laddr_D
+ simp:socket_state.simps split:if_splits option.splits t_sock_addr.splits)
+done
+*)
+
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Delete_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,248 @@
+theory Delete_prop
+imports Main Flask Flask_type Init_prop Alive_prop Current_files_prop
+begin
+
+context flask begin
+
+lemma deleted_cons_I: "deleted obj s \<Longrightarrow> deleted obj (e # s)"
+by (case_tac e, auto)
+
+lemma not_deleted_cons_D: "\<not> deleted obj (e # s) \<Longrightarrow> \<not> deleted obj s"
+by (auto dest:deleted_cons_I)
+
+lemma cons_app_simp_aux:
+ "(a # b) @ c = a # (b @ c)" by auto
+
+lemma not_deleted_app_I:
+ "deleted obj s \<Longrightarrow> deleted obj (s' @ s)"
+apply (induct s', simp)
+by (simp add:cons_app_simp_aux deleted_cons_I)
+
+lemma not_deleted_app_D:
+ "\<not> deleted obj (s' @ s) \<Longrightarrow> \<not> deleted obj s"
+apply (rule notI)
+by (simp add:not_deleted_app_I)
+
+lemma not_deleted_init_file:
+ "\<lbrakk>\<not> deleted (O_file f) s; valid s; is_init_file f\<rbrakk> \<Longrightarrow> is_file s f"
+apply (induct s, simp add:is_file_nil)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a) prefer 6 apply (case_tac option)
+apply (auto simp:is_file_simps split:option.splits)
+done
+
+lemma not_deleted_init_dir:
+ "\<lbrakk>\<not> deleted (O_dir f) s; valid s; is_init_dir f\<rbrakk> \<Longrightarrow> is_dir s f"
+apply (induct s, simp add:is_dir_nil)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a) prefer 6 apply (case_tac option)
+apply (auto simp:is_dir_simps split:option.splits)
+done
+
+lemma not_deleted_init_proc:
+ "\<lbrakk>\<not> deleted (O_proc p) s; p \<in> init_procs\<rbrakk> \<Longrightarrow> p \<in> current_procs s"
+apply (induct s, simp)
+by (case_tac a, auto)
+
+(**** ???*)
+lemma current_fd_imp_current_proc:
+ "\<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 vd_cons, drule vt_grant_os, case_tac a)
+by (auto split:if_splits option.splits)
+
+lemma not_deleted_init_fd_aux:
+ "\<lbrakk>\<not> deleted (O_fd p fd) s; valid s; fd \<in> init_fds_of_proc p\<rbrakk>
+ \<Longrightarrow> fd \<in> current_proc_fds s p \<and> \<not> deleted (O_proc p) s"
+apply (induct s arbitrary: p, simp)
+apply (frule vd_cons, drule vt_grant_os)
+apply (case_tac a, auto dest:current_fd_imp_current_proc)
+done
+
+lemma not_deleted_init_fd2:
+ "\<lbrakk>\<not> deleted (O_fd p fd) s; fd \<in> init_fds_of_proc p; valid s\<rbrakk> \<Longrightarrow> \<not> deleted (O_proc p) s"
+by (auto dest:not_deleted_init_fd_aux)
+
+lemma not_deleted_init_fd1:
+ "\<lbrakk>\<not> deleted (O_fd p fd) s; valid s; fd \<in> init_fds_of_proc p\<rbrakk> \<Longrightarrow> fd \<in> current_proc_fds s p"
+by (auto dest:not_deleted_init_fd_aux)
+
+lemma not_deleted_init_tcp_aux:
+ "\<lbrakk>\<not> deleted (O_tcp_sock (p,fd)) s; valid s; is_init_tcp_sock (p,fd)\<rbrakk>
+ \<Longrightarrow> is_tcp_sock s (p,fd) \<and> \<not> deleted (O_proc p) s"
+apply (induct s arbitrary:p, simp add:is_tcp_sock_nil)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a) prefer 6 apply (case_tac option)
+by (auto simp:is_tcp_sock_simps split:option.splits t_socket_type.splits
+ dest:is_tcp_sock_imp_curernt_proc)
+
+lemma not_deleted_init_tcp1:
+ "\<lbrakk>\<not> deleted (O_tcp_sock (p,fd)) s; valid s; is_init_tcp_sock (p,fd)\<rbrakk>
+ \<Longrightarrow> is_tcp_sock s (p,fd)"
+by (auto dest:not_deleted_init_tcp_aux)
+
+lemma not_deleted_init_tcp2:
+ "\<lbrakk>\<not> deleted (O_tcp_sock (p,fd)) s; valid s; is_init_tcp_sock (p,fd)\<rbrakk>
+ \<Longrightarrow> \<not> deleted (O_proc p) s"
+by (auto dest:not_deleted_init_tcp_aux)
+
+lemma not_deleted_init_udp_aux:
+ "\<lbrakk>\<not> deleted (O_udp_sock (p,fd)) s; valid s; is_init_udp_sock (p,fd)\<rbrakk>
+ \<Longrightarrow> is_udp_sock s (p,fd) \<and> \<not> deleted (O_proc p) s"
+apply (induct s arbitrary:p, simp add:is_udp_sock_nil)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a) prefer 6 apply (case_tac option)
+by (auto simp:is_udp_sock_simps split:option.splits t_socket_type.splits
+ dest:is_udp_sock_imp_curernt_proc)
+
+lemma not_deleted_init_udp1:
+ "\<lbrakk>\<not> deleted (O_udp_sock (p,fd)) s; valid s; is_init_udp_sock (p,fd)\<rbrakk>
+ \<Longrightarrow> is_udp_sock s (p,fd)"
+by (auto dest:not_deleted_init_udp_aux)
+
+lemma not_deleted_init_udp2:
+ "\<lbrakk>\<not> deleted (O_udp_sock (p,fd)) s; valid s; is_init_udp_sock (p,fd)\<rbrakk>
+ \<Longrightarrow> \<not> deleted (O_proc p) s"
+by (auto dest:not_deleted_init_udp_aux)
+
+lemma not_deleted_init_shm:
+ "\<lbrakk>\<not> deleted (O_shm h) s; h \<in> init_shms\<rbrakk> \<Longrightarrow> h \<in> current_shms s"
+apply (induct s, simp)
+by (case_tac a, auto)
+
+lemma not_deleted_init_msgq:
+ "\<lbrakk>\<not> deleted (O_msgq q) s; q \<in> init_msgqs\<rbrakk> \<Longrightarrow> q \<in> current_msgqs s"
+apply (induct s, simp)
+by (case_tac a, auto)
+
+lemma current_msg_imp_current_msgq:
+ "\<lbrakk>m \<in> set (msgs_of_queue s q); valid s\<rbrakk> \<Longrightarrow> q \<in> current_msgqs s"
+apply (induct s)
+apply (simp add:init_msgs_valid)
+apply (frule vd_cons, drule vt_grant_os)
+apply (case_tac a, auto split:if_splits)
+done
+
+lemma not_deleted_init_msg:
+ "\<lbrakk>\<not> deleted (O_msg q m) s; valid s; m \<in> set (init_msgs_of_queue q)\<rbrakk> \<Longrightarrow> m \<in> set (msgs_of_queue s q)"
+apply (induct s, simp)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a, auto dest:current_msg_imp_current_msgq)
+apply (case_tac "msgs_of_queue s q", simp+)
+done
+
+lemma not_deleted_imp_alive:
+ "\<lbrakk>\<not> deleted obj s; valid s; init_alive obj\<rbrakk> \<Longrightarrow> alive s obj"
+apply (case_tac obj)
+apply (auto dest!:not_deleted_init_msg not_deleted_init_file not_deleted_init_dir not_deleted_init_proc
+ not_deleted_init_fd1 not_deleted_init_tcp1 not_deleted_init_udp1 not_deleted_init_shm
+ not_deleted_init_msgq
+ intro:is_file_in_current is_dir_in_current is_tcp_in_current is_udp_in_current
+ current_msg_imp_current_msgq)
+done
+
+lemma not_deleted_cur_file_app:
+ "\<lbrakk>\<not> deleted (O_file f) (s' @ s); valid (s' @ s); is_file s f\<rbrakk> \<Longrightarrow> is_file (s' @ s) f"
+apply (induct s', simp, simp add:cons_app_simp_aux)
+apply (frule vd_cons, frule vt_grant_os, simp)
+apply (case_tac a) prefer 6 apply (case_tac option)
+apply (auto simp:is_file_simps split:option.splits)
+done
+
+lemma not_deleted_cur_dir_app:
+ "\<lbrakk>\<not> deleted (O_dir f) (s' @ s); valid (s' @ s); is_dir s f\<rbrakk> \<Longrightarrow> is_dir (s' @ s) f"
+apply (induct s', simp, simp add:cons_app_simp_aux)
+apply (frule vd_cons, frule vt_grant_os, simp)
+apply (case_tac a) prefer 6 apply (case_tac option)
+apply (auto simp:is_dir_simps split:option.splits)
+done
+
+lemma not_deleted_cur_proc_app:
+ "\<lbrakk>\<not> deleted (O_proc p) (s' @ s); p \<in> current_procs s\<rbrakk> \<Longrightarrow> p \<in> current_procs (s' @ s)"
+apply (induct s', simp, simp add:cons_app_simp_aux)
+by (case_tac a, auto)
+
+lemma not_deleted_cur_fd_app:
+ "\<lbrakk>\<not> deleted (O_fd p fd) (s' @ s); valid (s' @ s); fd \<in> current_proc_fds s p\<rbrakk>
+ \<Longrightarrow> fd \<in> current_proc_fds (s' @ s) p"
+apply (induct s' arbitrary: p, simp, simp add:cons_app_simp_aux)
+apply (frule vd_cons, drule vt_grant_os)
+apply (case_tac a, auto dest:current_fd_imp_current_proc)
+done
+
+lemma not_deleted_cur_tcp_app:
+ "\<lbrakk>\<not> deleted (O_tcp_sock (p,fd)) (s' @ s); valid (s' @ s); is_tcp_sock s (p,fd)\<rbrakk>
+ \<Longrightarrow> is_tcp_sock (s' @ s) (p,fd)"
+apply (induct s' arbitrary:p, simp, simp add:cons_app_simp_aux)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a) prefer 6 apply (case_tac option)
+by (auto simp:is_tcp_sock_simps split:option.splits t_socket_type.splits
+ dest:is_tcp_sock_imp_curernt_proc)
+
+lemma not_deleted_cur_udp_app:
+ "\<lbrakk>\<not> deleted (O_udp_sock (p,fd)) (s' @ s); valid (s' @ s); is_udp_sock s (p,fd)\<rbrakk>
+ \<Longrightarrow> is_udp_sock (s' @ s) (p,fd)"
+apply (induct s' arbitrary:p, simp, simp add:cons_app_simp_aux)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a) prefer 6 apply (case_tac option)
+by (auto simp:is_udp_sock_simps split:option.splits t_socket_type.splits
+ dest:is_udp_sock_imp_curernt_proc)
+
+lemma not_deleted_cur_shm_app:
+ "\<lbrakk>\<not> deleted (O_shm h) (s' @ s); h \<in> current_shms s\<rbrakk> \<Longrightarrow> h \<in> current_shms (s' @ s)"
+apply (induct s', simp, simp add:cons_app_simp_aux)
+by (case_tac a, auto)
+
+lemma not_deleted_cur_msgq_app:
+ "\<lbrakk>\<not> deleted (O_msgq q) (s' @ s); q \<in> current_msgqs s\<rbrakk> \<Longrightarrow> q \<in> current_msgqs (s' @ s)"
+apply (induct s', simp, simp add:cons_app_simp_aux)
+by (case_tac a, auto)
+
+lemma not_deleted_cur_msg_app:
+ "\<lbrakk>\<not> deleted (O_msg q m) (s' @ s); valid (s' @ s); m \<in> set (msgs_of_queue s q)\<rbrakk>
+ \<Longrightarrow> m \<in> set (msgs_of_queue (s' @ s) q)"
+apply (induct s', simp, simp add:cons_app_simp_aux)
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac a, auto dest:current_msg_imp_current_msgq)
+apply (case_tac "msgs_of_queue (s' @ s) q", simp+)
+done
+
+lemma not_deleted_imp_alive_app:
+ "\<lbrakk>\<not> deleted obj (s' @ s); valid (s' @ s); alive s obj\<rbrakk> \<Longrightarrow> alive (s' @ s) obj"
+apply (case_tac obj)
+apply (auto dest!:not_deleted_cur_msg_app not_deleted_cur_file_app not_deleted_cur_dir_app
+ not_deleted_cur_proc_app not_deleted_cur_fd_app not_deleted_cur_tcp_app
+ not_deleted_cur_udp_app not_deleted_cur_shm_app not_deleted_cur_msgq_app
+ intro:is_file_in_current is_dir_in_current is_tcp_in_current is_udp_in_current
+ current_msg_imp_current_msgq)
+done
+
+lemma not_deleted_imp_alive_cons:
+ "\<lbrakk>\<not> deleted obj (e # s); valid (e # s); alive s obj\<rbrakk> \<Longrightarrow> alive (e # s) obj"
+using not_deleted_imp_alive_app[where s = s and s' = "[e]" and obj = obj]
+by auto
+
+(*
+
+lemma nodel_imp_un_deleted:
+ "no_del_event s \<Longrightarrow> \<not> deleted obj s"
+by (induct s, simp, case_tac a,auto)
+
+lemma nodel_exists_remains:
+ "\<lbrakk>no_del_event (s'@s); exists s obj\<rbrakk> \<Longrightarrow> exists (s'@s) obj"
+apply (drule_tac obj = obj in nodel_imp_un_deleted)
+by (simp add:not_deleted_imp_exists')
+
+lemma nodel_imp_exists:
+ "\<lbrakk>no_del_event s; exists [] obj\<rbrakk> \<Longrightarrow> exists s obj"
+apply (drule_tac obj = obj in nodel_imp_un_deleted)
+by (simp add:not_deleted_imp_exists)
+
+lemma no_del_event_cons_D:
+ "no_del_event (e # s) \<Longrightarrow> no_del_event s"
+by (case_tac e, auto)
+*)
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Dynamic_static.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,734 @@
+theory Dynamic_static
+imports Main Flask Static Init_prop Valid_prop Tainted_prop Delete_prop Co2sobj_prop S2ss_prop S2ss_prop2
+ Temp
+begin
+
+context tainting_s begin
+
+fun remove_create_flag :: "t_open_flags \<Rightarrow> t_open_flags"
+where
+ "remove_create_flag (mflag, oflags) = (mflag, oflags - {OF_CREAT})"
+
+fun all_procs :: "t_state \<Rightarrow> t_process set"
+where
+ "all_procs [] = init_procs"
+| "all_procs (Clone p p' fds shms # s) = insert p' (all_procs s)"
+| "all_procs (e # s) = all_procs s"
+
+definition brandnew_proc :: "t_state \<Rightarrow> t_process"
+where
+ "brandnew_proc s \<equiv> next_nat (all_procs s)"
+
+(*
+definition brandnew_proc :: "t_state \<Rightarrow> t_process"
+where
+ "brandnew_proc s \<equiv> next_nat ({p | p s'. p \<in> current_procs s' \<and> s' \<preceq> s})"
+ another approach:
+ brandnew_proc = next_nat (all_procs s),
+ where all_procs is a event-trace listener *)
+
+(*
+lemma brandnew_proc_prop1:
+ "\<lbrakk>s' \<preceq> s; valid s\<rbrakk> \<Longrightarrow> brandnew_proc s \<notin> current_procs s'"
+apply (frule vd_preceq, simp)
+apply (simp add:brandnew_proc_def)
+apply (auto)
+sorry
+
+lemma brandnew_proc_prop2:
+ "\<lbrakk>p \<in> current_procs s'; s' \<preceq> s; valid s\<rbrakk> \<Longrightarrow> brandnew_proc s \<noteq> p"
+by (auto dest:brandnew_proc_prop1)
+
+lemma brandnew_proc_prop3:
+ "\<lbrakk>p \<in> current_procs s; valid (e # s)\<rbrakk> \<Longrightarrow> brandnew_proc (e # s) \<noteq> p"
+apply (rule brandnew_proc_prop2, simp)
+apply (rule no_juniorI, simp+)
+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 shms # s) tp dp = (
+ if (tp = p')
+ then Clone p dp (fds \<inter> proc_file_fds s p) shms # Clone p p' fds shms # s
+ else Clone p p' fds shms # (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 opt # Open p f flags fd opt # (enrich_proc s tp dp)
+ else Open p f flags fd opt # (enrich_proc s tp dp))"
+| "enrich_proc (CloseFd p fd # s) tp dp = (
+ if (tp = p)
+ 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> init_procs \<longrightarrow> deleted (O_proc p) s"
+
+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"
+ and vd: "valid s" and f_in: "is_file s f" and f_in': "is_file s' f"
+ and sec: "sectxt_of_obj s' (O_file f) = sectxt_of_obj s (O_file f)"
+ shows "search_check s' (up, rp, tp) f"
+proof (cases f)
+ case Nil
+ with f_in vd have "False"
+ by (auto dest:root_is_dir')
+ thus ?thesis by simp
+next
+ case (Cons n pf)
+ from vd f_in obtain sf where sf: "cf2sfile s f = Some sf"
+ apply (drule_tac is_file_in_current, drule_tac current_file_has_sfile, simp)
+ apply (erule exE, simp)
+ done
+ then obtain psfs where psfs: "get_parentfs_ctxts s pf = Some psfs" using Cons
+ by (auto simp:cf2sfile_def split:option.splits if_splits)
+ from sf cf2sf f_in have sf': "cf2sfile s' f = Some sf" by (auto dest:is_file_in_current)
+ then obtain psfs' where psfs': "get_parentfs_ctxts s' pf = Some psfs'"using Cons
+ by (auto simp:cf2sfile_def split:option.splits if_splits)
+ with sf sf' psfs have psfs_eq: "set psfs' = set psfs" using Cons f_in f_in'
+ apply (simp add:cf2sfile_def split:option.splits)
+ apply (case_tac sf, simp)
+ done
+ show ?thesis using grant f_in f_in' psfs psfs' psfs_eq sec
+ apply (simp add:Cons split:option.splits)
+ by (case_tac a, simp)
+qed
+
+lemma proc_filefd_has_sfd: "\<lbrakk>fd \<in> proc_file_fds s p; valid s\<rbrakk> \<Longrightarrow> \<exists> sfd. cfd2sfd s p fd = Some sfd"
+apply (simp add:proc_file_fds_def)
+apply (auto dest: current_filefd_has_sfd)
+done
+
+lemma enrich_inherit_fds_check:
+ assumes grant: "inherit_fds_check s (up, nr, nt) p fds" and vd: "valid s"
+ and cp2sp: "\<forall> p. p \<in> current_procs s \<longrightarrow> cp2sproc s' p = cp2sproc s p"
+ and p_in: "p \<in> current_procs s" and p_in': "p \<in> current_procs s'"
+ and fd_in: "fds \<subseteq> proc_file_fds s p" and fd_in': "fds \<subseteq> proc_file_fds s' p"
+ shows "inherit_fds_check s' (up, nr, nt) p fds"
+proof-
+ have "\<And> fd. fd \<in> fds \<Longrightarrow> sectxt_of_obj s' (O_fd p fd) = sectxt_of_obj s (O_fd p fd)"
+ proof-
+ fix fd
+ assume fd_in_fds: "fd \<in> fds"
+ hence fd_in_cfds: "fd \<in> proc_file_fds s p"
+ and fd_in_cfds': "fd \<in> proc_file_fds s' p"
+ using fd_in fd_in' by auto
+ from p_in vd obtain sp where csp: "cp2sproc s p = Some sp"
+ by (drule_tac current_proc_has_sp, simp, erule_tac exE, simp)
+ with cp2sp have "cpfd2sfds s p = cpfd2sfds s' p"
+ apply (erule_tac x = p in allE)
+ by (auto simp:cp2sproc_def split:option.splits simp:p_in)
+ hence "cfd2sfd s p fd = cfd2sfd s' p fd" using fd_in_cfds fd_in_cfds'
+ apply (simp add:cpfd2sfds_def)
+ apply (frule proc_filefd_has_sfd, simp add:vd, erule exE)
+ apply (drule_tac x = sfd in eqset_imp_iff, simp)
+ (*
+ thm inherit_fds_check_def
+ thm sectxts_of_fds_def
+ thm cpfd2sfds_def
+ apply (
+ *)sorry
+ show "sectxt_of_obj s' (O_fd p fd) = sectxt_of_obj s (O_fd p fd)"
+ sorry
+ qed
+ hence "sectxts_of_fds s' p fds = sectxts_of_fds s p fds"
+ by (simp add:sectxts_of_fds_def)
+ thus ?thesis using grant
+ by (simp add:inherit_fds_check_def)
+qed
+
+
+lemma enrich_proc_aux1:
+ assumes vs': "valid s'"
+ and os: "os_grant s e" and grant: "grant s e" and vd: "valid s"
+ and alive: "\<forall> obj. alive s obj \<longrightarrow> alive s' obj"
+ and cp2sp: "\<forall> p. p \<in> current_procs s \<longrightarrow> cp2sproc s' p = cp2sproc s p"
+ and cf2sf: "\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile s' f = cf2sfile s f"
+ shows "valid (e # s')"
+
+sorry (*
+proof (cases e)
+ case (Execve p f fds)
+ have p_in: "p \<in> 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 \<subseteq> current_proc_fds s' p" using os alive
+ apply (auto simp:Execve)
+ by (erule_tac x = "O_fd p x" in allE, auto)
+ have "os_grant s' e" using p_in f_in fd_in by (simp add:Execve)
+ moreover have "grant s' e" apply (simp add:Execve)
+ 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 "\<exists> 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
+ show ?
+ proof-
+ have
+
+
+*)
+
+lemma enrich_proc_prop:
+ "\<lbrakk>valid s; is_created_proc s p; p' \<notin> all_procs s\<rbrakk>
+ \<Longrightarrow> valid (enrich_proc s p p') \<and>
+ (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 {})))"
+sorry (*
+proof (induct s)
+ case Nil
+ thus ?case by (auto simp:is_created_proc_def)
+next
+ case (Cons e s)
+ hence p1: "\<lbrakk>valid s; is_created_proc s p; p' \<notin> all_procs s\<rbrakk>
+ \<Longrightarrow> valid (enrich_proc s p p') \<and>
+ (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>
+ (alive s obj \<longrightarrow> alive (enrich_proc s p p') obj \<and> co2sobj (enrich_proc s p p') obj = co2sobj s obj)"
+ and p2: "valid (e # s)" and p3: "is_created_proc (e # s) p" and p4: "p' \<notin> all_procs (e # s)"
+ by auto
+ from p2 have vd: "valid s" and os: "os_grant s e" and grant: "grant s e"
+ by (auto dest:vd_cons vt_grant vt_grant_os)
+ from p4 have p4': "p' \<notin> all_procs s" by (case_tac e, auto)
+ from p1 p4' have a1: "is_created_proc s p \<Longrightarrow> valid (enrich_proc s p p')" by (auto simp:vd)
+ have c1: "valid (enrich_proc (e # s) p p')"
+ apply (case_tac e)
+ using a1 os p3
+ apply (auto simp:is_created_proc_def)
+ sorry
+ moreover have c2: "p' \<in> current_procs (enrich_proc (e # s) p p')"
+ sorry
+ moreover have c3: "co2sobj (enrich_proc (e # s) p p') (O_proc p') = co2sobj (enrich_proc (e # s) p p') (O_proc p)"
+ sorry
+ moreover have c4: "alive (e # s) obj \<longrightarrow>
+ alive (enrich_proc (e # s) p p') obj \<and> co2sobj (enrich_proc (e # s) p p') obj = co2sobj (e # s) obj"
+ sorry
+ ultimately show ?case by auto
+qed
+*)
+
+
+lemma "alive s obj \<Longrightarrow> alive (enrich_proc s p p') obj"
+apply (induct s, simp)
+apply (case_tac a, case_tac[!] obj) sorry (*
+apply (auto simp:is_file_def is_dir_def split:option.splits t_inode_tag.splits)
+thm is_file_other
+*)
+lemma enrich_proc_valid:
+ "\<lbrakk>p \<in> current_procs s; valid s; p \<in> init_procs \<longrightarrow> deleted (O_proc p) s; p' \<notin> current_procs s\<rbrakk> \<Longrightarrow> valid (enrich_proc s p p')" (*
+apply (induct s, simp)
+apply (frule vd_cons, frule vt_grant, frule vt_grant_os, case_tac a)
+apply (auto intro!:valid.intros(2))
+prefer 28
+
+
+
+end
+*)
+sorry
+
+
+(* for any created obj, we can enrich trace with events that create new objs with the same static-properties *)
+definition enriched:: "t_state \<Rightarrow> t_object set \<Rightarrow> t_state \<Rightarrow> bool"
+where
+ "enriched s objs s' \<equiv> \<forall> obj \<in> objs. \<exists> obj'. \<not> alive s obj' \<and> obj' \<notin> objs \<and>
+ alive s' obj' \<and> co2sobj s' obj' = co2sobj s' obj"
+
+definition reserved:: "t_state \<Rightarrow> t_object set \<Rightarrow> t_state \<Rightarrow> bool"
+where
+ "reserved s objs s' \<equiv> \<forall> obj. alive s obj \<longrightarrow> alive s' obj \<and> co2sobj s' obj = co2sobj s obj"
+
+definition enrichable :: "t_state \<Rightarrow> t_object set \<Rightarrow> bool"
+where
+ "enrichable s objs \<equiv> \<exists> s'. valid s' \<and> s2ss s' = s2ss s \<and> enriched s objs s' \<and> reserved s objs s'"
+
+fun is_created :: "t_state \<Rightarrow> t_object \<Rightarrow> bool"
+where
+ "is_created s (O_file f) = (\<forall> f' \<in> same_inode_files s f. init_alive (O_file f') \<longrightarrow> deleted (O_file f') s)"
+| "is_created s obj = (init_alive obj \<longrightarrow> deleted obj s)"
+
+definition is_inited :: "t_state \<Rightarrow> t_object \<Rightarrow> bool"
+where
+ "is_inited s obj \<equiv> init_alive obj \<and> \<not> deleted obj s"
+
+(*
+lemma is_inited_eq_not_created:
+ "is_inited s obj = (\<not> is_created s obj)"
+by (auto simp:is_created_def is_inited_def)
+*)
+
+lemma many_sq_imp_sms:
+ "\<lbrakk>S_msgq (Create, sec, sms) \<in> ss; ss \<in> static\<rbrakk> \<Longrightarrow> \<forall> sm \<in> (set sms). is_many_smsg sm"
+sorry
+
+(* recorded in our static world *)
+fun recorded :: "t_object \<Rightarrow> bool"
+where
+ "recorded (O_proc p) = True"
+| "recorded (O_file f) = True"
+| "recorded (O_dir f) = True"
+| "recorded (O_node n) = False" (* cause socket is temperary not considered *)
+| "recorded (O_shm h) = True"
+| "recorded (O_msgq q) = True"
+| "recorded _ = False"
+
+lemma cf2sfile_fi_init_file:
+ "\<lbrakk>cf2sfile s f = Some (Init f', sec, psec, asecs); is_file s f; valid s\<rbrakk>
+ \<Longrightarrow> is_init_file f \<and> \<not> deleted (O_file f) s"
+apply (simp add:cf2sfile_def sroot_def split:option.splits if_splits)
+apply (case_tac f, simp, drule root_is_dir', simp+)
+done
+
+lemma root_not_deleted:
+ "valid s \<Longrightarrow> \<not> deleted (O_dir []) s"
+apply (induct s, simp)
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+by auto
+
+lemma cf2sfile_fi_init_dir:
+ "\<lbrakk>cf2sfile s f = Some (Init f', sec, psec, asecs); is_dir s f; valid s\<rbrakk>
+ \<Longrightarrow> is_init_dir f \<and> \<not> deleted (O_dir f) s"
+apply (simp add:cf2sfile_def sroot_def split:option.splits if_splits)
+apply (case_tac f, simp add:root_is_init_dir root_not_deleted, simp)
+apply (drule file_dir_conflict, simp+)
+done
+
+lemma is_created_imp_many:
+ "\<lbrakk>is_created s obj; co2sobj s obj = Some sobj; alive s obj; valid s\<rbrakk> \<Longrightarrow> is_many sobj"
+apply (case_tac obj, auto simp:co2sobj.simps split:option.splits)
+apply (case_tac [!] a)
+apply (auto simp:cp2sproc_def ch2sshm_def cq2smsgq_def cf2sfiles_def same_inode_files_def
+ split:option.splits if_splits)
+apply (frule cf2sfile_fi_init_file, simp add:is_file_def, simp)
+apply (erule_tac x = f' in allE, simp)
+apply (frule cf2sfile_fi_init_dir, simp+)+
+done
+
+lemma anotherp_imp_manysp:
+ "\<lbrakk>cp2sproc s p = Some sp; co2sobj s (O_proc p') = co2sobj s (O_proc p); p' \<noteq> p;
+ p' \<in> current_procs s; p \<in> current_procs s\<rbrakk>
+ \<Longrightarrow> is_many_sproc sp"
+by (case_tac sp, auto simp:cp2sproc_def co2sobj.simps split:option.splits if_splits)
+
+lemma is_file_has_sfs:
+ "\<lbrakk>is_file s f; valid s; cf2sfile s f = Some sf\<rbrakk>
+ \<Longrightarrow> \<exists> sfs. co2sobj s (O_file f) = Some (S_file sfs (O_file f \<in> Tainted s)) \<and> sf \<in> sfs"
+apply (rule_tac x = "{sf' | f' sf'. cf2sfile s f' = Some sf' \<and> f' \<in> same_inode_files s f}" in exI)
+apply (auto simp:co2sobj.simps cf2sfiles_def tainted_eq_Tainted)
+apply (rule_tac x = f in exI, simp add:same_inode_files_prop9)
+done
+
+declare Product_Type.split_paired_Ex Product_Type.split_paired_All [simp del]
+
+lemma current_proc_in_s2ss:
+ "\<lbrakk>cp2sproc s p = Some sp; p \<in> current_procs s; valid s\<rbrakk>
+ \<Longrightarrow> S_proc sp (O_proc p \<in> Tainted s) \<in> s2ss s"
+apply (simp add:s2ss_def, rule_tac x = "O_proc p" in exI)
+apply (auto simp:co2sobj.simps tainted_eq_Tainted)
+done
+
+lemma current_file_in_s2ss:
+ "\<lbrakk>co2sobj s (O_file f) = Some (S_file sfs tagf); is_file s f; valid s\<rbrakk>
+ \<Longrightarrow> S_file sfs tagf \<in> s2ss s"
+by (simp add:s2ss_def, rule_tac x = "O_file f" in exI, simp)
+
+declare npctxt_execve.simps grant_execve.simps search_check.simps [simp del]
+
+
+lemma npctxt_execve_eq_sec:
+ "\<lbrakk>sectxt_of_obj (Execve p f fds # s) (O_proc p) = Some sec'; sectxt_of_obj s (O_proc p) = Some sec;
+ sectxt_of_obj s (O_file f) = Some fsec; valid (Execve p f fds # s)\<rbrakk>
+ \<Longrightarrow> npctxt_execve sec fsec = Some sec'"
+by (case_tac sec, case_tac fsec, auto simp:npctxt_execve.simps sectxt_of_obj_simps split:option.splits)
+
+lemma npctxt_execve_eq_cp2sproc:
+ "\<lbrakk>cp2sproc (Execve p f fds # s) p = Some (pi', sec', sfds', shms'); valid (Execve p f fds # s);
+ cp2sproc s p = Some (pi, sec, sfds, shms); cf2sfile s f = Some (fi, fsec, psec, asecs)\<rbrakk>
+ \<Longrightarrow> npctxt_execve sec fsec = Some sec'"
+apply (frule vt_grant_os, frule vd_cons)
+apply (rule npctxt_execve_eq_sec, auto simp:cp2sproc_def cf2sfile_def split:option.splits)
+apply (case_tac f, auto dest:root_is_dir')
+done
+
+lemma seach_check_eq_static:
+ "\<lbrakk>cf2sfile s f = Some sf; valid s; is_dir s f \<or> is_file s f\<rbrakk>
+ \<Longrightarrow> search_check_s sec sf (is_file s f) = search_check s sec f"
+apply (case_tac sf)
+apply (induct f)
+apply (auto simp:search_check_s_def search_check.simps cf2sfile_def sroot_def
+ root_sec_remains init_sectxt_prop sec_of_root_valid
+ dest!:root_is_dir' current_has_sec' split:option.splits)
+done
+
+lemma grant_execve_intro_execve:
+ "\<lbrakk>cp2sproc (Execve p f fds # s) p = Some (pi', sec', sfds', shms'); valid (Execve p f fds # s);
+ cp2sproc s p = Some (pi, sec, sfds, shms); cf2sfile s f = Some (fi, fsec, psec, asecs)\<rbrakk>
+ \<Longrightarrow> grant_execve sec fsec sec'"
+apply (frule vt_grant_os, frule vd_cons, frule vt_grant)
+apply (auto split:option.splits dest!:current_has_sec' simp del:grant_execve.simps simp add:cp2sproc_execve)
+apply (erule_tac x = aba in allE, erule_tac x = aca in allE, erule_tac x = bb in allE)
+apply (auto simp del:grant_execve.simps simp add:cp2sproc_def cf2sfile_def split:option.splits)
+apply (case_tac f, simp, drule root_is_dir', simp, simp, simp)
+apply (simp add:sectxt_of_obj_simps)
+done
+
+lemma search_check_intro_execve:
+ "\<lbrakk>cp2sproc s p = Some (pi, sec, sfds, shms); valid (Execve p f fds # s)\<rbrakk>
+ \<Longrightarrow> search_check s sec f"
+apply (frule vt_grant_os, frule vd_cons, frule vt_grant)
+apply (auto split:option.splits dest!:current_has_sec' simp del:grant_execve.simps simp add:cp2sproc_execve)
+apply (erule_tac x = aaa in allE, erule_tac x = ab in allE, erule_tac x = ba in allE)
+apply (auto simp add:cp2sproc_def cf2sfile_def split:option.splits)
+done
+
+lemma inherit_fds_check_intro_execve:
+ "\<lbrakk>cp2sproc (Execve p f fds # s) p = Some (pi', sec', sfds', shms'); valid (Execve p f fds # s)\<rbrakk>
+ \<Longrightarrow> inherit_fds_check s sec' p fds"
+apply (frule vt_grant_os, frule vd_cons, frule vt_grant)
+apply (auto split:option.splits dest!:current_has_sec' simp add:cp2sproc_execve)
+apply (erule_tac x = aba in allE, erule_tac x = aca in allE, erule_tac x = bb in allE)
+apply (auto simp add:cp2sproc_def cf2sfile_def split:option.splits)
+apply (simp add:sectxt_of_obj_simps)
+done
+
+lemma execve_sfds_subset:
+ "\<lbrakk>cp2sproc (Execve p f fds # s) p = Some (pi', sec', sfds', shms'); valid (Execve p f fds # s);
+ cp2sproc s p = Some (pi, sec, sfds, shms)\<rbrakk>
+ \<Longrightarrow> sfds' \<subseteq> sfds"
+apply (frule vt_grant_os)
+apply (auto simp:cp2sproc_def cpfd2sfds_execve split:option.splits dest!:current_has_sec')
+apply (simp add:cpfd2sfds_def)
+apply (rule_tac x = fd in bexI, auto simp:proc_file_fds_def)
+done
+
+lemma inherit_fds_check_imp_static:
+ "\<lbrakk>cp2sproc (Execve p f fds # s) p = Some (pi', sec', sfds', shms');
+ inherit_fds_check s sec' p fds; valid (Execve p f fds # s)\<rbrakk>
+ \<Longrightarrow> inherit_fds_check_s sec' sfds'"
+apply (frule vt_grant_os, frule vd_cons, frule vt_grant)
+apply (auto simp:cp2sproc_def cpfd2sfds_execve inherit_fds_check_def inherit_fds_check_s_def split:option.splits)
+sorry (*
+apply (erule_tac x = "(ad, ae, bc)" in ballE, auto simp:sectxts_of_sfds_def sectxts_of_fds_def)
+apply (erule_tac x = fd in ballE, auto simp:cfd2sfd_def split:option.splits)
+done *)
+
+lemma d2s_main_execve_grant_aux:
+ "\<lbrakk>cp2sproc (Execve p f fds # s) p = Some (pi', sec', sfds', shms'); valid (Execve p f fds # s);
+ cp2sproc s p = Some (pi, sec, sfds, shms); cf2sfile s f = Some (fi, fsec, psec, asecs)\<rbrakk>
+ \<Longrightarrow> (npctxt_execve sec fsec = Some sec') \<and> grant_execve sec fsec sec' \<and>
+ search_check_s sec (fi, fsec, psec, asecs) (is_file s f) \<and>
+ inherit_fds_check_s sec' sfds' \<and> sfds' \<subseteq> sfds"
+apply (rule conjI, erule_tac pi = pi and sec = sec and sfds = sfds and
+ shms = shms and fi = fi and fsec = fsec and psec = psec and
+ asecs = asecs in npctxt_execve_eq_cp2sproc, simp, simp, simp)
+apply (rule conjI, erule_tac pi = pi and sec = sec and sfds = sfds and
+ shms = shms and fi = fi and fsec = fsec and psec = psec and
+ asecs = asecs in grant_execve_intro_execve, simp, simp, simp)
+apply (rule conjI, drule_tac sec = sec in search_check_intro_execve, simp)
+apply (frule vd_cons, frule vt_grant_os)
+apply (drule_tac sec = sec in seach_check_eq_static, simp, simp, simp)
+apply (rule conjI, rule inherit_fds_check_imp_static, simp)
+apply (erule inherit_fds_check_intro_execve, simp, simp)
+apply (erule_tac pi = pi and sfds = sfds and shms = shms in execve_sfds_subset, simp+)
+done
+
+lemma d2s_main_execve:
+ "\<lbrakk>valid (Execve p f fds # s); s2ss s \<propto> static\<rbrakk> \<Longrightarrow> s2ss (Execve p f fds # s) \<propto> static"
+apply (frule vd_cons, frule vt_grant_os, clarsimp)
+apply (frule is_file_has_sfile', simp, erule exE, frule is_file_has_sfs, simp+, erule exE, erule conjE)
+apply (auto simp:s2ss_execve split:if_splits option.splits dest:current_proc_has_sp')
+apply (clarsimp simp add:init_ss_in_def init_ss_eq_def)
+apply (rule_tac x = "update_ss ss' (S_proc (ah, (ai, aj, bd), ak, be) (O_proc p \<in> Tainted s))
+ (S_proc (ad, (ae, af, bb), ag, bc) (O_proc p \<in> Tainted s \<or> O_file f \<in> Tainted s))" in bexI)
+apply (auto simp:update_ss_def elim:Set.subset_insertI2 simp:anotherp_imp_manysp)[1]
+apply (case_tac "ah = ad", case_tac "bc = {}", simp)
+apply (erule_tac sfs = sfs and fi = a and fsec = "(aa, ab,b)" and pfsec = ac and asecs = ba in s_execve,
+ auto intro:current_proc_in_s2ss current_file_in_s2ss split:option.splits dest:d2s_main_execve_grant_aux)[1]
+apply (simp add:cp2sproc_execve split:option.splits)
+apply (simp add:cp2sproc_def split:option.splits if_splits)
+
+apply (clarsimp simp add:init_ss_in_def init_ss_eq_def)
+apply (rule_tac x = "update_ss ss' (S_proc (ah, (ai, aj, bd), ak, be) (O_proc p \<in> Tainted s))
+ (S_proc (ad, (ae, af, bb), ag, bc) (O_proc p \<in> Tainted s \<or> O_file f \<in> Tainted s))" in bexI)
+apply (rule conjI, simp add:update_ss_def)
+apply (rule conjI, simp add:update_ss_def)
+apply (auto)[1]
+apply (simp add:update_ss_def)
+apply (rule conjI, rule impI)
+apply (rule subsetI, clarsimp)
+apply (erule impE)
+apply (erule set_mp, simp)
+apply (case_tac ah, simp+)
+apply (rule impI, rule subsetI, clarsimp)
+apply (erule set_mp, simp)
+
+apply (case_tac "ah = ad", case_tac "bc = {}", simp)
+apply (erule_tac sfs = sfs and fi = a and fsec = "(aa, ab,b)" and pfsec = ac and asecs = ba in s_execve,
+ auto intro:current_proc_in_s2ss current_file_in_s2ss split:option.splits dest:d2s_main_execve_grant_aux)[1]
+apply (simp add:cp2sproc_execve split:option.splits)
+apply (simp add:cp2sproc_def split:option.splits if_splits)
+done
+
+lemma co2sobj_eq_alive_proc_imp:
+ "\<lbrakk>co2sobj s obj = co2sobj s (O_proc p); alive s (O_proc p); valid s\<rbrakk>
+ \<Longrightarrow> \<exists> p'. obj = O_proc p'"
+by (auto simp add:co2sobj.simps split:option.splits dest:current_proc_has_sp' intro:co2sobj_sproc_imp)
+
+
+lemma enrichable_execve:
+ assumes p1: "\<And> objs. \<forall> obj \<in> objs. alive s obj \<and> is_created s obj \<and> recorded obj
+ \<Longrightarrow> \<exists> s'. valid s' \<and> s2ss s' = s2ss s \<and> enriched s objs s' \<and> reserved s objs s'"
+ and p2: "valid (e # s)" and p3: "\<forall>obj\<in>objs. alive (e # s) obj \<and> is_created (e # s) obj \<and> recorded obj"
+ and p4: "e = Execve p f fds"
+ shows "enrichable (e # s) objs"
+proof-
+ from p2 have os: "os_grant s e" and se: "grant s e" and vd: "valid s"
+ by (auto dest:vt_grant_os vd_cons vt_grant)
+ from p3 have recorded: "\<forall> obj \<in> objs. recorded obj" by auto
+ from p3 p4 p2 have p1': "\<forall> obj \<in> objs. alive s obj \<and> is_created s obj"
+ apply clarify
+ apply (erule_tac x = obj in ballE, simp add:alive_simps)
+ apply (case_tac obj, auto simp:same_inode_files_simps)
+ done
+ then obtain s' where a1: "valid s'" and a2: "s2ss s' = s2ss s" and a3: "enriched s objs s'"
+ and a4: "reserved s objs s'"
+ using p1 recorded by metis
+ show ?thesis
+ proof (cases "O_proc p \<in> objs")
+ case True
+ hence p_in: "p \<in> current_procs s'" using a4 os p4
+ by (auto simp:reserved_def elim:allE[where x = "O_proc p"])
+ with a1 a3 True obtain p' where b1: "\<not> alive s (O_proc p')" and b2: "O_proc p' \<notin> objs"
+ and b3: "alive s' (O_proc p')" and b4: "co2sobj s' (O_proc p') = co2sobj s' (O_proc p)"
+ apply (simp only:enriched_def)
+ apply (erule_tac x = "O_proc p" in ballE)
+ apply (erule exE|erule conjE)+
+ apply (frule co2sobj_eq_alive_proc_imp, auto)
+ done
+ have "valid (Execve p' f fds # e # s')"
+ sorry
+ moreover have "s2ss (Execve p' f fds # e # s') = s2ss (e # s)"
+ sorry
+ moreover have "enriched (e # s) objs (Execve p' f fds # e # s')"
+ sorry
+ moreover have "reserved (e # s) objs (Execve p' f fds # e # s')"
+ sorry
+ ultimately show ?thesis
+ apply (simp add:enrichable_def)
+ apply (rule_tac x = "Execve p' f fds # e # s'" in exI)
+ by auto
+ next
+ case False
+ from a4 os p4 have "p \<in> current_procs s'"
+ apply (simp add:reserved_def)
+ by (erule_tac x = "O_proc p" in allE, auto)
+ moreover from a4 os p4 have "is_file s' f"
+ apply (simp add:reserved_def)
+ by (erule_tac x = "O_file f" in allE, auto)
+ moreover from a4 os p4 vd have "fds \<subseteq> proc_file_fds s' p"
+ apply (rule_tac subsetI, clarsimp simp:reserved_def current_proc_fds.simps)
+ apply (erule_tac x = "O_fd p x" in allE, erule impE)
+ sorry
+ ultimately have "os_grant s' e"
+ by (simp add:p4)
+ moreover have "grant s' e"
+ sorry
+ ultimately have "valid (e # s')"
+ using a1 by (erule_tac valid.intros(2), simp+)
+ thus ?thesis
+ apply (simp add:enrichable_def)
+ apply (rule_tac x = "e # s'" in exI)
+ apply (simp)
+ sorry
+qed
+qed
+
+lemma s2d_main_execve:
+ "\<lbrakk>grant_execve pctxt fsec pctxt'; ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; S_file sfs tagf \<in> ss;
+ (fi, fsec, pfsec, asecs) \<in> sfs; npctxt_execve pctxt fsec = Some pctxt';
+ search_check_s pctxt (fi, fsec, pfsec, asecs) True; inherit_fds_check_s pctxt' fds'; fds' \<subseteq> fds; valid s;
+ s2ss s = ss\<rbrakk> \<Longrightarrow> \<exists>s. valid s \<and>
+ s2ss s = update_ss ss (S_proc (pi, pctxt, fds, shms) tagp) (S_proc (pi, pctxt', fds', {}) (tagp \<or> tagf))"
+apply (simp add:update_ss_def)
+thm update_ss_def
+sorry
+
+(*
+lemma s2d_main_execve:
+ "ss \<in> static \<Longrightarrow> \<exists> s. valid s \<and> s2ss s = ss"
+apply (erule static.induct)
+apply (rule_tac x = "[]" in exI, simp add:s2ss_nil_prop valid.intros)
+apply (erule exE|erule conjE)+
+apply (rule s2d_main_execve, simp+)
+
+apply (erule exE|erule conjE)+
+
+
+sorry
+*)
+
+
+(*********************** uppest-level 3 theorems ***********************)
+
+lemma enrichability:
+ "\<lbrakk>valid s; \<forall> obj \<in> objs. alive s obj \<and> is_created s obj \<and> recorded obj\<rbrakk>
+ \<Longrightarrow> enrichable s objs" sorry (*
+proof (induct s arbitrary:objs)
+ case Nil
+ hence "objs = {}"
+ apply (auto)
+ apply (erule_tac x = x in ballE)
+ apply (case_tac x)
+ apply (auto simp:init_alive_prop)
+ sorry (* done *)
+ thus ?case using Nil unfolding enrichable_def enriched_def reserved_def
+ by (rule_tac x = "[]" in exI, auto)
+next
+ case (Cons e s)
+ hence p1: "\<And> objs. \<forall> obj \<in> objs. alive s obj \<and> is_created s obj \<and> recorded obj
+ \<Longrightarrow> \<exists> s'. valid s' \<and> s2ss s' = s2ss s \<and> enriched s objs s' \<and> reserved s objs s'"
+ and p2: "valid (e # s)" and p3: "\<forall>obj\<in>objs. alive (e # s) obj \<and> is_created (e # s) obj \<and> recorded obj"
+ and os: "os_grant s e" and se: "grant s e" and vd: "valid s"
+ by (auto dest:vt_grant_os vd_cons vt_grant simp:enrichable_def)
+ show ?case
+ proof (cases e)
+ case (Execve p f fds)
+ hence p4: "e = Execve p f fds" by simp
+ from p3 have p5: "is_inited s (O_proc p) \<Longrightarrow> (O_proc p) \<notin> objs"
+ by (auto simp:is_created_def is_inited_def p4 elim!:ballE[where x = "O_proc p"])
+ show "enrichable (e # s) objs"
+ proof (cases "is_inited s (O_proc p)")
+ case True
+ with p5 have a1: "(O_proc p) \<notin> objs" by simp
+ with p3 p4 p2 have a2: "\<forall> obj \<in> objs. alive s obj \<and> is_created s obj" and a2': "\<forall> obj \<in> objs. recorded obj"
+ apply (auto simp:is_created_def alive_simps is_inited_def)
+ apply (erule_tac x = obj in ballE, auto simp:alive_simps split:t_object.splits)
+ done
+ then obtain s' where a3: "valid s'" and a4: "s2ss s' = s2ss s"
+ and a5: "enriched s objs s'" and a6: "reserved s objs s'"
+ using p1 apply (simp add:enrichable_def) sorry
+ from a5 p4 p2 a2' have a7: "enriched s objs (e # s')"
+ apply (clarsimp simp add:enriched_def co2sobj_execve)
+ apply (erule_tac x = obj in ballE, clarsimp)
+ apply (rule_tac x = obj' in exI, auto simp:co2sobj_execve alive_simps)
+ thm enriched_def
+
+
+
+obtain s' where p6:"enriched s objs s'"
+ apply (simp add: alive_simps enrichable_def)
+ apply auto apply (rule ballI, rule_tac x = obj in exI)
+
+ have p6:"enriched (e # s) objs (e # s)"
+ apply (simp add:enriched_def alive_simps)
+ apply auto apply (rule ballI, rule_tac x = obj in exI)
+
+ have "enrich (e # s) objs (e # s)"
+ apply (simp add:enrich_def p4)
+ sorry
+ moreover have "reserve (e # s) objs (e # s)"
+ sorry
+ ultimately show ?thesis using p2
+ apply (simp add:enrichable_def)
+ by (rule_tac x = "e # s" in exI, simp)
+ next
+
+
+thm enrichable_def
+ apply (simp add:enrichable_def p4)
+
+
+
+ apply auto
+ apply (auto simp:enrichable_def)
+apply (induct s)
+
+
+
+done
+
+qed
+*)
+
+lemma d2s_main:
+ "valid s \<Longrightarrow> s2ss s \<propto> static"
+apply (induct s, simp add:s2ss_nil_prop init_ss_in_def)
+apply (rule_tac x = "init_static_state" in bexI, simp, simp add:s_init)
+apply (frule vd_cons, frule vt_grant_os, simp)
+apply (case_tac a)
+apply (clarsimp simp add:s2ss_execve)
+apply (rule conjI, rule impI)
+
+
+
+sorry
+
+lemma s2d_main:
+ "ss \<in> static \<Longrightarrow> \<exists> s. valid s \<and> s2ss s = ss"
+apply (erule static.induct)
+apply (rule_tac x = "[]" in exI, simp add:s2ss_nil_prop valid.intros)
+
+apply (erule exE|erule conjE)+
+
+apply (simp add:update_ss_def)
+
+sorry
+
+lemma s2d_main':
+ "ss \<in> static \<Longrightarrow> \<exists> s. valid s \<and> s2ss s \<doteq> ss"
+apply (erule static.induct)
+apply (rule_tac x = "[]" in exI, simp add:s2ss_nil_prop valid.intros)
+
+apply (erule exE|erule conjE)+
+
+apply (simp add:update_ss_def)
+
+sorry
+
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/File_renaming.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,221 @@
+theory File_renaming
+imports Main OS_type_def My_list_prefix
+begin
+
+fun parent :: "'a list \<Rightarrow> ('a list) option"
+where
+ "parent [] = None"
+ | "parent (n#ns) = Some ns"
+
+definition file_after_rename :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
+ "file_after_rename oldf newf f \<equiv> (if (oldf \<preceq> f) then ((f \<setminus> oldf) @ newf) else f)"
+
+definition file_before_rename :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
+ "file_before_rename oldf newf f \<equiv> (if (newf \<preceq> f) then ((f \<setminus> newf) @ oldf) else f)"
+
+lemma list_diff_rev_prop0: "f \<setminus> f = []"
+by (induct f, auto simp:list_diff_rev_def)
+
+lemma list_diff_rev_prop0': "f \<setminus> [] = f"
+apply (simp add:list_diff_rev_def)
+apply (induct f rule:rev_induct, simp+)
+done
+
+lemma list_diff_rev_prop1_aux: "a # f\<^isub>3 \<setminus> f\<^isub>3 = [a]"
+by (induct f\<^isub>3, auto simp:list_diff_rev_def)
+
+lemma list_diff_rev_prop1: "f @ f' \<setminus> f' = f"
+apply (subgoal_tac "f @ f' \<setminus> [] @ f' = f \<setminus> []")
+apply (simp add:list_diff_rev_prop0')
+apply (simp only: list_diff_rev_ext_left[where ys = "[]"])
+done
+
+lemma list_diff_rev_prop2_aux: "f \<preceq> f' \<Longrightarrow> a # f' \<setminus> f = a # (f' \<setminus> f)"
+apply (erule no_juniorE)
+by (simp only:list_diff_rev_prop1 append_Cons[THEN sym])
+
+lemma list_diff_rev_prop2: "f \<preceq> f' \<Longrightarrow> (f' \<setminus> f) @ f = f'"
+apply (induct f', simp add:no_junior_def list_diff_rev_def)
+apply (case_tac "f = a # f'", simp add:list_diff_rev_prop0)
+apply (drule no_junior_noteq, simp)
+by (simp add:list_diff_rev_prop2_aux)
+
+lemma file_renaming_prop0:
+ "file_after_rename f\<^isub>2 f\<^isub>3 f\<^isub>2 = f\<^isub>3"
+apply (simp add:file_after_rename_def list_diff_rev_def)
+apply (induct f\<^isub>2, auto)
+done
+
+lemma file_renaming_prop0':
+ "file_before_rename f\<^isub>2 f\<^isub>3 f\<^isub>3 = f\<^isub>2"
+apply (simp add:file_before_rename_def list_diff_rev_def)
+apply (induct f\<^isub>3, auto)
+done
+
+lemma file_renaming_prop1:
+ "f\<^isub>2 \<preceq> f \<Longrightarrow> f\<^isub>3 \<preceq> file_after_rename f\<^isub>2 f\<^isub>3 f "
+by (simp add:file_after_rename_def)
+
+lemma file_renaming_prop1':
+ "f\<^isub>3 \<preceq> f \<Longrightarrow> f\<^isub>2 \<preceq> file_before_rename f\<^isub>2 f\<^isub>3 f"
+by (simp add:file_before_rename_def)
+
+lemma file_renaming_prop1'':
+ "\<not> f\<^isub>3 \<preceq> file_after_rename f\<^isub>2 f\<^isub>3 f \<Longrightarrow> \<not> f\<^isub>2 \<preceq> f"
+by (case_tac "f\<^isub>2 \<preceq> f", drule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop1, simp+)
+
+lemma file_renaming_prop2:
+ "\<lbrakk>f\<^isub>2 \<preceq> f; f \<noteq> f\<^isub>2\<rbrakk> \<Longrightarrow> file_after_rename f\<^isub>2 f\<^isub>3 f \<noteq> f\<^isub>3"
+apply (case_tac f, simp add:no_junior_def)
+apply (auto simp add:file_after_rename_def list_diff_rev_prop1_aux)
+apply (erule no_juniorE, simp add:list_diff_rev_prop1)
+done
+
+lemma file_renaming_prop2':
+ "\<lbrakk>f\<^isub>3 \<preceq> f; f \<noteq> f\<^isub>3\<rbrakk> \<Longrightarrow> file_before_rename f\<^isub>2 f\<^isub>3 f \<noteq> f\<^isub>2"
+apply (case_tac f, simp add:no_junior_def)
+apply (auto simp add:file_before_rename_def list_diff_rev_prop1_aux)
+apply (erule no_juniorE, simp add:list_diff_rev_prop1)
+done
+
+lemma file_renaming_prop3:
+ "\<lbrakk>f\<^isub>2 \<preceq> f; f\<^isub>2 \<noteq> f\<rbrakk> \<Longrightarrow> f\<^isub>3 \<preceq> file_after_rename f\<^isub>2 f\<^isub>3 f \<and> f\<^isub>3 \<noteq> file_after_rename f\<^isub>2 f\<^isub>3 f"
+apply (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop1, simp)
+by (drule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop2, simp+)
+
+lemma file_renaming_prop3':
+ "\<lbrakk>f\<^isub>3 \<preceq> f; f\<^isub>3 \<noteq> f\<rbrakk> \<Longrightarrow> f\<^isub>2 \<preceq> file_before_rename f\<^isub>2 f\<^isub>3 f \<and> f\<^isub>2 \<noteq> file_before_rename f\<^isub>2 f\<^isub>3 f"
+apply (frule_tac f\<^isub>2 = f\<^isub>2 in file_renaming_prop1', simp)
+by (drule_tac f\<^isub>2 = f\<^isub>2 in file_renaming_prop2', simp+)
+
+lemma file_renaming_prop4: "\<not> f\<^isub>2 \<preceq> f \<Longrightarrow> file_after_rename f\<^isub>2 f\<^isub>3 f = f"
+by (simp add:file_after_rename_def)
+
+lemma file_renaming_prop4': "\<not> f\<^isub>3 \<preceq> f \<Longrightarrow> file_before_rename f\<^isub>2 f\<^isub>3 f = f"
+by (simp add:file_before_rename_def)
+
+lemma file_renaming_prop5:
+ "f\<^isub>2 \<preceq> f\<^isub>1 \<Longrightarrow> file_before_rename f\<^isub>2 f\<^isub>3 (file_after_rename f\<^isub>2 f\<^isub>3 f\<^isub>1) = f\<^isub>1"
+apply (case_tac "f\<^isub>1 = f\<^isub>2")
+apply (simp add:file_renaming_prop0 file_renaming_prop0')
+apply (simp add:file_after_rename_def file_before_rename_def)
+by (simp add:list_diff_rev_prop1 list_diff_rev_prop2)
+
+lemma file_renaming_prop5':
+ "f\<^isub>3 \<preceq> f\<^isub>1 \<Longrightarrow> file_after_rename f\<^isub>2 f\<^isub>3 (file_before_rename f\<^isub>2 f\<^isub>3 f\<^isub>1) = f\<^isub>1"
+apply (case_tac "f\<^isub>1 = f\<^isub>3")
+apply (simp add:file_renaming_prop0' file_renaming_prop0)
+apply (simp add:file_after_rename_def file_before_rename_def)
+by (simp add:list_diff_rev_prop1 list_diff_rev_prop2)
+
+lemma file_renaming_prop6:
+ "f\<^isub>2 \<preceq> f \<Longrightarrow> file_after_rename f\<^isub>2 f\<^isub>3 (a # f) = (a # (file_after_rename f\<^isub>2 f\<^isub>3 f))"
+by (simp add:file_after_rename_def list_diff_rev_prop2_aux)
+
+lemma file_renaming_prop6':
+ "f\<^isub>3 \<preceq> f \<Longrightarrow> file_before_rename f\<^isub>2 f\<^isub>3 (a # f) = (a # (file_before_rename f\<^isub>2 f\<^isub>3 f))"
+by (simp add:file_before_rename_def list_diff_rev_prop2_aux)
+
+lemma file_renaming_prop7: "file_after_rename f\<^isub>2 f\<^isub>3 (a # f\<^isub>2) = a # f\<^isub>3"
+apply (auto simp:file_after_rename_def no_junior_def list_diff_rev_def)
+by (induct f\<^isub>2, auto)
+
+lemma file_renaming_prop7': "file_before_rename f\<^isub>2 f\<^isub>3 (a # f\<^isub>3) = a # f\<^isub>2"
+by (auto simp:file_before_rename_def no_junior_def list_diff_rev_prop1_aux)
+
+lemma file_renaming_prop8: "f\<^isub>2 \<preceq> f \<Longrightarrow> f\<^isub>3 \<preceq> (file_after_rename f\<^isub>2 f\<^isub>3 (a # f))"
+by (simp add:file_after_rename_def no_junior_def)
+
+lemma file_renaming_prop8': "f\<^isub>3 \<preceq> f \<Longrightarrow> f\<^isub>2 \<preceq> (file_before_rename f\<^isub>2 f\<^isub>3 (a # f))"
+by (simp add:file_before_rename_def no_junior_def)
+
+lemma no_junior_prop0: "f\<preceq> f' \<Longrightarrow> f \<preceq> (a # f')"
+by (simp add:no_junior_def)
+
+lemma file_renaming_prop9:
+ "f\<^isub>2 \<preceq> f\<^isub>1 \<Longrightarrow> file_before_rename f\<^isub>2 f\<^isub>3 (file_after_rename f\<^isub>2 f\<^isub>3 (a # f\<^isub>1)) = (a # f\<^isub>1)"
+by (drule_tac a = a in no_junior_prop0, simp add:file_renaming_prop5)
+
+lemma file_renaming_prop9':
+ "f\<^isub>3 \<preceq> f\<^isub>1 \<Longrightarrow> file_after_rename f\<^isub>2 f\<^isub>3 (file_before_rename f\<^isub>2 f\<^isub>3 (a # f\<^isub>1)) = (a # f\<^isub>1)"
+by (drule_tac a = a in no_junior_prop0, simp add:file_renaming_prop5')
+
+lemma file_renaming_prop10:
+ "\<lbrakk>a # f\<^isub>1 = file_after_rename fx fy f\<^isub>1'; file_after_rename fx fy f\<^isub>1' \<noteq> fy; fx \<preceq> f\<^isub>1'\<rbrakk> \<Longrightarrow> a # (file_before_rename fx fy f\<^isub>1) = f\<^isub>1'"
+apply (frule_tac f\<^isub>3 = fy in file_renaming_prop1)
+apply (subgoal_tac "fy \<preceq> f\<^isub>1") defer apply (rotate_tac 3, erule no_juniorE, case_tac zs, simp+)
+apply (drule_tac f = f\<^isub>1 and f\<^isub>2 = fx and f\<^isub>3 = fy and a = a in file_renaming_prop6')
+by (drule_tac f\<^isub>1 = f\<^isub>1' and f\<^isub>3 = fy in file_renaming_prop5, simp)
+
+lemma file_renaming_prop11:
+ "\<lbrakk>f\<^isub>2 \<preceq> f; file_after_rename f\<^isub>2 f\<^isub>3 f = f'\<rbrakk> \<Longrightarrow> f = file_before_rename f\<^isub>2 f\<^isub>3 f'"
+by (drule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop5, simp)
+
+lemma file_renaming_prop11':
+ "\<lbrakk>f\<^isub>3 \<preceq> f; file_before_rename f\<^isub>2 f\<^isub>3 f = f'\<rbrakk> \<Longrightarrow> f = file_after_rename f\<^isub>2 f\<^isub>3 f'"
+by (drule_tac f\<^isub>2 = f\<^isub>2 in file_renaming_prop5', simp)
+
+lemma file_renaming_prop12:
+ "\<lbrakk>\<not> f\<^isub>3 \<preceq> aa; f\<^isub>3 \<preceq> f; file_after_rename f\<^isub>2 f\<^isub>3 aa = f\<rbrakk> \<Longrightarrow> file_before_rename f\<^isub>2 f\<^isub>3 f = aa"
+apply (simp add:file_after_rename_def file_before_rename_def split:if_splits)
+by (auto, simp add:list_diff_rev_prop1 list_diff_rev_prop2)
+
+lemma file_renaming_prop12':
+ "\<lbrakk>\<not> f\<^isub>2 \<preceq> aa; f\<^isub>2 \<preceq> f; file_before_rename f\<^isub>2 f\<^isub>3 aa = f\<rbrakk> \<Longrightarrow> file_after_rename f\<^isub>2 f\<^isub>3 f = aa"
+apply (simp add:file_after_rename_def file_before_rename_def split:if_splits)
+by (auto, simp add:list_diff_rev_prop1 list_diff_rev_prop2)
+
+
+lemma file_renaming_prop13:
+ "\<lbrakk>f\<^isub>3 \<preceq> f; file_before_rename f\<^isub>2 f\<^isub>3 f = f\<rbrakk> \<Longrightarrow> f\<^isub>2 \<preceq> f"
+apply (simp add:file_before_rename_def)
+apply (rule_tac zs = "f \<setminus> f\<^isub>3" in no_juniorI)
+by simp
+
+(************ something *****************)
+
+
+lemma parent_is_parent: "parent f = Some pf \<Longrightarrow> f \<noteq> pf"
+by (case_tac f, auto)
+
+lemma parent_is_parent': "parent f = Some f \<Longrightarrow> false"
+by (drule parent_is_parent, simp)
+
+lemma parent_is_parent'': "parent f \<noteq> Some f"
+by (case_tac f, auto)
+
+lemma parent_is_parent3: "parent f = Some f' \<Longrightarrow> \<not> f \<preceq> f'"
+by (case_tac f, auto elim:no_juniorE)
+
+lemma parent_is_ancen: "parent f = Some f' \<Longrightarrow> f' \<preceq> f"
+by (case_tac f, auto simp:no_junior_def)
+
+lemma parent_is_ancen': "\<lbrakk>parent f = Some pf; f' \<preceq> pf\<rbrakk> \<Longrightarrow> f' \<preceq> f"
+by (case_tac f, auto simp:no_junior_def)
+
+lemma parent_ancen: "\<lbrakk>parent f = Some pf; f' \<preceq> f; f \<noteq> f'\<rbrakk> \<Longrightarrow> f' \<preceq> pf"
+apply (erule no_juniorE)
+apply (case_tac zs, simp)
+apply (rule_tac zs = list in no_juniorI)
+by auto
+
+lemma parent_rename_ancen: "\<lbrakk>parent f = Some pf; f\<^isub>3 \<preceq> pf; f \<noteq> f\<^isub>3\<rbrakk> \<Longrightarrow> parent (file_before_rename f\<^isub>2 f\<^isub>3 f) = Some (file_before_rename f\<^isub>2 f\<^isub>3 pf)"
+apply (simp add:file_before_rename_def)
+apply (case_tac f, simp, simp)
+by (drule_tac a = a in list_diff_rev_prop2_aux, simp)
+
+lemma no_junior_trans: "\<lbrakk>f \<preceq> f'; f' \<preceq> f''\<rbrakk> \<Longrightarrow> f \<preceq> f''"
+by (auto elim:no_juniorE)
+
+lemma no_junior_conf: "\<lbrakk>a \<preceq> c; b \<preceq> c\<rbrakk> \<Longrightarrow> a \<preceq> b \<or> b \<preceq> a"
+apply (simp add:no_junior_def)
+apply (induct c, auto)
+done
+
+lemma noJ_Anc: "x \<prec> y = (x \<preceq> y \<and> x \<noteq> y)"
+apply (simp add: no_junior_expand)
+by (auto simp:is_ancestor_def)
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Final_theorem.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,274 @@
+theory Final_theorem
+imports Main Flask Static Init_prop Valid_prop Tainted_prop Delete_prop Co2sobj_prop S2ss_prop S2ss_prop2 Dynamic_static
+begin
+
+context tainting_s begin
+
+lemma t2ts:
+ "obj \<in> tainted s \<Longrightarrow> co2sobj s obj = Some sobj \<Longrightarrow> tainted_s (s2ss s) sobj"
+apply (frule tainted_in_current, frule tainted_is_valid)
+apply (frule s2ss_included_sobj, simp)
+apply (case_tac sobj, simp_all)
+apply (case_tac [!] obj, simp_all add:co2sobj.simps split:option.splits if_splits)
+apply (drule dir_not_tainted, simp)
+apply (drule msgq_not_tainted, simp)
+apply (drule shm_not_tainted, simp)
+done
+
+lemma delq_imp_delqm:
+ "deleted (O_msgq q) s \<Longrightarrow> deleted (O_msg q m) s"
+apply (induct s, simp)
+by (case_tac a, auto)
+
+lemma tainted_s_subset_prop:
+ "\<lbrakk>tainted_s ss sobj; ss \<subseteq> ss'\<rbrakk> \<Longrightarrow> tainted_s ss' sobj"
+apply (case_tac sobj)
+apply auto
+done
+
+theorem static_complete:
+ assumes undel: "undeletable obj" and tbl: "taintable obj"
+ shows "taintable_s obj"
+proof-
+ from tbl obtain s where tainted: "obj \<in> tainted s"
+ by (auto simp:taintable_def)
+ hence vs: "valid s" by (simp add:tainted_is_valid)
+ hence static: "s2ss s \<propto> static" using d2s_main by auto
+ from tainted tbl vs obtain sobj where sobj: "co2sobj s obj = Some sobj"
+ apply (clarsimp simp add:taintable_def)
+ apply (frule tainted_in_current)
+ apply (case_tac obj, simp_all add:co2sobj.simps)
+ apply (frule current_proc_has_sp, simp, auto)
+ done
+ from undel vs have "\<not> deleted obj s" and init_alive: "init_alive obj"
+ by (auto simp:undeletable_def)
+ with vs sobj have "init_obj_related sobj obj"
+ apply (case_tac obj, case_tac [!] sobj)
+ apply (auto split:option.splits if_splits simp:co2sobj.simps cp2sproc_def ch2sshm_def cq2smsgq_def cm2smsg_def delq_imp_delqm)
+ apply (frule not_deleted_init_file, simp+)
+ apply (drule is_file_has_sfile', simp, erule exE)
+ apply (rule_tac x = sf in bexI)
+ apply (case_tac list, auto split:option.splits simp:is_init_file_props)[1]
+ apply (drule root_is_init_dir', simp)
+ apply (frule not_deleted_init_file, simp, simp)
+ apply (simp add:cf2sfile_def split:option.splits if_splits)
+ apply (simp add:cf2sfiles_def)
+ apply (rule_tac x = list in bexI, simp, simp add:same_inode_files_def not_deleted_init_file)
+
+ apply (frule not_deleted_init_dir, simp+)
+ apply (simp add:cf2sfile_def split:option.splits if_splits)
+ apply (case_tac list, simp add:sroot_def, simp)
+ apply (drule file_dir_conflict, simp+)
+ done
+ with tainted t2ts init_alive sobj static
+ show ?thesis unfolding taintable_s_def
+ apply (simp add:init_ss_in_def)
+ apply (erule bexE)
+ apply (simp add:init_ss_eq_def)
+ apply (rule_tac x = "ss'" in bexI)
+ apply (rule_tac x = "sobj" in exI)
+ thm tainted_s_subset_prop
+ by (auto intro:tainted_s_subset_prop)
+qed
+
+lemma cp2sproc_pi:
+ "\<lbrakk>cp2sproc s p = Some (Init p', sec, fds, shms); valid s\<rbrakk> \<Longrightarrow> p = p' \<and> \<not> deleted (O_proc p) s \<and> p \<in> init_procs"
+by (simp add:cp2sproc_def split:option.splits if_splits)
+
+lemma cq2smsgq_qi:
+ "\<lbrakk>cq2smsgq s q = Some (Init q', sec, sms); valid s\<rbrakk> \<Longrightarrow> q = q' \<and> \<not> deleted (O_msgq q) s \<and> q \<in> init_msgqs"
+by (simp add:cq2smsgq_def split:option.splits if_splits)
+
+lemma cm2smsg_mi:
+ "\<lbrakk>cm2smsg s q m = Some (Init m', sec, ttag); q \<in> init_msgqs; valid s\<rbrakk>
+ \<Longrightarrow> m = m' \<and> \<not> deleted (O_msg q m) s \<and> m \<in> set (init_msgs_of_queue q) \<and> q \<in> init_msgqs"
+by (clarsimp simp add:cm2smsg_def split:if_splits option.splits)
+
+lemma ch2sshm_hi:
+ "\<lbrakk>ch2sshm s h = Some (Init h', sec); valid s\<rbrakk> \<Longrightarrow> h = h' \<and> \<not> deleted (O_shm h) s \<and> h \<in> init_shms"
+by (clarsimp simp:ch2sshm_def split:if_splits option.splits)
+
+lemma root_not_deleted:
+ "\<lbrakk>deleted (O_dir []) s; valid s\<rbrakk> \<Longrightarrow> False"
+apply (induct s, simp)
+apply (frule vd_cons, frule vt_grant_os, case_tac a, auto)
+done
+
+lemma cf2sfile_fi:
+ "\<lbrakk>cf2sfile s f = Some (Init f', sec, psecopt, asecs); valid s\<rbrakk> \<Longrightarrow> f = f' \<and>
+ (if (is_file s f) then \<not> deleted (O_file f) s \<and> is_init_file f
+ else \<not> deleted (O_dir f) s \<and> is_init_dir f)"
+apply (case_tac f)
+by (auto simp:sroot_def cf2sfile_def root_is_init_dir dest!:root_is_dir' root_not_deleted
+ split:if_splits option.splits)
+
+lemma init_deled_imp_deled_s:
+ "\<lbrakk>deleted obj s; init_alive obj; sobj \<in> (s2ss s); valid s\<rbrakk> \<Longrightarrow> \<not> init_obj_related sobj obj"
+apply (rule notI)
+apply (clarsimp simp:s2ss_def)
+apply (case_tac obj, case_tac [!] obja, case_tac sobj)
+apply (auto split:option.splits if_splits dest!:cp2sproc_pi cq2smsgq_qi ch2sshm_hi cm2smsg_mi cf2sfile_fi simp:co2sobj.simps)
+apply (auto simp:cf2sfiles_def same_inode_files_def has_same_inode_prop1' is_file_def is_dir_def co2sobj.simps
+ split:option.splits t_inode_tag.splits dest!:cf2sfile_fi)
+done
+
+lemma deleted_imp_deletable_s:
+ "\<lbrakk>deleted obj s; init_alive obj; valid s\<rbrakk> \<Longrightarrow> deletable_s obj"
+apply (simp add:deletable_s_def)
+apply (frule d2s_main)
+apply (simp add:init_ss_in_def)
+apply (erule bexE)
+apply (rule_tac x = ss' in bexI)
+apply (auto simp add: init_ss_eq_def dest!:init_deled_imp_deled_s)
+apply (case_tac obj, case_tac [!] sobj)
+apply auto
+apply (erule set_mp)
+apply (simp)
+apply auto
+apply (rule_tac x = "(Init list, (aa, ab, b), ac, ba)" in bexI)
+apply auto
+done
+
+lemma init_related_imp_init_sobj:
+ "init_obj_related sobj obj \<Longrightarrow> is_init_sobj sobj"
+apply (case_tac sobj, case_tac [!] obj, auto)
+apply (rule_tac x = "(Init list, (aa, ab, b), ac, ba)" in bexI, auto)
+done
+
+theorem undeletable_s_complete:
+ assumes undel_s: "undeletable_s obj"
+ shows "undeletable obj"
+proof-
+ from undel_s have init_alive: "init_alive obj"
+ and alive_s: "\<forall> ss \<in> static. \<exists> sobj \<in> ss. init_obj_related sobj obj"
+ using undeletable_s_def by auto
+ have "\<not> (\<exists> s. valid s \<and> deleted obj s)"
+ proof
+ assume "\<exists> s. valid s \<and> deleted obj s"
+ then obtain s where vs: "valid s" and del: "deleted obj s" by auto
+ from vs have vss: "s2ss s \<propto> static" by (rule d2s_main)
+ with alive_s obtain sobj where in_ss: "sobj \<in> (s2ss s)"
+ and related: "init_obj_related sobj obj"
+ apply (simp add:init_ss_in_def init_ss_eq_def)
+ apply (erule bexE, erule_tac x= ss' in ballE)
+ apply (auto dest:init_related_imp_init_sobj)
+ done
+ from init_alive del vs have "deletable_s obj"
+ by (auto elim:deleted_imp_deletable_s)
+ with alive_s
+ show False by (auto simp:deletable_s_def)
+ qed
+ with init_alive show ?thesis
+ by (simp add:undeletable_def)
+qed
+
+theorem final_offer:
+ "\<lbrakk>undeletable_s obj; \<not> taintable_s obj; init_alive obj\<rbrakk> \<Longrightarrow> \<not> taintable obj"
+apply (erule swap)
+by (simp add:static_complete undeletable_s_complete)
+
+(************** static \<rightarrow> dynamic ***************)
+
+
+lemma set_eq_D:
+ "\<lbrakk>x \<in> S; {x. P x} = S\<rbrakk> \<Longrightarrow> P x"
+by auto
+
+lemma cqm2sms_prop1:
+ "\<lbrakk>cqm2sms s q queue = Some sms; sm \<in> set sms\<rbrakk> \<Longrightarrow> \<exists> m. cm2smsg s q m = Some sm"
+apply (induct queue arbitrary:sms)
+apply (auto simp:cqm2sms.simps split:option.splits)
+done
+
+lemma sq_sm_prop:
+ "\<lbrakk>sm \<in> set sms; cq2smsgq s q = Some (qi, qsec, sms); valid s\<rbrakk>
+ \<Longrightarrow> \<exists> m. cm2smsg s q m = Some sm"
+by (auto simp:cq2smsgq_def split: option.splits intro:cqm2sms_prop1)
+
+declare co2sobj.simps [simp add]
+
+lemma subseteq_D:
+ "\<lbrakk> S \<subseteq> {x. P x}; x \<in> S\<rbrakk> \<Longrightarrow> P x"
+by auto
+
+lemma "\<lbrakk>tainted_s ss sobj; ss \<in> static; is_init_sobj sobj\<rbrakk>
+ \<Longrightarrow> \<exists> s. valid s \<and> co2sobj s obj = Some sobj \<and> obj \<in> tainted s"
+apply (drule s2d_main')
+apply (erule exE, erule conjE, simp add:s2ss_def init_ss_eq_def, erule conjE)
+apply (rule_tac x = s in exI, simp)
+apply (case_tac sobj, simp_all only:tainted_s.simps)
+thm set_eq_D
+apply (simp split:option.splits)
+apply (erule conjE, drule_tac S = ss in set_eq_D, simp, (erule exE|erule conjE)+)
+apply (rule_tac x = obj in exI, simp)
+apply (case_tac obj, (simp split:option.splits if_splits)+)
+
+lemma tainted_s_imp_tainted:
+ "\<lbrakk>tainted_s ss sobj; ss \<in> static; init_obj_related sobj obj\<rbrakk>
+ \<Longrightarrow> \<exists> s. valid s \<and> co2sobj s obj = Some sobj \<and> obj \<in> tainted s"
+apply (drule s2d_main')
+apply (erule exE, erule conjE, simp add:s2ss_def init_ss_eq_def, erule conjE)
+apply (rule_tac x = s in exI, simp)
+apply (case_tac sobj, simp_all)
+apply (case_tac[!] obj, simp_all del:co2sobj.simps)
+apply (simp split:option.splits)
+apply (erule conjE, drule_tac S = ss in set_eq_D, simp, (erule exE|erule conjE)+)
+apply (rule_tac x = obj in exI, simp)
+apply (case_tac obj, (simp split:option.splits if_splits)+)
+
+apply (erule conjE, drule_tac S = ss in set_eq_D, simp, (erule exE|erule conjE)+)
+apply (rule_tac x = obj in exI, simp)
+apply (case_tac obj, (simp split:option.splits if_splits)+)
+sorry
+
+
+
+
+lemma tainted_s_imp_tainted:
+ "\<lbrakk>tainted_s ss sobj; ss \<in> static\<rbrakk> \<Longrightarrow> \<exists> s obj. valid s \<and> co2sobj s obj = Some sobj \<and> obj \<in> tainted s"
+apply (drule s2d_main)
+apply (erule exE, erule conjE, simp add:s2ss_def)
+apply (rule_tac x = s in exI, simp)
+apply (case_tac sobj, simp_all)
+apply (erule conjE, drule_tac S = ss in set_eq_D, simp, (erule exE|erule conjE)+)
+apply (rule_tac x = obj in exI, simp)
+apply (case_tac obj, (simp split:option.splits if_splits)+)
+
+apply (erule conjE, drule_tac S = ss in set_eq_D, simp, (erule exE|erule conjE)+)
+apply (rule_tac x = obj in exI, simp)
+apply (case_tac obj, (simp split:option.splits if_splits)+)
+done
+
+lemma has_same_inode_prop3:
+ "has_same_inode s f f' \<Longrightarrow> has_same_inode s f' f"
+by (auto simp:has_same_inode_def)
+
+theorem static_sound:
+ assumes tbl_s: "taintable_s obj"
+ shows "taintable obj"
+proof-
+ from tbl_s obtain ss sobj where static: "ss \<in> static"
+ and sobj: "tainted_s ss sobj" and related: "init_obj_related sobj obj"
+ and init_alive: "init_alive obj" by (auto simp:taintable_s_def)
+ from static sobj tainted_s_imp_tainted
+ obtain s obj' where co2sobj: "co2sobj s obj' = Some sobj"
+ and tainted': "obj' \<in> tainted s" and vs: "valid s" by blast
+
+ from co2sobj related vs
+ have eq:"obj = obj' \<or> (\<exists> f f'. obj = O_file f \<and> obj' = O_file f' \<and> has_same_inode s f f')"
+ apply (case_tac obj', case_tac [!] obj)
+ apply (auto split:option.splits if_splits dest!:cp2sproc_pi cq2smsgq_qi ch2sshm_hi cm2smsg_mi cf2sfile_fi)
+ apply (auto simp:cf2sfiles_def same_inode_files_def has_same_inode_def is_file_def is_dir_def
+ split:option.splits t_inode_tag.splits dest!:cf2sfile_fi)
+ done
+ with tainted' vs have tainted: "obj \<in> tainted s"
+ by (auto dest:has_same_inode_prop3 intro:has_same_inode_tainted)
+ from sobj related init_alive have "appropriate obj"
+ by (case_tac obj, case_tac [!] sobj, auto)
+ with vs init_alive tainted
+ show ?thesis by (auto simp:taintable_def)
+qed
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Finite_current.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,108 @@
+theory Finite_current
+imports Main Valid_prop Flask Flask_type Proc_fd_of_file_prop
+begin
+
+context flask begin
+
+lemma finite_cf: "valid \<tau> \<Longrightarrow> finite (current_files \<tau>)"
+apply (induct \<tau>)
+apply (simp add:current_files_def inum_of_file.simps)
+apply (rule_tac B = "init_files" in finite_subset)
+apply (clarsimp dest!:inof_has_file_tag, simp add:init_finite_sets)
+
+apply (frule vt_grant_os, frule vd_cons, simp, case_tac a)
+
+apply (auto simp:current_files_def os_grant.simps inum_of_file.simps split:if_splits option.splits)
+apply (rule_tac B = "insert list {f. \<exists>i. inum_of_file \<tau> f = Some i}" in finite_subset, clarsimp, simp)
+apply (rule_tac B = "{f. \<exists>i. inum_of_file \<tau> f = Some i}" in finite_subset, clarsimp, simp)
+apply (rule_tac B = "{f. \<exists>i. inum_of_file \<tau> f = Some i}" in finite_subset, clarsimp, simp)
+apply (rule_tac B = "{f. \<exists>i. inum_of_file \<tau> f = Some i}" in finite_subset, clarsimp, simp)
+apply (rule_tac B = "insert list {f. \<exists>i. inum_of_file \<tau> f = Some i}" in finite_subset, clarsimp, simp)
+done
+
+lemma finite_cp: "finite (current_procs \<tau>)"
+apply (induct \<tau>)
+apply (simp add:current_procs.simps init_finite_sets)
+apply (case_tac a, auto simp:current_procs.simps)
+done
+
+lemma finite_cfd: "valid \<tau> \<Longrightarrow> finite (current_proc_fds \<tau> p)"
+apply (induct \<tau> arbitrary:p)
+apply (simp add:current_proc_fds.simps init_finite_sets)
+apply (frule vd_cons, frule vt_grant_os, case_tac a, auto simp:current_proc_fds.simps)
+apply (erule finite_subset)
+apply (frule_tac s = \<tau> and p = nat in file_fds_subset_pfds)
+apply (erule finite_subset, simp)
+apply (erule finite_subset)
+apply (frule_tac s = \<tau> and p = nat1 in file_fds_subset_pfds)
+apply (erule finite_subset, simp)
+done
+
+lemma finite_pair: "\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow> finite {(x, y). x \<in> A \<and> y \<in> B}"
+by auto
+
+lemma finite_UN_I': "\<lbrakk>finite X; \<forall> x. x \<in> X \<longrightarrow> finite (f x)\<rbrakk> \<Longrightarrow> finite {(x, y). x \<in> X \<and> y \<in> f x}"
+apply (frule_tac B = f in finite_UN_I, simp)
+apply (drule_tac finite_pair, simp)
+apply (rule_tac B = "{(x, y). x \<in> X \<and> y \<in> (\<Union>a\<in>X. f a)}" in finite_subset, auto)
+done
+
+lemma finite_init_netobjs: "finite init_sockets"
+apply (subgoal_tac "finite {(p, fd). p \<in> init_procs \<and> fd \<in> init_fds_of_proc p}")
+apply (rule_tac B = "{(p, fd). p \<in> init_procs \<and> fd \<in> init_fds_of_proc p}" in finite_subset)
+apply (clarsimp dest!:init_socket_has_inode, simp)
+using init_finite_sets finite_UN_I'
+by (metis Collect_mem_eq SetCompr_Sigma_eq internal_split_def)
+
+lemma finite_cn_aux: "valid \<tau> \<Longrightarrow> finite {s. \<exists>i. inum_of_socket \<tau> s = Some i}"
+apply (induct \<tau>)
+apply (rule_tac B = "init_sockets" in finite_subset)
+apply (clarsimp simp:inum_of_socket.simps dest!:inos_has_sock_tag, simp add:finite_init_netobjs)
+
+apply (frule vd_cons, frule vt_grant_os, simp, case_tac a)
+apply (auto split:option.splits if_splits)
+apply (rule_tac B = "{s. \<exists>i. inum_of_socket \<tau> s = Some i}" in finite_subset, clarsimp split:if_splits, simp)
+apply (rule_tac B = "{s. \<exists>i. inum_of_socket \<tau> s = Some i} \<union> {(p, fd). \<exists> i. inum_of_socket \<tau> (nat1, fd) = Some i \<and> p = nat2 \<and> fd \<in> set}" in finite_subset, clarsimp split:if_splits)
+apply (simp only:finite_Un, rule conjI, simp)
+apply (rule_tac B = "{(p, fd). \<exists> i. inum_of_socket \<tau> (nat1, fd) = Some i \<and> p = nat2}" in finite_subset, clarsimp)
+apply (drule_tac h = "\<lambda> (p, fd). if (p = nat1) then (nat2, fd) else (p, fd)" in finite_imageI)
+apply (rule_tac B = "((\<lambda>(p, fd). if p = nat1 then (nat2, fd) else (p, fd)) ` {a. \<exists>i. inum_of_socket \<tau> a = Some i})" in finite_subset)
+apply (rule subsetI,erule CollectE, case_tac x, simp, (erule exE|erule conjE)+)
+unfolding image_def
+apply (rule CollectI, rule_tac x = "(nat1, b)" in bexI, simp+)
+apply (rule_tac B = "{s. \<exists>i. inum_of_socket \<tau> s = Some i}" in finite_subset, clarsimp split:if_splits, simp)+
+apply (rule_tac B = "insert (nat1, nat2) {s. \<exists>i. inum_of_socket \<tau> s = Some i}" in finite_subset, clarsimp, simp)
+apply (rule_tac B = "insert (nat1, nat4) {s. \<exists>i. inum_of_socket \<tau> s = Some i}" in finite_subset, clarsimp, simp)
+done
+
+lemma finite_cn: "valid \<tau> \<Longrightarrow> finite (current_sockets \<tau>)"
+apply (simp add:current_sockets_def inum_of_socket.simps)
+using finite_cn_aux[where \<tau> = \<tau>] by auto
+
+lemma finite_cm: "finite (current_msgqs \<tau>)"
+apply (induct \<tau>) defer
+apply (case_tac a, auto simp: init_finite_sets)
+done
+
+
+lemma finite_option: "finite {x. \<exists> y. f x = Some y} \<Longrightarrow> finite {y. \<exists> x. f x = Some y}"
+apply (drule_tac h = f in finite_imageI)
+apply (clarsimp simp only:image_def)
+apply (rule_tac f = Some in finite_imageD)
+apply (rule_tac B = "{y. \<exists>x. (\<exists>y. f x = Some y) \<and> y = f x}" in finite_subset)
+unfolding image_def
+apply auto
+done
+
+lemma finite_ci: "valid \<tau> \<Longrightarrow> finite (current_inode_nums \<tau>)"
+apply (simp add:current_inode_nums_def current_file_inums_def current_sock_inums_def)
+apply (rule conjI, drule finite_cf, simp add:current_files_def, erule finite_option)
+using finite_cn[where \<tau> = \<tau>]
+apply (simp add:current_sockets_def, drule_tac finite_option, simp)
+done
+
+end
+
+end
+
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Flask.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,1610 @@
+theory Flask
+imports Main Flask_type OS_type_def File_renaming
+begin
+
+(****** Policy-language-level rule match/auxiliary functions ******)
+
+fun patt_match :: "'a patt \<Rightarrow> 'a \<Rightarrow> bool"
+where
+ "patt_match (Single x) y = (x = y)"
+| "patt_match (Set s) x = (x \<in> s)"
+| "patt_match (Tilde s) x = (x \<notin> s)"
+| "patt_match Asterisk x = True"
+
+(* match a target with a patt, but including the SELF special
+ * patt_match_self patt target source
+ *)
+fun patt_match_self :: "(type_t patt) target_with_self_t \<Rightarrow> type_t \<Rightarrow> type_t \<Rightarrow> bool"
+where
+ "patt_match_self Self tt st = (tt = st)"
+| "patt_match_self (Not_self (Single p)) tt st = (tt = p)"
+| "patt_match_self (Not_self (Set p)) tt st = (tt \<in> p)"
+| "patt_match_self (Not_self (Tilde p)) tt st = (tt \<notin> p)"
+| "patt_match_self (Not_self Asterisk) tt st = True"
+
+(* list lookup: find the first one *)
+fun lookup :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a option"
+where
+ "lookup [] P = None"
+| "lookup (x # xs) P = (if (P x) then Some x else lookup xs P)"
+
+(* list member: if exists an element satisfy predicate P *)
+fun member :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
+where
+ "member [] P = False"
+| "member (x # xs) P = (if (P x) then True else member xs P)"
+
+(******* socket related aux functions ********)
+fun is_remote_request :: "t_event \<Rightarrow> bool"
+where
+ "is_remote_request (Accept p fd addr port fd' i) = True"
+| "is_remote_request (Connect p fd addr) = True"
+| "is_remote_request (SendSock p fd) = True"
+| "is_remote_request (RecvSock p fd) = True"
+| "is_remote_request e = False"
+
+fun ip_of_addr:: "t_sock_addr \<Rightarrow> t_inet_addr"
+where
+ "ip_of_addr (INET_ADDR saddr sport) = saddr"
+
+fun port_of_addr:: "t_sock_addr \<Rightarrow> t_socket_port"
+where
+ "port_of_addr (INET_ADDR saddr sport) = sport"
+
+fun after_bind:: "t_sock_state option \<Rightarrow> bool"
+where
+ "after_bind None = False"
+| "after_bind (Some SS_create) = False"
+| "after_bind _= True" (* after bind , it should have local addr *)
+
+fun after_conacc:: "t_sock_state option \<Rightarrow> bool"
+where
+ "after_conacc None = False"
+| "after_conacc (Some SS_create) = False"
+| "after_conacc (Some SS_bind) = False"
+| "after_conacc (Some SS_listen) = False"
+| "after_conacc _ = True" (* after connect or accept, it should have remote addr *)
+
+(* inet_addr_compare addr1 addr2 \<equiv> addr1 \<le> addr2; but in \<le>'s def,there is a \<exists>, is not a function*)
+fun subnet_addr :: "t_inet_addr \<Rightarrow> t_inet_addr \<Rightarrow> bool"
+where
+ "subnet_addr [] addr = True"
+| "subnet_addr (x#xs) [] = False"
+| "subnet_addr (x#xs) (y#ys) = ((x = y) \<and> subnet_addr xs ys)"
+
+definition port_in_range :: "t_socket_port \<Rightarrow> t_socket_port \<Rightarrow> t_socket_port \<Rightarrow> bool"
+where
+ "port_in_range minp maxp p \<equiv> ((p < maxp \<or> p = maxp) \<and> (p > minp \<or> p = minp))"
+
+(* initial value of initiate state constrains aux function *)
+definition bidirect_in_init :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b option) \<Rightarrow> bool"
+where
+ "bidirect_in_init S f \<equiv> (\<forall> a. a \<in> S \<longrightarrow> (\<exists> b. f a = Some b)) \<and>
+ (\<forall> a b. f a = Some b \<longrightarrow> a \<in> S)"
+
+fun is_file_itag :: "t_inode_tag \<Rightarrow> bool"
+where
+ "is_file_itag Tag_FILE = True"
+| "is_file_itag tag = False"
+
+fun is_dir_itag :: "t_inode_tag \<Rightarrow> bool"
+where
+ "is_dir_itag Tag_DIR = True"
+| "is_dir_itag tag = False"
+
+fun is_file_dir_itag :: "t_inode_tag \<Rightarrow> bool"
+where
+ "is_file_dir_itag Tag_FILE = True"
+| "is_file_dir_itag Tag_DIR = True"
+| "is_file_dir_itag tag = False"
+
+fun is_tcp_itag :: "t_inode_tag \<Rightarrow> bool"
+where
+ "is_tcp_itag Tag_TCP_SOCK = True"
+| "is_tcp_itag tag = False"
+
+fun is_udp_itag :: "t_inode_tag \<Rightarrow> bool"
+where
+ "is_udp_itag Tag_UDP_SOCK = True"
+| "is_udp_itag tag = False"
+
+fun is_sock_itag :: "t_inode_tag \<Rightarrow> bool"
+where
+ "is_sock_itag Tag_TCP_SOCK = True"
+| "is_sock_itag Tag_UDP_SOCK = True"
+| "is_sock_itag tag = False"
+
+
+locale init =
+ fixes
+ init_files :: "t_file set"
+ and init_procs :: "t_process set"
+ and init_nodes :: "t_node set"
+ and init_sockets :: "t_socket set"
+ and init_users :: "user_t set"
+
+ and init_file_of_proc_fd :: "t_process \<Rightarrow> t_fd \<rightharpoonup> t_file"
+ and init_fds_of_proc :: "t_process \<Rightarrow> t_fd set"
+ and init_oflags_of_proc_fd:: "t_process \<Rightarrow> t_fd \<rightharpoonup> t_open_flags"
+ and init_inum_of_file :: "t_file \<rightharpoonup> t_inode_num"
+ and init_inum_of_socket:: "t_socket \<rightharpoonup> t_inode_num"
+ and init_itag_of_inum:: "t_inode_num \<rightharpoonup> t_inode_tag"
+ and init_files_hung_by_del :: "t_file set"
+(*
+ and init_file_at_writing_by:: "t_file \<Rightarrow> (t_process \<times> t_fd) set"
+*)
+ and init_socket_state :: "t_socket \<rightharpoonup> t_sock_state"
+
+ and init_msgqs :: "t_msgq set"
+ and init_msgs_of_queue:: "t_msgq \<Rightarrow> t_msg list"
+(*
+ and init_shms :: "t_shm set"
+ and init_procs_of_shm :: "t_shm \<Rightarrow> (t_process \<times> t_shm_attach_flag) set"
+ and init_flag_of_proc_shm :: "t_process \<Rightarrow> t_shm \<rightharpoonup> t_shm_attach_flag"
+*)
+
+ assumes
+ parent_file_in_init: "\<And> f pf. \<lbrakk>parent f = Some pf; f \<in> init_files\<rbrakk> \<Longrightarrow> pf \<in> init_files"
+ and root_is_dir: "\<exists> i. init_inum_of_file [] = Some i \<and> init_itag_of_inum i = Some Tag_DIR"
+ and init_file_has_inum: "\<And> f. f \<in> init_files \<Longrightarrow> \<exists> im. init_inum_of_file f = Some im"
+ and inof_has_file_tag: "\<And> f im. init_inum_of_file f = Some im \<Longrightarrow> f \<in> init_files \<and> (\<exists> tag. init_itag_of_inum im = Some tag \<and> is_file_dir_itag tag)"
+ and init_file_no_son: "\<And> f im. \<lbrakk>init_inum_of_file f = Some im; init_itag_of_inum im = Some Tag_FILE\<rbrakk> \<Longrightarrow> \<not> (\<exists> f' \<in> init_files. parent f' = Some f)"
+ and init_parentf_is_dir: "\<And> f pf im. \<lbrakk>parent f = Some pf; f \<in> init_files; init_inum_of_file pf = Some im\<rbrakk> \<Longrightarrow> init_itag_of_inum im = Some Tag_DIR"
+
+ and init_files_hung_valid: "\<And> f. f \<in> init_files_hung_by_del \<Longrightarrow> f \<in> init_files \<and> (\<exists> p fd. init_file_of_proc_fd p fd = Some f)"
+ and init_files_hung_valid': "\<And> f im. \<lbrakk>f \<in> init_files_hung_by_del; init_inum_of_file f = Some im\<rbrakk> \<Longrightarrow>
+ init_itag_of_inum im = Some Tag_FILE \<or> (init_itag_of_inum im = Some Tag_DIR \<and> \<not> (\<exists> f'. f' \<in> init_files \<and> f \<prec> f'))"
+
+ and init_procfds_valid: "\<And> p fd. fd \<in> init_fds_of_proc p \<Longrightarrow> p \<in> init_procs \<and> ((\<exists> f \<in> init_files. init_file_of_proc_fd p fd = Some f) \<or> (p, fd) \<in> init_sockets)"
+ and init_filefd_valid: "\<And> p fd f. init_file_of_proc_fd p fd = Some f \<Longrightarrow> fd \<in> init_fds_of_proc p \<and> (\<exists> im. init_inum_of_file f = Some im \<and> init_itag_of_inum im = Some Tag_FILE) \<and> p \<in> init_procs \<and> (\<exists> flags. init_oflags_of_proc_fd p fd = Some flags) \<and> (p, fd) \<notin> init_sockets"
+ and init_fileflag_valid: "\<And> p fd flags. init_oflags_of_proc_fd p fd = Some flags \<Longrightarrow> \<exists> f. init_file_of_proc_fd p fd = Some f"
+
+(*
+ and init_procs_has_shm: "\<And> p h flag. (p,flag) \<in> init_procs_of_shm h \<Longrightarrow> p \<in> init_procs \<and> h \<in> init_shms \<and> init_flag_of_proc_shm p h = Some flag"
+ and init_shmflag_has_proc: "\<And> p h flag. init_flag_of_proc_shm p h = Some flag \<Longrightarrow> (p, flag) \<in> init_procs_of_shm h"
+*)
+ and init_msgs_distinct: "\<And> q. distinct (init_msgs_of_queue q)"
+ and init_msgs_valid: "\<And> m q. m \<in> set (init_msgs_of_queue q) \<Longrightarrow> q \<in> init_msgqs"
+
+ and init_socket_has_inode: "\<And> p fd. (p, fd) \<in> init_sockets \<Longrightarrow> \<exists> im. init_inum_of_socket (p, fd) = Some im \<and> p \<in> init_procs \<and> fd \<in> init_fds_of_proc p \<and> init_file_of_proc_fd p fd = None"
+ and inos_has_sock_tag: "\<And> s im. init_inum_of_socket s = Some im \<Longrightarrow> s \<in> init_sockets \<and> (\<exists> tag. init_itag_of_inum im = Some tag \<and> is_sock_itag tag)"
+(*
+ and init_netobj_has_state: "bidirect_in_init init_netobjs init_netobj_state"
+ and init_netobj_has_socktype:"bidirect_in_init init_netobjs init_netobj_socktype"
+ and init_netobj_has_laddr: "\<And> s. after_bind (init_netobj_state s) \<Longrightarrow> init_netobj_localaddr s \<noteq> None"
+ and init_netobj_has_raddr: "\<And> s. after_conacc (init_netobj_state s) \<Longrightarrow> init_netobj_remoteaddr s \<noteq> None"
+*)
+
+ and init_finite_sets: "finite init_files \<and> finite init_procs \<and> (\<forall> p. finite (init_fds_of_proc p)) \<and> finite init_msgqs \<and> finite init_users"
+
+begin
+
+definition is_init_file:: "t_file \<Rightarrow> bool"
+where
+ "is_init_file f \<equiv> (case (init_inum_of_file f) of
+ Some i \<Rightarrow> (case (init_itag_of_inum i) of
+ Some Tag_FILE \<Rightarrow> True
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+
+definition is_init_dir:: "t_file \<Rightarrow> bool"
+where
+ "is_init_dir f \<equiv> (case (init_inum_of_file f) of
+ Some i \<Rightarrow> (case (init_itag_of_inum i) of
+ Some Tag_DIR \<Rightarrow> True
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+
+definition is_init_tcp_sock:: "t_socket \<Rightarrow> bool"
+where
+ "is_init_tcp_sock s \<equiv> (case (init_inum_of_socket s) of
+ Some i \<Rightarrow> (case (init_itag_of_inum i) of
+ Some Tag_TCP_SOCK \<Rightarrow> True
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+
+definition is_init_udp_sock:: "t_socket \<Rightarrow> bool"
+where
+ "is_init_udp_sock s \<equiv> (case (init_inum_of_socket s) of
+ Some i \<Rightarrow> (case (init_itag_of_inum i) of
+ Some Tag_UDP_SOCK \<Rightarrow> True
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+
+definition init_same_inode_files :: "t_file \<Rightarrow> t_file set"
+where
+ "init_same_inode_files f \<equiv> if is_init_file f then {f'. init_inum_of_file f = init_inum_of_file f'} else {}"
+
+fun init_alive :: "t_object \<Rightarrow> bool"
+where
+ "init_alive (O_proc p) = (p \<in> init_procs)"
+| "init_alive (O_file f) = (is_init_file f)"
+| "init_alive (O_dir f) = (is_init_dir f)"
+| "init_alive (O_fd p fd) = (fd \<in> init_fds_of_proc p)"
+| "init_alive (O_node n) = (n \<in> init_nodes)"
+| "init_alive (O_tcp_sock k) = (is_init_tcp_sock k)"
+| "init_alive (O_udp_sock k) = (is_init_udp_sock k)"
+(*
+| "init_alive (O_shm h) = (h \<in> init_shms)" *)
+| "init_alive (O_msgq q) = (q \<in> init_msgqs)"
+| "init_alive (O_msg q m) = (m \<in> set (init_msgs_of_queue q) \<and> q \<in> init_msgqs)"
+
+
+(************ system listeners, event-related ***********)
+
+fun file_of_proc_fd :: "t_state \<Rightarrow> t_process \<Rightarrow> t_fd \<Rightarrow> t_file option"
+where
+ "file_of_proc_fd [] p' fd' = (init_file_of_proc_fd p' fd')"
+| "file_of_proc_fd (Open p f flags fd iopt # \<tau>) p' fd' =
+ (if (p' = p \<and> fd = fd') then Some f else file_of_proc_fd \<tau> p' fd')"
+| "file_of_proc_fd (CloseFd p fd # \<tau>) p' fd' =
+ (if (p = p' \<and> fd = fd') then None else file_of_proc_fd \<tau> p' fd')"
+(*deleted: if (f\<^isub>3 \<preceq> f\<^isub>1) then None else *)
+(*
+| "file_of_proc_fd (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) p' fd =
+ (case (file_of_proc_fd \<tau> p' fd) of
+ Some f\<^isub>1 \<Rightarrow> (if (f\<^isub>2 \<preceq> f\<^isub>1) then Some (file_after_rename f\<^isub>2 f\<^isub>3 f\<^isub>1)
+ else Some f\<^isub>1)
+ | None \<Rightarrow> None )"
+*)
+| "file_of_proc_fd (Execve p f fds # \<tau>) p' fd =
+ (if (p' = p \<and> fd \<in> fds) then file_of_proc_fd \<tau> p fd
+ else if (p' = p) then None
+ else file_of_proc_fd \<tau> p' fd) "
+| "file_of_proc_fd (Clone p p' fds # \<tau>) p'' fd =
+ (if (p'' = p' \<and> fd \<in> fds) then file_of_proc_fd \<tau> p fd
+ else if (p'' = p') then None
+ else file_of_proc_fd \<tau> p'' fd)"
+| "file_of_proc_fd (Kill p\<^isub> p\<^isub>' # \<tau>) p'' fd =
+ (if (p'' = p\<^isub>') then None else file_of_proc_fd \<tau> p'' fd)"
+| "file_of_proc_fd (Exit p # \<tau>) p' fd =
+ (if (p = p') then None else file_of_proc_fd \<tau> p' fd)"
+| "file_of_proc_fd (_ # \<tau>) p fd = file_of_proc_fd \<tau> p fd"
+
+definition proc_fd_of_file :: "t_state \<Rightarrow> t_file \<Rightarrow> (t_process \<times> t_fd) set"
+where
+ "proc_fd_of_file \<tau> f = {(p, fd) | p fd. file_of_proc_fd \<tau> p fd = Some f}"
+
+definition proc_file_fds :: "t_state \<Rightarrow> t_process \<Rightarrow> t_fd set"
+where
+ "proc_file_fds s p \<equiv> {fd. \<exists> f. file_of_proc_fd s p fd = Some f}"
+
+definition init_proc_file_fds:: "t_process \<Rightarrow> t_fd set"
+where
+ "init_proc_file_fds p \<equiv> {fd. \<exists> f. init_file_of_proc_fd p fd = Some f}"
+
+(****** files and directories ******)
+
+fun files_hung_by_del :: "t_state \<Rightarrow> t_file set"
+where
+ "files_hung_by_del [] = init_files_hung_by_del"
+| "files_hung_by_del (UnLink p f # \<tau>) = (
+ if (proc_fd_of_file \<tau> f = {}) then files_hung_by_del \<tau>
+ else insert f (files_hung_by_del \<tau>))"
+(* | files_hung_by_del (Rmdir p f) is not need, for open can only apply to file not dir *)
+| "files_hung_by_del (CloseFd p fd # \<tau>) = (
+ case (file_of_proc_fd \<tau> p fd) of
+ Some f \<Rightarrow> (if (proc_fd_of_file \<tau> f = {(p, fd)})
+ then files_hung_by_del \<tau> - {f}
+ else files_hung_by_del \<tau>)
+ | None \<Rightarrow> files_hung_by_del \<tau> )"
+(*
+| "files_hung_by_del (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) = {file_after_rename f\<^isub>2 f\<^isub>3 f\<^isub>1| f\<^isub>1. f\<^isub>1 \<in> files_hung_by_del \<tau>}" (* for rename(2) does not infect on fd of the file, see man -S 2 rename *)
+*)
+| "files_hung_by_del (e # \<tau>) = files_hung_by_del \<tau>"
+declare files_hung_by_del.simps [simp del]
+
+fun inum_of_file :: "t_state \<Rightarrow> (t_file \<rightharpoonup> t_inode_num)"
+where
+ "inum_of_file [] = init_inum_of_file"
+| "inum_of_file (Open p f flags fd (Some inum) # \<tau>) = (inum_of_file \<tau>) (f := Some inum)"
+| "inum_of_file (CloseFd p fd # \<tau>) = (
+ case (file_of_proc_fd \<tau> p fd) of
+ Some f \<Rightarrow> ( if ((proc_fd_of_file \<tau> f = {(p, fd)}) \<and> (f \<in> files_hung_by_del \<tau>))
+ then (inum_of_file \<tau>) (f := None)
+ else inum_of_file \<tau> )
+ | _ \<Rightarrow> (inum_of_file \<tau>) )"
+| "inum_of_file (UnLink p f # \<tau>) = (
+ if (proc_fd_of_file \<tau> f = {})
+ then (inum_of_file \<tau>) (f := None)
+ else inum_of_file \<tau> )"
+| "inum_of_file (Rmdir p f # \<tau>) = (inum_of_file \<tau>) (f := None)"
+| "inum_of_file (Mkdir p f ino # \<tau>) = (inum_of_file \<tau>) (f:= Some ino)"
+(*
+| "inum_of_file (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>) = (inum_of_file \<tau>) (f\<^isub>2 := inum_of_file \<tau> f\<^isub>1)"
+*)
+(*
+| "inum_of_file (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) = (\<lambda> f.
+ if (f\<^isub>2 \<preceq> f) then None
+ else if (f\<^isub>3 \<preceq> f) then inum_of_file \<tau> (file_before_rename f\<^isub>2 f\<^isub>3 f)
+ else inum_of_file \<tau> f )"
+*)
+| "inum_of_file (e # \<tau>) = inum_of_file \<tau>"
+
+definition current_files :: "t_state \<Rightarrow> t_file set"
+where
+ "current_files \<tau> = {f. \<exists> i. inum_of_file \<tau> f = Some i}"
+
+(* experimentary: for the delete obj's (file or socket) inode num, it is unnecessary for itag_of_inum to be NONE, cause this is the situation blocked by inum_of_file / inum_of_socket !!! *)
+(* delete a file, just recycle its inode num, but not destroy its inode \<dots>, so it is irrelative to these events, like closeFd, Rmdir, UnLink \<dots>*)
+fun itag_of_inum :: "t_state \<Rightarrow> (t_inode_num \<rightharpoonup> t_inode_tag)"
+where
+ "itag_of_inum [] = init_itag_of_inum"
+| "itag_of_inum (Open p f flags fd (Some ino) # \<tau>) = (itag_of_inum \<tau>) (ino := Some Tag_FILE)"
+| "itag_of_inum (Mkdir p f ino # \<tau>) = (itag_of_inum \<tau>) (ino := Some Tag_DIR)"
+| "itag_of_inum (CreateSock p af st fd ino # \<tau>) =
+ (case st of
+ STREAM \<Rightarrow> (itag_of_inum \<tau>) (ino := Some Tag_TCP_SOCK)
+ | DGRAM \<Rightarrow> (itag_of_inum \<tau>) (ino := Some Tag_UDP_SOCK) )"
+| "itag_of_inum (Accept p fd addr lport' fd' ino # \<tau>) =
+ (itag_of_inum \<tau>) (ino := Some Tag_TCP_SOCK)"
+| "itag_of_inum (_ # \<tau>) = itag_of_inum \<tau>"
+
+definition is_file:: "t_state \<Rightarrow> t_file \<Rightarrow> bool"
+where
+ "is_file \<tau> f \<equiv> (case (inum_of_file \<tau> f) of
+ Some i \<Rightarrow> (case (itag_of_inum \<tau> i) of
+ Some Tag_FILE \<Rightarrow> True
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+
+definition is_dir:: "t_state \<Rightarrow> t_file \<Rightarrow> bool"
+where
+ "is_dir \<tau> f \<equiv> (case (inum_of_file \<tau> f) of
+ Some i \<Rightarrow> (case (itag_of_inum \<tau> i) of
+ Some Tag_DIR \<Rightarrow> True
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+
+definition dir_is_empty :: "t_state \<Rightarrow> t_file \<Rightarrow> bool"
+where
+ "dir_is_empty \<tau> f \<equiv> is_dir \<tau> f \<and> \<not> (\<exists> f'. f' \<in> current_files \<tau> \<and> f \<prec> f')"
+
+(*
+definition same_inode_files :: "t_state \<Rightarrow> t_file \<Rightarrow> t_file set"
+where
+ "same_inode_files s f \<equiv> if is_file s f then {f'. inum_of_file s f = inum_of_file s f'} else {}"
+
+definition has_same_inode :: "t_state \<Rightarrow> t_file \<Rightarrow> t_file \<Rightarrow> bool"
+where
+ "has_same_inode s fa fb \<equiv> fb \<in> same_inode_files s fa"
+*)
+
+fun inum_of_socket :: "t_state \<Rightarrow> (t_socket \<rightharpoonup> t_inode_num)"
+where
+ "inum_of_socket [] = init_inum_of_socket"
+| "inum_of_socket (CloseFd p fd # \<tau>) =
+ (inum_of_socket \<tau>) ((p, fd):= None)"
+| "inum_of_socket (CreateSock p af st fd ino # \<tau>) =
+ (inum_of_socket \<tau>) ((p, fd) := Some ino)"
+| "inum_of_socket (Accept p fd addr lport' fd' ino # \<tau>) =
+ (inum_of_socket \<tau>) ((p, fd') := Some ino)"
+| "inum_of_socket (Clone p p' fds # \<tau>) =
+ (\<lambda> (p'', fd). if (p'' = p' \<and> fd \<in> fds) then inum_of_socket \<tau> (p, fd)
+ else if (p'' = p') then None
+ else inum_of_socket \<tau> (p'',fd))"
+| "inum_of_socket (Execve p f fds # \<tau>) =
+ (\<lambda> (p', fd). if (p' = p \<and> fd \<in> fds) then inum_of_socket \<tau> (p, fd)
+ else if (p' = p) then None
+ else inum_of_socket \<tau> (p', fd))"
+| "inum_of_socket (Kill p\<^isub>1 p\<^isub>2 # \<tau>) =
+ (\<lambda> (p, fd). if (p = p\<^isub>2) then None else inum_of_socket \<tau> (p, fd) )"
+| "inum_of_socket (Exit p # \<tau>) =
+ (\<lambda> (p', fd). if (p' = p) then None else inum_of_socket \<tau> (p', fd))"
+| "inum_of_socket (_ # \<tau>) = inum_of_socket \<tau>"
+
+definition current_sockets :: "t_state \<Rightarrow> t_socket set"
+where
+ "current_sockets \<tau> = {s. \<exists> i. inum_of_socket \<tau> s = Some i}"
+
+definition is_tcp_sock :: "t_state \<Rightarrow> t_socket \<Rightarrow> bool"
+where
+ "is_tcp_sock \<tau> s \<equiv> (case (inum_of_socket \<tau> s) of
+ Some i \<Rightarrow> (case (itag_of_inum \<tau> i) of
+ Some Tag_TCP_SOCK \<Rightarrow> True
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+
+definition is_udp_sock :: "t_state \<Rightarrow> t_socket \<Rightarrow> bool"
+where
+ "is_udp_sock \<tau> s \<equiv> (case (inum_of_socket \<tau> s) of
+ Some i \<Rightarrow> (case (itag_of_inum \<tau> i) of
+ Some Tag_UDP_SOCK \<Rightarrow> True
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+
+fun current_procs :: "t_state \<Rightarrow> t_process set"
+where
+ "current_procs [] = init_procs"
+| "current_procs (Clone p p' fds # \<tau>) = (insert p' (current_procs \<tau>))"
+| "current_procs (Exit p # \<tau>) = (current_procs \<tau> - {p})"
+| "current_procs (Kill p p' # \<tau>) = (current_procs \<tau> - {p'})"
+| "current_procs (_ # \<tau>) = current_procs \<tau>"
+
+fun current_proc_fds :: "t_state \<Rightarrow> t_process \<Rightarrow> t_fd set"
+where
+ "current_proc_fds [] = init_fds_of_proc"
+| "current_proc_fds (Open p f flags fd iopt # \<tau>) =
+ (current_proc_fds \<tau>) (p:= insert fd (current_proc_fds \<tau> p))"
+| "current_proc_fds (CreateSock p sf st newfd newi # \<tau>) =
+ (current_proc_fds \<tau>) (p:= insert newfd (current_proc_fds \<tau> p))"
+| "current_proc_fds (Accept p fd raddr port fd' ino # \<tau>) =
+ (current_proc_fds \<tau>) (p:= insert fd' (current_proc_fds \<tau> p))"
+| "current_proc_fds (CloseFd p fd # \<tau>) =
+ (current_proc_fds \<tau>) (p:= (current_proc_fds \<tau> p) - {fd})"
+| "current_proc_fds (Clone p p' fds # \<tau>) =
+ (current_proc_fds \<tau>) (p':= fds)"
+| "current_proc_fds (Execve p f fds # \<tau>) =
+ (current_proc_fds \<tau>) (p:= fds)"
+| "current_proc_fds (Exit p # \<tau>) = (current_proc_fds \<tau>) (p := {})"
+| "current_proc_fds (Kill p p' # \<tau>) = (current_proc_fds \<tau>) (p' := {})"
+| "current_proc_fds (_ # \<tau>) = current_proc_fds \<tau>"
+
+(*
+fun current_shms :: "t_state \<Rightarrow> t_shm set"
+where
+ "current_shms [] = init_shms"
+| "current_shms (CreateShM p newshm # \<tau>) = insert newshm (current_shms \<tau>)"
+| "current_shms (DeleteShM p s # \<tau>) = (current_shms \<tau>) - {s}"
+| "current_shms (e # \<tau>) = current_shms \<tau> "
+
+fun flag_of_proc_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_shm \<Rightarrow> t_shm_attach_flag option"
+where
+ "flag_of_proc_shm [] = init_flag_of_proc_shm"
+| "flag_of_proc_shm (Attach p h flag # s) = (\<lambda> p' h'.
+ if (p' = p \<and> h' = h) then Some flag else flag_of_proc_shm s p' h')"
+| "flag_of_proc_shm (Detach p h # s) = (\<lambda> p' h'.
+ if (p' = p \<and> h' = h) then None else flag_of_proc_shm s p' h')"
+| "flag_of_proc_shm (CreateShM p h # s) = (\<lambda> p' h'.
+ if (h' = h) then None else flag_of_proc_shm s p' h')"
+| "flag_of_proc_shm (DeleteShM p h # s) = (\<lambda> p' h'.
+ if (h' = h) then None else flag_of_proc_shm s p' h')"
+| "flag_of_proc_shm (Clone p p' fds shms # s) = (\<lambda> p'' h.
+ if (p'' = p' \<and> h \<in> shms) then flag_of_proc_shm s p h
+ else if (p'' = p') then None
+ else flag_of_proc_shm s p'' h)"
+| "flag_of_proc_shm (Execve p f fds # s) = (\<lambda> p' h.
+ if (p' = p) then None else flag_of_proc_shm s p' h)"
+| "flag_of_proc_shm (Kill p p' # s) = (\<lambda> p'' h.
+ if (p'' = p') then None else flag_of_proc_shm s p'' h)"
+| "flag_of_proc_shm (Exit p # s) = (\<lambda> p' h.
+ if (p' = p) then None else flag_of_proc_shm s p' h)"
+| "flag_of_proc_shm (e # s) = flag_of_proc_shm s"
+
+fun procs_of_shm :: "t_state \<Rightarrow> t_shm \<Rightarrow> (t_process \<times> t_shm_attach_flag) set"
+where
+ "procs_of_shm [] = init_procs_of_shm"
+| "procs_of_shm (CreateShM p h # \<tau>) =
+ (procs_of_shm \<tau>) (h := {})" (* creator may not be the sharer of the shm *)
+| "procs_of_shm (Attach p h flag # \<tau>) =
+ (procs_of_shm \<tau>) (h := insert (p, flag) (procs_of_shm \<tau> h))"
+| "procs_of_shm (Detach p h # \<tau>) =
+ (procs_of_shm \<tau>) (h := (procs_of_shm \<tau> h) - {(p,flag) | flag. (p, flag) \<in> procs_of_shm \<tau> h})"
+| "procs_of_shm (DeleteShM p h # \<tau>) = (procs_of_shm \<tau>) (h := {})"
+| "procs_of_shm (Clone p p' fds shms # \<tau>) =
+ (\<lambda> h. if (h \<in> shms)
+ then (procs_of_shm \<tau> h) \<union> {(p', flag) | flag. (p, flag) \<in> procs_of_shm \<tau> h}
+ else procs_of_shm \<tau> h)"
+| "procs_of_shm (Execve p f fds # \<tau>) =
+ (\<lambda> h. procs_of_shm \<tau> h - {(p, flag) | flag. (p, flag) \<in> procs_of_shm \<tau> h})"
+| "procs_of_shm (Kill p p' # s) =
+ (\<lambda> h. (procs_of_shm s h) - {(p', flag) | flag. (p', flag) \<in> procs_of_shm s h})"
+| "procs_of_shm (Exit p # s) =
+ (\<lambda> h. (procs_of_shm s h) - {(p, flag) | flag. (p, flag) \<in> procs_of_shm s h})"
+| "procs_of_shm (e # \<tau>) = procs_of_shm \<tau>"
+
+(*
+inductive info_flow_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
+where
+ "p \<in> current_procs s \<Longrightarrow> info_flow_shm s p p"
+| "\<lbrakk>info_flow_shm s p p'; (p', SHM_RDWR) \<in> procs_of_shm s h; (p'', flag) \<in> procs_of_shm s h; p' \<noteq> p''\<rbrakk>
+ \<Longrightarrow> info_flow_shm s p p''"
+*)
+definition one_flow_shm :: "t_state \<Rightarrow> t_shm \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
+where
+ "one_flow_shm s h from to \<equiv> from \<noteq> to \<and> (from, SHM_RDWR) \<in> procs_of_shm s h \<and>
+ (\<exists> flag. (to, flag) \<in> procs_of_shm s h)"
+
+fun self_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
+where
+ "self_shm s p p' = (p = p' \<and> p' \<in> current_procs s)"
+
+definition info_flow_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
+where
+ "info_flow_shm s from to \<equiv> (self_shm s from to) \<or> (\<exists> h. one_flow_shm s h from to)"
+*)
+
+fun current_msgqs :: "t_state \<Rightarrow> t_msg set"
+where
+ "current_msgqs [] = init_msgqs"
+| "current_msgqs (CreateMsgq p q # \<tau>) = insert q (current_msgqs \<tau>)"
+| "current_msgqs (RemoveMsgq p q # \<tau>) = (current_msgqs \<tau>) - {q}"
+| "current_msgqs (_ # \<tau>) = current_msgqs \<tau>"
+
+fun msgs_of_queue :: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_msg list"
+where
+ "msgs_of_queue [] = init_msgs_of_queue"
+| "msgs_of_queue (CreateMsgq p q # \<tau>) = (msgs_of_queue \<tau>)(q := [])"
+| "msgs_of_queue (SendMsg p q m # \<tau>) = (msgs_of_queue \<tau>)(q := msgs_of_queue \<tau> q @ [m])"
+| "msgs_of_queue (RecvMsg p q m # \<tau>) = (msgs_of_queue \<tau>)(q := tl (msgs_of_queue \<tau> q))"
+| "msgs_of_queue (RemoveMsgq p q # \<tau>) = (msgs_of_queue \<tau>)(q := [])"
+| "msgs_of_queue (_ # \<tau>) = msgs_of_queue \<tau>"
+
+definition current_file_inums :: "t_state \<Rightarrow> t_inode_num set"
+where
+ "current_file_inums \<tau> \<equiv> {im. \<exists> f. inum_of_file \<tau> f = Some im}"
+
+definition current_sock_inums :: "t_state \<Rightarrow> t_inode_num set"
+where
+ "current_sock_inums \<tau> = {im. \<exists> s. inum_of_socket \<tau> s = Some im}"
+
+definition current_inode_nums :: "t_state \<Rightarrow> nat set"
+where
+ "current_inode_nums \<tau> = current_file_inums \<tau> \<union> current_sock_inums \<tau>"
+
+fun flags_of_proc_fd :: "t_state \<Rightarrow> t_process \<Rightarrow> t_fd \<Rightarrow> t_open_flags option"
+where
+ "flags_of_proc_fd [] p fd = init_oflags_of_proc_fd p fd"
+| "flags_of_proc_fd (Open p f flags fd iopt # \<tau>) p' fd' =
+ (if (p = p' \<and> fd = fd') then Some flags else flags_of_proc_fd \<tau> p' fd')"
+| "flags_of_proc_fd (CloseFd p fd # \<tau>) p' fd' =
+ (if (p = p' \<and> fd = fd') then None else flags_of_proc_fd \<tau> p' fd')"
+| "flags_of_proc_fd (CreateSock p af st fd ino # \<tau>) p' fd' =
+ (if (p = p' \<and> fd = fd') then None else flags_of_proc_fd \<tau> p' fd')"
+| "flags_of_proc_fd (Accept p fd addr lport' fd' ino # \<tau>) p' fd'' =
+ (if (p = p' \<and> fd' = fd'') then None else flags_of_proc_fd \<tau> p' fd'')"
+| "flags_of_proc_fd (Clone p p' fds # \<tau>) p'' fd =
+ (if (p' = p'' \<and> fd \<in> fds) then flags_of_proc_fd \<tau> p fd else flags_of_proc_fd \<tau> p'' fd)"
+| "flags_of_proc_fd (Execve p f fds # \<tau>) p' fd =
+ (if (p' = p \<and> fd \<in> fds) then flags_of_proc_fd \<tau> p fd else flags_of_proc_fd \<tau> p' fd)"
+| "flags_of_proc_fd (Kill p\<^isub>1 p\<^isub>2 # \<tau>) p fd =
+ (if (p = p\<^isub>2) then None else flags_of_proc_fd \<tau> p fd)"
+| "flags_of_proc_fd (Exit p # \<tau>) p' fd' =
+ (if (p = p') then None else flags_of_proc_fd \<tau> p' fd')"
+| "flags_of_proc_fd (e # \<tau>) p fd = flags_of_proc_fd \<tau> p fd"
+
+(*
+fun file_at_writing_by:: "t_state \<Rightarrow> t_file \<Rightarrow> (t_process \<times> t_fd) set"
+where
+ "file_at_writing_by [] f = init_file_at_writing_by f"
+| "file_at_writing_by (Open p f flags fd (Some inum) # \<tau>) f' = (
+ if (f' = f)
+ then ( if (is_write_flag flags)
+ then {(p, fd)}
+ else {} )
+ else file_at_writing_by \<tau> f' )"
+| "file_at_writing_by (Open p f flags fd None # \<tau>) f' = (
+ if (f' = f \<and> is_write_flag flags)
+ then insert (p, fd) (file_at_writing_by \<tau> f)
+ else file_at_writing_by \<tau> f' )"
+| "file_at_writing_by (Mkdir p f inum # \<tau>) f' =
+ (if (f' = f) then {} else file_at_writing_by \<tau> f')"
+| "file_at_writing_by (LinkHard p f f' # \<tau>) f'' =
+ (if (f'' = f') then file_at_writing_by \<tau> f else file_at_writing_by \<tau> f'')"
+| "file_at_writing_by (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = (
+ if (f\<^isub>3 \<preceq> f) then file_at_writing_by \<tau> (file_before_rename f\<^isub>2 f\<^isub>3 f)
+ else file_at_writing_by \<tau> f )"
+| "file_at_writing_by (CloseFd p fd # \<tau>) f =
+ (file_at_writing_by \<tau> f - {(p, fd)})"
+| "file_at_writing_by (Clone p p' fds shms # \<tau>) f =
+ (file_at_writing_by \<tau> f) \<union> {(p',fd)| fd. fd \<in> fds \<and> (p, fd) \<in> file_at_writing_by \<tau> f}"
+| "file_at_writing_by (Execve p f fds # \<tau>) f' =
+ (if (f' = f) then {} else file_at_writing_by \<tau> f')"
+| "file_at_writing_by (Exit p # \<tau>) f =
+ (file_at_writing_by \<tau> f - {(p, fd)| fd. (p, fd) \<in> file_at_writing_by \<tau> f})"
+| "file_at_writing_by (Kill p p' # \<tau>) f =
+ (file_at_writing_by \<tau> f - {(p', fd)| fd. (p', fd) \<in> file_at_writing_by \<tau> f})"
+| "file_at_writing_by (e # \<tau>) f = file_at_writing_by \<tau> f"
+*)
+
+definition current_users :: "t_state \<Rightarrow> user_t set"
+where
+ "current_users \<tau> \<equiv> init_users"
+
+fun update_with_shuthow :: "t_sock_state option \<Rightarrow> t_shutdown_how \<Rightarrow> t_sock_state option"
+where
+ "update_with_shuthow (Some (SS_trans Trans_Recv)) SHUT_RD = Some (SS_trans Trans_No)"
+| "update_with_shuthow (Some (SS_trans Trans_RS)) SHUT_RD = Some (SS_trans Trans_Send)"
+| "update_with_shuthow (Some (SS_trans Trans_Send)) SHUT_RD = Some (SS_trans Trans_Send)"
+| "update_with_shuthow (Some (SS_trans Trans_No)) SHUT_RD = Some (SS_trans Trans_No)"
+| "update_with_shuthow (Some (SS_trans Trans_Recv)) SHUT_WR = Some (SS_trans Trans_Recv)"
+| "update_with_shuthow (Some (SS_trans Trans_RS)) SHUT_WR = Some (SS_trans Trans_Recv)"
+| "update_with_shuthow (Some (SS_trans Trans_Send)) SHUT_WR = Some (SS_trans Trans_No)"
+| "update_with_shuthow (Some (SS_trans Trans_No)) SHUT_WR = Some (SS_trans Trans_No)"
+| "update_with_shuthow (Some (SS_trans Trans_Recv)) SHUT_RDWR = Some (SS_trans Trans_No)"
+| "update_with_shuthow (Some (SS_trans Trans_RS)) SHUT_RDWR = Some (SS_trans Trans_No)"
+| "update_with_shuthow (Some (SS_trans Trans_Send)) SHUT_RDWR = Some (SS_trans Trans_No)"
+| "update_with_shuthow (Some (SS_trans Trans_No)) SHUT_RDWR = Some (SS_trans Trans_No)"
+| "update_with_shuthow opt1 how = None"
+
+fun socket_state :: "t_state \<Rightarrow> t_socket \<Rightarrow> t_sock_state option"
+where
+ "socket_state [] = init_socket_state"
+| "socket_state (CloseFd p fd # \<tau>) =
+ (socket_state \<tau>) ((p, fd):= None)"
+| "socket_state (CreateSock p af st fd ino # \<tau>) =
+ (socket_state \<tau>) ((p, fd) := Some SS_create)"
+| "socket_state (Bind p fd addr # \<tau>) =
+ (socket_state \<tau>) ((p, fd) := Some SS_bind)"
+| "socket_state (Connect p fd addr # \<tau>) =
+ (socket_state \<tau>) ((p, fd) := Some (SS_trans Trans_RS))"
+| "socket_state (Listen p fd # \<tau>) =
+ (socket_state \<tau>) ((p, fd) := Some SS_listen)"
+| "socket_state (Accept p fd addr lport' fd' ino # \<tau>) =
+ (socket_state \<tau>) ((p, fd') := Some (SS_trans Trans_RS))"
+| "socket_state (Shutdown p fd sh # \<tau>) =
+ (socket_state \<tau>) ((p, fd) := update_with_shuthow (socket_state \<tau> (p, fd)) sh)"
+| "socket_state (Clone p p' fds # \<tau>) =
+ (\<lambda> (p'', fd). if (p'' = p' \<and> fd \<in> fds) then socket_state \<tau> (p, fd)
+ else if (p'' = p') then None
+ else socket_state \<tau> (p'', fd))"
+| "socket_state (Execve p f fds # \<tau>) =
+ (\<lambda> (p', fd). if (p' = p \<and> fd \<in> fds) then socket_state \<tau> (p, fd)
+ else if (p' = p) then None
+ else socket_state \<tau> (p', fd))"
+| "socket_state (Kill p\<^isub>1 p\<^isub>2 # \<tau>) =
+ (\<lambda> (p, fd). if (p = p\<^isub>2) then None else socket_state \<tau> (p, fd))"
+| "socket_state (Exit p # \<tau>) =
+ (\<lambda> (p', fd'). if (p = p') then None else socket_state \<tau> (p', fd'))"
+| "socket_state (e # \<tau>) = socket_state \<tau>"
+
+definition socket_in_trans :: "t_state \<Rightarrow> t_socket \<Rightarrow> bool"
+where
+ "socket_in_trans \<tau> s \<equiv> (case (socket_state \<tau> s) of
+ Some st \<Rightarrow> (case st of
+ SS_trans para \<Rightarrow> True
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+
+definition socket_in_sending :: "t_state \<Rightarrow> t_socket \<Rightarrow> bool"
+where
+ "socket_in_sending \<tau> s \<equiv> (socket_state \<tau> s = Some (SS_trans Trans_Send)) \<or> (socket_state \<tau> s = Some (SS_trans Trans_RS))"
+
+definition socket_in_recving :: "t_state \<Rightarrow> t_socket \<Rightarrow> bool"
+where
+ "socket_in_recving \<tau> s \<equiv> (socket_state \<tau> s = Some (SS_trans Trans_Recv)) \<or> (socket_state \<tau> s = Some (SS_trans Trans_RS))"
+
+
+(******* admissable check by the OS *******)
+fun os_grant :: "t_state \<Rightarrow> t_event \<Rightarrow> bool"
+where
+ "os_grant \<tau> (Open p f flags fd inumopt) = (
+ case inumopt of
+ Some ino \<Rightarrow>
+ (p \<in> current_procs \<tau> \<and> f \<notin> current_files \<tau> \<and>
+ (\<exists> pf. parent f = Some pf \<and> is_dir \<tau> pf \<and> pf \<notin> files_hung_by_del \<tau>) \<and>
+ is_creat_flag flags \<and> fd \<notin> (current_proc_fds \<tau> p) \<and> ino \<notin> (current_inode_nums \<tau>))
+ | None \<Rightarrow>
+ (p \<in> current_procs \<tau> \<and> \<not> is_creat_excl_flag flags \<and> is_file \<tau> f \<and>
+ fd \<notin> (current_proc_fds \<tau> p)) )"
+(*(if (is_dir \<tau> f) then (is_dir_flag flags \<and> \<not> is_write_flag flags) else (\<not> is_dir_flag flags)),
+ here open is not applied to directories, readdir is for that, but it is not security-related *)
+| "os_grant \<tau> (CloseFd p fd) =
+ (p \<in> current_procs \<tau> \<and> fd \<in> current_proc_fds \<tau> p)"
+| "os_grant \<tau> (UnLink p f) =
+ (p \<in> current_procs \<tau> \<and> is_file \<tau> f \<and> f \<notin> files_hung_by_del \<tau>)"
+| "os_grant \<tau> (Rmdir p f) =
+ (p \<in> current_procs \<tau> \<and> dir_is_empty \<tau> f \<and> f \<notin> files_hung_by_del \<tau> \<and> f \<noteq> [])" (* root file cannot be deleted *)
+| "os_grant \<tau> (Mkdir p f ino) =
+ (p \<in> current_procs \<tau> \<and> f \<notin> current_files \<tau> \<and>
+ (\<exists> pf. parent f = Some pf \<and> is_dir \<tau> pf \<and> pf \<notin> files_hung_by_del \<tau>) \<and>
+ ino \<notin> (current_inode_nums \<tau>))"
+(*
+| "os_grant \<tau> (LinkHard p f\<^isub>1 f\<^isub>2) =
+ (p \<in> current_procs \<tau> \<and> is_file \<tau> f\<^isub>1 \<and> f\<^isub>2 \<notin> current_files \<tau> \<and>
+ (\<exists> pf\<^isub>2. parent f\<^isub>2 = Some pf\<^isub>2 \<and> is_dir \<tau> pf\<^isub>2 \<and> pf\<^isub>2 \<notin> files_hung_by_del \<tau>))"
+*)
+| "os_grant \<tau> (Truncate p f len) =
+ (p \<in> current_procs \<tau> \<and> is_file \<tau> f)"
+(*
+| "os_grant \<tau> (FTruncate p fd len) =
+ (p \<in> current_procs \<tau> \<and> fd \<in> current_proc_fds \<tau> p \<and>
+ (\<exists> f. file_of_proc_fd \<tau> p fd = Some f \<and> f \<in> current_files \<tau> \<and> is_file \<tau> f) \<and>
+ (\<exists> flags. flags_of_proc_fd \<tau> p fd = Some flags \<and> is_write_flag flags))"
+*)
+| "os_grant \<tau> (ReadFile p fd) =
+ (p \<in> current_procs \<tau> \<and> fd \<in> current_proc_fds \<tau> p \<and>
+ (\<exists> f. file_of_proc_fd \<tau> p fd = Some f \<and> f \<in> current_files \<tau>) \<and>
+ (\<exists> flags. flags_of_proc_fd \<tau> p fd = Some flags \<and> is_read_flag flags))"
+| "os_grant \<tau> (WriteFile p fd) =
+ (p \<in> current_procs \<tau> \<and> fd \<in> current_proc_fds \<tau> p \<and>
+ (\<exists> f. file_of_proc_fd \<tau> p fd = Some f \<and> f \<in> current_files \<tau>) \<and>
+ (\<exists> flags. flags_of_proc_fd \<tau> p fd = Some flags \<and> is_write_flag flags))"
+| "os_grant \<tau> (Execve p f fds) =
+ (p \<in> current_procs \<tau> \<and> is_file \<tau> f \<and> fds \<subseteq> proc_file_fds \<tau> p)" (* file_at_writing_by \<tau> f = {} \<and> fds \<subseteq> current_proc_fds \<tau> p *)
+(*
+| "os_grant \<tau> (Rename p f\<^isub>2 f\<^isub>3) =
+ (p \<in> current_procs \<tau> \<and> f\<^isub>2 \<in> current_files \<tau> \<and> \<not>(f\<^isub>2 \<preceq> f\<^isub>3) \<and> f\<^isub>3 \<notin> current_files \<tau> \<and>
+ (\<exists> pf\<^isub>3. parent f\<^isub>3 = Some pf\<^isub>3 \<and> pf\<^isub>3 \<in> current_files \<tau> \<and> is_dir \<tau> pf\<^isub>3 \<and>
+ pf\<^isub>3 \<notin> files_hung_by_del \<tau>) )"
+*)
+| "os_grant \<tau> (CreateMsgq p q) =
+ (p \<in> current_procs \<tau> \<and> q \<notin> (current_msgqs \<tau>))"
+| "os_grant \<tau> (SendMsg p q m) =
+ (p \<in> current_procs \<tau> \<and> q \<in> current_msgqs \<tau> \<and> m \<notin> (set (msgs_of_queue \<tau> q)))"
+| "os_grant \<tau> (RecvMsg p q m) =
+ (p \<in> current_procs \<tau> \<and> q \<in> current_msgqs \<tau> \<and> m = hd (msgs_of_queue \<tau> q) \<and> msgs_of_queue \<tau> q \<noteq> [])"
+| "os_grant \<tau> (RemoveMsgq p q) =
+ (p \<in> current_procs \<tau> \<and> q \<in> current_msgqs \<tau>)"
+(*
+| "os_grant \<tau> (CreateShM p h) =
+ (p \<in> current_procs \<tau> \<and> h \<notin> (current_shms \<tau>))"
+| "os_grant \<tau> (Attach p h flag) =
+ (p \<in> current_procs \<tau> \<and> h \<in> current_shms \<tau> \<and> \<not> (\<exists> flag. (p, flag) \<in> procs_of_shm \<tau> h))"
+| "os_grant \<tau> (Detach p h) =
+ (p \<in> current_procs \<tau> \<and> h \<in> current_shms \<tau> \<and> (\<exists> flag. (p, flag) \<in> procs_of_shm \<tau> h))"
+| "os_grant \<tau> (DeleteShM p h) =
+ (p \<in> current_procs \<tau> \<and> h \<in> current_shms \<tau>)"
+*)
+| "os_grant \<tau> (CreateSock p af st fd ino) =
+ (p \<in> current_procs \<tau> \<and> af = AF_INET \<and> fd \<notin> (current_proc_fds \<tau> p) \<and>
+ ino \<notin> (current_inode_nums \<tau>))"
+| "os_grant \<tau> (Accept p fd addr port fd' ino) =
+ (p \<in> current_procs \<tau> \<and> fd \<in> current_proc_fds \<tau> p \<and> (p, fd) \<in> current_sockets \<tau> \<and>
+ fd' \<notin> (current_proc_fds \<tau> p) \<and> socket_state \<tau> (p, fd) = Some SS_listen \<and>
+ is_tcp_sock \<tau> (p, fd) \<and> ino \<notin> (current_inode_nums \<tau>))"
+| "os_grant \<tau> (Bind p fd addr) =
+ (p \<in> current_procs \<tau> \<and> fd \<in> current_proc_fds \<tau> p \<and> (p, fd) \<in> current_sockets \<tau> \<and>
+ socket_state \<tau> (p, fd) = Some SS_create)" (* ?? !! (case addr of INET_ADDR ip port \<Rightarrow> ip \<in> local_ipaddrs)) *)
+| "os_grant \<tau> (Connect p fd addr) =
+ (p \<in> current_procs \<tau> \<and> fd \<in> current_proc_fds \<tau> p \<and> (p, fd) \<in> current_sockets \<tau> \<and>
+ socket_state \<tau> (p, fd) = Some SS_bind)"
+| "os_grant \<tau> (Listen p fd) =
+ (p \<in> current_procs \<tau> \<and> fd \<in> current_proc_fds \<tau> p \<and> is_tcp_sock \<tau> (p, fd) \<and>
+ socket_state \<tau> (p, fd) = Some SS_bind)" (* Listen and Accept should only be used for TCP sever side *)
+| "os_grant \<tau> (SendSock p fd) =
+ (p \<in> current_procs \<tau> \<and> fd \<in> current_proc_fds \<tau> p \<and> (p, fd) \<in> current_sockets \<tau> \<and>
+ socket_in_sending \<tau> (p, fd))"
+| "os_grant \<tau> (RecvSock p fd) =
+ (p \<in> current_procs \<tau> \<and> fd \<in> current_proc_fds \<tau> p \<and> (p, fd) \<in> current_sockets \<tau> \<and>
+ socket_in_recving \<tau> (p, fd))"
+| "os_grant \<tau> (Shutdown p fd sh) =
+ (p \<in> current_procs \<tau> \<and> fd \<in> current_proc_fds \<tau> p \<and> (p, fd) \<in> current_sockets \<tau> \<and>
+ socket_in_trans \<tau> (p, fd))"
+(*
+| "os_grant \<tau> (ChangeOwner p u) = (p \<in> current_procs \<tau> \<and> u \<in> current_users \<tau>)"
+*)
+| "os_grant \<tau> (Clone p p' fds) =
+ (p \<in> current_procs \<tau> \<and> p' \<notin> (current_procs \<tau>) \<and> fds \<subseteq> proc_file_fds \<tau> p)" (* current_proc_fds \<rightarrow> proc_file_fds *)
+| "os_grant \<tau> (Kill p\<^isub>1 p\<^isub>2) =
+ (p\<^isub>1 \<in> current_procs \<tau> \<and> p\<^isub>2 \<in> current_procs \<tau>)" (* a process can kill itself right? *)
+| "os_grant \<tau> (Exit p) =
+ (p \<in> current_procs \<tau>)"
+| "os_grant \<tau> (Ptrace p\<^isub>1 p\<^isub>2) =
+ (p\<^isub>1 \<in> current_procs \<tau> \<and> p\<^isub>2 \<in> current_procs \<tau>)" (* a process can trace itself right? *)
+
+fun alive :: "t_state \<Rightarrow> t_object \<Rightarrow> bool"
+where
+ "alive s (O_proc p) = (p \<in> current_procs s)"
+| "alive s (O_file f) = (is_file s f)"
+| "alive s (O_dir f) = (is_dir s f)"
+| "alive s (O_fd p fd) = (fd \<in> current_proc_fds s p)"
+| "alive s (O_node n) = (n \<in> init_nodes)"
+| "alive s (O_tcp_sock k) = (is_tcp_sock s k)"
+| "alive s (O_udp_sock k) = (is_udp_sock s k)"
+(*
+| "alive s (O_shm h) = (h \<in> current_shms s)"
+*)
+| "alive s (O_msgq q) = (q \<in> current_msgqs s)"
+| "alive s (O_msg q m) = (m \<in> set (msgs_of_queue s q) \<and> q \<in> current_msgqs s)"
+
+(* deleted objects should be "taintable" objects, objects like fd should not be in consideration *)
+(* actually, deleted should be renamed as "vanished", cause "exit/closefd" *)
+fun deleted :: "t_object \<Rightarrow> t_event list \<Rightarrow> bool"
+where
+ "deleted obj [] = False"
+| "deleted obj (Kill p p' # s) = ((obj = O_proc p') \<or> (\<exists> fd \<in> current_proc_fds s p'. obj = O_fd p' fd) \<or>
+ (\<exists> fd. obj = O_tcp_sock (p', fd) \<and> is_tcp_sock s (p',fd)) \<or>
+ (\<exists> fd. obj = O_udp_sock (p', fd) \<and> is_udp_sock s (p',fd)) \<or>
+ deleted obj s)"
+| "deleted obj (CloseFd p fd # s) = ((obj = O_fd p fd) \<or> (obj = O_tcp_sock (p,fd) \<and> is_tcp_sock s (p,fd)) \<or>
+ (obj = O_udp_sock (p,fd) \<and> is_udp_sock s (p, fd)) \<or>
+ (\<exists> f. obj = O_file f \<and> proc_fd_of_file s f = {(p, fd)} \<and>
+ f \<in> files_hung_by_del s) \<or> deleted obj s)"
+| "deleted obj (Execve p f fds # s) = ((\<exists> fd \<in> current_proc_fds s p. obj = O_fd p fd \<and> fd \<notin> fds) \<or>
+ (\<exists> fd. obj = O_tcp_sock (p, fd) \<and> is_tcp_sock s (p,fd) \<and> fd \<notin> fds) \<or>
+ (\<exists> fd. obj = O_udp_sock (p, fd) \<and> is_udp_sock s (p,fd) \<and> fd \<notin> fds) \<or>
+ deleted obj s)"
+| "deleted obj (Clone p p' fds # s) = ((\<exists> fd \<in> current_proc_fds s p. obj = O_fd p' fd \<and> fd \<notin> fds) \<or>
+ deleted obj s)"
+| "deleted obj (UnLink p f # s) = ((proc_fd_of_file s f = {} \<and> obj = O_file f) \<or> deleted obj s)"
+| "deleted obj (Rmdir p f # s) = ((obj = O_dir f) \<or> deleted obj s)"
+| "deleted obj (Exit p # s) = ((obj = O_proc p) \<or> (\<exists> fd \<in> current_proc_fds s p. obj = O_fd p fd) \<or>
+ (\<exists> fd. obj = O_tcp_sock (p, fd) \<and> is_tcp_sock s (p,fd)) \<or>
+ (\<exists> fd. obj = O_udp_sock (p, fd) \<and> is_udp_sock s (p,fd)) \<or> deleted obj s)"
+(*
+| "deleted obj (Rename p f f' # s) =
+ (case obj of
+ O_file f'' \<Rightarrow> f \<preceq> f'' \<or> deleted obj s
+ | O_dir f'' \<Rightarrow> f \<preceq> f'' \<or> deleted obj s
+ | _ \<Rightarrow> deleted obj s)"
+*)
+| "deleted obj (RemoveMsgq p q # s) = (obj = O_msgq q \<or> (\<exists> m. obj = O_msg q m) \<or> deleted obj s)"
+(*
+| "deleted obj (DeleteShM p h # s) = (obj = O_shm h \<or> deleted obj s)"
+*)
+| "deleted obj (RecvMsg p q m # s) = (obj = O_msg q m \<or> deleted obj s)"
+| "deleted obj (_ # s) = deleted obj s"
+
+end
+
+
+
+locale flask = init +
+ fixes
+ (***************** Policy-specific Parameters *************)
+ type_transition_rules :: "type_transition_rule_list_t"
+ and allowed_rules :: "allowed_rule_list_t"
+ and role_allowed_rules :: "role_allow_rule_list_t"
+ and role_declaration_rules :: "role_declarations_list_t"
+ and role_transition_rules:: "role_transition_rule_list_t"
+ and constrains_rules :: "constrains_list_t"
+
+ and init_type_of_obj :: "t_object \<rightharpoonup> type_t"
+ and init_user_of_obj :: "t_object \<rightharpoonup> user_t"
+ and init_role_of_proc :: "t_process \<rightharpoonup> role_t"
+
+ assumes
+
+ init_obj_has_type: "\<And> obj. init_alive obj \<Longrightarrow> \<exists> t. init_type_of_obj obj = Some t"
+ and init_type_has_obj: "\<And> obj t. init_type_of_obj obj = Some t \<Longrightarrow> init_alive obj"
+ and init_obj_has_user: "\<And> obj. init_alive obj \<Longrightarrow> \<exists> u. init_user_of_obj obj = Some u"
+ and init_user_has_obj: "\<And> obj u. init_user_of_obj obj = Some u \<Longrightarrow> init_alive obj \<and> u \<in> init_users"
+ and init_proc_has_role: "bidirect_in_init init_procs init_role_of_proc"
+
+begin
+
+(*********** policy-specified computations, not-event-related ************)
+
+(* security_transition_sid : type_transition_rules
+ * It's a system-call offered by security-server? So access-control-checked too?
+ * according to selinux, it deals only with execve and file-creation
+ * is_execve: if is the execve case
+ * is_file: if is the creation of file/dir
+ * change from is_execve to is_file, because the new message by default use
+ * the type of the sending process too.
+ *)
+
+definition type_transition :: "type_t \<Rightarrow> type_t \<Rightarrow> security_class_t \<Rightarrow> bool \<Rightarrow> type_t"
+where
+ "type_transition st rt tc IS_file \<equiv> (
+ case (lookup type_transition_rules
+ (\<lambda> (stp, rtp, tc', dt). patt_match stp st \<and> patt_match rtp rt \<and> tc = tc')) of
+ None \<Rightarrow> (if IS_file then rt else st)
+ | Some (stp,rtp,tc',dt) \<Rightarrow> dt)"
+
+(* allowed_rule_list *)
+definition TE_allow :: "type_t \<Rightarrow> type_t \<Rightarrow> security_class_t \<Rightarrow> permission_t \<Rightarrow> bool"
+where
+ "TE_allow st tt tc perm \<equiv> member allowed_rules
+ (\<lambda> (stp, ttp, tc', pp). tc = tc' \<and> patt_match stp st \<and> patt_match_self ttp tt st \<and>
+ patt_match pp perm)"
+
+(* role_allow_rule_list_t *)
+definition role_transition :: "role_t \<Rightarrow> type_t \<Rightarrow> role_t option"
+where
+ "role_transition r t \<equiv>
+ (case (lookup role_transition_rules (\<lambda> (pr, ft, nr). pr = r \<and> t = ft)) of
+ None \<Rightarrow> None
+ | Some (pr,ft,nr) \<Rightarrow> Some nr)"
+
+definition role_decl_check :: "role_t \<Rightarrow> type_t \<Rightarrow> bool"
+where
+ "role_decl_check r t \<equiv> member role_declaration_rules (\<lambda> (role, types). r = role \<and> t \<in> types)"
+
+definition role_allow :: "role_t \<Rightarrow> role_t \<Rightarrow> bool"
+where
+ "role_allow r nr \<equiv> member role_allowed_rules (\<lambda> (from, to). r \<in> from \<and> nr \<in> to)"
+
+(* here don't use lookup, because constrains should all be satisfied,
+ * while transitions always choose the "first" one
+ *)
+definition constrains_satisfied ::
+ "security_context_t \<Rightarrow> security_context_t \<Rightarrow> security_class_t \<Rightarrow> permission_t \<Rightarrow> bool"
+where
+ "constrains_satisfied sctxt tctxt c p \<equiv>
+ (fold (\<lambda> (cset, pset, P) res. res \<and> (if (c \<in> cset \<and> p \<in> pset)
+ then P sctxt tctxt else True))
+ constrains_rules True)"
+
+(* main interface for TE_allow check and constrains check *)
+fun permission_check::
+ "security_context_t \<Rightarrow> security_context_t \<Rightarrow> security_class_t \<Rightarrow> permission_t \<Rightarrow> bool"
+where
+ "permission_check (su,sr,st) (tu,tr,tt) c p =
+ ((TE_allow st tt c p) \<and> (constrains_satisfied (su,sr,st) (tu,tr,tt) c p))"
+declare permission_check.simps [simp del]
+
+(* no changeowner ??? : by adding "relabel obj sectxt" event or "login" event?
+ * or this explanation: the running of the sever will not grant any login but running
+ * of the server-programs!
+ *)
+
+fun user_of_obj :: "t_state \<Rightarrow> t_object \<Rightarrow> user_t option"
+where
+ "user_of_obj [] = init_user_of_obj"
+| "user_of_obj (Clone p p' fds # s) =
+ (\<lambda> obj. case obj of
+ O_proc p'' \<Rightarrow> if (p'' = p') then user_of_obj s (O_proc p) else user_of_obj s obj
+ | O_fd p'' fd \<Rightarrow> if (p'' = p' \<and> fd \<in> fds) then user_of_obj s (O_fd p fd)
+ else if (p'' = p') then None
+ else user_of_obj s obj
+ | O_tcp_sock (p'',fd) \<Rightarrow> if (p'' = p' \<and> fd \<in> fds) then user_of_obj s (O_tcp_sock (p, fd))
+ else if (p'' = p') then None
+ else user_of_obj s obj
+ | O_udp_sock (p'',fd) \<Rightarrow> if (p'' = p' \<and> fd \<in> fds) then user_of_obj s (O_udp_sock (p, fd))
+ else if (p'' = p') then None
+ else user_of_obj s obj
+ | _ \<Rightarrow> user_of_obj s obj)"
+| "user_of_obj (Open p f flags fd iopt # s) =
+ (case iopt of
+ None \<Rightarrow> (user_of_obj s) ((O_fd p fd) := user_of_obj s (O_proc p))
+ | _ \<Rightarrow> ((user_of_obj s) ((O_fd p fd) := user_of_obj s (O_proc p)))
+ ((O_file f) := user_of_obj s (O_proc p)))"
+| "user_of_obj (Mkdir p f inum # s) =
+ (user_of_obj s) ((O_dir f) := user_of_obj s (O_proc p))"
+(*
+| "user_of_obj (LinkHard p f f' # s) =
+ (user_of_obj s) ((O_file f') := user_of_obj s (O_proc p))"
+*)
+(*
+| "user_of_obj (Rename p f2 f3 # s) =
+ (\<lambda> obj. case obj of
+ O_file f \<Rightarrow> if (f3 \<preceq> f) then user_of_obj s (O_file (file_before_rename f2 f3 f))
+ else user_of_obj s obj
+ | O_dir f \<Rightarrow> if (f3 \<preceq> f) then user_of_obj s (O_dir (file_before_rename f2 f3 f))
+ else user_of_obj s obj
+ | _ \<Rightarrow> user_of_obj s obj)"
+*)
+| "user_of_obj (CreateMsgq p q # s) =
+ (user_of_obj s) ((O_msgq q) := user_of_obj s (O_proc p))"
+| "user_of_obj (SendMsg p q m # s) =
+ (user_of_obj s) ((O_msg q m) := user_of_obj s (O_proc p))"
+(*
+| "user_of_obj (CreateShM p h # s) =
+ (user_of_obj s) ((O_shm h) := user_of_obj s (O_proc p))"
+*)
+| "user_of_obj (CreateSock p af st fd inum # s) = (\<lambda> obj.
+ case obj of
+ O_fd p' fd' \<Rightarrow> if (p' = p \<and> fd' = fd) then user_of_obj s (O_proc p)
+ else user_of_obj s obj
+ | O_tcp_sock (p', fd') \<Rightarrow> if (p' = p \<and> fd' = fd \<and> st = STREAM) then user_of_obj s (O_proc p)
+ else user_of_obj s obj
+ | O_udp_sock (p', fd') \<Rightarrow> if (p' = p \<and> fd' = fd \<and> st = DGRAM) then user_of_obj s (O_proc p)
+ else user_of_obj s obj
+ | _ \<Rightarrow> user_of_obj s obj )"
+| "user_of_obj (Accept p fd addr port fd' inum # s) = (\<lambda> obj.
+ case obj of
+ O_fd p' fd'' \<Rightarrow> if (p' = p \<and> fd'' = fd') then user_of_obj s (O_proc p)
+ else user_of_obj s obj
+ | O_tcp_sock (p', fd'') \<Rightarrow> if (p' = p \<and> fd'' = fd') then user_of_obj s (O_proc p)
+ else user_of_obj s obj
+ | _ \<Rightarrow> user_of_obj s obj )"
+| "user_of_obj (_ # s) = user_of_obj s"
+
+fun type_of_obj :: "t_state \<Rightarrow> (t_object \<Rightarrow> type_t option)"
+where
+ "type_of_obj [] = init_type_of_obj"
+| "type_of_obj (Clone p p' fds # s) = (\<lambda> obj.
+ case obj of
+ O_proc p'' \<Rightarrow> if (p'' = p') then type_of_obj s (O_proc p) else type_of_obj s obj
+ | O_fd p'' fd \<Rightarrow> if (p'' = p' \<and> fd \<in> fds) then type_of_obj s (O_fd p fd)
+ else if (p'' = p') then None
+ else type_of_obj s obj
+ | O_tcp_sock (p'', fd) \<Rightarrow> if (p'' = p' \<and> fd \<in> fds) then type_of_obj s (O_tcp_sock (p, fd))
+ else if (p'' = p') then None
+ else type_of_obj s obj
+ | O_udp_sock (p'', fd) \<Rightarrow> if (p'' = p' \<and> fd \<in> fds) then type_of_obj s (O_udp_sock (p, fd))
+ else if (p'' = p') then None
+ else type_of_obj s obj
+ | _ \<Rightarrow> type_of_obj s obj )"
+| "type_of_obj (Execve p f fds # s) = (type_of_obj s) ((O_proc p) :=
+ (case (type_of_obj s (O_proc p), type_of_obj s (O_file f)) of
+ (Some tp, Some tf) \<Rightarrow> Some (type_transition tp tf C_process False)
+ | _ \<Rightarrow> None) )"
+| "type_of_obj (Open p f flags fd (Some inum) # s) = ((type_of_obj s) ((O_file f) :=
+ (case (parent f) of
+ Some pf \<Rightarrow> (case (type_of_obj s (O_proc p), type_of_obj s (O_dir pf)) of
+ (Some tp, Some tpf) \<Rightarrow> Some (type_transition tp tpf C_file True)
+ | _ \<Rightarrow> None)
+ | _ \<Rightarrow> None) )) ((O_fd p fd) :=
+ type_of_obj s (O_proc p))"
+| "type_of_obj (Open p f flags fd None # s) = (type_of_obj s) ((O_fd p fd) :=
+ type_of_obj s (O_proc p))"
+| "type_of_obj (Mkdir p f inum # s) = (type_of_obj s) ((O_dir f) :=
+ (case (parent f) of
+ Some pf \<Rightarrow> (case (type_of_obj s (O_proc p), type_of_obj s (O_dir pf)) of
+ (Some tp, Some tpf) \<Rightarrow> Some (type_transition tp tpf C_dir True)
+ | _ \<Rightarrow> None)
+ | _ \<Rightarrow> None) )"
+(*
+| "type_of_obj (LinkHard p f f' # s) = (type_of_obj s) ((O_file f') :=
+ (case (parent f') of
+ Some pf \<Rightarrow> (case (type_of_obj s (O_proc p), type_of_obj s (O_dir pf)) of
+ (Some tp, Some tpf) \<Rightarrow> Some (type_transition tp tpf C_file True)
+ | _ \<Rightarrow> None)
+ | _ \<Rightarrow> None) )"
+*)
+(*
+| "type_of_obj (Rename p f f' # s) = (\<lambda> obj.
+ case obj of
+ O_file f'' \<Rightarrow> (if (f' \<preceq> f'') then type_of_obj s (O_file (file_before_rename f f' f''))
+ else type_of_obj s (O_file f''))
+ | O_dir f'' \<Rightarrow> (if (f' \<preceq> f'') then type_of_obj s (O_file (file_before_rename f f' f''))
+ else type_of_obj s (O_file f''))
+ | _ \<Rightarrow> type_of_obj s obj )"
+*)
+| "type_of_obj (CreateMsgq p q # s) = (type_of_obj s) ((O_msgq q) :=
+ (case (type_of_obj s (O_proc p)) of
+ Some tp \<Rightarrow> Some tp
+ | _ \<Rightarrow> None) )"
+| "type_of_obj (SendMsg p q m # s) = (type_of_obj s) ((O_msg q m) :=
+ (case (type_of_obj s (O_proc p), type_of_obj s (O_msgq q)) of
+ (Some tp, Some tq) \<Rightarrow> Some (type_transition tp tq C_msg False)
+ | _ \<Rightarrow> None) )"
+(*
+| "type_of_obj (CreateShM p h # s) = (type_of_obj s) ((O_shm h) :=
+ (case (type_of_obj s (O_proc p)) of
+ Some tp \<Rightarrow> Some tp
+ | _ \<Rightarrow> None) )"
+*)
+| "type_of_obj (CreateSock p af st fd inum # s) = (\<lambda> obj.
+ case obj of
+ O_fd p' fd' \<Rightarrow> if (p' = p \<and> fd' = fd) then type_of_obj s (O_proc p)
+ else type_of_obj s obj
+ | O_tcp_sock (p', fd') \<Rightarrow> if (p' = p \<and> fd' = fd \<and> st = STREAM) then type_of_obj s (O_proc p)
+ else type_of_obj s obj
+ | O_udp_sock (p', fd') \<Rightarrow> if (p' = p \<and> fd' = fd \<and> st = DGRAM) then type_of_obj s (O_proc p)
+ else type_of_obj s obj
+ | _ \<Rightarrow> type_of_obj s obj )"
+| "type_of_obj (Accept p fd addr port fd' inum # s) = (\<lambda> obj.
+ case obj of
+ O_fd p' fd'' \<Rightarrow> if (p' = p \<and> fd'' = fd') then type_of_obj s (O_proc p)
+ else type_of_obj s obj
+ | O_tcp_sock (p', fd'') \<Rightarrow> if (p' = p \<and> fd'' = fd') then type_of_obj s (O_proc p)
+ else type_of_obj s obj
+ | _ \<Rightarrow> type_of_obj s obj )"
+| "type_of_obj (_ # s) = type_of_obj s"
+
+fun role_of_proc :: "t_state \<Rightarrow> (t_process \<Rightarrow> role_t option)"
+where
+ "role_of_proc [] = init_role_of_proc"
+| "role_of_proc (Clone p p' fds # s) =
+ (role_of_proc s) (p' := role_of_proc s p)"
+| "role_of_proc (Execve p f fds # s) = (role_of_proc s) (p :=
+ (case (role_of_proc s p, type_of_obj s (O_file f)) of
+ (Some r, Some ft) \<Rightarrow> role_transition r ft
+ | _ \<Rightarrow> None) )"
+| "role_of_proc (_ # s) = role_of_proc s"
+
+fun role_of_obj :: "t_state \<Rightarrow> t_object \<Rightarrow> role_t option"
+where
+ "role_of_obj s (O_proc p) = role_of_proc s p"
+| "role_of_obj s _ = Some R_object" (* object_r *)
+
+definition sectxt_of_obj :: "t_state \<Rightarrow> t_object \<Rightarrow> security_context_t option"
+where
+ "sectxt_of_obj s obj = (
+ case (user_of_obj s obj, role_of_obj s obj, type_of_obj s obj) of
+ (Some u, Some r, Some t) \<Rightarrow> Some (u, r, t)
+ | _ \<Rightarrow> None)"
+
+fun init_role_of_obj :: "t_object \<Rightarrow> role_t option"
+where
+ "init_role_of_obj (O_proc p) = init_role_of_proc p"
+| "init_role_of_obj _ = Some R_object"
+
+definition init_sectxt_of_obj :: "t_object \<Rightarrow> security_context_t option"
+where
+ "init_sectxt_of_obj obj \<equiv>
+ (case (init_user_of_obj obj, init_role_of_obj obj, init_type_of_obj obj) of
+ (Some u, Some r, Some t) \<Rightarrow> Some (u,r,t)
+ | _ \<Rightarrow> None)"
+
+definition sec_of_root :: "security_context_t"
+where
+ "sec_of_root \<equiv> (case (init_user_of_obj (O_dir []), init_type_of_obj (O_dir [])) of
+ (Some u, Some t) \<Rightarrow> (u, R_object, t))"
+
+(******* flask granting ********
+ * involves three kinds of rules:
+ * 1. allow rule of types, allowed_rule_list_t, main
+ * 2. allow rule of roles, role_allow_rule_list_t, mainly for execve call
+ * 3. constrains, section 3.4.5, user-specifically defined.
+ * Passed all these 3, then grant Yes, else No
+ *)
+
+definition search_check_allp :: "security_context_t \<Rightarrow> security_context_t set \<Rightarrow> bool"
+where
+ "search_check_allp pctxt sectxts = (\<forall> sec \<in> sectxts. permission_check pctxt sec C_dir P_search)"
+
+definition search_check_file :: "security_context_t \<Rightarrow> security_context_t \<Rightarrow> bool"
+where
+ "search_check_file sctxt fctxt \<equiv> permission_check sctxt fctxt C_file P_search"
+
+definition search_check_dir :: "security_context_t \<Rightarrow> security_context_t \<Rightarrow> bool"
+where
+ "search_check_dir sctxt fctxt \<equiv> permission_check sctxt fctxt C_dir P_search"
+
+fun get_parentfs_ctxts :: "t_state \<Rightarrow> t_file \<Rightarrow> (security_context_t list) option"
+where
+ "get_parentfs_ctxts s [] =
+ (case (sectxt_of_obj s (O_dir [])) of
+ Some ctxt \<Rightarrow> Some [ctxt]
+ | _ \<Rightarrow> None)"
+| "get_parentfs_ctxts s (f#pf) =
+ (case (get_parentfs_ctxts s pf, sectxt_of_obj s (O_dir (f#pf))) of
+ (Some ctxts, Some ctxt) \<Rightarrow> Some (ctxt#ctxts)
+ | _ \<Rightarrow> None)"
+
+definition search_check_ctxt::
+ "security_context_t \<Rightarrow> security_context_t \<Rightarrow> security_context_t set \<Rightarrow> bool \<Rightarrow> bool"
+where
+ "search_check_ctxt pctxt fctxt asecs if_file \<equiv> (
+ if if_file
+ then search_check_file pctxt fctxt \<and> search_check_allp pctxt asecs
+ else search_check_dir pctxt fctxt \<and> search_check_allp pctxt asecs )"
+
+fun search_check :: "t_state \<Rightarrow> security_context_t \<Rightarrow> t_file \<Rightarrow> bool"
+where
+ "search_check s pctxt [] =
+ (case (sectxt_of_obj s (O_dir [])) of
+ Some fctxt \<Rightarrow> search_check_ctxt pctxt fctxt {} False
+ | _ \<Rightarrow> False)"
+| "search_check s pctxt (f#pf) =
+ (if (is_file s (f#pf))
+ then (case (sectxt_of_obj s (O_file (f#pf)), get_parentfs_ctxts s pf) of
+ (Some fctxt, Some pfctxts) \<Rightarrow> search_check_ctxt pctxt fctxt (set pfctxts) True
+ | _ \<Rightarrow> False)
+ else (case (sectxt_of_obj s (O_dir (f#pf)), get_parentfs_ctxts s pf) of
+ (Some fctxt, Some pfctxts) \<Rightarrow> search_check_ctxt pctxt fctxt (set pfctxts) False
+ | _ \<Rightarrow> False))"
+
+(* this means we should prove every current fd has a security context! *)
+definition sectxts_of_fds :: "t_state \<Rightarrow> t_process \<Rightarrow> t_fd set \<Rightarrow> security_context_t set"
+where
+ "sectxts_of_fds s p fds \<equiv> {ctxt. \<exists> fd \<in> fds. sectxt_of_obj s (O_fd p fd) = Some ctxt}"
+
+definition inherit_fds_check_ctxt:: "security_context_t \<Rightarrow> security_context_t set \<Rightarrow> bool"
+where
+ "inherit_fds_check_ctxt p fds \<equiv> (\<forall> ctxt \<in> fds. permission_check p ctxt C_fd P_inherit)"
+
+definition inherit_fds_check :: "t_state \<Rightarrow> security_context_t \<Rightarrow> t_process \<Rightarrow> t_fd set \<Rightarrow> bool"
+where
+ "inherit_fds_check s npsectxt p fds \<equiv> inherit_fds_check_ctxt npsectxt (sectxts_of_fds s p fds)"
+
+fun npctxt_execve :: "security_context_t \<Rightarrow> security_context_t \<Rightarrow> security_context_t option"
+where
+ "npctxt_execve (pu,pr,pt) (fu,fr,ft) =
+ (case (role_transition pr ft) of
+ Some nr \<Rightarrow> Some (pu, nr, type_transition pt ft C_process False)
+ | _ \<Rightarrow> None)"
+
+fun nfctxt_create :: "security_context_t \<Rightarrow> security_context_t \<Rightarrow> security_class_t \<Rightarrow> security_context_t"
+where
+ "nfctxt_create (pu,pr,pt) (fu,fr,ft) cls = (pu, R_object, type_transition pt ft cls True)"
+
+fun grant_execve ::
+ "security_context_t \<Rightarrow> security_context_t \<Rightarrow> security_context_t \<Rightarrow> bool"
+where
+ "grant_execve (up,rp,tp) (uf,rf,tf) (up',nr,nt) =
+ (role_decl_check nr nt \<and> role_allow rp nr \<and>
+ permission_check (up,rp,tp) (uf,rf,tf) C_file P_execute \<and>
+ permission_check (up,nr,nt) (uf,rf,tf) C_file P_entrypoint \<and>
+ permission_check (up,rp,tp) (up,nr,nt) C_process P_transition \<and>
+ permission_check (up,nr,nt) (uf,rf,tf) C_process P_execute)"
+
+(*
+definition sectxts_of_shms :: "t_state \<Rightarrow> t_shm set \<Rightarrow> security_context_t set"
+where
+ "sectxts_of_shms s shms \<equiv> {ctxt. \<exists> h \<in> shms. sectxt_of_obj s (O_shm h) = Some ctxt}"
+
+definition inherit_shms_check_ctxt:: "security_context_t \<Rightarrow> security_context_t set \<Rightarrow> bool"
+where
+ "inherit_shms_check_ctxt p shms \<equiv> (\<forall> ctxt \<in> shms. permission_check p ctxt C_shm P_inherit)"
+
+definition inherit_shms_check :: "t_state \<Rightarrow> security_context_t \<Rightarrow> t_shm set \<Rightarrow> bool"
+where
+ "inherit_shms_check s npsectxt shms \<equiv> inherit_shms_check_ctxt npsectxt (sectxts_of_shms s shms)"
+*)
+
+fun perm_of_mflag :: "t_open_must_flag \<Rightarrow> permission_t set"
+where
+ "perm_of_mflag OF_RDONLY = {P_read}"
+| "perm_of_mflag OF_WRONLY = {P_write}"
+| "perm_of_mflag OF_RDWR = {P_read, P_write}"
+
+definition perm_of_oflag :: "t_open_option_flag set \<Rightarrow> permission_t set"
+where
+ "perm_of_oflag oflags \<equiv>
+ (case (OF_APPEND \<in> oflags, OF_CREAT \<in> oflags) of
+ (True, True) \<Rightarrow> {P_append, P_create}
+ | (True, _ ) \<Rightarrow> {P_append}
+ | (_ , True) \<Rightarrow> {P_create}
+ | _ \<Rightarrow> {})"
+
+definition perms_of_flags :: "t_open_flags \<Rightarrow> permission_t set"
+where
+ "perms_of_flags flags \<equiv>
+ (case flags of
+ (mflag,oflags) \<Rightarrow> (perm_of_mflag mflag \<union> perm_of_oflag oflags))"
+
+(*
+definition class_of_flags :: "t_open_flags \<Rightarrow> security_class_t"
+where
+ "class_of_flags flags \<equiv>
+ (case flags of
+ (mflag, oflags) \<Rightarrow> if (OF_DIRECTORY \<in> oflags) then C_dir else C_file)"
+
+definition obj_of_flags :: "t_open_flags \<Rightarrow> t_file \<Rightarrow> t_object"
+where
+ "obj_of_flags flags f \<equiv>
+ (case flags of
+ (mflag, oflags) \<Rightarrow> if (OF_DIRECTORY \<in> oflags) then O_dir f else O_file f)"
+*)
+
+definition oflags_check :: "t_open_flags \<Rightarrow> security_context_t \<Rightarrow> security_context_t \<Rightarrow> bool"
+where
+ "oflags_check flags pctxt fctxt \<equiv>
+ \<forall> perm \<in> (perms_of_flags flags). permission_check pctxt fctxt C_file perm"
+
+fun grant :: "t_state \<Rightarrow> t_event \<Rightarrow> bool"
+where
+ "grant s (Execve p f fds) =
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_file f)) of
+ (Some (up, rp, tp), Some (uf, rf, tf)) \<Rightarrow>
+ (case (npctxt_execve (up,rp,tp) (uf,rf,tf)) of
+ Some (pu,nr,nt) \<Rightarrow> (
+ search_check s (up,rp,tp) f \<and>
+ grant_execve (up,rp,tp) (uf,rf,tf) (up,nr,nt) \<and>
+ inherit_fds_check s (up,nr,nt) p fds)
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False )"
+| "grant s (Clone p p' fds) =
+ (case (sectxt_of_obj s (O_proc p)) of
+ Some (up, rp, tp) \<Rightarrow>
+ (permission_check (up,rp,tp) (up,rp,tp) C_process P_fork \<and>
+ inherit_fds_check s (up,rp,tp) p fds)
+ | _ \<Rightarrow> False )"
+| "grant s (Kill p p') =
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_proc p')) of
+ (Some (su,sr,st), Some (tu,tr,tt)) \<Rightarrow>
+ permission_check (su,sr,st) (tu,tr,tt) C_process P_sigkill
+ | _ \<Rightarrow> False)"
+| "grant s (Ptrace p p') =
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_proc p')) of
+ (Some (su,sr,st), Some (tu,tr,tt)) \<Rightarrow>
+ permission_check (su,sr,st) (tu,tr,tt) C_process P_ptrace
+ | _ \<Rightarrow> False)"
+| "grant s (Exit p) = True"
+| "grant s (Open p f flags fd None) =
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_file f)) of
+ (Some (pu,pr,pt), Some (fu,fr,ft)) \<Rightarrow>
+ search_check s (pu,pr,pt) f \<and>
+ oflags_check flags (pu,pr,pt) (fu,fr,ft) \<and>
+ permission_check (pu,pr,pt) (pu,pr,pt) C_fd P_setattr
+ | _ \<Rightarrow>False)"
+| "grant s (Open p f flags fd (Some inum)) =
+ (case (parent f) of
+ Some pf \<Rightarrow> (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of
+ (Some (pu,pr,pt), Some (pfu,pfr,pft)) \<Rightarrow>
+ (search_check s (pu,pr,pt) pf \<and>
+ oflags_check flags (pu,pr,pt) (nfctxt_create (pu,pr,pt) (pfu,pfr,pft) C_file) \<and>
+ permission_check (pu,pr,pt) (pfu,pfr,pft) C_dir P_add_name \<and>
+ permission_check (pu,pr,pt) (pu,pr,pt) C_fd P_setattr)
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+| "grant s (ReadFile p fd) =
+ (case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow>
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_fd p fd),
+ sectxt_of_obj s (O_file f)) of
+ (Some (pu,pr,pt), Some (fdu,fdr,fdt), Some (fu,fr,ft)) \<Rightarrow>
+ (permission_check (pu,pr,pt) (fdu,fdr,fdt) C_fd P_setattr \<and>
+ permission_check (pu,pr,pt) (fu,fr,ft) C_file P_read)
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+| "grant s (WriteFile p fd) =
+ (case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow>
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_fd p fd),
+ sectxt_of_obj s (O_file f)) of
+ (Some (pu,pr,pt), Some (fdu,fdr,fdt), Some (fu,fr,ft)) \<Rightarrow>
+ (permission_check (pu,pr,pt) (fdu,fdr,fdt) C_fd P_setattr \<and>
+ permission_check (pu,pr,pt) (fu,fr,ft) C_file P_write)
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+| "grant s (CloseFd p fd) = True"
+| "grant s (UnLink p f) =
+ (case (parent f) of
+ Some pf \<Rightarrow>
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf),
+ sectxt_of_obj s (O_file f)) of
+ (Some (pu,pr,pt), Some (du,dr,dt), Some (fu,fr,ft)) \<Rightarrow>
+ (search_check s (pu,pr,pt) f \<and>
+ permission_check (pu,pr,pt) (fu,fr,ft) C_file P_unlink \<and>
+ permission_check (pu,pr,pt) (du,dr,dt) C_dir P_remove_name)
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+| "grant s (Rmdir p f) =
+ (case (parent f) of
+ Some pf \<Rightarrow>
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf),
+ sectxt_of_obj s (O_dir f)) of
+ (Some (pu,pr,pt), Some (du,dr,dt), Some (fu,fr,ft)) \<Rightarrow>
+ (search_check s (pu,pr,pt) f \<and>
+ permission_check (pu,pr,pt) (fu,fr,ft) C_dir P_rmdir \<and>
+ permission_check (pu,pr,pt) (du,dr,dt) C_dir P_remove_name)
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+| "grant s (Mkdir p f inum) =
+ (case (parent f) of
+ Some pf \<Rightarrow>
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of
+ (Some (pu,pr,pt), Some (du,dr,dt)) \<Rightarrow>
+ (search_check s (pu,pr,pt) f \<and>
+ permission_check (pu,pr,pt) (nfctxt_create (pu,pr,pt) (du,dr,dt) C_dir) C_dir P_create \<and>
+ permission_check (pu,pr,pt) (du,dr,dt) C_dir P_add_name)
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+(*
+| "grant s (LinkHard p f f') =
+ (case (parent f') of
+ Some pf \<Rightarrow>
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf),
+ sectxt_of_obj s (O_file f)) of
+ (Some (pu,pr,pt), Some (du,dr,dt), Some (fu,fr,ft)) \<Rightarrow>
+ (search_check s (pu,pr,pt) f \<and> search_check s (pu,pr,pt) pf \<and>
+ permission_check (pu,pr,pt) (fu,fr,ft) C_file P_link \<and>
+ permission_check (pu,pr,pt) (du,dr,dt) C_dir P_add_name)
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+*)
+| "grant s (Truncate p f len) =
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_file f)) of
+ (Some (pu,pr,pt), Some (fu,fr,ft)) \<Rightarrow>
+ (search_check s (pu,pr,pt) f \<and>
+ permission_check (pu,pr,pt) (fu,fr,ft) C_file P_setattr)
+ | _ \<Rightarrow> False)"
+(*
+| "grant s (Rename p f f') =
+ (case (parent f, parent f') of
+ (Some pf, Some pf') \<Rightarrow>
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf),
+ sectxt_of_obj s (O_file f), sectxt_of_obj s (O_dir pf')) of
+ (Some (pu,pr,pt), Some (pfu,pfr,pft), Some (fu,fr,ft),
+ Some (pfu',pfr',pft')) \<Rightarrow>
+ (search_check s (pu,pr,pt) f \<and>
+ permission_check (pu,pr,pt) (pfu,pfr,pft) C_dir P_remove_name \<and>
+ (if (is_file s f)
+ then permission_check (pu,pr,pt) (fu,fr,ft) C_file P_rename
+ else permission_check (pu,pr,pt) (fu,fr,ft) C_dir P_reparent) \<and>
+ search_check s (pu,pr,pt) pf' \<and>
+ permission_check (pu,pr,pt) (pfu',pfr',pft') C_dir P_add_name)
+ | _ \<Rightarrow> False)
+ | _ \<Rightarrow> False)"
+*)
+| "grant s (CreateMsgq p q) =
+ (case (sectxt_of_obj s (O_proc p)) of
+ Some (pu,pr,pt) \<Rightarrow>
+ (permission_check (pu,pr,pt) (pu,pr,pt) C_msgq P_associate \<and>
+ permission_check (pu,pr,pt) (pu,pr,pt) C_msgq P_create)
+ | _ \<Rightarrow> False)"
+| "grant s (SendMsg p q m) =
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_msgq q)) of
+ (Some (pu,pr,pt), Some (qu,qr,qt)) \<Rightarrow>
+ (permission_check (pu,pr,pt) (qu,qr,qt) C_msgq P_enqueue \<and>
+ permission_check (pu,pr,pt) (qu,qr,qt) C_msgq P_write \<and>
+ permission_check (pu,pr,pt) (pu,pr,pt) C_msg P_send)
+ | _ \<Rightarrow> False)"
+| "grant s (RecvMsg p q m) =
+ (case (sectxt_of_obj s (O_proc p),sectxt_of_obj s (O_msgq q),sectxt_of_obj s (O_msg q m)) of
+ (Some (pu,pr,pt), Some (qu,qr,qt), Some (mu,mr,mt)) \<Rightarrow>
+ permission_check (pu,pr,pt) (qu,qr,qt) C_msgq P_read \<and>
+ permission_check (pu,pr,pt) (mu,mr,mt) C_msg P_receive
+ | _ \<Rightarrow> False)"
+| "grant s (RemoveMsgq p q) =
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_msgq q)) of
+ (Some (pu,pr,pt), Some (qu,qr,qt)) \<Rightarrow>
+ permission_check (pu,pr,pt) (qu,qr,qt) C_msgq P_destroy
+ | _ \<Rightarrow> False)"
+(*
+| "grant s (CreateShM p h) =
+ (case (sectxt_of_obj s (O_proc p)) of
+ Some (pu,pr,pt) \<Rightarrow>
+ (permission_check (pu,pr,pt) (pu,pr,pt) C_shm P_associate \<and>
+ permission_check (pu,pr,pt) (pu,pr,pt) C_shm P_create)
+ | _ \<Rightarrow> False)"
+| "grant s (Attach p h flag) =
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_shm h)) of
+ (Some (pu,pr,pt), Some (hu,hr,ht)) \<Rightarrow>
+ if flag = SHM_RDONLY
+ then permission_check (pu,pr,pt) (hu,hr,ht) C_shm P_read
+ else permission_check (pu,pr,pt) (hu,hr,ht) C_shm P_read \<and>
+ permission_check (pu,pr,pt) (hu,hr,ht) C_shm P_write
+ | _ \<Rightarrow> False)"
+| "grant s (Detach p h) = True" (*?*)
+| "grant s (DeleteShM p h) =
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_shm h)) of
+ (Some (pu,pr,pt), Some (hu,hr,ht)) \<Rightarrow>
+ permission_check (pu,pr,pt) (hu,hr,ht) C_shm P_destroy
+ | _ \<Rightarrow> False)"
+*)
+| "grant s _ = False" (* socket event is currently not in consideration *)
+
+inductive valid :: "t_state \<Rightarrow> bool"
+where
+ "valid []"
+| "\<lbrakk>valid s; os_grant s e; grant s e\<rbrakk> \<Longrightarrow> valid (e # s)"
+
+(* tobj: object that can be tainted *)
+fun tobj_in_init :: "t_object \<Rightarrow> bool"
+where
+ "tobj_in_init (O_proc p) = (p \<in> init_procs)"
+| "tobj_in_init (O_file f) = (is_init_file f)"
+(* directory is not taintable
+| "tobj_in_init (O_dir f) = (f \<in> init_files)"
+*)
+| "tobj_in_init (O_node n) = (n \<in> init_nodes)"
+| "tobj_in_init (O_msg q m) = (m \<in> set (init_msgs_of_queue q) \<and> q \<in> init_msgqs)"
+| "tobj_in_init _ = False" (* other kind of obj cannot be tainted *)
+
+fun no_del_event:: "t_event list \<Rightarrow> bool"
+where
+ "no_del_event [] = True"
+| "no_del_event (Kill p p' # \<tau>) = False"
+| "no_del_event (CloseFd p fd # \<tau>) = False"
+| "no_del_event (UnLink p f # \<tau>) = False"
+| "no_del_event (Rmdir p f # \<tau>) = False"
+(*
+| "no_del_event (Rename p f f' # \<tau>) = False"
+*)
+| "no_del_event (RemoveMsgq p q # \<tau>) = False"
+(*
+| "no_del_event (RecvMsg p q m # \<tau>) = False"
+*)
+| "no_del_event (Exit p # s) = False"
+| "no_del_event (_ # \<tau>) = no_del_event \<tau>"
+
+end
+
+locale tainting = flask +
+
+fixes seeds :: "t_object set"
+assumes
+ seeds_in_init: "\<And> obj. obj \<in> seeds \<Longrightarrow> tobj_in_init obj"
+(*
+ and same_inode_in_seeds: "\<And> f f'. \<lbrakk>O_file f \<in> seeds; has_same_inode [] f f'\<rbrakk> \<Longrightarrow> O_file f' \<in> seeds"
+ and flow_shm_in_seeds: "\<And> p p'. \<lbrakk>O_proc p \<in> seeds; info_flow_shm [] p p'\<rbrakk> \<Longrightarrow> O_proc p' \<in> seeds"
+*)
+begin
+
+(*
+inductive tainted :: "t_object \<Rightarrow> t_state \<Rightarrow> bool" ("_ \<in> tainted _" [100, 100] 100)
+where
+ t_init: "obj \<in> seeds \<Longrightarrow> obj \<in> tainted []"
+| t_clone: "\<lbrakk>O_proc p \<in> tainted s; valid (Clone p p' fds shms # s)\<rbrakk>
+ \<Longrightarrow> O_proc p' \<in> tainted (Clone p p' fds shms # s)"
+| t_execve: "\<lbrakk>O_file f \<in> tainted s; valid (Execve p f fds # s)\<rbrakk>
+ \<Longrightarrow> O_proc p \<in> tainted (Execve p f fds # s)"
+| t_ptrace: "\<lbrakk>O_proc p \<in> tainted s; valid (Ptrace p p' # s); info_flow_shm s p' p''\<rbrakk>
+ \<Longrightarrow> O_proc p'' \<in> tainted (Ptrace p p' # s)"
+| t_ptrace':"\<lbrakk>O_proc p' \<in> tainted s; valid (Ptrace p p' # s); info_flow_shm s p p''\<rbrakk>
+ \<Longrightarrow> O_proc p'' \<in> tainted (Ptrace p p' # s)"
+| t_cfile: "\<lbrakk>O_proc p \<in> tainted s; valid (Open p f flags fd (Some inum) # s)\<rbrakk>
+ \<Longrightarrow> O_file f \<in> tainted (Open p f flags fd (Some inum) # s)"
+| t_read: "\<lbrakk>O_file f \<in> tainted s; valid (ReadFile p fd # s);
+ file_of_proc_fd s p fd = Some f; info_flow_shm s p p'\<rbrakk>
+ \<Longrightarrow> O_proc p' \<in> tainted (ReadFile p fd # s)"
+| t_write: "\<lbrakk>O_proc p \<in> tainted s; valid (WriteFile p fd # s);
+ file_of_proc_fd s p fd = Some f; has_same_inode s f f'\<rbrakk>
+ \<Longrightarrow> O_file f' \<in> tainted (WriteFile p fd # s)"
+(* directory is not tainted
+| t_mkdir: "\<lbrakk>O_proc p \<in> tainted s; valid (Mkdir p f inum # s)\<rbrakk>
+ \<Longrightarrow> O_dir f \<in> tainted (Mkdir p f inum # s)"
+ *)
+| t_link: "\<lbrakk>O_file f \<in> tainted s; valid (LinkHard p f f' # s)\<rbrakk>
+ \<Longrightarrow> O_file f' \<in> tainted (LinkHard p f f' # s)"
+| t_trunc: "\<lbrakk>O_proc p \<in> tainted s; valid (Truncate p f len # s); len > 0; has_same_inode s f f'\<rbrakk>
+ \<Longrightarrow> O_file f' \<in> tainted (Truncate p f len # s)"
+(*
+| t_rename: "\<lbrakk>O_file f'' \<in> tainted s; valid (Rename p f f' # s);f \<preceq> f''\<rbrakk>
+ \<Longrightarrow> O_file (file_after_rename f f' f'') \<in> tainted (Rename p f f' # s)"
+| t_rename':"\<lbrakk>O_dir f'' \<in> tainted s; valid (Rename p f f' # s);f \<preceq> f''\<rbrakk>
+ \<Longrightarrow> O_dir (file_after_rename f f' f'') \<in> tainted (Rename p f f' # s)"
+*)
+| t_attach1:"\<lbrakk>O_proc p \<in> tainted s; valid (Attach p h SHM_RDWR # s); (p', flag') \<in> procs_of_shm s h; info_flow_shm s p' p''\<rbrakk>
+ \<Longrightarrow> O_proc p'' \<in> tainted (Attach p h SHM_RDWR # s)"
+| t_attach2:"\<lbrakk>O_proc p' \<in> tainted s; valid (Attach p h flag # s); (p', SHM_RDWR) \<in> procs_of_shm s h; info_flow_shm s p p''\<rbrakk>
+ \<Longrightarrow> O_proc p'' \<in> tainted (Attach p h flag # s)"
+| t_sendmsg:"\<lbrakk>O_proc p \<in> tainted s; valid (SendMsg p q m # s)\<rbrakk>
+ \<Longrightarrow> O_msg q m \<in> tainted (SendMsg p q m # s)"
+| t_recvmsg:"\<lbrakk>O_msg q m \<in> tainted s; valid (RecvMsg p q m # s); info_flow_shm s p p'\<rbrakk>
+ \<Longrightarrow> O_proc p' \<in> tainted (RecvMsg p q m # s)"
+| t_remain: "\<lbrakk>obj \<in> tainted s; valid (e # s); alive (e # s) obj\<rbrakk>
+ \<Longrightarrow> obj \<in> tainted (e # s)"
+*)
+
+fun tainted :: "t_state \<Rightarrow> t_object set"
+where
+ "tainted [] = seeds"
+| "tainted (Clone p p' fds # s) =
+ (if (O_proc p) \<in> tainted s then tainted s \<union> {O_proc p'} else tainted s)"
+| "tainted (Execve p f fds # s) =
+ (if (O_file f) \<in> tainted s then tainted s \<union> {O_proc p} else tainted s)"
+| "tainted (Kill p p' # s) = tainted s - {O_proc p'}"
+| "tainted (Ptrace p p' # s) =
+ (if (O_proc p \<in> tainted s)
+ then tainted s \<union> {O_proc p'}
+ else if (O_proc p' \<in> tainted s)
+ then tainted s \<union> {O_proc p}
+ else tainted s)"
+| "tainted (Exit p # s) = tainted s - {O_proc p}"
+| "tainted (Open p f flags fd opt # s) =
+ (case opt of
+ Some inum \<Rightarrow> (if (O_proc p) \<in> tainted s
+ then tainted s \<union> {O_file f}
+ else tainted s)
+ | _ \<Rightarrow> tainted s)"
+| "tainted (ReadFile p fd # s) =
+ (case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> if (O_file f) \<in> tainted s
+ then tainted s \<union> {O_proc p}
+ else tainted s
+ | None \<Rightarrow> tainted s)"
+| "tainted (WriteFile p fd # s) =
+ (case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> if (O_proc p) \<in> tainted s
+ then tainted s \<union> {O_file f}
+ else tainted s
+ | None \<Rightarrow> tainted s)"
+| "tainted (CloseFd p fd # s) =
+ (case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> ( if ((proc_fd_of_file s f = {(p,fd)}) \<and> (f \<in> files_hung_by_del s))
+ then tainted s - {O_file f} else tainted s )
+ | _ \<Rightarrow> tainted s)"
+| "tainted (UnLink p f # s) =
+ (if (proc_fd_of_file s f = {}) then tainted s - {O_file f} else tainted s)"
+| "tainted (Truncate p f len # s) =
+ (if (len > 0 \<and> O_proc p \<in> tainted s)
+ then tainted s \<union> {O_file f}
+ else tainted s)"
+| "tainted (SendMsg p q m # s) =
+ (if (O_proc p \<in> tainted s) then tainted s \<union> {O_msg q m} else tainted s)"
+| "tainted (RecvMsg p q m # s) =
+ (if (O_msg q m \<in> tainted s)
+ then (tainted s \<union> {O_proc p}) - {O_msg q m}
+ else tainted s)"
+| "tainted (RemoveMsgq p q # s) = tainted s - {O_msg q m| m. O_msg q m \<in> tainted s}"
+| "tainted (e # s) = tainted s"
+
+
+(* reasonable tainted object, fd and sock and msgq and msg is excluded
+ * this is different from tainted, which already excluded some of them,
+ * tainted not excluded msg, which does not have the meaning of "tainteable", but
+ * but acting as a media to "pass" the virus. We normally will not say that
+ * a message is tableable or not.
+ *)
+fun appropriate :: "t_object \<Rightarrow> bool"
+where
+ "appropriate (O_proc p) = (p \<in> init_procs)"
+| "appropriate (O_file f) = (is_init_file f)"
+| "appropriate (O_dir f) = False"
+| "appropriate (O_fd p fd) = False"
+| "appropriate (O_node n) = False" (* cause socket is temperary not considered *)
+| "appropriate (O_tcp_sock k) = False"
+| "appropriate (O_udp_sock k) = False"
+| "appropriate (O_msgq q) = False"
+| "appropriate (O_msg q m) = False"
+
+definition taintable:: "t_object \<Rightarrow> bool"
+where
+ "taintable obj \<equiv> init_alive obj \<and> appropriate obj \<and> (\<exists> s. obj \<in> tainted s)"
+
+definition undeletable :: "t_object \<Rightarrow> bool"
+where
+ "undeletable obj \<equiv> init_alive obj \<and> \<not> (\<exists> s. valid s \<and> deleted obj s)"
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Flask_type.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,241 @@
+theory Flask_type
+imports Main
+begin
+
+
+(************* types that for modeling Access-Control part of Selinux *************)
+
+(* Table 4-5 of "Selinux by example" *)
+datatype permission_t =
+(* C_file *)
+ P_append (* the O_append flag for open *)
+| P_create (* create new file *)
+| P_entrypoint (* file can be used as the entry point of new domain via a domain transition *)
+(* | P_execmod execute memory-mapped files that have modified *)
+| P_execute (* corresponds to x access in standard linux *)
+(* | P_execute_no_trans: execute file in calller's domain, without domain transition *)
+(* | P_getattr : stat, ls, ioctls, ... not-related to security *)
+(*| P_link (* create hard link to file *) *)
+| P_read (* corresponds to r access in standard linux *)
+| P_rename (* rename a hard link *)
+| P_setattr (* chmod, ioctls ... *)
+| P_unlink (* remove hard link (delete *)
+| P_write (* corresponds to w access in standard linux *)
+(* P_lock; P_mounton; P_quotaon; P_relabelfrom; P_relabelto *)
+
+(* C_dir *)
+| P_add_name (* add file in directory *)
+| P_remove_name (* like write for file *)
+(*
+| P_reparent (* like rename for file *) *)
+| P_search
+| P_rmdir (* like unlink for file *)
+
+(* C_fd *)
+(* P_create; P_getattr; P_setattr*)
+| P_inherit (* inherit across execve,
+we simplify that execve would pass all fds to new "process", but flask checks the inherit right, if denied the fd is closed. *)
+(* P_receive: ignored, ? I can find a reasonable explanation *)
+
+(* C_process *)
+| P_fork
+| P_ptrace
+| P_transition
+| P_share (* allow state sharing for execve, clone: fds and ipc
+ * for clone, we might have two flags, share Fd and IPC
+ * for execve, as default, fds are shared, IPCs are detached.
+ * But fds not passed access-control are closed
+ *)
+(*| P_siginh: inheritance of signal state, pending signals, this is not needed, as
+we assume all signal are delivered instantly *)
+(* P_rlimitnh : inheritance of resource limits *)
+(* P_dyntransition; P_execheap; P_execmem; P_execstack; get_attr : not security related;
+ * P_getcap; P_getsched *)
+| P_sigkill (* signal *)
+
+(* C_msgq, C_shm *)
+| P_associate
+| P_destroy
+| P_enqueue
+| P_receive
+| P_send
+(* create, read, write *)
+
+type_synonym access_vector_t = "permission_t set"
+
+(************ Object Classes *************)
+datatype security_class_t =
+ C_process
+
+(* file-related object classes, 4.3.1 of Book *)
+| C_file
+| C_fd
+| C_dir
+(* C_blk file C_chr_file C_fifo_file C_filesystem C_lnk_file C_sock_file *)
+
+(* network, 4.3.2, using only tcp/udp as a demonstration *)
+| C_node (* Host by IP address or range of addresses *)
+| C_netif (* network interface, eth0 *)
+| C_tcp_socket
+| C_udp_socket
+(* netlink_socket unix_socket rawip_socket *)
+
+(* IPCs, 4.3.3 *)
+| C_msg (* messages within a message queue *)
+| C_msgq (* message queues *)
+(*
+| C_shm (* shared memory segment *) *)
+(* C_semaphores: only for sync, not tainting related *)
+
+(*other classes, 4.3.4*)
+(* C_capability; C_security:security server; *)
+(* | C_system the system itself *)
+
+(********* TE rules **********
+ * 1. type declaration rules
+ * 2. type transition rules
+ * 3. type member rules
+ * 4. type change rules
+ * 5. access vector rules
+ * 6. cloning rules
+ * 7. access vector assertions
+ *)
+
+(* 1. type declaration rule *)
+datatype type_t =
+ T_syslogd
+| T_syslogd_exec
+| T_devlog
+
+type_synonym attribute_t = "type_t set"
+
+(* for the asterisk/wildcard in the rules lists*)
+datatype 'a patt =
+ Single "'a"
+| Set "'a set"
+| Tilde "'a set" (* complement *)
+| Asterisk
+
+(* 2. type transition rule, applied only when: execve & createfile
+ * (creating subj_type, related obj_typ, type class, default type)
+ * related obj_typ: proc \<rightarrow> executable; file \<rightarrow> parent-file
+ * if not in the list, use subj_type for proc and parent_type for file
+ * if multiple in the list, use the first one (they use the last one, we can say our list is the rev of their list, for "lookup" is easy)
+ * if using attributes, it can be transformed into the simple form of this list,
+ now we use "patt" to deal with attributes
+ *)
+type_synonym type_transition_rule_list_t =
+ "(type_t patt \<times> type_t patt \<times> security_class_t \<times> type_t) list"
+
+
+(* for type member rules:
+ * ?? I don't know the ideal yet, page 13 bottom of left column
+ *
+type_synonym type_member_rule_list_t =
+ "(type_t patt \<times> type_t patt \<times> security_class_t \<times> type_t) list"
+ *)
+
+(* for type change rules (relabeling to diff security context):
+ * ?? I don't know what's its meaning yet, page 13 up of the right column
+ *
+type_synonym type_change_rule_list_t =
+ "(type_t patt \<times> type_t patt \<times> security_class_t \<times> type_t) list"
+ *)
+
+datatype 'a target_with_self_t =
+ Not_self 'a
+| Self
+
+(* for access vector rules(verify a role-changed if allowed or not):
+ * Any role-changed should be verified by this list!
+ * here we only do the allow rule, audit and notify rules are ignored
+ * (source type, target type, class, permissionS)
+ * permissionS might be access_vector_t, which is already be model as a "set" of "pattern"
+ * Self in the target type means this rule should be applied to the source process itself
+ * if multiple, we search the whole list to mimic their "the union" ideal
+ *)
+type_synonym allowed_rule_list_t =
+ "(type_t patt \<times> (type_t patt) target_with_self_t \<times> security_class_t \<times> permission_t patt) list"
+
+datatype role_t =
+ R_object
+| R_user
+| R_login
+
+(* role declarations:
+ * (r, ts): ts are the types/domains that can be associated with r
+ *)
+type_synonym role_declarations_list_t =
+ "(role_t \<times> type_t set) list"
+
+(* role transition rules(only happens when execve call):
+ * (r, t, r'): the process with role r executes a file with type t, its role will be changed to r'
+ *)
+type_synonym role_transition_rule_list_t =
+ "(role_t \<times> type_t \<times> role_t) list"
+
+(* role allow rules:
+ * (from_role, to_role) changes when "Execve" calls!
+ * What's the difference between this list and the role_transition_rule_list?
+ * It is: if a process execute a file, the new role will be determined by
+ * role-transition, but this new role should be in role-allow(if not deny),
+ * and the new type should be compatible with the new role, which is checked
+ * by role-declarations(if not deny).
+ *)
+type_synonym role_allow_rule_list_t =
+ "(role_t set \<times> role_t set) list"
+
+datatype user_t =
+ U_system
+| U_sds
+| U_pal
+
+(* user declarations:
+ * (u, roles) roles are associated with user u
+ * here MLS are not in consideration
+ *)
+type_synonym user_declarations_list_t =
+ "(user_t \<times> role_t set)"
+
+type_synonym security_context_t =
+ "(user_t \<times> role_t \<times> type_t)"
+
+(* constrains :
+ * (classes, perms, expression) list
+ * when flask do a permission check, constrains will fetch the
+ * associated two security-contexts (source & target), flask
+ * makes sure that appling expression to these two contexts,
+ * will get True. If False, deny the request.
+ *)
+type_synonym constrains_list_t =
+ "(security_class_t set \<times> permission_t set \<times> (security_context_t \<Rightarrow> security_context_t \<Rightarrow> bool)) list"
+
+(* comments:
+ * 1. sid is not needed, it is a run-time index for security-context,
+ * what we do is policy-analysis, so no this "dynamic" sid as mid-ware;
+ * 2. "seeds" of tainting cannot be solid objects in the initial states,
+ should be extended as a "condition/predicate" to describe the "seeds"
+ *)
+
+(* "temporary unused definitions :
+
+datatype security_id_t =
+ SID nat
+| SECSID_NULL
+| SECSID_WILD
+
+
+(* this list mimic the action of "security_compute_av",
+ * if in the list, means func returns "grant"
+ * if not in, means "denied"
+ * !!! this list is a "work result" of policy by selinux's
+ * pre-processing func "security_compute_av"
+ * in my formalisation, this is not needed, cause the specification of
+ * our "selinux" contains already the policy specification.
+ * we can directly get the access-vector for sid/sectxt based on allow-rules & constrains
+ *)
+type_synonym avc_t = "(security_id_t \<times> security_id_t \<times> security_class_t \<times> access_vector_t) list"
+
+ *)
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Info_flow_shm_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,831 @@
+theory Info_flow_shm_prop
+imports Main Flask_type Flask My_list_prefix Init_prop Valid_prop Delete_prop Current_prop
+begin
+
+context flask begin
+
+(*********** simpset for one_flow_shm **************)
+
+lemma one_flow_not_self:
+ "one_flow_shm s h p p \<Longrightarrow> False"
+by (simp add:one_flow_shm_def)
+
+lemma one_flow_shm_attach:
+ "valid (Attach p h flag # s) \<Longrightarrow> one_flow_shm (Attach p h flag # s) = (\<lambda> h' pa pb.
+ if (h' = h)
+ then (pa = p \<and> pb \<noteq> p \<and> flag = SHM_RDWR \<and> (\<exists> flagb. (pb, flagb) \<in> procs_of_shm s h)) \<or>
+ (pb = p \<and> pa \<noteq> p \<and> (pa, SHM_RDWR) \<in> procs_of_shm s h) \<or>
+ (one_flow_shm s h pa pb)
+ else one_flow_shm s h' pa pb )"
+apply (rule ext, rule ext, rule ext, frule vd_cons, frule vt_grant_os)
+by (auto simp add: one_flow_shm_def)
+
+lemma one_flow_shm_detach:
+ "valid (Detach p h # s) \<Longrightarrow> one_flow_shm (Detach p h # s) = (\<lambda> h' pa pb.
+ if (h' = h)
+ then (pa \<noteq> p \<and> pb \<noteq> p \<and> one_flow_shm s h' pa pb)
+ else one_flow_shm s h' pa pb)"
+apply (rule ext, rule ext, rule ext, frule vt_grant_os)
+by (auto simp:one_flow_shm_def)
+
+lemma one_flow_shm_deleteshm:
+ "valid (DeleteShM p h # s) \<Longrightarrow> one_flow_shm (DeleteShM p h # s) = (\<lambda> h' pa pb.
+ if (h' = h)
+ then False
+ else one_flow_shm s h' pa pb)"
+apply (rule ext, rule ext, rule ext, frule vt_grant_os)
+by (auto simp: one_flow_shm_def)
+
+lemma one_flow_shm_clone:
+ "valid (Clone p p' fds shms # s) \<Longrightarrow> one_flow_shm (Clone p p' fds shms # s) = (\<lambda> h pa pb.
+ if (pa = p' \<and> pb \<noteq> p' \<and> h \<in> shms)
+ then (if (pb = p) then (flag_of_proc_shm s p h = Some SHM_RDWR) else one_flow_shm s h p pb)
+ else if (pb = p' \<and> pa \<noteq> p' \<and> h \<in> shms)
+ then (if (pa = p) then (flag_of_proc_shm s p h = Some SHM_RDWR) else one_flow_shm s h pa p)
+ else one_flow_shm s h pa pb)"
+apply (rule ext, rule ext, rule ext, frule vt_grant_os, frule vd_cons, clarsimp)
+apply (frule_tac p = p' in procs_of_shm_prop2', simp)
+apply (auto simp:one_flow_shm_def intro:procs_of_shm_prop4 flag_of_proc_shm_prop1)
+done
+
+lemma one_flow_shm_execve:
+ "valid (Execve p f fds # s) \<Longrightarrow> one_flow_shm (Execve p f fds # s) = (\<lambda> h pa pb.
+ pa \<noteq> p \<and> pb \<noteq> p \<and> one_flow_shm s h pa pb )"
+apply (rule ext, rule ext, rule ext, frule vt_grant_os, frule vd_cons)
+by (auto simp:one_flow_shm_def)
+
+lemma one_flow_shm_kill:
+ "valid (Kill p p' # s) \<Longrightarrow> one_flow_shm (Kill p p' # s) = (\<lambda> h pa pb.
+ pa \<noteq> p' \<and> pb \<noteq> p' \<and> one_flow_shm s h pa pb )"
+apply (rule ext, rule ext, rule ext, frule vt_grant_os, frule vd_cons)
+by (auto simp:one_flow_shm_def)
+
+lemma one_flow_shm_exit:
+ "valid (Exit p # s) \<Longrightarrow> one_flow_shm (Exit p # s) = (\<lambda> h pa pb.
+ pa \<noteq> p \<and> pb \<noteq> p \<and> one_flow_shm s h pa pb )"
+apply (rule ext, rule ext, rule ext, frule vt_grant_os, frule vd_cons)
+by (auto simp:one_flow_shm_def)
+
+lemma one_flow_shm_other:
+ "\<lbrakk>valid (e # s);
+ \<forall> p h flag. e \<noteq> Attach p h flag;
+ \<forall> p h. e \<noteq> Detach p h;
+ \<forall> p h. e \<noteq> DeleteShM p h;
+ \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms;
+ \<forall> p f fds. e \<noteq> Execve p f fds;
+ \<forall> p p'. e \<noteq> Kill p p';
+ \<forall> p. e \<noteq> Exit p
+ \<rbrakk> \<Longrightarrow> one_flow_shm (e # s) = one_flow_shm s"
+apply (rule ext, rule ext, rule ext, frule vt_grant_os, frule vd_cons)
+apply (case_tac e, auto simp:one_flow_shm_def dest:procs_of_shm_prop2)
+apply (drule procs_of_shm_prop1, auto)
+done
+
+lemmas one_flow_shm_simps = one_flow_shm_other one_flow_shm_attach one_flow_shm_detach one_flow_shm_deleteshm
+ one_flow_shm_clone one_flow_shm_execve one_flow_shm_kill one_flow_shm_exit
+
+type_synonym t_edge_shm = "t_process \<times> t_shm \<times> t_process"
+fun Fst:: "t_edge_shm \<Rightarrow> t_process" where "Fst (a, b, c) = a"
+fun Snd:: "t_edge_shm \<Rightarrow> t_shm" where "Snd (a, b, c) = b"
+fun Trd:: "t_edge_shm \<Rightarrow> t_process" where "Trd (a, b, c) = c"
+
+fun edge_related:: "t_edge_shm list \<Rightarrow> t_process \<Rightarrow> t_shm \<Rightarrow> bool"
+where
+ "edge_related [] p h = False"
+| "edge_related ((from, shm, to) # path) p h =
+ (if (((p = from) \<or> (p = to)) \<and> (h = shm)) then True
+ else edge_related path p h)"
+
+inductive path_by_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_edge_shm list \<Rightarrow> t_process \<Rightarrow> bool"
+where
+ pbs1: "p \<in> current_procs s \<Longrightarrow> path_by_shm s p [] p"
+| pbs2: "\<lbrakk>path_by_shm s p path p'; one_flow_shm s h p' p''; p'' \<notin> set (map Fst path)\<rbrakk>
+ \<Longrightarrow> path_by_shm s p ((p', h, p'')# path) p''"
+
+
+lemma one_step_path: "\<lbrakk>one_flow_shm s h p p'; valid s\<rbrakk> \<Longrightarrow> path_by_shm s p [(p, h, p')] p'"
+apply (rule_tac path = "[]" and p = p in path_by_shm.intros(2))
+apply (rule path_by_shm.intros(1))
+apply (auto intro:procs_of_shm_prop2 simp:one_flow_shm_def)
+done
+
+lemma pbs_prop1:
+ "path_by_shm s p path p' \<Longrightarrow> ((path = []) = (p = p')) \<and> (path \<noteq> [] \<longrightarrow> p \<in> set (map Fst path))"
+apply (erule path_by_shm.induct, simp)
+apply (auto simp:one_flow_shm_def)
+done
+
+lemma pbs_prop2:
+ "path_by_shm s p path p' \<Longrightarrow> (path = []) = (p = p')"
+by (simp add:pbs_prop1)
+
+lemma pbs_prop2':
+ "path_by_shm s p path p \<Longrightarrow> path = []"
+by (simp add:pbs_prop2)
+
+lemma pbs_prop3:
+ "\<lbrakk>path_by_shm s p path p'; path \<noteq> []\<rbrakk> \<Longrightarrow> p \<in> set (map Fst path)"
+by (drule pbs_prop1, auto)
+
+lemma pbs_prop4[rule_format]:
+ "path_by_shm s p path p'\<Longrightarrow> path \<noteq> [] \<longrightarrow> p' \<in> set (map Trd path)"
+by (erule path_by_shm.induct, auto)
+
+lemma pbs_prop5[rule_format]:
+ "path_by_shm s p path p' \<Longrightarrow> path \<noteq> [] \<longrightarrow> p' \<notin> set (map Fst path)"
+by (erule path_by_shm.induct, auto simp:one_flow_shm_def)
+
+lemma pbs_prop6_aux:
+ "path_by_shm s pa pathac pc \<Longrightarrow> valid s \<longrightarrow> (\<forall> pb \<in> set (map Fst pathac). \<exists> pathab pathbc. path_by_shm s pa pathab pb \<and> path_by_shm s pb pathbc pc \<and> pathac = pathbc @ pathab)"
+apply (erule path_by_shm.induct)
+apply simp
+apply clarify
+apply (case_tac "pb = p'", simp)
+apply (rule_tac x = path in exI, simp)
+apply (erule one_step_path, simp)
+apply (erule_tac x = pb in ballE, simp_all, clarsimp)
+apply (rule_tac x = pathab in exI, simp)
+apply (erule pbs2, auto)
+done
+
+lemma pbs_prop6:
+ "\<lbrakk>path_by_shm s pa pathac pc; pb \<in> set (map Fst pathac); valid s\<rbrakk>
+ \<Longrightarrow> \<exists> pathab pathbc. path_by_shm s pa pathab pb \<and> path_by_shm s pb pathbc pc \<and> pathac = pathbc @ pathab"
+by (drule pbs_prop6_aux, auto)
+
+lemma pbs_prop7_aux:
+ "path_by_shm s pa pathac pc \<Longrightarrow> valid s \<longrightarrow> (\<forall> pb \<in> set (map Trd pathac). \<exists> pathab pathbc. path_by_shm s pa pathab pb \<and> path_by_shm s pb pathbc pc \<and> pathac = pathbc @ pathab)"
+apply (erule path_by_shm.induct)
+apply simp
+apply clarify
+apply (case_tac "pb = p''", simp)
+apply (rule_tac x = "(p',h,p'') # path" in exI, simp)
+apply (rule conjI, erule pbs2, simp+)
+apply (rule pbs1, clarsimp simp:one_flow_shm_def procs_of_shm_prop2)
+apply (erule_tac x = pb in ballE, simp_all, clarsimp)
+apply (rule_tac x = pathab in exI, simp)
+apply (erule pbs2, auto)
+done
+
+lemma pbs_prop7:
+ "\<lbrakk>path_by_shm s pa pathac pc; pb \<in> set (map Trd pathac); valid s\<rbrakk>
+ \<Longrightarrow> \<exists> pathab pathbc. path_by_shm s pa pathab pb \<and> path_by_shm s pb pathbc pc \<and> pathac = pathbc @ pathab"
+by (drule pbs_prop7_aux, drule mp, simp, erule_tac x = pb in ballE, auto)
+
+lemma pbs_prop8:
+ "path_by_shm s p path p' \<Longrightarrow> (set (map Fst path) - {p}) = (set (map Trd path) - {p'})"
+proof (induct rule:path_by_shm.induct)
+ case (pbs1 p s)
+ thus ?case by simp
+next
+ case (pbs2 s p path p' h p'')
+ assume p1:"path_by_shm s p path p'" and p2: "set (map Fst path) - {p} = set (map Trd path) - {p'}"
+ and p3: "one_flow_shm s h p' p''" and p4: "p'' \<notin> set (map Fst path)"
+ show "set (map Fst ((p', h, p'') # path)) - {p} = set (map Trd ((p', h, p'') # path)) - {p''}"
+ (is "?left = ?right")
+ proof (cases "path = []")
+ case True
+ with p1 have "p = p'" by (drule_tac pbs_prop2, simp)
+ thus ?thesis using True
+ using p2 by (simp)
+ next
+ case False
+ with p1 have a1: "p \<noteq> p'" by (drule_tac pbs_prop2, simp)
+ from p3 have a2: "p' \<noteq> p''" by (simp add:one_flow_shm_def)
+ from p1 False have a3: "p' \<in> set (map Trd path)" by (drule_tac pbs_prop4, simp+)
+ from p4 p1 False have a4: "p \<noteq> p''" by (drule_tac pbs_prop3, auto)
+ with p2 a2 p4 have a5: "p'' \<notin> set (map Trd path)" by auto
+
+ have "?left = (set (map Fst path) - {p}) \<union> {p'}" using a1 by auto
+ moreover have "... = (set (map Trd path) - {p'}) \<union> {p'}"
+ using p2 by auto
+ moreover have "... = set (map Trd path)" using a3 by auto
+ moreover have "... = set (map Trd path) - {p''}" using a5 by simp
+ moreover have "... = ?right" by simp
+ ultimately show ?thesis by simp
+ qed
+qed
+
+lemma pbs_prop9_aux[rule_format]:
+ "path_by_shm s p path p' \<Longrightarrow> h \<in> set (map Snd path) \<and> valid s \<longrightarrow> (\<exists> pa pb patha pathb. path_by_shm s p patha pa \<and> path_by_shm s pb pathb p' \<and> one_flow_shm s h pa pb \<and> path = pathb @ [(pa, h, pb)] @ patha \<and> h \<notin> set (map Snd patha))"
+apply (erule path_by_shm.induct, simp)
+apply (rule impI, case_tac "h \<in> set (map Snd path)", simp_all)
+apply (erule exE|erule conjE)+
+apply (rule_tac x = pa in exI, rule_tac x = pb in exI, rule_tac x = patha in exI, simp)
+apply (rule pbs2, auto)
+apply (rule_tac x = p' in exI, rule_tac x = p'' in exI, rule_tac x = path in exI, simp)
+apply (rule pbs1, clarsimp simp:one_flow_shm_def procs_of_shm_prop2)
+done
+
+lemma pbs_prop9:
+ "\<lbrakk>h \<in> set (map Snd path); path_by_shm s p path p'; valid s\<rbrakk>
+ \<Longrightarrow> \<exists> pa pb patha pathb. path_by_shm s p patha pa \<and> path_by_shm s pb pathb p' \<and>
+ one_flow_shm s h pa pb \<and> path = pathb @ [(pa, h, pb)] @ patha \<and> h \<notin> set (map Snd patha)"
+by (rule pbs_prop9_aux, auto)
+
+lemma path_by_shm_trans_aux[rule_format]:
+ "path_by_shm s p' path' p'' \<Longrightarrow> valid s \<longrightarrow> (\<forall> p path. path_by_shm s p path p' \<longrightarrow> (\<exists> path''. path_by_shm s p path'' p''))"
+proof (induct rule:path_by_shm.induct)
+ case (pbs1 p s)
+ thus ?case
+ by (clarify, rule_tac x = path in exI, simp)
+next
+ case (pbs2 s p path p' h p'')
+ hence p1: "path_by_shm s p path p'" and p2: "one_flow_shm s h p' p''"
+ and p3: "valid s \<longrightarrow> (\<forall>pa path. path_by_shm s pa path p \<longrightarrow> (\<exists>path''. path_by_shm s pa path'' p'))"
+ and p4: "p'' \<notin> set (map Fst path)" by auto
+ show ?case
+ proof clarify
+ fix pa path'
+ assume p5: "path_by_shm s pa path' p" and p6: "valid s"
+ with p3 obtain path'' where a1: "path_by_shm s pa path'' p'" by auto
+ have p3': "\<forall>pa path. path_by_shm s pa path p \<longrightarrow> (\<exists>path''. path_by_shm s pa path'' p')"
+ using p3 p6 by simp
+ show "\<exists>path''. path_by_shm s pa path'' p''"
+ proof (cases "p'' \<in> set (map Fst path'')")
+ case True
+ then obtain res where "path_by_shm s pa res p''" using a1 pbs_prop6 p6 by blast
+ thus ?thesis by auto
+ next
+ case False
+ with p2 a1 show ?thesis
+ apply (rule_tac x = "(p', h, p'') # path''" in exI)
+ apply (rule path_by_shm.intros(2), auto)
+ done
+ qed
+ qed
+qed
+
+lemma path_by_shm_trans:
+ "\<lbrakk>path_by_shm s p path p'; path_by_shm s p' path' p''; valid s\<rbrakk> \<Longrightarrow> \<exists> path''. path_by_shm s p path'' p''"
+by (drule_tac p' = p' and p'' = p'' in path_by_shm_trans_aux, auto)
+
+lemma path_by_shm_intro1_prop:
+ "\<not> path_by_shm s p [] p \<Longrightarrow> p \<notin> current_procs s"
+by (auto dest:path_by_shm.intros(1))
+
+lemma path_by_shm_intro3:
+ "\<lbrakk>path_by_shm s p path from; (from, SHM_RDWR) \<in> procs_of_shm s h; (to, flag) \<in> procs_of_shm s h;
+ to \<notin> set (map Fst path); from \<noteq> to\<rbrakk>
+ \<Longrightarrow> path_by_shm s p ((from, h, to)#path) to"
+apply (rule path_by_shm.intros(2), simp_all)
+by (auto simp:one_flow_shm_def)
+
+lemma path_by_shm_intro4:
+ "\<lbrakk>(p, flag) \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> path_by_shm s p [] p"
+by (drule procs_of_shm_prop2, simp, simp add:path_by_shm.intros(1))
+
+lemma path_by_shm_intro5:
+ "\<lbrakk>(from, SHM_RDWR) \<in> procs_of_shm s h; (to,flag) \<in> procs_of_shm s h; valid s; from \<noteq> to\<rbrakk>
+ \<Longrightarrow> path_by_shm s from [(from, h, to)] to"
+apply (rule_tac p' = "from" and h = h in path_by_shm.intros(2))
+apply (rule path_by_shm.intros(1), simp add:procs_of_shm_prop2)
+apply (simp add:one_flow_shm_def, rule_tac x = flag in exI, auto)
+done
+
+(* p'' \<notin> set (map Fst path): not duplicated target process;
+ * p1 - ha \<rightarrow> p2; p2 - hb \<rightarrow> p3; p3 - ha \<rightarrow> p4; so path_by_shm p1 [(p3,ha,p4), (p2,hb,p3),(p1,ha,p2)] p4,
+ * but this could be also path_by_shm p1 [(p1,ha,p4)] p4, so the former one is redundant! *)
+
+inductive path_by_shm':: "t_state \<Rightarrow> t_process \<Rightarrow> t_edge_shm list \<Rightarrow> t_process \<Rightarrow> bool"
+where
+ pbs1': "p \<in> current_procs s \<Longrightarrow> path_by_shm' s p [] p"
+| pbs2': "\<lbrakk>path_by_shm' s p path p'; one_flow_shm s h p' p''; p'' \<notin> set (map Fst path);
+ h \<notin> set (map Snd path)\<rbrakk>
+ \<Longrightarrow> path_by_shm' s p ((p', h, p'')# path) p''"
+
+lemma pbs_prop10:
+ "\<lbrakk>path_by_shm s p path p'; one_flow_shm s h p' p''; valid s\<rbrakk> \<Longrightarrow> \<exists>path'. path_by_shm s p path' p''"
+apply (case_tac "p'' \<in> set (map Fst path)")
+apply (drule_tac pb = p'' in pbs_prop6, simp+)
+apply ((erule exE|erule conjE)+, rule_tac x = pathab in exI, simp)
+apply (rule_tac x = "(p', h, p'') # path" in exI)
+apply (erule pbs2, simp+)
+done
+
+lemma pbs'_imp_pbs[rule_format]:
+ "path_by_shm' s p path p' \<Longrightarrow> valid s \<longrightarrow> (\<exists> path'. path_by_shm s p path' p')"
+apply (erule path_by_shm'.induct)
+apply (rule impI, rule_tac x = "[]" in exI, simp add:pbs1)
+apply (rule impI, clarsimp)
+apply (erule pbs_prop10, simp+)
+done
+
+lemma pbs_imp_pbs'[rule_format]:
+ "path_by_shm s p path p' \<Longrightarrow> valid s \<longrightarrow> (\<exists> path'. path_by_shm' s p path' p')"
+apply (erule path_by_shm.induct)
+apply (rule impI, rule_tac x = "[]" in exI, erule pbs1')
+apply (rule impI, simp, erule exE) (*
+apply ( erule exE, case_tac "h \<in> set (map Snd path)")
+apply (drule_tac s = s and p = p and p' = p' in pbs_prop9, simp+) defer
+apply (rule_tac x = "(p', h, p'') # path'" in exI)
+apply (erule pbs2', simp+)
+apply ((erule exE|erule conjE)+)
+apply (rule_tac x = "(pa, h, p'') # patha" in exI)
+apply (erule pbs2', auto simp:one_flow_shm_def)
+done*)
+sorry
+
+
+lemma pbs'_eq_pbs:
+ "valid s \<Longrightarrow> (\<exists> path'. path_by_shm' s p path' p') = (\<exists> path. path_by_shm s p path p')"
+by (rule iffI, auto intro:pbs_imp_pbs' pbs'_imp_pbs)
+
+definition flow_by_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
+where
+ "flow_by_shm s p p' \<equiv> \<exists> path. path_by_shm s p path p'"
+
+lemma flow_by_shm_intro':
+ "valid s \<Longrightarrow> flow_by_shm s p p' = (\<exists> path. path_by_shm' s p path p')"
+by (auto simp:flow_by_shm_def pbs'_eq_pbs)
+
+lemma one_step_flows: "\<lbrakk>one_flow_shm s h p p'; valid s\<rbrakk> \<Longrightarrow> flow_by_shm s p p'"
+by (drule one_step_path, auto simp:flow_by_shm_def)
+
+lemma flow_by_shm_trans:
+ "\<lbrakk>flow_by_shm s p p'; flow_by_shm s p' p''; valid s\<rbrakk> \<Longrightarrow> flow_by_shm s p p''"
+by (auto simp:flow_by_shm_def intro!:path_by_shm_trans)
+
+lemma flow_by_shm_intro1_prop:
+ "\<not> flow_by_shm s p p \<Longrightarrow> p \<notin> current_procs s"
+by (auto dest:path_by_shm.intros(1) simp:flow_by_shm_def)
+
+lemma flow_by_shm_intro1:
+ "p \<in> current_procs s \<Longrightarrow> flow_by_shm s p p"
+by (auto dest:path_by_shm.intros(1) simp:flow_by_shm_def)
+
+lemma flow_by_shm_intro2:
+ "\<lbrakk>flow_by_shm s p p'; one_flow_shm s h p' p''; valid s\<rbrakk> \<Longrightarrow> flow_by_shm s p p''"
+by (auto intro:flow_by_shm_trans dest:one_step_flows)
+
+lemma flow_by_shm_intro3:
+ "\<lbrakk>flow_by_shm s p from; (from, SHM_RDWR) \<in> procs_of_shm s h; (to, flag) \<in> procs_of_shm s h; from \<noteq> to; valid s\<rbrakk>
+ \<Longrightarrow> flow_by_shm s p to"
+apply (rule flow_by_shm_intro2)
+by (auto simp:one_flow_shm_def)
+
+lemma flow_by_shm_intro4:
+ "\<lbrakk>(p, flag) \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> flow_by_shm s p p"
+by (drule procs_of_shm_prop2, simp, simp add:flow_by_shm_intro1)
+
+lemma flow_by_shm_intro5:
+ "\<lbrakk>(from, SHM_RDWR) \<in> procs_of_shm s h; (to,flag) \<in> procs_of_shm s h; valid s; from \<noteq> to\<rbrakk>
+ \<Longrightarrow> flow_by_shm s from to"
+apply (rule_tac p' = "from" and h = h in flow_by_shm_intro2)
+apply (rule flow_by_shm_intro1, simp add:procs_of_shm_prop2)
+apply (simp add:one_flow_shm_def, rule_tac x = flag in exI, auto)
+done
+
+lemma flow_by_shm_intro6:
+ "path_by_shm s p path p' \<Longrightarrow> flow_by_shm s p p'"
+by (auto simp:flow_by_shm_def)
+
+(********* simpset for inductive Info_flow_shm **********)
+lemma path_by_shm_detach1_aux:
+ "path_by_shm s' pa path pb \<Longrightarrow> valid (Detach p h # s) \<and> (s' = Detach p h # s)
+ \<longrightarrow> \<not> edge_related path p h \<and> path_by_shm s pa path pb"
+apply (erule path_by_shm.induct, simp)
+apply (rule impI, rule path_by_shm.intros(1), simp+)
+by (auto simp:one_flow_shm_def split:if_splits intro:path_by_shm_intro3)
+
+lemma path_by_shm_detach1:
+ "\<lbrakk>path_by_shm (Detach p h # s) pa path pb; valid (Detach p h # s)\<rbrakk>
+ \<Longrightarrow> \<not> edge_related path p h \<and> path_by_shm s pa path pb"
+by (auto dest:path_by_shm_detach1_aux)
+
+lemma path_by_shm_detach2_aux[rule_format]:
+ "path_by_shm s pa path pb \<Longrightarrow> valid (Detach p h # s) \<and> \<not> edge_related path p h
+ \<longrightarrow> path_by_shm (Detach p h # s) pa path pb"
+apply (induct rule:path_by_shm.induct)
+apply (rule impI, rule path_by_shm.intros(1), simp)
+apply (rule impI, erule conjE, simp split:if_splits)
+apply (rule path_by_shm.intros(2), simp)
+apply (simp add:one_flow_shm_detach)
+apply (rule impI, simp+)
+done
+
+lemma path_by_shm_detach2:
+ "\<lbrakk>valid (Detach p h # s); \<not> edge_related path p h; path_by_shm s pa path pb\<rbrakk>
+ \<Longrightarrow> path_by_shm (Detach p h # s) pa path pb"
+by (auto intro!:path_by_shm_detach2_aux)
+
+lemma path_by_shm_detach:
+ "valid (Detach p h # s) \<Longrightarrow>
+ path_by_shm (Detach p h # s) pa path pb = (\<not> edge_related path p h \<and> path_by_shm s pa path pb)"
+by (auto dest:path_by_shm_detach1 path_by_shm_detach2)
+
+lemma flow_by_shm_detach:
+ "valid (Detach p h # s) \<Longrightarrow>
+ flow_by_shm (Detach p h # s) pa pb = (\<exists> path. \<not> edge_related path p h \<and> path_by_shm s pa path pb)"
+by (auto dest:path_by_shm_detach simp:flow_by_shm_def)
+
+lemma path_by_shm'_attach1_aux:
+ "path_by_shm' s' pa path pb \<Longrightarrow> valid s' \<and> (s' = Attach p h flag # s) \<longrightarrow>
+ (path_by_shm' s pa path pb) \<or>
+ (\<exists> path1 path2 p'. path_by_shm' s pa path1 p' \<and> path_by_shm' s p path2 pb \<and>
+ (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path = path2 @ [(p', h, p)] @ path1 ) \<or>
+ (\<exists> path1 path2 p' flag'. path_by_shm' s pa path1 p \<and> path_by_shm' s p' path2 pb \<and>
+ (p', flag') \<in> procs_of_shm s h \<and> path = path2 @ [(p, h, p')] @ path1 \<and> flag = SHM_RDWR)"
+apply (erule path_by_shm'.induct)
+apply (simp, rule impI, rule pbs1', simp)
+apply (rule impI, erule impE, clarsimp)
+apply (erule disjE)
+apply (clarsimp simp:one_flow_shm_attach split:if_splits)
+apply (erule disjE, clarsimp)
+apply (erule_tac x = path in allE, clarsimp)
+apply (erule impE, rule pbs1', erule procs_of_shm_prop2, erule vd_cons, simp)
+apply (erule disjE, clarsimp)
+apply (rule_tac x = path in exI, rule_tac x = "[]" in exI, rule_tac x = p' in exI, simp)
+apply (rule pbs1', drule vt_grant_os, clarsimp)
+apply (drule_tac p = pa and p' = p' and p'' = p'' in pbs2', simp+)
+apply (drule_tac p = pa and p' = p' and p'' = p'' in pbs2', simp+)
+
+apply (erule disjE)
+apply ((erule exE|erule conjE)+, clarsimp split:if_splits simp:one_flow_shm_attach)
+apply (clarsimp simp:one_flow_shm_attach split:if_splits)
+apply (erule disjE, clarsimp)
+apply (clarsimp)
+
+
+apply (erule conjE)+
+
+
+
+apply (erule conjE, clarsimp simp only:one_flow_shm_attach split:if_splits)
+apply simp
+
+
+
+lemma path_by_shm_attach1_aux:
+ "path_by_shm s' pa path pb \<Longrightarrow> valid s' \<and> (s' = Attach p h flag # s) \<longrightarrow>
+ path_by_shm s pa path pb \<or>
+ (if (pa = p \<and> flag = SHM_RDWR)
+ then \<exists> p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and>
+ path_by_shm s p' path' pb \<and> path = path' @ [(p, h, p')]
+ else if (pb = p)
+ then \<exists> p' path'. path_by_shm s pa path' p' \<and> path = (p', h, p) # path' \<and>
+ (p', SHM_RDWR) \<in> procs_of_shm s h
+ else (\<exists> p' flag' patha pathb. path_by_shm s pa patha p \<and> flag = SHM_RDWR \<and>
+ (p', flag') \<in> procs_of_shm s h \<and> path_by_shm s p' pathb pb \<and>
+ path = pathb @ [(p, h, p')] @ patha) \<or>
+ (\<exists> p' patha pathb. path_by_shm s pa patha p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and>
+ path_by_shm s p pathb pb \<and> path = pathb @ [(p', h, p)] @ patha))"
+proof (induct rule:path_by_shm.induct)
+ case (pbs1 proc \<tau>)
+ show ?case
+ proof (rule impI)
+ assume pre: "valid \<tau> \<and> \<tau> = Attach p h flag # s"
+ from pbs1 pre have "proc \<in> current_procs s" by simp
+ thus "path_by_shm s proc [] proc \<or>
+ (if proc = p \<and> flag = SHM_RDWR
+ then \<exists>p' flagb path'.
+ (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' path' proc \<and> [] = path' @ [(p, h, p')]
+ else if proc = p
+ then \<exists>p' path'.
+ path_by_shm s proc path' p' \<and> [] = (p', h, p) # path' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h
+ else (\<exists>p' flag' patha pathb.
+ path_by_shm s proc patha p \<and>
+ flag = SHM_RDWR \<and>
+ (p', flag') \<in> procs_of_shm s h \<and>
+ path_by_shm s p' pathb proc \<and> [] = pathb @ [(p, h, p')] @ patha) \<or>
+ (\<exists>p' patha pathb.
+ path_by_shm s proc patha p' \<and>
+ (p', SHM_RDWR) \<in> procs_of_shm s h \<and>
+ path_by_shm s p pathb proc \<and> [] = pathb @ [(p', h, p)] @ patha))"
+ by (auto intro:path_by_shm.intros)
+ qed
+next
+ case (pbs2 \<tau> pa path pb h' pc)
+ thus ?case
+ proof (rule_tac impI)
+ assume p1:"path_by_shm \<tau> pa path pb" and p2: "valid \<tau> \<and> \<tau> = Attach p h flag # s \<longrightarrow>
+ path_by_shm s pa path pb \<or>
+ (if pa = p \<and> flag = SHM_RDWR
+ then \<exists>p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' path' pb \<and> path = path' @ [(p, h, p')]
+ else if pb = p
+ then \<exists>p' path'. path_by_shm s pa path' p' \<and> path = (p', h, p) # path' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h
+ else (\<exists>p' flag' pathaa pathb. path_by_shm s pa pathaa p \<and> flag = SHM_RDWR \<and>
+ (p', flag') \<in> procs_of_shm s h \<and> path_by_shm s p' pathb pb \<and>
+ path = pathb @ [(p, h, p')] @ pathaa) \<or>
+ (\<exists>p' pathaa pathb. path_by_shm s pa pathaa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and>
+ path_by_shm s p pathb pb \<and> path = pathb @ [(p', h, p)] @ pathaa))"
+ and p3: "one_flow_shm \<tau> h' pb pc" and p4: "valid \<tau> \<and> \<tau> = Attach p h flag # s"
+
+ from p2 and p4 have p2': "
+ path_by_shm s pa path pb \<or>
+ (if pa = p \<and> flag = SHM_RDWR
+ then \<exists>p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' path' pb \<and> path = path' @ [(p, h, p')]
+ else if pb = p
+ then \<exists>p' path'. path_by_shm s pa path' p' \<and> path = (p', h, p) # path' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h
+ else (\<exists>p' flag' pathaa pathb. path_by_shm s pa pathaa p \<and> flag = SHM_RDWR \<and>
+ (p', flag') \<in> procs_of_shm s h \<and> path_by_shm s p' pathb pb \<and>
+ path = pathb @ [(p, h, p')] @ pathaa) \<or>
+ (\<exists>p' pathaa pathb. path_by_shm s pa pathaa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and>
+ path_by_shm s p pathb pb \<and> path = pathb @ [(p', h, p)] @ pathaa))"
+ by (erule_tac impE, simp)
+ from p4 have p5: "valid s" and p6: "os_grant s (Attach p h flag)" by (auto intro:vd_cons dest:vt_grant_os)
+ from p6 have "p \<in> current_procs s" by simp hence p7:"path_by_shm s p [] p" by (erule_tac path_by_shm.intros)
+ from p3 p4 have p8: "if (h' = h)
+ then (pb = p \<and> pc \<noteq> p \<and> flag = SHM_RDWR \<and> (\<exists> flagb. (pc, flagb) \<in> procs_of_shm s h)) \<or>
+ (pc = p \<and> pb \<noteq> p \<and> (pb, SHM_RDWR) \<in> procs_of_shm s h) \<or>
+ (one_flow_shm s h pb pc)
+ else one_flow_shm s h' pb pc" by (auto simp add:one_flow_shm_attach)
+
+
+(*
+ have "\<And> flagb. (pc, flagb) \<in> procs_of_shm s h
+ \<Longrightarrow> \<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' [] pc"
+ apply (rule_tac x= pc in exI, rule_tac x = flagb in exI, frule procs_of_shm_prop2)
+ by (simp add:p5, simp add:path_by_shm.intros(1))
+ hence p10: "\<not> path_by_shm s p path pc \<Longrightarrow> (\<exists>p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' pc) \<or>
+ path_by_shm s pa pc"
+ using p2' p7 p8 p5
+ by (auto split:if_splits dest:path_by_shm.intros(2))
+ (* apply (rule_tac x = pb in exI, simp add:one_flow_flows, rule_tac x = flagb in exI, simp)+ *) *)
+
+ from p1 have a0: "(path = []) = (pa = pb)" using pbs_prop2 by simp
+ have a1:"\<lbrakk>pa = p; flag = SHM_RDWR; \<not> path_by_shm s pa path pb\<rbrakk> \<Longrightarrow>
+ \<exists>p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' path' pb \<and> path = path' @ [(p, h, p')]"
+ using p2' by auto
+ have b1: "\<lbrakk>pa = p; flag = SHM_RDWR; \<not> path_by_shm s pa path pc\<rbrakk> \<Longrightarrow>
+ \<exists>p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' path' pc \<and>
+ (pb, h', pc) # path = path' @ [(p, h, p')]"
+
+
+ using p8 a1 p7 p5 a0
+ apply (auto split:if_splits elim:path_by_shm_intro4)
+ apply (rule_tac x = pb in exI, rule conjI, rule_tac x = SHM_RDWR in exI, simp)
+ apply (rule_tac x = pc in exI, rule conjI, rule_tac x = flagb in exI, simp)
+ apply (rule_tac x = "[]" in exI, rule conjI)
+apply (erule path_by_shm_intro4, simp)
+
+ apply (case_tac "path_by_shm s pa path pb", simp) defer
+ apply (drule a1, simp+, clarsimp)
+ apply (rule conjI, rule_tac x = flagb in exI, simp)
+ apply (rule path_by_shm_
+ using p2' p8 p5
+ apply (auto split:if_splits dest!:pbs_prop2' simp:path_by_shm_intro4)
+ apply (drule pbs_prop2', simp)
+ apply (erule_tac x = pc in allE, simp add:path_by_shm_intro4)
+
+ apply (drule_tac x = "pc" in allE)
+
+ apply simp
+
+ sorry
+ moreover have "pc = p \<Longrightarrow> (\<exists>p' path'. path_by_shm s pa path' p' \<and>
+ (pb, h', pc) # path = path' @ [(p', h, p)] \<and> (p', SHM_RDWR) \<in> procs_of_shm s h) \<or>
+ (path_by_shm s pa path pc \<and> \<not> edge_related path p h)"
+ using p2' p7 p8 p5
+ sorry (*
+ apply (auto split:if_splits intro:path_by_shm_intro3 simp:one_flow_shm_def) *)
+ moreover have "\<lbrakk>pc \<noteq> p; pa \<noteq> p \<or> flag \<noteq> SHM_RDWR\<rbrakk> \<Longrightarrow>
+ (\<exists>p' flag' pathaa pathb. path_by_shm s pa pathaa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and>
+ path_by_shm s p' pathb pc \<and> (pb, h', pc) # path = pathaa @ [(p, h, p')] @ pathb) \<or>
+ (\<exists>p' pathaa pathb. path_by_shm s pa pathaa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and>
+ path_by_shm s p pathb pc \<and> (pb, h', pc) # path = pathaa @ [(p', h, p)] @ pathb) \<or>
+ (path_by_shm s pa path pc \<and> \<not> edge_related path p h)"
+ using p2' p7 p8 p5 (*
+ apply (auto split:if_splits intro:path_by_shm_intro3 simp:one_flow_shm_def)
+ apply (rule_tac x = pc in exI, simp add:path_by_shm_intro4)
+ apply (rule_tac x = flagb in exI, simp)
+ done *)
+ sorry
+ ultimately
+ show "if (pb, h', pc) # path = [] then pa = pc \<and> pa \<in> current_procs s
+ else path_by_shm s pa ((pb, h', pc) # path) pc \<and> \<not> edge_related ((pb, h', pc) # path) p h \<or>
+ (if pa = p \<and> flag = SHM_RDWR
+ then \<exists>p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and>
+ path_by_shm s p' path' pc \<and> (pb, h', pc) # path = path' @ [(p, h, p')]
+ else if pc = p
+ then \<exists>p' path'. path_by_shm s pa path' p' \<and>
+ (pb, h', pc) # path = (p', h, p) # path' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h
+ else (\<exists>p' flag' pathaa pathb. path_by_shm s pa pathaa p \<and> flag = SHM_RDWR \<and>
+ (p', flag') \<in> procs_of_shm s h \<and>
+ path_by_shm s p' pathb pc \<and> (pb, h', pc) # path = pathb @ [(p, h, p')] @ pathaa) \<or>
+ (\<exists>p' pathaa pathb. path_by_shm s pa pathaa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and>
+ path_by_shm s p pathb pc \<and> (pb, h', pc) # path = pathb @ [(p', h, p)] @ pathaa))"
+ apply (auto split:if_splits)
+ using p7 by auto
+ qed
+qed
+
+lemma path_by_shm_attach1:
+ "\<lbrakk>valid (Attach p h flag # s); path_by_shm (Attach p h flag # s) pa pb\<rbrakk>
+ \<Longrightarrow> (if path_by_shm s pa pb then True else
+ (if (pa = p \<and> flag = SHM_RDWR)
+ then (\<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' pb)
+ else if (pb = p)
+ then (\<exists> p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s pa p')
+ else (\<exists> p' flag'. path_by_shm s pa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and>
+ path_by_shm s p' pb) \<or>
+ (\<exists> p'. path_by_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s p pb)
+ ) )"
+apply (drule_tac p = p and h = h and flag = flag in path_by_shm_attach1_aux)
+by auto
+
+lemma path_by_shm_attach_aux[rule_format]:
+ "path_by_shm s pa pb \<Longrightarrow> valid (Attach p h flag # s) \<longrightarrow> path_by_shm (Attach p h flag # s) pa pb"
+apply (erule path_by_shm.induct)
+apply (rule impI, rule path_by_shm.intros(1), simp)
+apply (rule impI, simp, rule_tac h = ha in path_by_shm.intros(2), simp)
+apply (auto simp add:one_flow_shm_simps)
+done
+
+lemma path_by_shm_attach2:
+ "\<lbrakk>valid (Attach p h flag # s); if path_by_shm s pa pb then True else
+ (if (pa = p \<and> flag = SHM_RDWR)
+ then (\<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' pb)
+ else if (pb = p)
+ then (\<exists> p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s pa p')
+ else (\<exists> p' flag'. path_by_shm s pa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and>
+ path_by_shm s p' pb) \<or>
+ (\<exists> p'. path_by_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s p pb))\<rbrakk>
+ \<Longrightarrow> path_by_shm (Attach p h flag # s) pa pb"
+apply (frule vt_grant_os, frule vd_cons)
+apply (auto split:if_splits intro:path_by_shm_intro3 simp:one_flow_shm_def intro:path_by_shm_attach_aux)
+apply (rule_tac p' = p' in Info_flow_trans)
+apply (rule_tac p' = p and h = h in path_by_shm.intros(2))
+apply (rule path_by_shm.intros(1), simp)
+apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
+apply (rule conjI, rule notI, simp, rule_tac x = flagb in exI, simp)
+apply (simp add:path_by_shm_attach_aux)
+
+apply (rule_tac p' = p' in Info_flow_trans)
+apply (rule_tac p' = p in Info_flow_trans)
+apply (simp add:path_by_shm_attach_aux)
+apply (rule_tac p' = p and h = h in path_by_shm.intros(2))
+apply (rule path_by_shm.intros(1), simp)
+apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
+apply (rule conjI, rule notI, simp, rule_tac x = flag' in exI, simp)
+apply (simp add:path_by_shm_attach_aux)
+
+apply (rule_tac p' = p in Info_flow_trans)
+apply (rule_tac p' = p' in Info_flow_trans)
+apply (simp add:path_by_shm_attach_aux)
+apply (rule_tac p' = p' and h = h in path_by_shm.intros(2))
+apply (rule path_by_shm.intros(1), simp add:procs_of_shm_prop2)
+apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
+apply (rule notI, simp)
+apply (simp add:path_by_shm_attach_aux)
+
+apply (rule_tac p' = p in Info_flow_trans)
+apply (rule_tac p' = p' in Info_flow_trans)
+apply (simp add:path_by_shm_attach_aux)
+apply (rule_tac p' = p' and h = h in path_by_shm.intros(2))
+apply (rule path_by_shm.intros(1), simp add:procs_of_shm_prop2)
+apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
+apply (rule notI, simp)
+apply (simp add:path_by_shm_attach_aux)
+done
+
+lemma path_by_shm_attach:
+ "valid (Attach p h flag # s) \<Longrightarrow> path_by_shm (Attach p h flag # s) = (\<lambda> pa pb.
+ path_by_shm s pa pb \<or>
+ (if (pa = p \<and> flag = SHM_RDWR)
+ then (\<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' pb)
+ else if (pb = p)
+ then (\<exists> p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s pa p')
+ else (\<exists> p' flag'. path_by_shm s pa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and>
+ path_by_shm s p' pb) \<or>
+ (\<exists> p'. path_by_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s p pb)
+ ) )"
+apply (rule ext, rule ext, rule iffI)
+apply (drule_tac pa = pa and pb = pb in path_by_shm_attach1, simp)
+apply (auto split:if_splits)[1]
+apply (drule_tac pa = pa and pb = pb in path_by_shm_attach2)
+apply (auto split:if_splits)
+done
+
+
+
+lemma info_flow_shm_detach:
+ "valid (Detach p h # s) \<Longrightarrow> info_flow_shm (Detach p h # s) = (\<lambda> pa pb.
+ self_shm s pa pb \<or> ((p = pa \<or> p = pb) \<and> (\<exists> h'. h' \<noteq> h \<and> one_flow_shm s h' pa pb)) \<or>
+ (pa \<noteq> p \<and> pb \<noteq> p \<and> info_flow_shm s pa pb) )"
+apply (rule ext, rule ext, frule vt_grant_os)
+by (auto simp:info_flow_shm_def one_flow_shm_def)
+
+lemma info_flow_shm_deleteshm:
+ "valid (DeleteShM p h # s) \<Longrightarrow> info_flow_shm (DeleteShM p h # s) = (\<lambda> pa pb.
+ self_shm s pa pb \<or> (\<exists> h'. h' \<noteq> h \<and> one_flow_shm s h' pa pb) )"
+apply (rule ext, rule ext, frule vt_grant_os)
+by (auto simp:info_flow_shm_def one_flow_shm_def)
+
+lemma info_flow_shm_clone:
+ "valid (Clone p p' fds shms # s) \<Longrightarrow> info_flow_shm (Clone p p' fds shms # s) = (\<lambda> pa pb.
+ (pa = p' \<and> pb = p') \<or> (pa = p' \<and> pb \<noteq> p' \<and> (\<exists> h \<in> shms. one_flow_shm s h p pb)) \<or>
+ (pb = p' \<and> pa \<noteq> p' \<and> (\<exists> h \<in> shms. one_flow_shm s h pa p)) \<or>
+ (pa \<noteq> p' \<and> pb \<noteq> p' \<and> info_flow_shm s pa pb))"
+apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons, clarsimp)
+apply (frule_tac p = p' in procs_of_shm_prop2', simp)
+sorry (*
+apply (auto simp:info_flow_shm_def one_flow_shm_def)
+done *)
+
+lemma info_flow_shm_execve:
+ "valid (Execve p f fds # s) \<Longrightarrow> info_flow_shm (Execve p f fds # s) = (\<lambda> pa pb.
+ (pa = p \<and> pb = p) \<or> (pa \<noteq> p \<and> pb \<noteq> p \<and> info_flow_shm s pa pb) )"
+apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons)
+by (auto simp:info_flow_shm_def one_flow_shm_def)
+
+lemma info_flow_shm_kill:
+ "valid (Kill p p' # s) \<Longrightarrow> info_flow_shm (Kill p p' # s) = (\<lambda> pa pb.
+ pa \<noteq> p' \<and> pb \<noteq> p' \<and> info_flow_shm s pa pb )"
+apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons)
+by (auto simp:info_flow_shm_def one_flow_shm_def)
+
+lemma info_flow_shm_exit:
+ "valid (Exit p # s) \<Longrightarrow> info_flow_shm (Exit p # s) = (\<lambda> pa pb.
+ pa \<noteq> p \<and> pb \<noteq> p \<and> info_flow_shm s pa pb )"
+apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons)
+by (auto simp:info_flow_shm_def one_flow_shm_def)
+
+lemma info_flow_shm_other:
+ "\<lbrakk>valid (e # s);
+ \<forall> p h flag. e \<noteq> Attach p h flag;
+ \<forall> p h. e \<noteq> Detach p h;
+ \<forall> p h. e \<noteq> DeleteShM p h;
+ \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms;
+ \<forall> p f fds. e \<noteq> Execve p f fds;
+ \<forall> p p'. e \<noteq> Kill p p';
+ \<forall> p. e \<noteq> Exit p
+ \<rbrakk> \<Longrightarrow> info_flow_shm (e # s) = info_flow_shm s"
+apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons)
+apply (case_tac e, auto simp:info_flow_shm_def one_flow_shm_def dest:procs_of_shm_prop2)
+apply (erule_tac x = h in allE, simp)
+apply (drule procs_of_shm_prop1, auto)
+done
+
+
+(*
+lemma info_flow_shm_prop1:
+ "\<lbrakk>info_flow_shm s p p'; p \<noteq> p'; valid s\<rbrakk>
+ \<Longrightarrow> \<exists> h h' flag. (p, SHM_RDWR) \<in> procs_of_shm s h \<and> (p', flag) \<in> procs_of_shm s h'"
+by (induct rule: info_flow_shm.induct, auto)
+
+lemma info_flow_shm_cases:
+ "\<lbrakk>info_flow_shm \<tau> pa pb; \<And>p s. \<lbrakk>s = \<tau> ; pa = p; pb = p; p \<in> current_procs s\<rbrakk> \<Longrightarrow> P;
+ \<And>s p p' h p'' flag. \<lbrakk>s = \<tau>; pa = p; pb = p''; info_flow_shm s p p'; (p', SHM_RDWR) \<in> procs_of_shm s h;
+ (p'', flag) \<in> procs_of_shm s h\<rbrakk>\<Longrightarrow> P\<rbrakk>
+ \<Longrightarrow> P"
+by (erule info_flow_shm.cases, auto)
+
+definition one_flow_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
+where
+ "one_flow_shm s p p' \<equiv> p \<noteq> p' \<and> (\<exists> h flag. (p, SHM_RDWR) \<in> procs_of_shm s h \<and> (p', flag) \<in> procs_of_shm s h)"
+
+inductive flows_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
+where
+ "p \<in> current_procs s \<Longrightarrow> flows_shm s p p"
+| "\<lbrakk>flows_shm s p p'; one_flow_shm s p' p''\<rbrakk> \<Longrightarrow> flows_shm s p p''"
+
+definition attached_procs :: "t_state \<Rightarrow> t_shm \<Rightarrow> t_process set"
+where
+ "attached_procs s h \<equiv> {p. \<exists> flag. (p, flag) \<in> procs_of_shm s h}"
+
+definition flowed_procs:: "t_state \<Rightarrow> t_shm \<Rightarrow> t_process set"
+where
+ "flowed_procs s h \<equiv> {p'. \<exists> p \<in> attached_procs s h. flows_shm s p p'}"
+
+inductive flowed_shm:: "t_state \<Rightarrow> t_process \<Rightarrow> t_shm set"
+
+fun Info_flow_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process set"
+where
+ "Info_flow_shm [] = (\<lambda> p. {p'. flows_shm [] p p'})"
+| "Info_flow_shm (Attach p h flag # s) = (\<lambda> p'.
+ if (p' = p) then flowed_procs s h
+ else if ()
+ "
+
+
+lemma info_flow_shm_attach:
+ "valid (Attach p h flag # s) \<Longrightarrow> info_flow_shm (Attach p h flag # s) = (\<lambda> pa pb. (info_flow_shm s pa pb) \<or>
+ (if (pa = p)
+ then (if (flag = SHM_RDWR)
+ then (\<exists> flag. (pb, flag) \<in> procs_of_shm s h)
+ else (pb = p))
+ else (if (pb = p)
+ then (pa, SHM_RDWR) \<in> procs_of_shm s h
+ else info_flow_shm s pa pb)) )"
+apply (frule vd_cons, frule vt_grant_os, rule ext, rule ext)
+apply (case_tac "info_flow_shm s pa pb", simp)
+
+thm info_flow_shm.cases
+apply (auto split:if_splits intro:info_flow_shm.intros elim:info_flow_shm_cases)
+apply (erule info_flow_shm_cases, simp, simp split:if_splits)
+apply (rule_tac p = pa and p' = p' in info_flow_shm.intros(2), simp+)
+apply (rule notI, erule info_flow_shm.cases, simp+)
+pr 5
+*)
+lemmas info_flow_shm_simps = info_flow_shm_other (* info_flow_shm_attach *) info_flow_shm_detach info_flow_shm_deleteshm
+ info_flow_shm_clone info_flow_shm_execve info_flow_shm_kill info_flow_shm_exit
+
+
+
+
+
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Init_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,758 @@
+(*<*)
+theory Init_prop
+imports Main OS_type_def Flask Flask_type Static_type Static
+begin
+(*>*)
+
+context init begin
+
+lemma init_files_prop1: "init_inum_of_file f = Some im \<Longrightarrow> f \<in> init_files"
+by (simp add:inof_has_file_tag)
+
+lemma init_files_prop2: "finite init_files"
+by (simp add:init_finite_sets)
+
+lemma init_files_prop3: "f \<in> init_files \<Longrightarrow> init_inum_of_file f \<noteq> None"
+by (auto dest:init_file_has_inum)
+
+lemma init_files_prop4: "(f \<in> init_files) = (f \<in> current_files [])"
+apply (simp add:current_files_def, rule iffI)
+using init_files_prop1 init_files_prop3 by auto
+
+lemmas init_files_props = init_file_has_inum init_files_prop1 init_files_prop2 init_files_prop3 init_files_prop4
+
+lemma init_inumof_prop1: "init_inum_of_file f = Some im \<Longrightarrow> \<exists> tag. init_itag_of_inum im = Some tag"
+by (auto dest:inof_has_file_tag)
+
+lemma init_inumof_prop2: "init_inum_of_file f = Some im \<Longrightarrow> init_itag_of_inum im \<noteq> None"
+by (auto dest:inof_has_file_tag)
+
+lemma init_inumof_prop3: "\<lbrakk>init_inum_of_file f = Some im; init_itag_of_inum im = Some tag\<rbrakk> \<Longrightarrow> is_file_dir_itag tag"
+by (auto dest:inof_has_file_tag)
+
+lemmas init_inum_of_file_props = init_files_prop1 init_inumof_prop1 init_inumof_prop2 init_inumof_prop3
+
+lemma init_inumos_prop1: "init_inum_of_socket s = Some im \<Longrightarrow> s \<in> init_sockets"
+by (auto dest:inos_has_sock_tag)
+
+lemma init_inumos_prop2: "init_inum_of_socket s = Some im \<Longrightarrow> init_itag_of_inum im = Some Tag_TCP_SOCK \<or> init_itag_of_inum im = Some Tag_UDP_SOCK"
+apply (auto dest!:inos_has_sock_tag)
+apply (case_tac tag, simp+)
+done
+
+lemma init_inumos_prop3: "init_inum_of_socket s = Some im \<Longrightarrow> init_itag_of_inum im \<noteq> None"
+by (auto dest:inos_has_sock_tag)
+
+lemma init_inumos_prop4: "init_inum_of_socket s = Some im \<Longrightarrow> \<exists> tag. init_itag_of_inum im = Some tag \<and> is_sock_itag tag"
+by (auto dest!:inos_has_sock_tag)
+
+lemmas init_inum_of_socket_props = init_inumos_prop1 init_inumos_prop2 init_inumos_prop3 init_inumos_prop4
+
+lemma init_sockets_prop1: "(p, fd) \<in> init_sockets \<Longrightarrow> p \<in> init_procs"
+by (auto dest: init_socket_has_inode)
+
+lemma init_sockets_prop2: "(p, fd) \<in> init_sockets \<Longrightarrow> fd \<in> init_fds_of_proc p"
+by (auto dest:init_socket_has_inode)
+
+lemma init_sockets_prop3: "s \<in> init_sockets \<Longrightarrow> \<exists> im. init_inum_of_socket s = Some im"
+by (case_tac s, auto dest:init_socket_has_inode)
+
+lemma init_sockets_prop4: "s \<in> init_sockets \<Longrightarrow> init_inum_of_socket s \<noteq> None"
+by (simp add:init_sockets_prop3)
+
+lemma init_sockets_prop5: "s \<in> init_sockets = (s \<in> current_sockets [])"
+apply (simp add:current_sockets_def, rule iffI)
+using init_sockets_prop4 inos_has_sock_tag apply auto
+apply (case_tac s, auto)
+done
+
+lemma init_socket_prop6: "(p, fd) \<in> init_sockets \<Longrightarrow> init_file_of_proc_fd p fd = None"
+by (auto dest: init_socket_has_inode)
+
+lemmas init_sockets_props = init_sockets_prop1 init_sockets_prop2 init_sockets_prop3 init_sockets_prop4 init_sockets_prop5
+
+lemma is_init_file_prop1: "is_init_file f \<Longrightarrow> f \<in> init_files"
+by (auto simp add:is_init_file_def init_inum_of_file_props split:option.splits)
+
+lemma is_init_file_prop2: "is_init_file f \<Longrightarrow> \<not> is_init_dir f"
+by (auto simp add:is_init_file_def is_init_dir_def split:option.splits t_inode_tag.splits)
+
+lemmas is_init_file_props = is_init_file_prop1 is_init_file_prop2
+
+lemma is_init_dir_prop1: "is_init_dir f \<Longrightarrow> f \<in> init_files"
+by (auto simp add:is_init_dir_def is_dir_def init_inum_of_file_props split:option.splits)
+
+lemma is_init_dir_prop2: "is_init_dir f \<Longrightarrow> \<not> is_init_file f"
+by (auto simp add:is_init_dir_def is_init_file_def split:option.splits t_inode_tag.splits)
+
+lemmas is_init_dir_props = is_init_dir_prop1 is_init_dir_prop2
+
+lemma is_file_nil: "is_file [] = is_init_file"
+by (auto simp:is_init_file_def is_file_def init_inum_of_file_props intro!:ext split:option.splits)
+
+lemma is_dir_nil: "is_dir [] = is_init_dir"
+by (auto simp:is_init_dir_def is_dir_def init_inum_of_file_props intro!:ext split:option.splits)
+
+lemma is_udp_sock_nil:
+ "is_udp_sock [] k = is_init_udp_sock k"
+by (auto simp:is_udp_sock_def is_init_udp_sock_def split:option.splits)
+
+lemma is_init_udp_sock_prop1: "is_init_udp_sock s \<Longrightarrow> s \<in> init_sockets"
+apply (auto simp add:is_init_udp_sock_def is_udp_sock_def init_inum_of_socket_props
+ dest:init_socket_has_inode split:option.splits)
+done
+
+lemma is_init_udp_sock_prop2: "is_init_udp_sock s \<Longrightarrow> \<not> is_init_tcp_sock s"
+apply (auto simp add:is_init_udp_sock_def is_init_tcp_sock_def
+ dest:init_socket_has_inode split:option.splits t_inode_tag.splits)
+done
+
+lemma is_init_udp_sock_prop3:
+ "is_init_udp_sock (p, fd) \<Longrightarrow> p \<in> init_procs"
+by (auto simp:is_init_udp_sock_def split:option.splits t_inode_tag.splits
+ dest:init_socket_has_inode inos_has_sock_tag)
+
+lemma is_init_udp_sock_prop4:
+ "is_init_udp_sock (p, fd) \<Longrightarrow> fd \<in> init_fds_of_proc p"
+by (auto simp:is_init_udp_sock_def split:option.splits t_inode_tag.splits
+ dest:init_socket_has_inode inos_has_sock_tag)
+
+lemma is_init_udp_sock_prop5:
+ "is_init_udp_sock (p, fd) \<Longrightarrow> init_file_of_proc_fd p fd = None"
+by (auto dest:is_init_udp_sock_prop1 intro:init_socket_prop6)
+
+lemmas is_init_udp_sock_props = is_init_udp_sock_prop1 is_init_udp_sock_prop2 is_init_udp_sock_prop3
+ is_init_udp_sock_prop4 is_init_udp_sock_prop5
+
+lemma is_tcp_sock_nil:
+ "is_tcp_sock [] k = is_init_tcp_sock k"
+by (auto simp:is_tcp_sock_def is_init_tcp_sock_def split:option.splits)
+
+lemma is_init_tcp_sock_prop1: "is_init_tcp_sock s \<Longrightarrow> s \<in> init_sockets"
+apply (auto simp add:is_init_tcp_sock_def is_tcp_sock_def init_inum_of_socket_props
+ dest:init_socket_has_inode split:option.splits)
+done
+
+lemma is_init_tcp_sock_prop2: "is_init_tcp_sock s \<Longrightarrow> \<not> is_init_udp_sock s"
+apply (auto simp add:is_init_tcp_sock_def is_init_udp_sock_def
+ dest:init_socket_has_inode split:option.splits t_inode_tag.splits)
+done
+
+lemma is_init_tcp_sock_prop3:
+ "is_init_tcp_sock (p, fd) \<Longrightarrow> p \<in> init_procs"
+by (auto simp:is_init_tcp_sock_def split:option.splits t_inode_tag.splits
+ dest:init_socket_has_inode inos_has_sock_tag)
+
+lemma is_init_tcp_sock_prop4:
+ "is_init_tcp_sock (p, fd) \<Longrightarrow> fd \<in> init_fds_of_proc p"
+by (auto simp:is_init_tcp_sock_def split:option.splits t_inode_tag.splits
+ dest:init_socket_has_inode inos_has_sock_tag)
+
+lemma is_init_tcp_sock_prop5:
+ "is_init_tcp_sock (p, fd) \<Longrightarrow> init_file_of_proc_fd p fd = None"
+by (auto dest:is_init_tcp_sock_prop1 intro:init_socket_prop6)
+
+lemmas is_init_tcp_sock_props = is_init_tcp_sock_prop1 is_init_tcp_sock_prop2 is_init_tcp_sock_prop3
+ is_init_tcp_sock_prop4 is_init_tcp_sock_prop5
+
+lemma init_parent_file_prop1:
+ "\<lbrakk>parent f = Some pf; f \<in> init_files\<rbrakk> \<Longrightarrow> is_init_dir pf"
+apply (frule parent_file_in_init, simp, frule_tac f = pf in init_files_prop3)
+apply (clarsimp, drule_tac im = y in init_parentf_is_dir, simp+)
+by (simp add:is_init_dir_def)
+
+lemma init_parent_file_prop1':
+ "a # f \<in> init_files \<Longrightarrow> is_init_dir f"
+by (rule_tac pf = f in init_parent_file_prop1, auto)
+
+lemma init_parent_file_prop2:
+ "\<lbrakk>parent f = Some pf; is_init_file f\<rbrakk> \<Longrightarrow> is_init_dir pf"
+by (rule init_parent_file_prop1, simp, simp add: is_init_file_props)
+
+lemma init_parent_file_prop2':
+ "is_init_file (f#pf) \<Longrightarrow> is_init_dir pf"
+apply (rule init_parent_file_prop2)
+by auto
+
+lemma init_parent_file_prop3:
+ "\<lbrakk>parent f = Some pf; is_init_dir f\<rbrakk> \<Longrightarrow> is_init_dir pf"
+by (rule init_parent_file_prop1, simp, simp add: is_init_dir_props)
+
+lemma init_parent_file_prop3':
+ "is_init_dir (f#pf) \<Longrightarrow> is_init_dir pf"
+apply (rule init_parent_file_prop3)
+by auto
+
+lemma parent_file_in_init': "a # f \<in> init_files \<Longrightarrow> f \<in> init_files"
+by (subgoal_tac "parent (a # f) = Some f", drule parent_file_in_init, auto)
+
+lemmas init_parent_file_props = parent_file_in_init init_parent_file_prop1 parent_file_in_init' init_parent_file_prop1' init_parent_file_prop2 init_parent_file_prop2' init_parent_file_prop3 init_parent_file_prop3'
+
+lemma root_in_filesystem: "[] \<in> init_files"
+using init_files_prop1 root_is_dir by auto
+
+lemma root_is_init_dir: "is_init_dir []"
+using root_is_dir
+by (auto simp add:is_init_dir_def split:option.splits)
+
+lemma root_is_init_dir': "is_init_file [] \<Longrightarrow> False"
+using root_is_dir
+by (auto simp:is_init_file_def split:option.splits)
+
+
+lemma init_files_hung_prop1: "f \<in> init_files_hung_by_del \<Longrightarrow> f \<in> init_files"
+by (auto dest:init_files_hung_valid)
+
+lemma init_files_hung_prop2: "f \<in> init_files_hung_by_del \<Longrightarrow> \<exists> p fd. init_file_of_proc_fd p fd = Some f"
+by (auto dest:init_files_hung_valid)
+
+lemmas init_files_hung_by_del_props = init_files_hung_prop1 init_files_hung_prop2 init_files_hung_valid'
+
+
+lemma init_fds_of_proc_prop1: "fd \<in> init_fds_of_proc p \<Longrightarrow> p \<in> init_procs"
+by (auto dest!:init_procfds_valid)
+
+lemma init_fds_of_proc_prop2: "fd \<in> init_fds_of_proc p \<Longrightarrow> (\<exists> f \<in> init_files. init_file_of_proc_fd p fd = Some f) \<or> (p, fd) \<in> init_sockets"
+by (auto dest:init_procfds_valid)
+
+lemmas init_fds_of_proc_props = init_fds_of_proc_prop1 init_fds_of_proc_prop2
+
+lemma init_filefd_prop1: "init_file_of_proc_fd p fd = Some f \<Longrightarrow> f \<in> init_files"
+by (auto dest!:init_filefd_valid intro:init_files_prop1)
+
+lemma init_filefd_prop2: "init_file_of_proc_fd p fd = Some f \<Longrightarrow> p \<in> init_procs"
+by (auto dest:init_filefd_valid)
+
+lemma init_filefd_prop3: "init_file_of_proc_fd p fd = Some f \<Longrightarrow> fd \<in> init_fds_of_proc p"
+by (auto dest:init_filefd_valid)
+
+lemma init_filefd_prop4: "init_file_of_proc_fd p fd = Some f \<Longrightarrow> \<exists> flags. init_oflags_of_proc_fd p fd = Some flags"
+by (auto dest:init_filefd_valid)
+
+lemma init_filefd_prop5: "init_file_of_proc_fd p fd = Some f \<Longrightarrow> is_init_file f"
+by (auto dest:init_filefd_valid simp:is_init_file_def)
+
+lemma init_filefd_prop6: "init_file_of_proc_fd p fd = Some f \<Longrightarrow> \<not> is_init_tcp_sock (p, fd)"
+by (auto dest!:init_filefd_valid is_init_tcp_sock_prop1)
+
+lemma init_filefd_prop7: "init_file_of_proc_fd p fd = Some f \<Longrightarrow> \<not> is_init_udp_sock (p, fd)"
+by (auto dest!:init_filefd_valid is_init_udp_sock_prop1)
+
+lemma init_filefd_prop8: "init_file_of_proc_fd p fd = Some f \<Longrightarrow> (p, fd) \<notin> init_sockets"
+by (auto dest!:init_filefd_valid)
+
+lemmas init_file_of_proc_fd_props = init_filefd_prop1 init_filefd_prop2 init_filefd_prop3 init_filefd_prop4 init_filefd_prop5 init_filefd_prop6 init_filefd_prop7 init_filefd_prop8
+
+lemma init_oflags_prop1: "init_oflags_of_proc_fd p fd = Some flags \<Longrightarrow> p \<in> init_procs"
+by (auto dest:init_fileflag_valid init_file_of_proc_fd_props)
+
+lemma init_oflags_prop2: "init_oflags_of_proc_fd p fd = Some flags \<Longrightarrow> fd \<in> init_fds_of_proc p"
+by (auto dest:init_fileflag_valid init_file_of_proc_fd_props)
+
+lemmas init_oflags_of_proc_fd_props = init_oflags_prop1 init_oflags_prop2 init_fileflag_valid
+
+(*
+lemma init_socketstate_prop1: "s \<in> init_sockets \<Longrightarrow> init_socket_state s \<noteq> None"
+using init_socket_has_state
+by (case_tac s, simp add:bidirect_in_init_def)
+
+lemma init_socketstate_prop2: "s \<in> init_sockets \<Longrightarrow> \<exists> t. init_socket_state s = Some t"
+using init_socket_has_state
+by (case_tac s, simp add:bidirect_in_init_def)
+
+lemma init_socketstate_prop3: "init_socket_state s = Some t \<Longrightarrow> s \<in> init_sockets"
+using init_socket_has_state
+by (case_tac s, simp add:bidirect_in_init_def)
+
+lemmas init_socket_state_props = init_socketstate_prop1 init_socketstate_prop2 init_socketstate_prop3
+*)
+
+lemma init_inum_sock_file_noninter: "\<lbrakk>init_inum_of_socket s = Some im; init_inum_of_file f = Some im\<rbrakk> \<Longrightarrow> False"
+apply (frule init_inumof_prop1, erule exE, drule init_inumof_prop3, simp)
+apply (frule init_inumos_prop2)
+apply (case_tac tag, simp+)
+done
+
+lemma init_parent_file_has_inum: "\<lbrakk>parent f = Some pf; init_inum_of_file f = Some im\<rbrakk> \<Longrightarrow> \<exists> im. init_inum_of_file pf = Some im"
+by (drule init_files_prop1, drule parent_file_in_init, simp, simp add:init_files_props)
+
+lemma init_file_has_no_son': "\<lbrakk>init_itag_of_inum im = Some Tag_FILE; init_inum_of_file f = Some im; parent f' = Some f\<rbrakk> \<Longrightarrow> init_inum_of_file f' = None"
+apply (drule init_file_no_son, simp)
+by (case_tac "init_inum_of_file f'", auto dest:init_files_prop1)
+
+lemma init_parent_file_is_dir': "\<lbrakk>parent f = Some pf; init_inum_of_file f = Some im; init_inum_of_file pf = Some ipm\<rbrakk> \<Longrightarrow> init_itag_of_inum ipm = Some Tag_DIR"
+by (drule init_parentf_is_dir, auto dest:init_files_prop1)
+
+lemma init_file_hung_has_no_son: "\<lbrakk>f \<in> init_files_hung_by_del; parent f' = Some f; init_inum_of_file f' = Some im\<rbrakk> \<Longrightarrow> False"
+apply (frule init_files_hung_prop1, drule init_file_has_inum, erule exE)
+apply (drule init_files_hung_valid', simp)
+apply (frule init_parent_file_is_dir', simp+)
+apply (drule init_files_prop1)
+apply (erule_tac x = f' in allE, simp)
+by (case_tac f', simp_all add:no_junior_def)
+
+(*
+lemma same_inode_nil_prop:
+ "same_inode_files [] f = init_same_inode_files f"
+by (simp add:same_inode_files_def init_same_inode_files_def is_file_nil)
+
+lemma init_same_inode_prop1:
+ "f \<in> init_files \<Longrightarrow> \<forall> f' \<in> init_same_inode_files f. f' \<in> init_files"
+apply (simp add:init_same_inode_files_def)
+apply (drule init_files_prop3)
+apply (auto simp:init_files_prop1)
+done
+
+lemma init_same_inode_prop2:
+ "\<lbrakk>f' \<in> init_same_inode_files f; f \<in> init_files\<rbrakk> \<Longrightarrow> f' \<in> init_files"
+by (drule init_same_inode_prop1, simp)
+
+lemma init_same_inode_prop3:
+ "f' \<in> init_same_inode_files f \<Longrightarrow> f \<in> init_same_inode_files f'"
+by (auto simp add:init_same_inode_files_def is_init_file_def split:if_splits)
+
+lemma init_same_inode_prop4:
+ "\<lbrakk>f' \<in> init_same_inode_files f; f' \<in> init_files\<rbrakk> \<Longrightarrow> f \<in> init_files"
+apply (drule init_same_inode_prop3)
+by (simp add:init_same_inode_prop2)
+*)
+
+end
+
+context flask begin
+
+lemma init_alive_prop: "init_alive obj = alive [] obj"
+apply (case_tac obj, simp_all add:is_init_file_props is_init_dir_props is_init_tcp_sock_props
+ is_init_udp_sock_props init_files_props init_sockets_props is_file_nil is_dir_nil
+ is_tcp_sock_nil is_udp_sock_nil)
+done
+
+lemma init_alive_proc: "p \<in> init_procs \<Longrightarrow> init_alive (O_proc p)" by simp
+lemma init_alive_file: "is_init_file f \<Longrightarrow> init_alive (O_file f)" by simp
+lemma init_alive_dir: "is_init_dir f \<Longrightarrow> init_alive (O_dir f)" by simp
+lemma init_alive_fd: "fd \<in> init_fds_of_proc p \<Longrightarrow> init_alive (O_fd p fd)" by simp
+lemma init_alive_tcp: "is_init_tcp_sock s \<Longrightarrow> init_alive (O_tcp_sock s)" by simp
+lemma init_alive_udp: "is_init_udp_sock s \<Longrightarrow> init_alive (O_udp_sock s)" by simp
+lemma init_alive_node: "n \<in> init_nodes \<Longrightarrow> init_alive (O_node n)" by simp
+lemma init_alive_msgq: "q \<in> init_msgqs \<Longrightarrow> init_alive (O_msgq q)" by simp
+lemma init_alive_msg: "\<lbrakk>m \<in> set (init_msgs_of_queue q); q \<in> init_msgqs\<rbrakk>
+ \<Longrightarrow> init_alive (O_msg q m)" by simp
+
+lemmas init_alive_intros = init_alive_proc init_alive_file init_alive_dir init_alive_fd
+ init_alive_tcp init_alive_udp init_alive_node init_alive_msgq init_alive_msg
+
+lemma init_file_type_prop1: "is_init_file f \<Longrightarrow> \<exists> t. init_type_of_obj (O_file f) = Some t"
+using init_obj_has_type
+by (auto simp:is_init_file_def split:option.splits)
+
+lemma init_file_type_prop2: "is_init_file f \<Longrightarrow> init_type_of_obj (O_file f) \<noteq> None"
+by (simp add:init_file_type_prop1)
+
+lemma init_file_type_prop3: "init_type_of_obj (O_file f) = Some t \<Longrightarrow> f \<in> init_files"
+apply (drule init_type_has_obj)
+by (simp add:is_init_file_def init_inum_of_file_props split:option.splits)
+
+lemma init_file_type_prop4: "init_type_of_obj (O_file f) = Some t \<Longrightarrow> is_init_file f"
+apply (drule init_type_has_obj)
+by (simp add:is_init_file_def init_inum_of_file_props split:option.splits)
+
+lemmas init_file_types_props = init_file_type_prop1 init_file_type_prop2 init_file_type_prop3 init_file_type_prop4
+
+lemma init_dir_type_prop1: "is_init_dir f \<Longrightarrow> \<exists> t. init_type_of_obj (O_dir f) = Some t"
+using init_obj_has_type
+by (auto simp:is_init_dir_def split:option.splits)
+
+lemma init_dir_type_prop2: "is_init_dir f \<Longrightarrow> init_type_of_obj (O_dir f) \<noteq> None"
+by (simp add:init_dir_type_prop1)
+
+lemma init_dir_type_prop3: "init_type_of_obj (O_dir f) = Some t \<Longrightarrow> f \<in> init_files"
+apply (drule init_type_has_obj)
+by (simp add:is_init_dir_def init_inum_of_file_props split:option.splits)
+
+lemma init_dir_type_prop4: "init_type_of_obj (O_dir f) = Some t \<Longrightarrow> is_init_dir f"
+apply (drule init_type_has_obj)
+by (simp add:is_init_dir_def init_inum_of_file_props split:option.splits)
+
+lemmas init_dir_types_props = init_dir_type_prop1 init_dir_type_prop2 init_dir_type_prop3 init_dir_type_prop4
+
+lemma init_procrole_prop1: "init_role_of_proc p = Some r \<Longrightarrow> p \<in> init_procs"
+using init_proc_has_role
+by (auto simp:bidirect_in_init_def)
+
+lemma init_procrole_prop2: "p \<in> init_procs \<Longrightarrow> \<exists> r. init_role_of_proc p = Some r"
+using init_proc_has_role
+by (auto simp:bidirect_in_init_def)
+
+lemma init_procrole_prop3: "p \<in> init_procs \<Longrightarrow> init_role_of_proc p \<noteq> None"
+using init_proc_has_role
+by (auto simp:bidirect_in_init_def)
+
+lemmas init_role_of_proc_props = init_procrole_prop1 init_procrole_prop2 init_procrole_prop3
+
+lemma init_file_user_prop1: "is_init_file f \<Longrightarrow> \<exists> t. init_user_of_obj (O_file f) = Some t"
+apply (drule init_alive_file)
+by (drule init_obj_has_user, auto)
+
+lemma init_file_user_prop2: "is_init_file f \<Longrightarrow> init_user_of_obj (O_file f) \<noteq> None"
+by (simp add:init_file_user_prop1)
+
+lemma init_file_user_prop3: "init_user_of_obj (O_file f) = Some t \<Longrightarrow> f \<in> init_files"
+apply (drule init_user_has_obj)
+by (simp add:is_init_file_def init_inum_of_file_props split:option.splits)
+
+lemma init_file_user_prop4: "init_user_of_obj (O_file f) = Some t \<Longrightarrow> is_init_file f"
+apply (drule init_user_has_obj)
+by (simp add:is_init_file_def init_inum_of_file_props split:option.splits)
+
+lemma init_file_user_prop5: "init_user_of_obj (O_file f) = Some u \<Longrightarrow> u \<in> init_users"
+by (simp add:init_user_has_obj)
+
+lemmas init_file_users_props = init_file_user_prop1 init_file_user_prop2 init_file_user_prop3 init_file_user_prop4 init_file_user_prop5
+
+lemma init_dir_user_prop1: "is_init_dir f \<Longrightarrow> \<exists> t. init_user_of_obj (O_dir f) = Some t"
+apply (drule init_alive_dir)
+by (drule init_obj_has_user, auto)
+
+lemma init_dir_user_prop2: "is_init_dir f \<Longrightarrow> init_user_of_obj (O_dir f) \<noteq> None"
+by (simp add:init_dir_user_prop1)
+
+lemma init_dir_user_prop3: "init_user_of_obj (O_dir f) = Some t \<Longrightarrow> f \<in> init_files"
+apply (drule init_user_has_obj)
+by (simp add:is_init_dir_def init_inum_of_file_props split:option.splits)
+
+lemma init_dir_user_prop4: "init_user_of_obj (O_dir f) = Some t \<Longrightarrow> is_init_dir f"
+apply (drule init_user_has_obj)
+by (simp add:is_init_dir_def init_inum_of_file_props split:option.splits)
+
+lemma init_dir_user_prop5: "init_user_of_obj (O_dir f) = Some u \<Longrightarrow> u \<in> init_users"
+by (simp add:init_user_has_obj)
+
+lemmas init_dir_users_props = init_dir_user_prop1 init_dir_user_prop2 init_dir_user_prop3 init_dir_user_prop4 init_dir_user_prop5
+
+lemma init_file_dir_conflict: "\<lbrakk>is_init_file f; is_init_dir f\<rbrakk> \<Longrightarrow> False"
+by (auto simp:is_init_file_def is_init_dir_def split:option.splits t_inode_tag.splits)
+
+lemma init_file_dir_conflict1: "is_init_file f \<Longrightarrow> \<not> is_init_dir f"
+by (auto simp:is_init_file_def is_init_dir_def split:option.splits t_inode_tag.splits)
+
+lemma init_file_dir_conflict2: "is_init_dir f \<Longrightarrow> \<not> is_init_file f"
+by (auto simp:is_init_file_def is_init_dir_def split:option.splits t_inode_tag.splits)
+
+end
+
+context tainting begin
+
+lemma tainted_nil_prop:
+ "(x \<in> tainted []) = (x \<in> seeds)"
+by auto
+
+end
+
+context tainting_s begin
+
+lemma init_file_has_ctxt:
+ "is_init_file f \<Longrightarrow> \<exists> sec. init_sectxt_of_obj (O_file f) = Some sec"
+apply (simp add:init_sectxt_of_obj_def split:option.splits)
+apply (rule conjI, rule init_obj_has_user, simp add:is_init_file_props)
+by (simp add:init_file_types_props)
+
+lemma init_file_has_ctxt':
+ "init_sectxt_of_obj (O_file f) = None \<Longrightarrow> \<not> is_init_file f"
+by (rule notI, drule init_file_has_ctxt, simp)
+
+lemma init_dir_has_ctxt:
+ "is_init_dir f \<Longrightarrow> \<exists> sec. init_sectxt_of_obj (O_dir f) = Some sec"
+apply (simp add:init_sectxt_of_obj_def split:option.splits)
+apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props)
+by (simp add:init_dir_types_props)
+
+lemma init_dir_has_ctxt':
+ "init_sectxt_of_obj (O_dir f) = None \<Longrightarrow> \<not> is_init_dir f"
+by (rule notI, drule init_dir_has_ctxt, simp)
+
+lemma init_proc_has_ctxt:
+ "p \<in> init_procs \<Longrightarrow> \<exists> sec. init_sectxt_of_obj (O_proc p) = Some sec"
+apply (simp add:init_sectxt_of_obj_def split:option.splits)
+apply (rule conjI, rule init_obj_has_user, simp)
+apply (frule init_alive_proc, drule init_obj_has_type)
+by (drule init_procrole_prop2, auto)
+
+lemma init_proc_has_ctxt':
+ "init_sectxt_of_obj (O_proc p) = None \<Longrightarrow> p \<notin> init_procs"
+by (rule notI, drule init_proc_has_ctxt, simp)
+
+lemma init_fd_has_ctxt:
+ "fd \<in> init_fds_of_proc p \<Longrightarrow> \<exists> sec. init_sectxt_of_obj (O_fd p fd) = Some sec"
+apply (simp add:init_sectxt_of_obj_def split:option.splits)
+apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props)
+apply (drule init_alive_intros)
+apply (drule init_obj_has_type, clarsimp)
+done
+
+lemma init_fd_has_ctxt':
+ "init_sectxt_of_obj (O_fd p fd) = None \<Longrightarrow> fd \<notin> init_fds_of_proc p"
+by (rule notI, drule init_fd_has_ctxt, simp)
+
+lemma init_node_has_ctxt:
+ "n \<in> init_nodes \<Longrightarrow> \<exists> sec. init_sectxt_of_obj (O_node n) = Some sec"
+apply (simp add:init_sectxt_of_obj_def split:option.splits)
+apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props)
+apply (drule init_alive_intros)
+apply (drule init_obj_has_type, clarsimp)
+done
+
+lemma init_node_has_ctxt':
+ "init_sectxt_of_obj (O_node n) = None \<Longrightarrow> n \<notin> init_nodes"
+by (rule notI, drule init_node_has_ctxt, simp)
+
+lemma init_tcp_has_ctxt:
+ "is_init_tcp_sock s \<Longrightarrow> \<exists> sec. init_sectxt_of_obj (O_tcp_sock s) = Some sec"
+apply (simp add:init_sectxt_of_obj_def split:option.splits)
+apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props)
+apply (drule init_alive_intros)
+apply (drule init_obj_has_type, clarsimp)
+done
+
+lemma init_tcp_has_ctxt':
+ "init_sectxt_of_obj (O_tcp_sock s) = None \<Longrightarrow> \<not> is_init_tcp_sock s"
+by (rule notI, drule init_tcp_has_ctxt, simp)
+
+lemma init_udp_has_ctxt:
+ "is_init_udp_sock s \<Longrightarrow> \<exists> sec. init_sectxt_of_obj (O_udp_sock s) = Some sec"
+apply (simp add:init_sectxt_of_obj_def split:option.splits)
+apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props)
+by (drule init_alive_intros, drule init_obj_has_type, clarsimp)
+
+lemma init_udp_has_ctxt':
+ "init_sectxt_of_obj (O_udp_sock s) = None \<Longrightarrow> \<not> is_init_udp_sock s"
+by (rule notI, drule init_udp_has_ctxt, simp)
+
+lemma init_msgq_has_ctxt:
+ "q \<in> init_msgqs \<Longrightarrow> \<exists> sec. init_sectxt_of_obj (O_msgq q) = Some sec"
+apply (simp add:init_sectxt_of_obj_def split:option.splits)
+apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props)
+by (drule init_alive_intros, drule init_obj_has_type, clarsimp)
+
+lemma init_msgq_has_ctxt':
+ "init_sectxt_of_obj (O_msgq q) = None \<Longrightarrow> q \<notin> init_msgqs"
+by (rule notI, drule init_msgq_has_ctxt, simp)
+
+lemma init_msg_has_ctxt:
+ "\<lbrakk>m \<in> set (init_msgs_of_queue q); q \<in> init_msgqs\<rbrakk> \<Longrightarrow> \<exists> sec. init_sectxt_of_obj (O_msg q m) = Some sec"
+apply (simp add:init_sectxt_of_obj_def split:option.splits)
+apply (rule conjI, rule init_obj_has_user, simp add:is_init_dir_props)
+by (drule init_alive_intros, simp, drule init_obj_has_type, clarsimp)
+
+lemma init_msg_has_ctxt':
+ "init_sectxt_of_obj (O_msg q m) = None \<Longrightarrow> m \<notin> set (init_msgs_of_queue q) \<or> q \<notin> init_msgqs"
+by (auto dest:init_msg_has_ctxt)
+
+lemma init_rootf_has_ctxt:
+ "\<exists> sec. init_sectxt_of_obj (O_dir []) = Some sec"
+apply (rule init_dir_has_ctxt, simp add:is_init_dir_def split:option.splits)
+using root_is_dir by auto
+
+lemma init_rootf_has_ctxt':
+ "init_sectxt_of_obj (O_dir []) = None \<Longrightarrow> False"
+using init_rootf_has_ctxt by auto
+
+lemmas init_has_ctxt = init_file_has_ctxt init_dir_has_ctxt init_proc_has_ctxt init_fd_has_ctxt
+ init_node_has_ctxt init_tcp_has_ctxt init_udp_has_ctxt init_msgq_has_ctxt
+ init_msg_has_ctxt init_rootf_has_ctxt
+
+lemmas init_has_ctxt' = init_file_has_ctxt' init_dir_has_ctxt' init_proc_has_ctxt' init_fd_has_ctxt'
+ init_node_has_ctxt' init_tcp_has_ctxt' init_udp_has_ctxt' init_msgq_has_ctxt'
+ init_msg_has_ctxt' init_rootf_has_ctxt'
+
+lemma sec_of_root_valid:
+ "init_sectxt_of_obj (O_dir []) = Some sec_of_root"
+using init_rootf_has_ctxt
+by (auto simp:init_sectxt_of_obj_def sec_of_root_def split:option.splits)
+
+lemma sec_of_root_is_tuple:
+ "\<exists> u t. sec_of_root = (u, R_object, t)"
+using sec_of_root_valid
+by (auto simp:sec_of_root_def init_sectxt_of_obj_def split:option.splits)
+
+lemma sroot_valid:
+ "init_cf2sfile [] = Some sroot"
+by (simp add:init_cf2sfile_def)
+
+lemma sroot_valid':
+ "cf2sfile s [] = Some sroot"
+by (simp add:cf2sfile_def)
+
+lemma init_sectxt_prop:
+ "sectxt_of_obj [] obj = init_sectxt_of_obj obj"
+apply (auto simp:init_sectxt_of_obj_def sectxt_of_obj_def split:option.splits)
+apply (case_tac [!] obj, simp+)
+done
+
+lemma init_sectxt_prop2:
+ "init_sectxt_of_obj obj = Some sec \<Longrightarrow> init_alive obj"
+by (case_tac obj, auto simp:init_sectxt_of_obj_def split:option.splits dest:init_type_has_obj)
+
+lemma init_dir_has_seclist:
+ "is_init_dir f \<Longrightarrow> \<exists> seclist. get_parentfs_ctxts [] f = Some seclist"
+apply (induct f)
+apply (simp only:get_parentfs_ctxts.simps init_sectxt_prop)
+using init_rootf_has_ctxt apply (auto)[1]
+apply (frule init_parent_file_prop3', simp del:get_parentfs_ctxts.simps)
+apply (erule exE, drule init_dir_has_ctxt)
+by (auto simp add:init_sectxt_prop)
+
+lemma is_init_file_dir_prop1:
+ "is_init_dir f \<Longrightarrow> \<not> is_init_file f"
+by (auto simp:is_init_dir_def is_init_file_def split:option.splits t_inode_tag.splits)
+
+lemma is_init_file_dir_prop2:
+ "is_init_file f \<Longrightarrow> \<not> is_init_dir f"
+by (auto simp:is_init_dir_def is_init_file_def split:option.splits t_inode_tag.splits)
+
+lemma is_init_file_dir_prop3:
+ "\<lbrakk>is_init_dir f; is_init_file f\<rbrakk> \<Longrightarrow> False"
+by (auto simp:is_init_dir_def is_init_file_def split:option.splits t_inode_tag.splits)
+
+lemma is_init_file_dir_prop4:
+ "\<lbrakk>is_init_file f; is_init_dir f\<rbrakk> \<Longrightarrow> False"
+by (auto simp:is_init_dir_def is_init_file_def split:option.splits t_inode_tag.splits)
+
+lemmas is_init_file_dir_props = is_init_file_dir_prop1 is_init_file_dir_prop2 is_init_file_dir_prop3 is_init_file_dir_prop4
+
+lemma init_dir_has_sfile:
+ "is_init_dir f \<Longrightarrow> \<exists> sf. init_cf2sfile f = Some sf"
+apply (case_tac f)
+using init_rootf_has_ctxt apply (auto)[1]
+apply (simp add:sec_of_root_valid sroot_valid sroot_def)
+apply (simp, frule init_parent_file_prop3')
+apply (frule_tac f = list in init_dir_has_seclist)
+apply (frule_tac f = list in init_dir_has_ctxt)
+apply (frule_tac f = "a # list" in init_dir_has_ctxt)
+apply ((erule exE)+, case_tac sec, auto simp:init_cf2sfile_def split:option.splits)
+by (auto simp:is_init_file_def is_init_dir_def split:option.splits t_inode_tag.splits)
+
+lemma init_file_has_sfile:
+ "is_init_file f \<Longrightarrow> \<exists> sf. init_cf2sfile f = Some sf"
+apply (case_tac f)
+apply (simp, drule root_is_init_dir', simp)
+apply (simp, frule init_parent_file_prop2')
+apply (frule_tac f = list in init_dir_has_seclist)
+apply (frule_tac f = list in init_dir_has_ctxt)
+apply (frule_tac f = "a # list" in init_file_has_ctxt)
+by ((erule exE)+, case_tac sec, auto simp:init_cf2sfile_def)
+
+lemma init_proc_has_sproc:
+ "p \<in> init_procs \<Longrightarrow> \<exists> sp. init_cp2sproc p = Some sp"
+apply (frule init_proc_has_ctxt, erule exE)
+apply (simp add:init_cp2sproc_def)
+by (case_tac sec, simp+)
+
+lemma init_cqm2sms_has_sms_aux:
+ "\<forall> m \<in> set ms. init_sectxt_of_obj (O_msg q m) \<noteq> None \<Longrightarrow> (\<exists> sms. init_cqm2sms q ms = Some sms)"
+by (induct ms, auto split:option.splits simp:init_cm2smsg_def)
+
+lemma init_cqm2sms_has_sms:
+ "q \<in> init_msgqs \<Longrightarrow> \<exists> sms. init_cqm2sms q (init_msgs_of_queue q) = Some sms"
+apply (rule init_cqm2sms_has_sms_aux)
+using init_msg_has_ctxt by auto
+
+lemma init_msgq_has_smsgq:
+ "q \<in> init_msgqs \<Longrightarrow> \<exists> sq. init_cq2smsgq q = Some sq"
+apply (frule init_msgq_has_ctxt, erule exE, drule init_cqm2sms_has_sms, erule exE)
+apply (simp add:init_cq2smsgq_def)
+by (case_tac sec, simp+)
+
+lemma cf2sfile_nil_prop:
+ "f \<in> init_files \<Longrightarrow> cf2sfile [] f = init_cf2sfile f"
+apply (case_tac f)
+apply (simp add:init_sectxt_prop cf2sfile_def init_cf2sfile_def)
+apply (auto simp:init_sectxt_prop cf2sfile_def init_cf2sfile_def split:option.splits dest!:init_has_ctxt')
+apply (auto simp:is_init_file_def is_init_dir_def is_file_nil split:option.splits t_inode_tag.splits
+ dest:init_file_has_inum inof_has_file_tag)
+done
+
+lemma init_sec_file_dir:
+ "\<lbrakk>init_sectxt_of_obj (O_file f) = Some x; init_sectxt_of_obj (O_dir f) = Some y\<rbrakk> \<Longrightarrow> False"
+apply (drule init_sectxt_prop2)+
+apply (auto intro:init_file_dir_conflict)
+done
+
+lemma cf2sfile_nil_prop3:
+ "is_init_file f \<Longrightarrow> cf2sfile [] f = init_cf2sfile f"
+by (simp add:is_init_file_prop1 cf2sfile_nil_prop)
+
+lemma cf2sfile_nil_prop4:
+ "is_init_dir f \<Longrightarrow> cf2sfile [] f = init_cf2sfile f"
+apply (frule init_file_dir_conflict2)
+by (simp add:is_init_file_prop1 is_init_dir_prop1 cf2sfile_nil_prop)
+
+lemma cfd2sfd_nil_prop:
+ "init_file_of_proc_fd p fd = Some f \<Longrightarrow> cfd2sfd [] p fd = init_cfd2sfd p fd"
+apply (simp add:cfd2sfd_def init_sectxt_prop init_cfd2sfd_def)
+apply (frule init_filefd_prop5, drule init_filefd_prop1, drule cf2sfile_nil_prop)
+by (auto split:option.splits)
+
+lemma cpfd2sfds_nil_prop:
+ "cpfd2sfds [] p = init_cfds2sfds p"
+apply (simp only:cpfd2sfds_def init_cfds2sfds_def proc_file_fds_def init_proc_file_fds_def)
+apply (rule set_eqI, rule iffI)
+apply (drule CollectD, erule bexE, drule CollectD, erule exE)
+apply (rule CollectI, rule_tac x = fd in bexI) defer
+apply (rule CollectI, rule_tac x = f in exI, simp)
+apply (drule CollectD, erule bexE, drule CollectD, erule exE)
+apply (rule CollectI, rule_tac x = fd in bexI) defer
+apply (rule CollectI, rule_tac x = f in exI)
+using cfd2sfd_nil_prop
+by auto
+
+lemma cp2sproc_nil_prop:
+ "p \<in> init_procs \<Longrightarrow> cp2sproc [] p = init_cp2sproc p"
+by (auto simp add:init_cp2sproc_def cp2sproc_def init_sectxt_prop cpfd2sfds_nil_prop
+ split:option.splits)
+
+lemma msg_has_sec_imp_init:
+ "init_sectxt_of_obj (O_msg q m) = Some sec \<Longrightarrow> q \<in> init_msgqs \<and> m \<in> set (init_msgs_of_queue q)"
+apply (simp add:init_sectxt_of_obj_def split:option.splits)
+by (drule init_type_has_obj, simp)
+
+lemma msgq_has_sec_imp_init:
+ "init_sectxt_of_obj (O_msgq q) = Some sec \<Longrightarrow> q \<in> init_msgqs"
+apply (simp add:init_sectxt_of_obj_def split:option.splits)
+by (drule init_type_has_obj, simp)
+
+lemma cm2smsg_nil_prop:
+ "cm2smsg [] q m = init_cm2smsg q m"
+by (auto simp add:init_sectxt_prop cm2smsg_def init_cm2smsg_def split:option.splits
+ dest: msg_has_sec_imp_init elim:tainted.cases)
+
+lemma cqm2sms_nil_prop:
+ "cqm2sms [] q ms = init_cqm2sms q ms"
+apply (induct ms, simp)
+by (auto simp add:cm2smsg_def init_sectxt_prop tainted_nil_prop msg_has_sec_imp_init init_cm2smsg_def
+ split:option.splits)
+
+lemma cq2smsga_nil_prop:
+ "cq2smsgq [] q = init_cq2smsgq q"
+by (auto simp add:cq2smsgq_def init_cq2smsgq_def init_sectxt_prop cqm2sms_nil_prop
+ intro:msgq_has_sec_imp_init split:option.splits)
+
+lemma co2sobj_nil_prop:
+ "init_alive obj \<Longrightarrow> co2sobj [] obj = init_obj2sobj obj"
+apply (case_tac obj)
+apply (auto simp add:cf2sfile_nil_prop cq2smsga_nil_prop cqm2sms_nil_prop tainted_nil_prop
+ cp2sproc_nil_prop is_init_dir_prop1 is_init_file_prop1
+ is_init_udp_sock_prop1 is_init_tcp_sock_prop1
+ cm2smsg_nil_prop
+ split:option.splits)
+done
+
+lemma s2ss_nil_prop:
+ "s2ss [] = init_static_state"
+using co2sobj_nil_prop init_alive_prop
+by (auto simp add:s2ss_def init_static_state_def)
+
+end
+
+(*<*)
+end
+(*>*)
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/List_Prefix.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,382 @@
+(* Title: HOL/Library/List_Prefix.thy
+ Author: Tobias Nipkow and Markus Wenzel, TU Muenchen
+*)
+
+header {* List prefixes and postfixes *}
+
+theory List_Prefix
+imports List Main
+begin
+
+subsection {* Prefix order on lists *}
+
+instantiation list :: (type) "{order, bot}"
+begin
+
+definition
+ prefix_def: "xs \<le> ys \<longleftrightarrow> (\<exists>zs. ys = xs @ zs)"
+
+definition
+ strict_prefix_def: "xs < ys \<longleftrightarrow> xs \<le> ys \<and> xs \<noteq> (ys::'a list)"
+
+definition
+ "bot = []"
+
+instance proof
+qed (auto simp add: prefix_def strict_prefix_def bot_list_def)
+
+end
+
+lemma prefixI [intro?]: "ys = xs @ zs ==> xs \<le> ys"
+ unfolding prefix_def by blast
+
+lemma prefixE [elim?]:
+ assumes "xs \<le> ys"
+ obtains zs where "ys = xs @ zs"
+ using assms unfolding prefix_def by blast
+
+lemma strict_prefixI' [intro?]: "ys = xs @ z # zs ==> xs < ys"
+ unfolding strict_prefix_def prefix_def by blast
+
+lemma strict_prefixE' [elim?]:
+ assumes "xs < ys"
+ obtains z zs where "ys = xs @ z # zs"
+proof -
+ from `xs < ys` obtain us where "ys = xs @ us" and "xs \<noteq> ys"
+ unfolding strict_prefix_def prefix_def by blast
+ with that show ?thesis by (auto simp add: neq_Nil_conv)
+qed
+
+lemma strict_prefixI [intro?]: "xs \<le> ys ==> xs \<noteq> ys ==> xs < (ys::'a list)"
+ unfolding strict_prefix_def by blast
+
+lemma strict_prefixE [elim?]:
+ fixes xs ys :: "'a list"
+ assumes "xs < ys"
+ obtains "xs \<le> ys" and "xs \<noteq> ys"
+ using assms unfolding strict_prefix_def by blast
+
+
+subsection {* Basic properties of prefixes *}
+
+theorem Nil_prefix [iff]: "[] \<le> xs"
+ by (simp add: prefix_def)
+
+theorem prefix_Nil [simp]: "(xs \<le> []) = (xs = [])"
+ by (induct xs) (simp_all add: prefix_def)
+
+lemma prefix_snoc [simp]: "(xs \<le> ys @ [y]) = (xs = ys @ [y] \<or> xs \<le> ys)"
+proof
+ assume "xs \<le> ys @ [y]"
+ then obtain zs where zs: "ys @ [y] = xs @ zs" ..
+ show "xs = ys @ [y] \<or> xs \<le> ys"
+ by (metis append_Nil2 butlast_append butlast_snoc prefixI zs)
+next
+ assume "xs = ys @ [y] \<or> xs \<le> ys"
+ then show "xs \<le> ys @ [y]"
+ by (metis order_eq_iff order_trans prefixI)
+qed
+
+lemma Cons_prefix_Cons [simp]: "(x # xs \<le> y # ys) = (x = y \<and> xs \<le> ys)"
+ by (auto simp add: prefix_def)
+
+lemma less_eq_list_code [code]:
+ "([]\<Colon>'a\<Colon>{equal, ord} list) \<le> xs \<longleftrightarrow> True"
+ "(x\<Colon>'a\<Colon>{equal, ord}) # xs \<le> [] \<longleftrightarrow> False"
+ "(x\<Colon>'a\<Colon>{equal, ord}) # xs \<le> y # ys \<longleftrightarrow> x = y \<and> xs \<le> ys"
+ by simp_all
+
+lemma same_prefix_prefix [simp]: "(xs @ ys \<le> xs @ zs) = (ys \<le> zs)"
+ by (induct xs) simp_all
+
+lemma same_prefix_nil [iff]: "(xs @ ys \<le> xs) = (ys = [])"
+ by (metis append_Nil2 append_self_conv order_eq_iff prefixI)
+
+lemma prefix_prefix [simp]: "xs \<le> ys ==> xs \<le> ys @ zs"
+ by (metis order_le_less_trans prefixI strict_prefixE strict_prefixI)
+
+lemma append_prefixD: "xs @ ys \<le> zs \<Longrightarrow> xs \<le> zs"
+ by (auto simp add: prefix_def)
+
+theorem prefix_Cons: "(xs \<le> y # ys) = (xs = [] \<or> (\<exists>zs. xs = y # zs \<and> zs \<le> ys))"
+ by (cases xs) (auto simp add: prefix_def)
+
+theorem prefix_append:
+ "(xs \<le> ys @ zs) = (xs \<le> ys \<or> (\<exists>us. xs = ys @ us \<and> us \<le> zs))"
+ apply (induct zs rule: rev_induct)
+ apply force
+ apply (simp del: append_assoc add: append_assoc [symmetric])
+ apply (metis append_eq_appendI)
+ done
+
+lemma append_one_prefix:
+ "xs \<le> ys ==> length xs < length ys ==> xs @ [ys ! length xs] \<le> ys"
+ unfolding prefix_def
+ by (metis Cons_eq_appendI append_eq_appendI append_eq_conv_conj
+ eq_Nil_appendI nth_drop')
+
+theorem prefix_length_le: "xs \<le> ys ==> length xs \<le> length ys"
+ by (auto simp add: prefix_def)
+
+lemma prefix_same_cases:
+ "(xs\<^isub>1::'a list) \<le> ys \<Longrightarrow> xs\<^isub>2 \<le> ys \<Longrightarrow> xs\<^isub>1 \<le> xs\<^isub>2 \<or> xs\<^isub>2 \<le> xs\<^isub>1"
+ unfolding prefix_def by (metis append_eq_append_conv2)
+
+lemma set_mono_prefix: "xs \<le> ys \<Longrightarrow> set xs \<subseteq> set ys"
+ by (auto simp add: prefix_def)
+
+lemma take_is_prefix: "take n xs \<le> xs"
+ unfolding prefix_def by (metis append_take_drop_id)
+
+lemma map_prefixI: "xs \<le> ys \<Longrightarrow> map f xs \<le> map f ys"
+ by (auto simp: prefix_def)
+
+lemma prefix_length_less: "xs < ys \<Longrightarrow> length xs < length ys"
+ by (auto simp: strict_prefix_def prefix_def)
+
+lemma strict_prefix_simps [simp, code]:
+ "xs < [] \<longleftrightarrow> False"
+ "[] < x # xs \<longleftrightarrow> True"
+ "x # xs < y # ys \<longleftrightarrow> x = y \<and> xs < ys"
+ by (simp_all add: strict_prefix_def cong: conj_cong)
+
+lemma take_strict_prefix: "xs < ys \<Longrightarrow> take n xs < ys"
+ apply (induct n arbitrary: xs ys)
+ apply (case_tac ys, simp_all)[1]
+ apply (metis order_less_trans strict_prefixI take_is_prefix)
+ done
+
+lemma not_prefix_cases:
+ assumes pfx: "\<not> ps \<le> ls"
+ obtains
+ (c1) "ps \<noteq> []" and "ls = []"
+ | (c2) a as x xs where "ps = a#as" and "ls = x#xs" and "x = a" and "\<not> as \<le> xs"
+ | (c3) a as x xs where "ps = a#as" and "ls = x#xs" and "x \<noteq> a"
+proof (cases ps)
+ case Nil then show ?thesis using pfx by simp
+next
+ case (Cons a as)
+ note c = `ps = a#as`
+ show ?thesis
+ proof (cases ls)
+ case Nil then show ?thesis by (metis append_Nil2 pfx c1 same_prefix_nil)
+ next
+ case (Cons x xs)
+ show ?thesis
+ proof (cases "x = a")
+ case True
+ have "\<not> as \<le> xs" using pfx c Cons True by simp
+ with c Cons True show ?thesis by (rule c2)
+ next
+ case False
+ with c Cons show ?thesis by (rule c3)
+ qed
+ qed
+qed
+
+lemma not_prefix_induct [consumes 1, case_names Nil Neq Eq]:
+ assumes np: "\<not> ps \<le> ls"
+ and base: "\<And>x xs. P (x#xs) []"
+ and r1: "\<And>x xs y ys. x \<noteq> y \<Longrightarrow> P (x#xs) (y#ys)"
+ and r2: "\<And>x xs y ys. \<lbrakk> x = y; \<not> xs \<le> ys; P xs ys \<rbrakk> \<Longrightarrow> P (x#xs) (y#ys)"
+ shows "P ps ls" using np
+proof (induct ls arbitrary: ps)
+ case Nil then show ?case
+ by (auto simp: neq_Nil_conv elim!: not_prefix_cases intro!: base)
+next
+ case (Cons y ys)
+ then have npfx: "\<not> ps \<le> (y # ys)" by simp
+ then obtain x xs where pv: "ps = x # xs"
+ by (rule not_prefix_cases) auto
+ show ?case by (metis Cons.hyps Cons_prefix_Cons npfx pv r1 r2)
+qed
+
+
+subsection {* Parallel lists *}
+
+definition
+ parallel :: "'a list => 'a list => bool" (infixl "\<parallel>" 50) where
+ "(xs \<parallel> ys) = (\<not> xs \<le> ys \<and> \<not> ys \<le> xs)"
+
+lemma parallelI [intro]: "\<not> xs \<le> ys ==> \<not> ys \<le> xs ==> xs \<parallel> ys"
+ unfolding parallel_def by blast
+
+lemma parallelE [elim]:
+ assumes "xs \<parallel> ys"
+ obtains "\<not> xs \<le> ys \<and> \<not> ys \<le> xs"
+ using assms unfolding parallel_def by blast
+
+theorem prefix_cases:
+ obtains "xs \<le> ys" | "ys < xs" | "xs \<parallel> ys"
+ unfolding parallel_def strict_prefix_def by blast
+
+theorem parallel_decomp:
+ "xs \<parallel> ys ==> \<exists>as b bs c cs. b \<noteq> c \<and> xs = as @ b # bs \<and> ys = as @ c # cs"
+proof (induct xs rule: rev_induct)
+ case Nil
+ then have False by auto
+ then show ?case ..
+next
+ case (snoc x xs)
+ show ?case
+ proof (rule prefix_cases)
+ assume le: "xs \<le> ys"
+ then obtain ys' where ys: "ys = xs @ ys'" ..
+ show ?thesis
+ proof (cases ys')
+ assume "ys' = []"
+ then show ?thesis by (metis append_Nil2 parallelE prefixI snoc.prems ys)
+ next
+ fix c cs assume ys': "ys' = c # cs"
+ then show ?thesis
+ by (metis Cons_eq_appendI eq_Nil_appendI parallelE prefixI
+ same_prefix_prefix snoc.prems ys)
+ qed
+ next
+ assume "ys < xs" then have "ys \<le> xs @ [x]" by (simp add: strict_prefix_def)
+ with snoc have False by blast
+ then show ?thesis ..
+ next
+ assume "xs \<parallel> ys"
+ with snoc obtain as b bs c cs where neq: "(b::'a) \<noteq> c"
+ and xs: "xs = as @ b # bs" and ys: "ys = as @ c # cs"
+ by blast
+ from xs have "xs @ [x] = as @ b # (bs @ [x])" by simp
+ with neq ys show ?thesis by blast
+ qed
+qed
+
+lemma parallel_append: "a \<parallel> b \<Longrightarrow> a @ c \<parallel> b @ d"
+ apply (rule parallelI)
+ apply (erule parallelE, erule conjE,
+ induct rule: not_prefix_induct, simp+)+
+ done
+
+lemma parallel_appendI: "xs \<parallel> ys \<Longrightarrow> x = xs @ xs' \<Longrightarrow> y = ys @ ys' \<Longrightarrow> x \<parallel> y"
+ by (simp add: parallel_append)
+
+lemma parallel_commute: "a \<parallel> b \<longleftrightarrow> b \<parallel> a"
+ unfolding parallel_def by auto
+
+
+subsection {* Postfix order on lists *}
+
+definition
+ postfix :: "'a list => 'a list => bool" ("(_/ >>= _)" [51, 50] 50) where
+ "(xs >>= ys) = (\<exists>zs. xs = zs @ ys)"
+
+lemma postfixI [intro?]: "xs = zs @ ys ==> xs >>= ys"
+ unfolding postfix_def by blast
+
+lemma postfixE [elim?]:
+ assumes "xs >>= ys"
+ obtains zs where "xs = zs @ ys"
+ using assms unfolding postfix_def by blast
+
+lemma postfix_refl [iff]: "xs >>= xs"
+ by (auto simp add: postfix_def)
+lemma postfix_trans: "\<lbrakk>xs >>= ys; ys >>= zs\<rbrakk> \<Longrightarrow> xs >>= zs"
+ by (auto simp add: postfix_def)
+lemma postfix_antisym: "\<lbrakk>xs >>= ys; ys >>= xs\<rbrakk> \<Longrightarrow> xs = ys"
+ by (auto simp add: postfix_def)
+
+lemma Nil_postfix [iff]: "xs >>= []"
+ by (simp add: postfix_def)
+lemma postfix_Nil [simp]: "([] >>= xs) = (xs = [])"
+ by (auto simp add: postfix_def)
+
+lemma postfix_ConsI: "xs >>= ys \<Longrightarrow> x#xs >>= ys"
+ by (auto simp add: postfix_def)
+lemma postfix_ConsD: "xs >>= y#ys \<Longrightarrow> xs >>= ys"
+ by (auto simp add: postfix_def)
+
+lemma postfix_appendI: "xs >>= ys \<Longrightarrow> zs @ xs >>= ys"
+ by (auto simp add: postfix_def)
+lemma postfix_appendD: "xs >>= zs @ ys \<Longrightarrow> xs >>= ys"
+ by (auto simp add: postfix_def)
+
+lemma postfix_is_subset: "xs >>= ys ==> set ys \<subseteq> set xs"
+proof -
+ assume "xs >>= ys"
+ then obtain zs where "xs = zs @ ys" ..
+ then show ?thesis by (induct zs) auto
+qed
+
+lemma postfix_ConsD2: "x#xs >>= y#ys ==> xs >>= ys"
+proof -
+ assume "x#xs >>= y#ys"
+ then obtain zs where "x#xs = zs @ y#ys" ..
+ then show ?thesis
+ by (induct zs) (auto intro!: postfix_appendI postfix_ConsI)
+qed
+
+lemma postfix_to_prefix [code]: "xs >>= ys \<longleftrightarrow> rev ys \<le> rev xs"
+proof
+ assume "xs >>= ys"
+ then obtain zs where "xs = zs @ ys" ..
+ then have "rev xs = rev ys @ rev zs" by simp
+ then show "rev ys <= rev xs" ..
+next
+ assume "rev ys <= rev xs"
+ then obtain zs where "rev xs = rev ys @ zs" ..
+ then have "rev (rev xs) = rev zs @ rev (rev ys)" by simp
+ then have "xs = rev zs @ ys" by simp
+ then show "xs >>= ys" ..
+qed
+
+lemma distinct_postfix: "distinct xs \<Longrightarrow> xs >>= ys \<Longrightarrow> distinct ys"
+ by (clarsimp elim!: postfixE)
+
+lemma postfix_map: "xs >>= ys \<Longrightarrow> map f xs >>= map f ys"
+ by (auto elim!: postfixE intro: postfixI)
+
+lemma postfix_drop: "as >>= drop n as"
+ unfolding postfix_def
+ apply (rule exI [where x = "take n as"])
+ apply simp
+ done
+
+lemma postfix_take: "xs >>= ys \<Longrightarrow> xs = take (length xs - length ys) xs @ ys"
+ by (clarsimp elim!: postfixE)
+
+lemma parallelD1: "x \<parallel> y \<Longrightarrow> \<not> x \<le> y"
+ by blast
+
+lemma parallelD2: "x \<parallel> y \<Longrightarrow> \<not> y \<le> x"
+ by blast
+
+lemma parallel_Nil1 [simp]: "\<not> x \<parallel> []"
+ unfolding parallel_def by simp
+
+lemma parallel_Nil2 [simp]: "\<not> [] \<parallel> x"
+ unfolding parallel_def by simp
+
+lemma Cons_parallelI1: "a \<noteq> b \<Longrightarrow> a # as \<parallel> b # bs"
+ by auto
+
+lemma Cons_parallelI2: "\<lbrakk> a = b; as \<parallel> bs \<rbrakk> \<Longrightarrow> a # as \<parallel> b # bs"
+ by (metis Cons_prefix_Cons parallelE parallelI)
+
+lemma not_equal_is_parallel:
+ assumes neq: "xs \<noteq> ys"
+ and len: "length xs = length ys"
+ shows "xs \<parallel> ys"
+ using len neq
+proof (induct rule: list_induct2)
+ case Nil
+ then show ?case by simp
+next
+ case (Cons a as b bs)
+ have ih: "as \<noteq> bs \<Longrightarrow> as \<parallel> bs" by fact
+ show ?case
+ proof (cases "a = b")
+ case True
+ then have "as \<noteq> bs" using Cons by simp
+ then show ?thesis by (rule Cons_parallelI2 [OF True ih])
+ next
+ case False
+ then show ?thesis by (rule Cons_parallelI1)
+ qed
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/My_list_prefix.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,402 @@
+(*<*)
+theory My_list_prefix
+imports List_Prefix
+begin
+(*>*)
+
+(* cmp:: 1:complete equal; 2:less; 3:greater; 4: len equal,but ele no equal *)
+fun cmp :: "'a list \<Rightarrow> 'a list \<Rightarrow> nat"
+where
+ "cmp [] [] = 1" |
+ "cmp [] (e#es) = 2" |
+ "cmp (e#es) [] = 3" |
+ "cmp (e#es) (a#as) = (let r = cmp es as in
+ if (e = a) then r else 4)"
+
+(* list_com:: fetch the same ele of the same left order into a new list*)
+fun list_com :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
+ "list_com [] ys = []" |
+ "list_com xs [] = []" |
+ "list_com (x#xs) (y#ys) = (if x = y then x#(list_com xs ys) else [])"
+
+(* list_com_rev:: by the right order of list_com *)
+definition list_com_rev :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "\<bullet>" 50)
+where
+ "xs \<bullet> ys \<equiv> rev (list_com (rev xs) (rev ys))"
+
+(* list_diff:: list substract, once different return tailer *)
+fun list_diff :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
+ "list_diff [] xs = []" |
+ "list_diff (x#xs) [] = x#xs" |
+ "list_diff (x#xs) (y#ys) = (if x = y then list_diff xs ys else (x#xs))"
+
+(* list_diff_rev:: list substract with rev order*)
+definition list_diff_rev :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "\<setminus>" 51)
+where
+ "xs \<setminus> ys \<equiv> rev (list_diff (rev xs) (rev ys))"
+
+(* xs <= ys:: \<exists>zs. ys = xs @ zs *)
+(* no_junior:: xs is ys' tail,or equal *)
+definition no_junior :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<preceq>" 50)
+where
+ "xs \<preceq> ys \<equiv> rev xs \<le> rev ys"
+
+(* < :: xs <= ys \<and> xs \<noteq> ys *)
+(* is_ancestor:: xs is ys' tail, but no equal *)
+definition is_ancestor :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<prec>" 50)
+where
+ "xs \<prec> ys \<equiv> rev xs < rev ys"
+
+lemma list_com_diff [simp]: "(list_com xs ys) @ (list_diff xs ys) = xs" (is "?P xs ys")
+ by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_com_diff_rev [simp]: "(xs \<setminus> ys) @ (xs \<bullet> ys) = xs"
+ apply (simp only:list_com_rev_def list_diff_rev_def)
+ by (fold rev_append, simp)
+
+lemma list_com_commute: "list_com xs ys = list_com ys xs" (is "?P xs ys")
+ by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_com_ido: "xs \<le> ys \<longrightarrow> list_com xs ys = xs" (is "?P xs ys")
+ by (rule_tac P = ?P in cmp.induct, auto+)
+
+lemma list_com_rev_ido [simp]: "xs \<preceq> ys \<Longrightarrow> xs \<bullet> ys = xs"
+ by (cut_tac list_com_ido, auto simp: no_junior_def list_com_rev_def)
+
+lemma list_com_rev_commute [iff]: "(xs \<bullet> ys) = (ys \<bullet> xs)"
+ by (simp only:list_com_rev_def list_com_commute)
+
+lemma list_com_rev_ido1 [simp]: "xs \<preceq> ys \<Longrightarrow> ys \<bullet> xs = xs"
+ by simp
+
+lemma list_diff_le: "(list_diff xs ys = []) = (xs \<le> ys)" (is "?P xs ys")
+ by (rule_tac P = ?P in cmp.induct, simp+)
+
+
+lemma list_diff_rev_le: "(xs \<setminus> ys = []) = (xs \<preceq> ys)"
+ by (auto simp:list_diff_rev_def no_junior_def list_diff_le)
+
+lemma list_diff_lt: "(list_diff xs ys = [] \<and> list_diff ys xs \<noteq> []) = (xs < ys)" (is "?P xs ys")
+ by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_diff_rev_lt: "(xs \<setminus> ys = [] \<and> ys \<setminus> xs \<noteq> []) = (xs \<prec> ys)"
+ by (auto simp: list_diff_rev_def list_diff_lt is_ancestor_def)
+
+
+(* xs diff ys result not [] \<Longrightarrow> \<exists> e \<in> xs. a \<in> ys. e \<noteq> a *)
+lemma list_diff_neq:
+ "\<forall> e es a as. list_diff xs ys = (e#es) \<and> list_diff ys xs = (a#as) \<longrightarrow> e \<noteq> a" (is "?P xs ys")
+ by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_diff_rev_neq_pre: "\<forall> e es a as. xs \<setminus> ys = rev (e#es) \<and> ys \<setminus> xs = rev (a#as) \<longrightarrow> e \<noteq> a"
+ apply (simp only:list_diff_rev_def, clarify)
+ apply (insert list_diff_neq, atomize)
+ by (erule_tac x = "rev xs" in allE, erule_tac x = "rev ys" in allE, blast)
+
+lemma list_diff_rev_neq: "\<forall> e es a as. xs \<setminus> ys = es @ [e] \<and> ys \<setminus> xs = as @ [a] \<longrightarrow> e \<noteq> a"
+ apply (rule_tac allI)+
+ apply (insert list_diff_rev_neq_pre, atomize)
+ apply (erule_tac x = "xs" in allE)
+ apply (erule_tac x = "ys" in allE)
+ apply (erule_tac x = "e" in allE)
+ apply (erule_tac x = "rev es" in allE)
+ apply (erule_tac x = "a" in allE)
+ apply (erule_tac x = "rev as" in allE)
+ by auto
+
+lemma list_com_self [simp]: "list_com zs zs = zs"
+ by (induct_tac zs, simp+)
+
+lemma list_com_rev_self [simp]: "zs \<bullet> zs = zs"
+ by (simp add:list_com_rev_def)
+
+lemma list_com_append [simp]: "(list_com (zs @ xs) (zs @ ys)) = (zs @ (list_com xs ys))"
+ by (induct_tac zs, simp+)
+
+lemma list_inter_append [simp]: "((xs @ zs) \<bullet> (ys @ zs)) = ((xs \<bullet> ys) @ zs)"
+ by (simp add:list_com_rev_def)
+
+lemma list_diff_djoin_pre:
+ "\<forall> e es a as. list_diff xs ys = e#es \<and> list_diff ys xs = a#as \<longrightarrow> (\<forall> zs zs'. (list_diff (xs @ zs) (ys @ zs') = [e]@es@zs))"
+ (is "?P xs ys")
+ by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_diff_djoin_rev_pre:
+ "\<forall> e es a as. xs \<setminus> ys = rev (e#es) \<and> ys \<setminus> xs = rev (a#as) \<longrightarrow> (\<forall> zs zs'. ((zs @ xs) \<setminus> (zs' @ ys) = rev ([e]@es@rev zs)))"
+ apply (simp only: list_diff_rev_def, clarify)
+ apply (insert list_diff_djoin_pre, atomize)
+ apply (erule_tac x = "rev xs" in allE)
+ apply (erule_tac x = "rev ys" in allE)
+ apply (erule_tac x = "e" in allE)
+ apply (erule_tac x = "es" in allE)
+ apply (erule_tac x = "a" in allE)
+ apply (erule_tac x = "as" in allE)
+ by simp
+
+lemma list_diff_djoin_rev:
+ "xs \<setminus> ys = es @ [e] \<and> ys \<setminus> xs = as @ [a] \<Longrightarrow> zs @ xs \<setminus> zs' @ ys = zs @ es @ [e]"
+ apply (insert list_diff_djoin_rev_pre [rule_format, simplified])
+ apply (clarsimp, atomize)
+ apply (erule_tac x = "xs" in allE)
+ apply (erule_tac x = "ys" in allE)
+ apply (erule_tac x = "rev es" in allE)
+ apply (erule_tac x = "e" in allE)
+ apply (erule_tac x = "rev as" in allE)
+ apply (erule_tac x = "a" in allE)
+ by auto
+
+lemmas list_diff_djoin_rev_simplified = conjI [THEN list_diff_djoin_rev, simp]
+
+lemmas list_diff_djoin = conjI [THEN list_diff_djoin_pre [rule_format], simp]
+
+lemma list_diff_ext_left [simp]: "(list_diff (zs @ xs) (zs @ ys) = (list_diff xs ys))"
+ by (induct_tac zs, simp+)
+
+lemma list_diff_rev_ext_left [simp]: "((xs @ zs \<setminus> ys @ zs) = (xs \<setminus> ys))"
+ by (auto simp: list_diff_rev_def)
+
+declare no_junior_def [simp]
+
+lemma no_juniorE: "\<lbrakk>xs \<preceq> ys; \<And> zs. ys = zs @ xs \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+proof -
+ assume h: "xs \<preceq> ys"
+ and h1: "\<And> zs. ys = zs @ xs \<Longrightarrow> R"
+ show "R"
+ proof -
+ from h have "rev xs \<le> rev ys" by (simp)
+ from this obtain zs where eq_rev: "rev ys = rev xs @ zs" by (auto simp:prefix_def)
+ show R
+ proof(rule h1 [where zs = "rev zs"])
+ from rev_rev_ident and eq_rev have "rev (rev (ys)) = rev zs @ rev (rev xs)"
+ by simp
+ thus "ys = rev zs @ xs" by simp
+ qed
+ qed
+qed
+
+lemma no_juniorI: "\<lbrakk>ys = zs @ xs\<rbrakk> \<Longrightarrow> xs \<preceq> ys"
+ by simp
+
+lemma no_junior_ident [simp]: "xs \<preceq> xs"
+ by simp
+
+lemma no_junior_expand: "xs \<preceq> ys = ((xs \<prec> ys) \<or> xs = ys)"
+ by (simp only:no_junior_def is_ancestor_def strict_prefix_def, blast)
+
+lemma no_junior_same_prefix: " e # \<tau> \<preceq> e' # \<tau>' \<Longrightarrow> \<tau> \<preceq> \<tau>'"
+apply (simp add:no_junior_def )
+apply (erule disjE, simp)
+apply (simp only:prefix_def)
+by (erule exE, rule_tac x = "[e] @ zs" in exI, auto)
+
+lemma no_junior_noteq: "\<lbrakk>\<tau> \<preceq> a # \<tau>'; \<tau> \<noteq> a # \<tau>'\<rbrakk> \<Longrightarrow> \<tau> \<preceq> \<tau>'"
+apply (erule no_juniorE)
+by (case_tac zs, simp+)
+
+lemma is_ancestor_app [simp]: "xs \<prec> ys \<Longrightarrow> xs \<prec> zs @ ys"
+ by (auto simp:is_ancestor_def strict_prefix_def)
+
+lemma is_ancestor_cons [simp]: "xs \<prec> ys \<Longrightarrow> xs \<prec> a # ys"
+ by (auto simp:is_ancestor_def strict_prefix_def)
+
+lemma no_junior_app [simp]: "xs \<preceq> ys \<Longrightarrow> xs \<preceq> zs @ ys"
+ by simp
+
+lemma is_ancestor_no_junior [simp]: "xs \<prec> ys \<Longrightarrow> xs \<preceq> ys"
+ by (simp add:is_ancestor_def)
+
+lemma is_ancestor_y [simp]: "ys \<prec> y#ys"
+ by (simp add:is_ancestor_def strict_prefix_def)
+
+lemma no_junior_cons [simp]: "xs \<preceq> ys \<Longrightarrow> xs \<prec> (y#ys)"
+ by (unfold no_junior_expand, auto)
+
+lemma no_junior_anti_sym: "\<lbrakk>xs \<preceq> ys; ys \<preceq> xs\<rbrakk> \<Longrightarrow> xs = ys"
+ by simp
+
+declare no_junior_def [simp del]
+
+(* djoin:: xs and ys is not the other's tail, not equal either *)
+definition djoin :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<asymp>" 50)
+where
+ "xs \<asymp> ys \<equiv> \<not> (xs \<preceq> ys \<or> ys \<preceq> xs)"
+
+(* dinj:: function f's returning list is not tailing when paras not equal *)
+definition dinj :: "('a \<Rightarrow> 'b list) \<Rightarrow> bool"
+where
+ "dinj f \<equiv> (\<forall> a b. a \<noteq> b \<longrightarrow> f a \<asymp> f b)"
+
+
+(* list_cmp:: list comparison: one is other's prefix or no equal at some position *)
+lemma list_cmp: "xs \<le> ys \<or> ys \<le> xs \<or> (\<exists> zs x y a b. xs = zs @ [a] @ x \<and> ys = zs @ [b] @ y \<and> a \<noteq> b)"
+proof(cases "list_diff xs ys")
+ assume " list_diff xs ys = []" with list_diff_le show ?thesis by blast
+next
+ fix e es
+ assume h: "list_diff xs ys = e # es"
+ show ?thesis
+ proof(cases "list_diff ys xs")
+ assume " list_diff ys xs = []" with list_diff_le show ?thesis by blast
+ next
+ fix a as assume h1: "list_diff ys xs = (a # as)"
+ have "xs = (list_com xs ys) @ [e] @ es \<and> ys = (list_com xs ys) @ [a] @ as \<and> e \<noteq> a"
+ apply (simp, fold h1, fold h)
+ apply (simp,subst list_com_commute, simp)
+ apply (rule_tac list_diff_neq[rule_format])
+ by (insert h1, insert h, blast)
+ thus ?thesis by blast
+ qed
+qed
+
+(* In fact, this is a case split *)
+lemma list_diff_ind: "\<lbrakk>list_diff xs ys = [] \<Longrightarrow> R; list_diff ys xs = [] \<Longrightarrow> R;
+ \<And> e es a as. \<lbrakk>list_diff xs ys = e#es; list_diff ys xs = a#as; e \<noteq> a\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+proof -
+ assume h1: "list_diff xs ys = [] \<Longrightarrow> R"
+ and h2: "list_diff ys xs = [] \<Longrightarrow> R"
+ and h3: "\<And> e es a as. \<lbrakk>list_diff xs ys = e#es; list_diff ys xs = a#as; e \<noteq> a\<rbrakk> \<Longrightarrow> R"
+ show R
+ proof(cases "list_diff xs ys")
+ assume "list_diff xs ys = []" from h1 [OF this] show R .
+ next
+ fix e es
+ assume he: "list_diff xs ys = e#es"
+ show R
+ proof(cases "list_diff ys xs")
+ assume "list_diff ys xs = []" from h2 [OF this] show R .
+ next
+ fix a as
+ assume ha: "list_diff ys xs = a#as" show R
+ proof(rule h3 [OF he ha])
+ from list_diff_neq [rule_format, OF conjI [OF he ha ]]
+ show "e \<noteq> a" .
+ qed
+ qed
+ qed
+qed
+
+lemma list_diff_rev_ind:
+ "\<lbrakk>xs \<setminus> ys = [] \<Longrightarrow> R; ys \<setminus> xs = [] \<Longrightarrow> R; \<And> e es a as. \<lbrakk>xs \<setminus> ys = es@[e]; ys \<setminus> xs = as@[a]; e \<noteq> a\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+proof -
+ fix xs ys R
+ assume h1: "xs \<setminus> ys = [] \<Longrightarrow> R"
+ and h2: "ys \<setminus> xs = [] \<Longrightarrow> R"
+ and h3: "\<And> e es a as. \<lbrakk>xs \<setminus> ys = es@[e]; ys \<setminus> xs = as@[a]; e \<noteq> a\<rbrakk> \<Longrightarrow> R"
+ show R
+ proof (rule list_diff_ind [where xs = "rev xs" and ys = "rev ys"])
+ assume "list_diff (rev xs) (rev ys) = []" thus R by (auto intro:h1 simp:list_diff_rev_def)
+ next
+ assume "list_diff (rev ys) (rev xs) = []" thus R by (auto intro:h2 simp:list_diff_rev_def)
+ next
+ fix e es a as
+ assume "list_diff (rev xs) (rev ys) = e # es"
+ and "list_diff (rev ys) (rev xs) = a # as"
+ and " e \<noteq> a"
+ thus R by (auto intro:h3 simp:list_diff_rev_def)
+ qed
+qed
+
+lemma djoin_diff_iff: "(xs \<asymp> ys) = (\<exists> e es a as. list_diff (rev xs) (rev ys) = e#es \<and> list_diff (rev ys) (rev xs) = a#as \<and> a \<noteq> e)"
+proof (rule list_diff_ind [where xs = "rev xs" and ys = "rev ys"])
+ assume "list_diff (rev xs) (rev ys) = []"
+ hence "xs \<preceq> ys" by (unfold no_junior_def, simp add:list_diff_le)
+ thus ?thesis
+ apply (auto simp:djoin_def no_junior_def)
+ by (fold list_diff_le, simp)
+next
+ assume "list_diff (rev ys) (rev xs) = []"
+ hence "ys \<preceq> xs" by (unfold no_junior_def, simp add:list_diff_le)
+ thus ?thesis
+ apply (auto simp:djoin_def no_junior_def)
+ by (fold list_diff_le, simp)
+next
+ fix e es a as
+ assume he: "list_diff (rev xs) (rev ys) = e # es"
+ and ha: "list_diff (rev ys) (rev xs) = a # as"
+ and hn: "e \<noteq> a"
+ show ?thesis
+ proof
+ from he ha hn
+ show
+ "\<exists>e es a as. list_diff (rev xs) (rev ys) = e # es \<and> list_diff (rev ys) (rev xs) = a # as \<and> a \<noteq> e"
+ by blast
+ next
+ from he ha hn
+ show "xs \<asymp> ys"
+ by (auto simp:djoin_def no_junior_def, fold list_diff_le, simp+)
+ qed
+qed
+
+lemma djoin_diff_rev_iff: "(xs \<asymp> ys) = (\<exists> e es a as. xs \<setminus> ys = es@[e] \<and> ys \<setminus> xs = as@[a] \<and> a \<noteq> e)"
+ apply (auto simp:djoin_diff_iff list_diff_rev_def)
+ apply (rule_tac x = e in exI, safe)
+ apply (rule_tac x = "rev es" in exI)
+ apply (rule_tac injD[where f = rev], simp+)
+ apply (rule_tac x = "a" in exI, safe)
+ apply (rule_tac x = "rev as" in exI)
+ apply (rule_tac injD[where f = rev], simp+)
+ done
+
+lemma djoin_revE: "\<lbrakk>xs \<asymp> ys; \<And>e es a as. \<lbrakk>xs \<setminus> ys = es@[e]; ys \<setminus> xs = as@[a]; a \<noteq> e\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+ by (unfold djoin_diff_rev_iff, blast)
+
+lemma djoin_append_left[simp, intro]: "xs \<asymp> ys \<Longrightarrow> (zs' @ xs) \<asymp> (zs @ ys)"
+ by (auto simp:djoin_diff_iff intro:list_diff_djoin[simplified])
+
+lemma djoin_cons_left[simp]: "xs \<asymp> ys \<Longrightarrow> (e # xs) \<asymp> (a # ys)"
+ by (drule_tac zs' = "[e]" and zs = "[a]" in djoin_append_left, simp)
+
+lemma djoin_simp_1 [simp]: "xs \<asymp> ys \<Longrightarrow> xs \<asymp> (zs @ ys)"
+ by (drule_tac djoin_append_left [where zs' = "[]"], simp)
+
+lemma djoin_simp_2 [simp]: "xs \<asymp> ys \<Longrightarrow> (zs' @ xs) \<asymp> ys"
+ by (drule_tac djoin_append_left [where zs = "[]"], simp)
+
+lemma djoin_append_right[simp]: "xs \<asymp> ys \<Longrightarrow> (xs @ zs) \<asymp> (ys @ zs)"
+ by (simp add:djoin_diff_iff)
+
+lemma djoin_cons_append[simp]: "xs \<asymp> ys \<Longrightarrow> (x # xs) \<asymp> (zs @ ys)"
+ by (subgoal_tac "[x] @ xs \<asymp> zs @ ys", simp, blast)
+
+lemma djoin_append_cons[simp]: "xs \<asymp> ys \<Longrightarrow> (zs @ xs) \<asymp> (y # ys)"
+ by (subgoal_tac "zs @ xs \<asymp> [y] @ ys", simp, blast)
+
+lemma djoin_neq [simp]: "xs \<asymp> ys \<Longrightarrow> xs \<noteq> ys"
+ by (simp only:djoin_diff_iff, clarsimp)
+
+lemma djoin_cons [simp]: "e \<noteq> a \<Longrightarrow> e # xs \<asymp> a # xs"
+ by (unfold djoin_diff_iff, simp)
+
+lemma djoin_append_e [simp]: "e \<noteq> a \<Longrightarrow> (zs @ [e] @ xs) \<asymp> (zs' @ [a] @ xs)"
+ by (unfold djoin_diff_iff, simp)
+
+lemma djoin_mono [simp]: "\<lbrakk>xs \<asymp> ys; xs \<preceq> xs'; ys \<preceq> ys'\<rbrakk> \<Longrightarrow> xs' \<asymp> ys'"
+proof(erule_tac djoin_revE,unfold djoin_diff_rev_iff)
+ fix e es a as
+ assume hx: "xs \<preceq> xs'"
+ and hy: "ys \<preceq> ys'"
+ and hmx: "xs \<setminus> ys = es @ [e]"
+ and hmy: "ys \<setminus> xs = as @ [a]"
+ and neq: "a \<noteq> e"
+ have "xs' \<setminus> ys' = ((xs' \<setminus> xs) @ es) @ [e] \<and> ys' \<setminus> xs' = ((ys' \<setminus> ys) @ as) @ [a] \<and> a \<noteq> e"
+ proof -
+ from hx have heqx: "(xs' \<setminus> xs) @ xs = xs'"
+ by (cut_tac list_com_diff_rev [of xs' xs], subgoal_tac "xs' \<bullet> xs = xs", simp+)
+ moreover from hy have heqy: "(ys' \<setminus> ys) @ ys = ys'"
+ by (cut_tac list_com_diff_rev [of ys' ys], subgoal_tac "ys' \<bullet> ys = ys", simp+)
+ moreover from list_diff_djoin_rev_simplified [OF hmx hmy]
+ have "((xs' \<setminus> xs) @ xs) \<setminus> ((ys' \<setminus> ys) @ ys) = (xs' \<setminus> xs) @ es @ [e]" by simp
+ moreover from list_diff_djoin_rev_simplified [OF hmy hmx]
+ have "((ys' \<setminus> ys) @ ys) \<setminus> ((xs' \<setminus> xs) @ xs) = (ys' \<setminus> ys) @ as @ [a]" by simp
+ ultimately show ?thesis by (simp add:neq)
+ qed
+ thus "\<exists>e es a as. xs' \<setminus> ys' = es @ [e] \<and> ys' \<setminus> xs' = as @ [a] \<and> a \<noteq> e" by blast
+qed
+
+lemmas djoin_append_e_simplified [simp] = djoin_append_e [simplified]
+
+(*<*)
+end
+(*>*)
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/New_obj_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,86 @@
+theory New_obj_prop
+imports Main Finite_current Flask_type Flask Static
+begin
+
+context tainting_s begin
+
+lemma nn_notin_aux: "finite s \<Longrightarrow> \<forall> a \<in> s. Max s \<ge> a "
+apply (erule finite.induct, simp)
+apply (rule ballI)
+apply (case_tac "aa = a", simp+)
+done
+
+lemma nn_notin: "finite s \<Longrightarrow> next_nat s \<notin> s"
+apply (drule nn_notin_aux)
+apply (simp add:next_nat_def)
+by (auto)
+
+
+(* lemmas of new created obj *)
+
+lemma np_notin_curp: "valid \<tau> \<Longrightarrow> new_proc \<tau> \<notin> current_procs \<tau>" using finite_cp
+by (simp add:new_proc_def nn_notin)
+
+lemma np_notin_curp': "\<lbrakk>new_proc \<tau> \<in> current_procs \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+apply (drule np_notin_curp, simp)
+done
+
+lemma ni_notin_curi: "valid \<tau> \<Longrightarrow> new_inode_num \<tau> \<notin> current_inode_nums \<tau>"
+apply (drule finite_ci)
+by (simp add:new_inode_num_def nn_notin)
+
+lemma ni_notin_curi': "\<lbrakk>new_inode_num \<tau> \<in> current_inode_nums \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+by (drule ni_notin_curi, simp)
+
+lemma nm_notin_curm: "valid \<tau> \<Longrightarrow> new_msgq \<tau> \<notin> current_msgqs \<tau>" using finite_cm
+by (simp add:new_msgq_def nn_notin)
+
+lemma nm_notin_curm': "\<lbrakk>new_msgq \<tau> \<in> current_msgqs \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+by (drule nm_notin_curm, simp)
+
+lemma nfd_notin_curfd: "valid \<tau> \<Longrightarrow> new_proc_fd \<tau> p \<notin> current_proc_fds \<tau> p"
+using finite_cfd[where p = p]
+apply (simp add:new_proc_fd_def nn_notin)
+done
+
+lemma nfd_notin_curfd': "\<lbrakk>new_proc_fd \<tau> p \<in> current_proc_fds \<tau> p; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+by (drule nfd_notin_curfd[where p = p], simp)
+
+lemma nim_notin_curim: "valid \<tau> \<Longrightarrow> new_inode_num \<tau> \<notin> current_inode_nums \<tau>"
+by (drule finite_ci, simp add:new_inode_num_def nn_notin)
+
+lemma nim_notin_curim': "\<lbrakk>new_inode_num \<tau> \<in> current_inode_nums \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+by (drule nim_notin_curim, simp)
+
+lemma len_fname_all: "length (fname_all_a len) = len"
+by (induct len, auto simp:fname_all_a.simps)
+
+lemma ncf_notin_curf: "valid \<tau> \<Longrightarrow> new_childf f \<tau> \<notin> current_files \<tau>"
+apply (drule finite_cf)
+apply (simp add:new_childf_def next_fname_def all_fname_under_dir_def)
+apply (rule notI)
+apply (subgoal_tac "(CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files \<tau>}))) \<in> {fn. fn # f \<in> current_files \<tau>}")
+defer apply simp
+apply (subgoal_tac "length (CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files \<tau>}))) \<in> fname_length_set {fn. fn # f \<in> current_files \<tau>}")
+defer apply (auto simp:fname_length_set_def image_def)[1]
+apply (subgoal_tac "finite (fname_length_set {fn. fn # f \<in> current_files \<tau>})")
+defer
+apply (simp add:fname_length_set_def)
+apply (rule finite_imageI)
+apply (drule_tac h = "\<lambda> f'. case f' of [] \<Rightarrow> '''' | fn # pf' \<Rightarrow> if (pf' = f) then fn else ''''" in finite_imageI)
+apply (rule_tac B = "(list_case [] (\<lambda>fn pf'. if pf' = f then fn else []) ` current_files \<tau>)" in finite_subset)
+unfolding image_def
+apply (clarsimp split:if_splits)
+apply (rule_tac x = "x # f" in bexI, simp+)
+
+apply (drule_tac s = "(fname_length_set {fn. fn # f \<in> current_files \<tau>})" in nn_notin_aux)
+apply (erule_tac x = "length (CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files \<tau>})))" in ballE)
+apply (simp add:len_fname_all, simp)
+done
+
+lemma ncf_parent: "valid \<tau> \<Longrightarrow> parent (new_childf f \<tau>) = Some f"
+by (simp add:new_childf_def)
+
+end
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/OS_type_def.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,270 @@
+theory OS_type_def
+imports Main Flask_type
+begin
+
+(****************************** os files ***************************)
+
+datatype t_open_must_flag =
+ OF_RDONLY
+| OF_WRONLY
+| OF_RDWR
+
+datatype t_open_option_flag =
+ OF_APPEND
+| OF_CREAT
+(*
+| OF_TRUNC
+*)
+| OF_EXCL
+(*
+| OF_DIRECTORY
+
+| OF_OTHER
+*)
+
+type_synonym t_open_flags = "t_open_must_flag \<times> t_open_option_flag set"
+
+definition is_read_flag :: "t_open_flags \<Rightarrow> bool"
+where
+ "is_read_flag flag \<equiv> let (mflag, oflag) = flag in (mflag = OF_RDONLY \<or> mflag = OF_RDWR)"
+
+definition is_write_flag :: "t_open_flags \<Rightarrow> bool"
+where
+ "is_write_flag flag \<equiv> let (mflag, oflag) = flag in (mflag = OF_WRONLY \<or> mflag = OF_RDWR)"
+
+definition is_excl_flag :: "t_open_flags \<Rightarrow> bool"
+where
+ "is_excl_flag flag \<equiv> let (mflag, oflag) = flag in (OF_EXCL \<in> oflag)"
+
+definition is_creat_flag :: "t_open_flags \<Rightarrow> bool"
+where
+ "is_creat_flag flag \<equiv> let (mflag, oflag) = flag in (OF_CREAT \<in> oflag)"
+
+definition is_creat_excl_flag :: "t_open_flags \<Rightarrow> bool"
+where
+ "is_creat_excl_flag flag \<equiv> let (mflag, oflag) = flag in (OF_CREAT \<in> oflag \<and> OF_EXCL \<in> oflag)"
+(*
+definition is_trunc_flag :: "t_open_flags \<Rightarrow> bool"
+where
+ "is_trunc_flag flag \<equiv> let (mflag, oflag) = flag in (OF_TRUNC \<in> oflag \<and> is_write_flag flag)"
+
+definition is_cloexec_flag :: "t_open_flags \<Rightarrow> bool"
+where
+ "is_cloexec_flag flag \<equiv> let (mflag, oflag) = flag in (OF_CLOEXEC \<in> oflag)"
+*)
+definition is_append_flag :: "t_open_flags \<Rightarrow> bool"
+where
+ "is_append_flag flag \<equiv> let (mflag, oflag) = flag in (OF_APPEND \<in> oflag)"
+
+(*
+definition is_dir_flag :: "t_open_flags \<Rightarrow> bool"
+where
+ "is_dir_flag flag \<equiv> let (mflag, oflag) = flag in (OF_DIRECTORY \<in> oflag)"
+*)
+
+
+type_synonym t_fname = string
+
+(* listing of file path:
+ * a#b#c:: b is a's parent and c is b's parent;
+ * [] represents the root file
+ *)
+type_synonym t_file = "t_fname list"
+
+type_synonym t_fd = "nat" (* given a fd, must has funs to tell its corresponding: "t_file \<times> t_open_flags \<times> t_inode" *)
+
+
+
+(****************************** os sockets ***************************)
+
+datatype t_shutdown_how =
+ SHUT_RD
+| SHUT_WR
+| SHUT_RDWR
+
+datatype t_zero_one =
+ Z
+| O
+
+type_synonym t_inet_addr = "t_zero_one list"
+
+type_synonym t_socket_port = nat
+
+type_synonym t_netif = string
+
+(*
+datatype t_socket_af =
+ AF_UNIX "t_file \<times> t_file"
+| AF_INET "t_inet_addr \<times> t_socket_port \<times> t_inet_addr \<times> t_socket_port \<times> t_netdev"
+*)
+
+datatype t_socket_af =
+(* AF_UNIX
+| *)AF_INET
+
+datatype t_socket_type =
+ STREAM
+| DGRAM
+
+datatype t_sock_addr =
+ INET_ADDR t_inet_addr t_socket_port(*
+| UNIX_ADDR t_file*)
+
+type_synonym t_node = t_inet_addr (* a endpoint in the netword *)
+
+datatype t_sock_trans_state =
+ Trans_Recv
+| Trans_Send
+| Trans_RS
+| Trans_No
+
+datatype t_sock_state =
+ SS_create
+| SS_bind
+| SS_connect
+| SS_listen
+| SS_accepted
+| SS_trans "t_sock_trans_state"
+
+
+(****************************** os other IPCs ***************************)
+(*
+datatype t_shm_attach_flag =
+ SHM_RDONLY
+| SHM_RDWR
+
+type_synonym t_shm = nat
+*)
+
+(* message queue is indexed by a nat, in OS
+ * but logically it is a list, created as [], updated as a queue
+ * so, it type should not be nat, but a t_msg list
+ *)
+type_synonym t_msgq = nat
+
+(* a message of ID nat, but it has meaning under "in what message queue" *)
+type_synonym t_msg = "nat"
+
+(****************************** os inodes ***************************)
+
+type_synonym t_inode_num = nat
+
+(*
+datatype t_sock_af =
+ Tag_AF_UNIX "(t_file option \<times> t_file option)" "t_socket_type"
+| Tag_AF_INET "(t_inet_addr \<times> t_socket_port) option \<times> (t_inet_addr \<times> t_socket_port) option \<times> t_netdev option" "t_socket_type"
+*)
+
+datatype t_inode_tag =
+ Tag_FILE (* this inode is a file, and its file is: "t_file" :: no need to do that and another reason is link may add two or more files to the inode, so it should be a list, can the maintainment of this list is tough *)
+| Tag_DIR (* do not contain chs anymore, for more infomation see svn log
+ "t_inode_num list" *) (* Tag_DIR inode number list: the inode number list of subchild files (direct son the file) , means: new created child file should add to this field. *)
+| Tag_TCP_SOCK (* information below is turn to seperative listener functions
+"(t_inet_addr \<times> t_socket_port) option \<times> (t_inet_addr \<times> t_socket_port) option \<times> t_netdev option" "t_socket_type" *)
+| Tag_UDP_SOCK
+
+(* type_synonym t_inode = "nat \<times> t_inode_tag" (* inode (id, tag) contains file internal id and directory tag attributes *)
+ // not needed any more: for inode num and its tag now are seperated by two funcs.
+*)
+
+
+type_synonym t_process = nat
+
+type_synonym t_trunc_len = nat
+
+type_synonym t_socket = "t_process \<times> t_fd"
+
+(* each constructor is defined according to each security-class *)
+datatype t_object =
+ O_proc "t_process"
+
+| O_file "t_file"
+| O_fd "t_process" "t_fd"
+| O_dir "t_file"
+
+| O_node "t_node"
+| O_tcp_sock "t_socket"
+| O_udp_sock "t_socket"
+(*
+| O_shm "t_shm"
+*)
+| O_msgq "t_msgq"
+| O_msg "t_msgq" "t_msg"
+
+(*
+datatype t_tainted_obj =
+ TO_proc "t_process"
+| TO_file "t_file"
+| TO_msg "t_msgq" "t_msg"
+*)
+
+
+(* not needed anymore
+datatype t_clone_share_flag =
+| CF_share_ipc
+| CF_share_fd
+| CF_share_both
+| CF_no_share *)
+
+(* unformalised system calls:
+ * 1. process management
+ * 1.1 wait: nothing to do with tainting relation
+ * 1.2 getpriority/setpriority/getsched/setsched: nothing to do with security
+ * 1.3 getcap/setcap: "at this time that will not be done"
+ * 1.4 getpgid/setpgid: ...
+ *)
+datatype t_event =
+ (* process management *)
+(* trace can be done trough tags of cloning and execving, but I ignore this. *)
+ Execve "t_process" "t_file" "t_fd set"
+(* | Uselib "t_process" "t_file" *)
+| Clone "t_process" "t_process" "t_fd set" (*"t_shm set"*) (* "t_clone_share_flag" is not needed, as event should record the fds and shms that passed to childprocess *)
+| Kill "t_process" "t_process"
+| Ptrace "t_process" "t_process"
+| Exit "t_process"
+
+| Open "t_process" "t_file" "t_open_flags" "t_fd" "nat option" (* the last nat option : if create file, the is Some inode_num else None *)
+| ReadFile "t_process" "t_fd"
+| WriteFile "t_process" "t_fd"
+(* | Sendfile "t_process" "t_fd" "t_fd"
+ can be modeled as readfile then writefile *)
+| CloseFd "t_process" "t_fd"
+| UnLink "t_process" "t_file"
+| Rmdir "t_process" "t_file"
+| Mkdir "t_process" "t_file" "nat" (* last nat: inode_num *)
+(*
+| LinkHard "t_process" "t_file" "t_file"
+*)
+| Truncate "t_process" "t_file" "t_trunc_len"
+(* | FTruncate "t_process" "t_fd" "t_trunc_len" *)
+(*
+| Rename "t_process" "t_file" "t_file"
+*)
+
+(* the message is in the msgq, so we must build a cache(list)
+ * to store the unreceived msg yet. And the os_grant should
+ * check if a message queue is empty before RecvMsg.
+ *)
+| CreateMsgq "t_process" "t_msgq"
+| SendMsg "t_process" "t_msgq" "t_msg"
+| RecvMsg "t_process" "t_msgq" "t_msg"
+| RemoveMsgq "t_process" "t_msgq"
+(*
+| CreateShM "t_process" "t_shm"
+| Attach "t_process" "t_shm" "t_shm_attach_flag"
+| Detach "t_process" "t_shm"
+| DeleteShM "t_process" "t_shm"
+*)
+| CreateSock "t_process" "t_socket_af" "t_socket_type" "t_fd" "nat" (* last nat: inode_num *)
+| Bind "t_process" "t_fd" "t_sock_addr"
+| Connect "t_process" "t_fd" "t_sock_addr"
+| Listen "t_process" "t_fd"
+| Accept "t_process" "t_fd" "t_sock_addr" "nat" "t_fd" "nat" (* addr: remote socket address; nat: port num; last nat: inode_num *)
+| SendSock "t_process" "t_fd"
+| RecvSock "t_process" "t_fd"
+| Shutdown "t_process" "t_fd" "t_shutdown_how"
+
+(* the state of the system is decided by the initial state and the event trace*)
+type_synonym t_state = "t_event list"
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Proc_fd_of_file_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,155 @@
+theory Proc_fd_of_file_prop
+imports Main Flask Flask_type Valid_prop Current_files_prop Current_sockets_prop
+begin
+
+context flask begin
+
+lemma proc_fd_in_procs: "\<lbrakk>file_of_proc_fd \<tau> p fd = Some f; valid \<tau>\<rbrakk> \<Longrightarrow> p \<in> current_procs \<tau>"
+apply (induct \<tau> arbitrary: f) defer
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:file_of_proc_fd.simps current_procs.simps os_grant.simps split:if_splits option.splits)
+by (drule init_filefd_valid, simp)
+
+lemma proc_fd_in_fds_aux: "\<forall> p f. file_of_proc_fd \<tau> p fd = Some f \<and> valid \<tau> \<longrightarrow> fd \<in> current_proc_fds \<tau> p"
+apply (induct \<tau>)
+apply (simp add:file_of_proc_fd.simps current_proc_fds.simps)
+apply (clarify, drule init_filefd_valid, simp)
+apply (clarify, frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:file_of_proc_fd.simps current_proc_fds.simps split:if_splits option.splits t_sock_addr.splits)
+done
+
+lemma proc_fd_in_fds: "\<lbrakk>file_of_proc_fd \<tau> p fd = Some f; valid \<tau>\<rbrakk> \<Longrightarrow> fd \<in> current_proc_fds \<tau> p"
+by (rule proc_fd_in_fds_aux[rule_format], simp+)
+
+lemma proc_fd_file_in_cur: "\<lbrakk>(p, fd) \<in> proc_fd_of_file \<tau> f; valid \<tau>\<rbrakk> \<Longrightarrow> f \<in> current_files \<tau>"
+by (auto simp:proc_fd_of_file_def intro:file_of_pfd_in_current)
+
+lemma proc_fd_file_in_cur': "\<lbrakk>proc_fd_of_file \<tau> f \<noteq> {}; valid \<tau>\<rbrakk> \<Longrightarrow> f \<in> current_files \<tau>"
+by (auto simp:proc_fd_file_in_cur)
+
+lemma proc_fd_file_in_cur'': "\<lbrakk>proc_fd_of_file \<tau> f = {(p,fd)}; valid \<tau>\<rbrakk> \<Longrightarrow> f \<in> current_files \<tau>"
+by (auto simp:proc_fd_file_in_cur')
+
+lemma procfd_of_file_imp_fpfd: "proc_fd_of_file \<tau> f = {(p, fd)} \<Longrightarrow> file_of_proc_fd \<tau> p fd = Some f"
+by (auto simp:proc_fd_of_file_def)
+
+lemma procfd_of_file_imp_fpfd': "proc_fd_of_file \<tau> f = {(p, fd)} \<Longrightarrow> file_of_proc_fd \<tau> p fd \<noteq> None"
+by (auto simp:proc_fd_of_file_def)
+
+lemma procfd_of_file_eq_fpfd'': "(p, fd) \<in> proc_fd_of_file \<tau> f = (file_of_proc_fd \<tau> p fd = Some f)"
+by (auto simp:proc_fd_of_file_def)
+
+lemma procfd_of_file_non_empty: "file_of_proc_fd \<tau> p fd = Some f \<Longrightarrow> proc_fd_of_file \<tau> f \<noteq> {}"
+by (auto simp:proc_fd_of_file_def)
+
+lemma file_of_proc_fd_in_curf: "\<lbrakk>file_of_proc_fd \<tau> p fd = Some f; valid \<tau>\<rbrakk> \<Longrightarrow> f \<in> current_files \<tau>"
+by (drule procfd_of_file_non_empty, simp add:proc_fd_file_in_cur')
+
+lemma file_fds_subset_pfds:
+ "valid s \<Longrightarrow> proc_file_fds s p \<subseteq> current_proc_fds s p"
+by (auto simp add:proc_file_fds_def intro:proc_fd_in_fds)
+
+lemma filefd_socket_conflict:
+ "\<lbrakk>file_of_proc_fd s p fd = Some f; (p, fd) \<in> current_sockets s; valid s\<rbrakk> \<Longrightarrow> False"
+apply (induct s arbitrary:p)
+apply (simp add:current_sockets_simps init_filefd_prop8)
+apply (frule vt_grant_os, frule vd_cons, frule file_fds_subset_pfds, case_tac a)
+apply (auto simp:current_sockets_simps split:if_splits option.splits
+ dest:cn_in_curp cn_in_curfd proc_fd_in_fds)
+done
+
+lemma is_tcp_in_current: "is_tcp_sock s sock \<Longrightarrow> sock \<in> current_sockets s"
+by (auto simp:is_tcp_sock_def current_sockets_def split:option.splits)
+
+lemma is_udp_in_current: "is_udp_sock s sock \<Longrightarrow> sock \<in> current_sockets s"
+by (auto simp:is_udp_sock_def current_sockets_def split:option.splits)
+
+lemma tcp_not_file_fd:
+ "\<lbrakk>is_tcp_sock s (p, fd); valid s\<rbrakk> \<Longrightarrow> file_of_proc_fd s p fd = None"
+apply (case_tac "file_of_proc_fd s p fd", simp)
+apply (drule is_tcp_in_current)
+apply (drule filefd_socket_conflict, simp+)
+done
+
+lemma udp_not_file_fd:
+ "\<lbrakk>is_udp_sock s (p, fd); valid s\<rbrakk> \<Longrightarrow> file_of_proc_fd s p fd = None"
+apply (case_tac "file_of_proc_fd s p fd", simp)
+apply (drule is_udp_in_current)
+apply (drule filefd_socket_conflict, simp+)
+done
+
+lemma tcp_notin_file_fds:
+ "\<lbrakk>is_tcp_sock s (p, fd); valid s\<rbrakk> \<Longrightarrow> fd \<notin> proc_file_fds s p"
+by (auto simp:proc_file_fds_def intro:tcp_not_file_fd)
+
+lemma udp_notin_file_fds:
+ "\<lbrakk>is_udp_sock s (p, fd); valid s\<rbrakk> \<Longrightarrow> fd \<notin> proc_file_fds s p"
+by (auto simp:proc_file_fds_def intro:udp_not_file_fd)
+
+(******************* rebuild proc_fd_of_file simpset ***********************)
+(*
+lemma proc_fd_of_file_open: "Open p f flags fd iopt # valid \<tau> \<Longrightarrow>
+ proc_fd_of_file (Open p f flags fd iopt # \<tau>) f' = (if (f' = f) then insert (p, fd) (proc_fd_of_file \<tau> f') else proc_fd_of_file \<tau> f')"
+apply (auto simp:proc_fd_of_file_def file_of_proc_fd.simps split:if_splits)
+apply (frule vd_cons, drule vt_grant_os, case_tac iopt)
+apply (drule proc_fd_in_fds, simp, simp add:os_grant.simps nfd_notin_curfd)+
+done
+
+lemma proc_fd_of_file_closefd: "proc_fd_of_file (CloseFd p fd # \<tau>) f = (if (file_of_proc_fd \<tau> p fd = Some f) then (proc_fd_of_file \<tau> f - {(p,fd)}) else proc_fd_of_file \<tau> f) "
+by (auto simp:proc_fd_of_file_def file_of_proc_fd.simps split:if_splits)
+
+lemma proc_fd_of_file_rename: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # valid \<tau>; f \<in> current_files (Rename p f\<^isub>2 f\<^isub>3 # \<tau>)\<rbrakk> \<Longrightarrow>
+ proc_fd_of_file (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = (if (f\<^isub>3 \<preceq> f) then proc_fd_of_file \<tau> (file_before_rename f\<^isub>2 f\<^isub>3 f) else proc_fd_of_file \<tau> f)"
+apply (frule vt_grant_os, frule vd_cons)
+apply (case_tac "f\<^isub>3 \<preceq> f")
+apply (subgoal_tac "f \<notin> current_files \<tau>") prefer 2 apply (rule notI)
+apply (clarsimp simp:os_grant.simps, drule_tac f = f\<^isub>3 and f' = f in ancenf_in_current, simp, simp, simp)
+apply (auto simp add:proc_fd_of_file_def)[1]
+
+apply (simp add:file_of_proc_fd.simps split:option.splits if_splits)
+apply (drule_tac f\<^isub>3 = f\<^isub>3 and f\<^isub>1 = aa and f\<^isub>2 = f\<^isub>2 in file_renaming_prop5, simp)
+apply (drule file_of_pfd_in_current, simp+)
+apply (simp add:file_of_proc_fd.simps)
+apply (rule conjI, rule impI, simp add:file_renaming_prop5')
+apply (rule impI, simp add:file_before_rename_def)
+
+apply (simp add:proc_fd_of_file_def split:if_splits)
+apply auto
+apply (simp add:file_of_proc_fd.simps split:option.splits if_splits)
+apply (drule_tac f\<^isub>3 = f\<^isub>3 and f\<^isub>2 = f\<^isub>2 and f = aa in file_renaming_prop1, simp)
+apply (simp add:current_files_simps)
+apply (erule exE| erule conjE)+
+apply (simp add:file_of_proc_fd.simps split:option.splits if_splits)
+apply (drule_tac f = f\<^isub>1 in rename_renaming_decom', simp+)
+apply (simp add:file_after_rename_def)
+done
+
+
+lemma proc_fd_of_file_kill: "proc_fd_of_file (Kill p\<^isub>1 p\<^isub>2 # \<tau>) f = {(p, fd). (p, fd) \<in> proc_fd_of_file \<tau> f \<and> p \<noteq> p\<^isub>2}"
+by (auto simp:proc_fd_of_file_def file_of_proc_fd.simps)
+
+lemma proc_fd_of_file_exit: "proc_fd_of_file (Exit p' # \<tau>) f = {(p, fd). (p, fd) \<in> proc_fd_of_file \<tau> f \<and> p \<noteq> p'}"
+by (auto simp:proc_fd_of_file_def file_of_proc_fd.simps)
+
+lemma proc_fd_of_file_clone: "Clone p\<^isub>1 p\<^isub>2 # valid \<tau> \<Longrightarrow> proc_fd_of_file (Clone p\<^isub>1 p\<^isub>2 # \<tau>) f = proc_fd_of_file \<tau> f \<union> {(p\<^isub>2, fd)| fd. (p\<^isub>1, fd) \<in> proc_fd_of_file \<tau> f}"
+apply (auto simp:proc_fd_of_file_def file_of_proc_fd.simps)
+apply (frule vd_cons, drule vt_grant_os)
+apply (drule proc_fd_in_procs, (simp add:os_grant.simps np_notin_curp)+)
+done
+
+lemma proc_fd_of_file_other: "\<lbrakk>e # valid \<tau>;
+ \<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ \<forall> p fd. e \<noteq> CloseFd p fd;
+ \<forall> p f f'. e \<noteq> Rename p f f';
+ \<forall> p p'. e \<noteq> Kill p p';
+ \<forall> p. e \<noteq> Exit p;
+ \<forall> p p'. e \<noteq> Clone p p'\<rbrakk> \<Longrightarrow> proc_fd_of_file (e # \<tau>) f = proc_fd_of_file \<tau> f"
+apply (case_tac e, auto simp:proc_fd_of_file_def file_of_proc_fd.simps)
+done
+
+lemmas proc_fd_of_file_simps = proc_fd_of_file_open proc_fd_of_file_closefd proc_fd_of_file_rename proc_fd_of_file_kill proc_fd_of_file_exit proc_fd_of_file_clone proc_fd_of_file_other
+*)
+end
+
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/ROOT Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,38 @@
+session "dynamic" = "HOL" +
+ options [document = false]
+ theories
+ List_Prefix
+ My_list_prefix
+ File_renaming
+ Flask_type
+ OS_type_def
+ Flask
+
+
+session "alive" = "dynamic" +
+ options [document = false]
+ theories
+ Static_type
+ Static
+ Valid_prop
+ Init_prop
+ Current_files_prop
+ Current_sockets_prop
+ Alive_prop
+ Proc_fd_of_file_prop
+ Finite_current
+ New_obj_prop
+
+session "co2sobj" = "alive" +
+ options [document = false]
+ theories
+ Sectxt_prop
+ Delete_prop
+ Tainted_prop
+ Co2sobj_prop
+
+session "s2ss" = "co2sobj" +
+ options [document = false]
+ theories
+ S2ss_prop
+ S2ss_prop2
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/S2ss_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,2555 @@
+(*<*)
+theory S2ss_prop
+imports Main Flask Flask_type Static Static_type Init_prop Tainted_prop Valid_prop Alive_prop Co2sobj_prop
+begin
+(*>*)
+
+context tainting_s begin
+
+(* simpset for s2ss*)
+
+lemma co2sobj_some_case:
+ "\<lbrakk>co2sobj s obj = Some sobj; \<And> p. obj = O_proc p \<Longrightarrow> P (O_proc p);
+ \<And> f. obj = O_file f \<Longrightarrow> P (O_file f); \<And> h. obj = O_shm h \<Longrightarrow> P (O_shm h);
+ \<And> f. obj = O_dir f \<Longrightarrow> P (O_dir f); \<And> q. obj = O_msgq q \<Longrightarrow> P (O_msgq q)\<rbrakk>
+ \<Longrightarrow> P obj"
+by (case_tac obj, auto)
+
+lemma co2sobj_execve_alive:
+ "\<lbrakk>alive s obj; co2sobj s obj = Some x; valid (Execve p f fds # s)\<rbrakk>
+ \<Longrightarrow> alive (Execve p f fds # s) obj"
+apply (erule co2sobj_some_case)
+by (auto simp:alive_simps simp del:alive.simps)
+
+lemma s2ss_execve:
+ "valid (Execve p f fds # s) \<Longrightarrow> s2ss (Execve p f fds # s) = (
+ if (\<exists> p'. p' \<noteq> p \<and> p' \<in> current_procs s \<and> co2sobj s (O_proc p') = co2sobj s (O_proc p))
+ then (case (cp2sproc (Execve p f fds # s) p) of
+ Some sp \<Rightarrow> s2ss s \<union> {S_proc sp (O_proc p \<in> Tainted s \<or> O_file f \<in> Tainted s)}
+ | _ \<Rightarrow> {} )
+ else (case (cp2sproc (Execve p f fds # s) p, cp2sproc s p) of
+ (Some sp, Some sp') \<Rightarrow> s2ss s - {S_proc sp' (O_proc p \<in> Tainted s)}
+ \<union> {S_proc sp (O_proc p \<in> Tainted s \<or> O_file f \<in> Tainted s)}
+ | _ \<Rightarrow> {} ) )"
+apply (frule vd_cons, frule vt_grant_os, simp split:if_splits)
+
+apply (rule conjI, rule impI, (erule exE|erule conjE)+)
+apply (frule_tac p = p in current_proc_has_sp, simp, erule exE)
+apply (frule_tac p = p' in current_proc_has_sp, simp, erule exE, simp)
+apply (subgoal_tac "p \<in> current_procs (Execve p f fds # s)")
+apply (drule_tac p = p and s = "Execve p f fds # s" in current_proc_has_sp, simp)
+apply (erule exE, simp)
+apply (simp add:s2ss_def, rule set_eqI, rule iffI)
+apply (drule CollectD, (erule exE|erule conjE)+)
+apply (case_tac "obj = O_proc p")
+apply (simp add:co2sobj_execve tainted_eq_Tainted split:if_splits)
+apply (simp add:co2sobj_execve, rule disjI2)
+apply (rule_tac x = obj in exI, case_tac obj, (simp add:alive_simps)+)[1]
+apply (simp, erule disjE)
+apply (rule_tac x = "O_proc p" in exI, simp add:tainted_eq_Tainted)
+apply (erule exE| erule conjE)+
+apply (case_tac "x = S_proc sp (O_proc p \<in> tainted s)")
+apply (rule_tac x = "O_proc p'" in exI)
+apply (simp add:alive_execve co2sobj_execve tainted_eq_Tainted cp2sproc_execve)
+apply (case_tac "obj = O_proc p", simp)
+apply (frule co2sobj_execve_alive, simp, simp)
+apply (frule_tac obj = obj in co2sobj_execve, simp)
+apply (rule_tac x = obj in exI, simp, simp)
+
+apply (erule conjE, frule current_proc_has_sp, simp, erule exE, rule impI, simp)
+apply (subgoal_tac "p \<in> current_procs (Execve p f fds # s)")
+apply (drule_tac p = p and s = "Execve p f fds # s" in current_proc_has_sp, simp)
+apply (erule exE, erule conjE, simp)
+apply (simp add:s2ss_def, rule set_eqI, rule iffI)
+apply (drule CollectD, (erule exE|erule conjE)+)
+apply (case_tac "obj = O_proc p", simp add:tainted_eq_Tainted)
+apply (rule disjI1, simp split:if_splits)
+apply (simp add:co2sobj_execve, rule disjI2)
+apply (rule conjI,rule_tac x = obj in exI, simp add:alive_simps split:t_object.splits)
+apply (rule notI, simp, case_tac obj)
+apply (erule_tac x = nat in allE, simp add:tainted_eq_Tainted, (simp split:option.splits)+)
+apply (erule disjE, simp)
+apply (rule_tac x = "O_proc p" in exI, simp add:tainted_eq_Tainted)
+apply (erule exE|erule conjE)+
+apply (rule_tac x = obj in exI)
+apply (drule_tac obj = obj in co2sobj_execve_alive, simp+)
+apply (frule_tac obj = obj in co2sobj_execve, simp, simp)
+apply (rule impI, simp add:tainted_eq_Tainted, simp)
+done
+
+lemma s2ss_clone_alive:
+ "\<lbrakk>co2sobj s obj = Some x; alive s obj; obj \<noteq> O_proc p'; valid (Clone p p' fds shms # s)\<rbrakk>
+ \<Longrightarrow> alive (Clone p p' fds shms # s) obj"
+by (erule co2sobj_some_case, auto simp:alive_clone)
+
+lemma s2ss_clone:
+ "valid (Clone p p' fds shms # s) \<Longrightarrow> s2ss (Clone p p' fds shms # s) = (
+ case (cp2sproc (Clone p p' fds shms # s) p') of
+ Some sp \<Rightarrow> s2ss s \<union> {S_proc sp (O_proc p \<in> Tainted s)}
+ | _ \<Rightarrow> {})"
+apply (frule vd_cons, frule vt_grant_os, split option.splits)
+apply (rule conjI, rule impI, drule current_proc_has_sp', simp, simp)
+apply (rule allI, rule impI, simp add:s2ss_def)
+apply (rule set_eqI, rule iffI, drule CollectD, (erule exE|erule conjE)+)
+apply (case_tac "obj = O_proc p'", simp add:tainted_eq_Tainted)
+apply (case_tac "O_proc p' \<in> Tainted s", drule Tainted_in_current, simp+)
+apply (rule disjI1, simp split:if_splits)
+apply (simp add:tainted_eq_Tainted, rule disjI2)
+apply (frule co2sobj_clone, simp)
+apply (rule_tac x = obj in exI, simp add:alive_simps split:t_object.splits)
+
+apply (simp, erule disjE, simp)
+apply (rule_tac x = "O_proc p'" in exI, simp add:tainted_eq_Tainted)
+apply (rule impI, rule notI, drule Tainted_in_current, simp+)
+apply (erule exE| erule conjE)+
+apply (case_tac "obj = O_proc p'", simp)
+apply (rule_tac x = obj in exI)
+apply (frule s2ss_clone_alive, simp+)
+apply (auto simp:co2sobj_clone alive_simps)
+done
+
+definition s2ss_shm_no_backup:: "t_state \<Rightarrow> t_process \<Rightarrow> t_static_state"
+where
+ "s2ss_shm_no_backup s pfrom \<equiv> {S_proc sp False | sp p. info_flow_shm s pfrom p \<and> cp2sproc s p = Some sp \<and>
+ (\<not> (\<exists> p'. \<not> info_flow_shm s pfrom p' \<and> p' \<in> current_procs s \<and> co2sobj s (O_proc p') = Some (S_proc sp False)))}"
+
+definition update_s2ss_shm:: "t_state \<Rightarrow> t_process \<Rightarrow> t_static_state"
+where
+ "update_s2ss_shm s pfrom \<equiv> s2ss s
+ \<union> {S_proc sp True| sp p. info_flow_shm s pfrom p \<and> cp2sproc s p = Some sp}
+ - (s2ss_shm_no_backup s pfrom)"
+
+lemma s2ss_shm_no_bk_elim:
+ "\<lbrakk>S_proc sp False \<notin> s2ss_shm_no_backup s pfrom; co2sobj s (O_proc p) = Some (S_proc sp False);
+ valid s; info_flow_shm s pfrom p\<rbrakk>
+ \<Longrightarrow> \<exists> p'. \<not> info_flow_shm s pfrom p' \<and> p' \<in> current_procs s \<and> co2sobj s (O_proc p') = Some (S_proc sp False)"
+apply (auto simp:s2ss_shm_no_backup_def co2sobj.simps tainted_eq_Tainted split:option.splits)
+apply (erule_tac x = p in allE, auto)
+apply (rule_tac x = p' in exI, auto)
+done
+
+lemma s2ss_shm_no_bk_intro1:
+ "\<lbrakk>co2sobj s' obj = Some x; \<forall> p. obj \<noteq> O_proc p\<rbrakk> \<Longrightarrow> x \<notin> s2ss_shm_no_backup s pfrom"
+apply (case_tac obj)
+apply (auto simp:co2sobj.simps s2ss_shm_no_backup_def split:option.splits)
+done
+
+lemma s2ss_shm_no_bk_intro2:
+ "\<lbrakk>co2sobj s' obj = Some x; obj \<in> Tainted s'; valid s'\<rbrakk> \<Longrightarrow> x \<notin> s2ss_shm_no_backup s pfrom"
+apply (case_tac obj)
+
+apply (auto simp:co2sobj.simps s2ss_shm_no_backup_def tainted_eq_Tainted split:option.splits)
+done
+
+lemma s2ss_shm_no_bk_intro3:
+ "\<lbrakk>co2sobj s (O_proc p) = Some x; \<not> info_flow_shm s pfrom p; p \<in> current_procs s
+ \<rbrakk> \<Longrightarrow> x \<notin> s2ss_shm_no_backup s pfrom"
+apply (auto simp add:s2ss_shm_no_backup_def split:option.splits)
+apply (rule_tac x = p in exI, simp)
+done
+
+lemma s2ss_shm_no_bk_intro4:
+ "\<lbrakk>co2sobj s (O_proc p) = Some x; info_flow_shm s pfrom p;
+ \<not> info_flow_shm s pfrom p'; p' \<in> current_procs s; co2sobj s (O_proc p') = Some x\<rbrakk>
+ \<Longrightarrow> x \<notin> s2ss_shm_no_backup s pfrom"
+apply (rule notI)
+apply (auto simp add:s2ss_shm_no_backup_def co2sobj.simps split:option.splits)
+done
+
+lemma Tainted_ptrace':
+ "valid s \<Longrightarrow> Tainted (Ptrace p p' # s) =
+ (if (O_proc p \<in> Tainted s \<and> O_proc p' \<notin> Tainted s)
+ then Tainted s \<union> {O_proc p'' | p''. info_flow_shm s p' p''}
+ else if (O_proc p' \<in> Tainted s \<and> O_proc p \<notin> Tainted s)
+ then Tainted s \<union> {O_proc p'' | p''. info_flow_shm s p p''}
+ else Tainted s)"
+using info_flow_shm_Tainted by auto
+
+lemma co2sobj_some_caseD:
+ "\<lbrakk>co2sobj s obj = Some sobj; \<And> p. \<lbrakk>co2sobj s obj = Some sobj; obj = O_proc p\<rbrakk> \<Longrightarrow> P (O_proc p);
+ \<And> f. \<lbrakk>co2sobj s obj = Some sobj; obj = O_file f\<rbrakk> \<Longrightarrow> P (O_file f);
+ \<And> h. \<lbrakk>co2sobj s obj = Some sobj; obj = O_shm h\<rbrakk> \<Longrightarrow> P (O_shm h);
+ \<And> f. \<lbrakk>co2sobj s obj = Some sobj; obj = O_dir f\<rbrakk> \<Longrightarrow> P (O_dir f);
+ \<And> q. \<lbrakk>co2sobj s obj = Some sobj; obj = O_msgq q\<rbrakk> \<Longrightarrow> P (O_msgq q)\<rbrakk>
+ \<Longrightarrow> P obj"
+by (case_tac obj, auto)
+
+lemma s2ss_ptrace1_aux: "x \<notin> {x. P x} \<Longrightarrow> \<not> P x" by simp
+
+lemma s2ss_ptrace1:
+ "\<lbrakk>valid (Ptrace p p' # s); O_proc p \<in> Tainted s; O_proc p' \<notin> Tainted s\<rbrakk>
+ \<Longrightarrow> s2ss (Ptrace p p' # s) = update_s2ss_shm s p'"
+unfolding update_s2ss_shm_def s2ss_def s2ss_shm_no_backup_def
+apply (frule vd_cons, rule set_eqI, rule iffI)
+apply (drule CollectD, (erule exE|erule conjE)+)
+apply (erule co2sobj_some_caseD)
+apply (rule DiffI)
+apply (case_tac "O_proc pa \<in> Tainted s")
+apply (rule UnI1, simp, rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted split:option.splits)
+apply (case_tac "info_flow_shm s p' pa")
+apply (rule UnI2, rule CollectI, simp only:co2sobj.simps split:option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (rule_tac x = a in exI, rule_tac x = pa in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted split:option.splits)
+apply (rule UnI1, simp)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted split:option.splits)
+apply (rule notI, clarsimp simp:cp2sproc_other tainted_eq_Tainted split:option.splits)
+apply (erule_tac x = pa in allE, simp)
+
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (simp)
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (simp split:option.splits)
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (simp split:option.splits)
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (simp split:option.splits)
+
+apply (erule DiffE, drule s2ss_ptrace1_aux, erule UnE)
+apply (erule CollectE, (erule exE|erule conjE)+, rule CollectI)
+apply (erule co2sobj_some_caseD)
+apply (case_tac "O_proc pa \<in> Tainted s")
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp add:tainted_eq_Tainted cp2sproc_other split:option.splits)
+apply (case_tac "info_flow_shm s p' pa", simp only:co2sobj.simps split:option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (drule Meson.not_exD, erule_tac x = a in allE, drule Meson.not_exD, erule_tac x = pa in allE)
+apply (simp add:tainted_eq_Tainted, (erule exE|erule conjE)+)
+apply (rule_tac x = "O_proc p'a" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted split:option.splits)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted split:option.splits)
+
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+
+apply (erule CollectE, (erule exE|erule conjE)+, rule CollectI)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted info_shm_flow_in_procs)
+done
+
+lemma s2ss_ptrace2:
+ "\<lbrakk>valid (Ptrace p p' # s); O_proc p' \<in> Tainted s; O_proc p \<notin> Tainted s\<rbrakk>
+ \<Longrightarrow> s2ss (Ptrace p p' # s) = update_s2ss_shm s p"
+unfolding update_s2ss_shm_def s2ss_def s2ss_shm_no_backup_def
+apply (frule vd_cons, rule set_eqI, rule iffI)
+apply (drule CollectD, (erule exE|erule conjE)+)
+apply (erule co2sobj_some_caseD)
+apply (rule DiffI)
+apply (case_tac "O_proc pa \<in> Tainted s")
+apply (rule UnI1, simp, rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted split:option.splits)
+apply (case_tac "info_flow_shm s p pa")
+apply (rule UnI2, rule CollectI, simp only:co2sobj.simps split:option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (rule_tac x = a in exI, rule_tac x = pa in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted split:option.splits)
+apply (rule UnI1, simp)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted split:option.splits)
+apply (rule notI, clarsimp simp:cp2sproc_other tainted_eq_Tainted split:option.splits)
+apply (erule_tac x = pa in allE, simp)
+
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (simp)
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (simp split:option.splits)
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (simp split:option.splits)
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (simp split:option.splits)
+
+apply (erule DiffE, drule s2ss_ptrace1_aux, erule UnE)
+apply (erule CollectE, (erule exE|erule conjE)+, rule CollectI)
+apply (erule co2sobj_some_caseD)
+apply (case_tac "O_proc pa \<in> Tainted s")
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp add:tainted_eq_Tainted cp2sproc_other split:option.splits)
+apply (case_tac "info_flow_shm s p pa", simp only:co2sobj.simps split:option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (drule Meson.not_exD, erule_tac x = a in allE, drule Meson.not_exD, erule_tac x = pa in allE)
+apply (simp add:tainted_eq_Tainted, (erule exE|erule conjE)+)
+apply (rule_tac x = "O_proc p'a" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted split:option.splits)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted split:option.splits)
+
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_ptrace alive_simps simp del:co2sobj.simps)[1]
+
+apply (erule CollectE, (erule exE|erule conjE)+, rule CollectI)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted info_shm_flow_in_procs)
+done
+
+lemma s2ss_ptrace3:
+ "\<lbrakk>valid (Ptrace p p' # s); (O_proc p' \<in> Tainted s) = (O_proc p \<in> Tainted s)\<rbrakk>
+ \<Longrightarrow> s2ss (Ptrace p p' # s) = s2ss s"
+unfolding s2ss_def
+apply (frule vd_cons, rule set_eqI, rule iffI)
+apply (erule CollectE, (erule exE|erule conjE)+, rule CollectI)
+apply (rule_tac x = obj in exI)
+apply (frule alive_other, simp+)
+apply (frule_tac obj = obj in co2sobj_ptrace, simp)
+apply (auto simp add:tainted_eq_Tainted split:t_object.splits option.splits if_splits
+ intro:info_flow_shm_Tainted)[1]
+
+apply (erule CollectE, (erule exE|erule conjE)+, rule CollectI)
+apply (rule_tac x = obj in exI)
+apply (frule alive_other, simp+)
+apply (frule_tac obj = obj in co2sobj_ptrace, simp)
+apply (auto simp add:tainted_eq_Tainted split:t_object.splits option.splits if_splits
+ intro:info_flow_shm_Tainted)
+done
+
+lemma s2ss_ptrace:
+ "valid (Ptrace p p' # s) \<Longrightarrow> s2ss (Ptrace p p' # s) = (
+ if (O_proc p \<in> Tainted s \<and> O_proc p' \<notin> Tainted s)
+ then update_s2ss_shm s p'
+ else if (O_proc p' \<in> Tainted s \<and> O_proc p \<notin> Tainted s)
+ then update_s2ss_shm s p
+ else s2ss s )"
+apply (frule vt_grant_os, frule vd_cons)
+apply (simp add:s2ss_ptrace2 s2ss_ptrace1 split:if_splits)
+by (auto dest:s2ss_ptrace3)
+
+lemma s2ss_kill:
+ "valid (Kill p p' # s) \<Longrightarrow> s2ss (Kill p p' # s) = (
+ if (\<exists> p''. p'' \<in> current_procs s \<and> p'' \<noteq> p' \<and> co2sobj s (O_proc p'') = co2sobj s (O_proc p'))
+ then s2ss s
+ else (case (co2sobj s (O_proc p')) of
+ Some sp \<Rightarrow> s2ss s - {sp}
+ | _ \<Rightarrow> {}))"
+apply (frule vt_grant_os, frule vd_cons)
+unfolding s2ss_def
+apply (simp split:if_splits, rule conjI)
+apply (rule impI, (erule exE|erule conjE)+)
+apply (split option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (simp split: option.splits, (erule conjE)+)
+apply (rule set_eqI, rule iffI, erule CollectE, (erule exE|erule conjE)+, rule CollectI)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_kill tainted_eq_Tainted alive_kill split:t_object.splits if_splits)
+apply (erule CollectE, erule exE, erule conjE, rule CollectI)
+apply (erule co2sobj_some_caseD)
+apply (case_tac "pa = p'")
+apply (rule_tac x = "O_proc p''" in exI)
+apply (simp add:cp2sproc_kill tainted_eq_Tainted alive_kill
+ split:t_object.splits if_splits option.splits)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp add:cp2sproc_kill tainted_eq_Tainted alive_kill
+ split:t_object.splits if_splits option.splits)
+apply (rule_tac x = obj in exI, frule alive_kill, simp add:co2sobj_kill del:co2sobj.simps)+
+
+apply (rule impI, erule conjE, frule current_proc_has_sp, simp, erule exE, simp)
+apply (rule set_eqI, rule iffI)
+apply (erule CollectE, erule exE, erule conjE, rule DiffI)
+apply (rule CollectI, rule_tac x = obj in exI)
+apply (simp add:co2sobj_kill tainted_eq_Tainted alive_kill split:t_object.splits if_splits)
+apply (rule notI, simp, case_tac obj)
+apply (erule_tac x = nat in allE)
+apply (simp add:co2sobj_kill cp2sproc_kill tainted_eq_Tainted split:option.splits)
+apply (simp split:option.splits)+
+apply (erule co2sobj_some_caseD)
+apply (case_tac "pa = p'")
+apply (rule_tac x = "O_proc p''" in exI)
+apply (simp add:cp2sproc_kill tainted_eq_Tainted alive_kill
+ split:t_object.splits if_splits option.splits)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp add:cp2sproc_kill tainted_eq_Tainted alive_kill
+ split:t_object.splits if_splits option.splits)
+apply (rule_tac x = obj in exI, frule alive_kill, simp add:co2sobj_kill del:co2sobj.simps)+
+done
+
+lemma s2ss_exit:
+ "valid (Exit p # s) \<Longrightarrow> s2ss (Exit p # s) = (
+ if (\<exists> p'. p' \<in> current_procs s \<and> p' \<noteq> p \<and> co2sobj s (O_proc p') = co2sobj s (O_proc p))
+ then s2ss s
+ else (case (co2sobj s (O_proc p)) of
+ Some sp \<Rightarrow> s2ss s - {sp}
+ | _ \<Rightarrow> {}))"
+apply (frule vt_grant_os, frule vd_cons)
+unfolding s2ss_def
+apply (simp split:if_splits, rule conjI)
+apply (rule impI, (erule exE|erule conjE)+)
+apply (split option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (simp split: option.splits, (erule conjE)+)
+apply (rule set_eqI, rule iffI, erule CollectE, (erule exE|erule conjE)+, rule CollectI)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_exit tainted_eq_Tainted alive_exit split:t_object.splits if_splits)
+apply (erule CollectE, erule exE, erule conjE, rule CollectI)
+apply (erule co2sobj_some_caseD)
+apply (case_tac "pa = p")
+apply (rule_tac x = "O_proc p'" in exI)
+apply (simp add:cp2sproc_exit tainted_eq_Tainted alive_exit
+ split:t_object.splits if_splits option.splits)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp add:cp2sproc_exit tainted_eq_Tainted alive_exit
+ split:t_object.splits if_splits option.splits)
+apply (rule_tac x = obj in exI, frule alive_exit, simp add:co2sobj_exit del:co2sobj.simps)+
+
+apply (rule impI, frule current_proc_has_sp, simp, erule exE, simp)
+apply (rule set_eqI, rule iffI)
+apply (erule CollectE, erule exE, erule conjE, rule DiffI)
+apply (rule CollectI, rule_tac x = obj in exI)
+apply (simp add:co2sobj_exit tainted_eq_Tainted alive_exit split:t_object.splits if_splits)
+apply (rule notI, simp, case_tac obj)
+apply (erule_tac x = nat in allE)
+apply (simp add:co2sobj_exit cp2sproc_exit tainted_eq_Tainted split:option.splits)
+apply (simp split:option.splits)+
+apply (erule co2sobj_some_caseD)
+apply (case_tac "pa = p")
+apply (rule_tac x = "O_proc p'" in exI)
+apply (simp add:cp2sproc_exit tainted_eq_Tainted alive_exit
+ split:t_object.splits if_splits option.splits)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp add:cp2sproc_exit tainted_eq_Tainted alive_exit
+ split:t_object.splits if_splits option.splits)
+apply (rule_tac x = obj in exI, frule alive_exit, simp add:co2sobj_exit del:co2sobj.simps)+
+done
+
+lemma alive_has_sobj':
+ "\<lbrakk>co2sobj s obj = None; valid s\<rbrakk> \<Longrightarrow> \<not> alive s obj"
+apply (case_tac obj)
+apply (auto split:option.splits)
+oops
+
+declare co2sobj.simps [simp del]
+
+lemma co2sobj_open_none:
+ "\<lbrakk>valid (Open p f flag fd None # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (Open p f flag fd None # s) obj = (
+ if (obj = O_proc p)
+ then (case (cp2sproc (Open p f flag fd None # s) p) of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons)
+apply (frule_tac obj = obj in co2sobj_open, simp add:alive_open)
+apply (auto split:t_object.splits option.splits dest!:current_proc_has_sp')
+done
+
+lemma co2sobj_open_some:
+ "\<lbrakk>valid (Open p f flag fd (Some i) # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (Open p f flag fd (Some i) # s) obj = (
+ if (obj = O_proc p)
+ then (case (cp2sproc (Open p f flag fd (Some i) # s) p) of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else if (obj = O_file f)
+ then (case (cf2sfile (Open p f flag fd (Some i) # s) f) of
+ Some sf \<Rightarrow> Some (S_file {sf} (O_proc p \<in> Tainted s))
+ | _ \<Rightarrow> None)
+ else co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons)
+apply (frule_tac obj = obj in co2sobj_open, simp add:alive_open)
+apply (auto split:t_object.splits option.splits dest!:current_proc_has_sp')
+done
+
+lemma alive_co2sobj_some_open_none:
+ "\<lbrakk>co2sobj s obj = Some sobj; alive s obj; valid (Open p f flag fd None # s)\<rbrakk>
+ \<Longrightarrow> alive (Open p f flag fd None # s) obj"
+by (erule co2sobj_some_caseD, auto simp:alive_simps is_file_simps is_dir_simps)
+
+lemma alive_co2sobj_some_open_none':
+ "\<lbrakk>co2sobj (Open p f flag fd None # s) obj = Some sobj; alive (Open p f flag fd None # s) obj;
+ valid (Open p f flag fd None # s)\<rbrakk> \<Longrightarrow> alive s obj"
+by (erule co2sobj_some_caseD, auto simp:alive_simps is_file_simps is_dir_simps)
+
+lemma co2sobj_proc_obj:
+ "\<lbrakk>co2sobj s obj = Some x; co2sobj s (O_proc p) = Some x\<rbrakk>
+ \<Longrightarrow> \<exists> p'. obj = O_proc p'"
+by (case_tac obj, auto simp:co2sobj.simps split:option.splits)
+
+lemma s2ss_open_none:
+ "valid (Open p f flag fd None # s) \<Longrightarrow> s2ss (Open p f flag fd None # s) = (
+ case (co2sobj s (O_proc p), co2sobj (Open p f flag fd None # s) (O_proc p)) of
+ (Some sp, Some sp') \<Rightarrow>
+ if (\<exists> p'. p' \<in> current_procs s \<and> p' \<noteq> p \<and> co2sobj s (O_proc p') = Some sp)
+ then s2ss s \<union> {sp'}
+ else s2ss s - {sp} \<union> {sp'}
+ | _ \<Rightarrow> {} )"
+unfolding s2ss_def
+apply (frule vt_grant_os, frule vd_cons)
+apply (case_tac "co2sobj s (O_proc p)", simp add:co2sobj.simps split:option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (case_tac "co2sobj (Open p f flag fd None # s) (O_proc p)")
+apply (simp add:co2sobj.simps split:option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (rule set_eqI, rule iffI, erule CollectE, erule exE, erule conjE, simp)
+apply (frule_tac obj = obj in alive_co2sobj_some_open_none', simp, simp)
+apply (rule conjI, rule impI, erule exE, (erule conjE)+)
+apply (rule Meson.disj_comm, rule disjCI, case_tac "obj = O_proc p", simp)
+apply (rule_tac x = obj in exI, simp add:co2sobj_open_none)
+apply (rule impI)
+apply (case_tac "obj = O_proc p", simp)
+apply (rule Meson.disj_comm, rule disjCI, rule conjI)
+apply (rule_tac x = obj in exI, simp add:co2sobj_open_none)
+apply (simp add:co2sobj_open_none split:option.splits)
+apply (rule notI)
+apply (frule co2sobj_proc_obj, simp, erule exE)
+apply (erule_tac x = p' in allE, simp)
+
+apply (simp split:if_splits)
+apply (erule disjE, rule_tac x = "O_proc p" in exI, simp)
+apply (erule exE, erule conjE, case_tac "obj = O_proc p")
+apply (rule_tac x = "O_proc p'" in exI, simp add:co2sobj_open_none)
+apply (rule_tac x = obj in exI, simp add:co2sobj_open_none alive_co2sobj_some_open_none)
+apply (erule disjE, rule_tac x = "O_proc p" in exI, simp)
+apply (erule conjE, erule exE, erule conjE, case_tac "obj = O_proc p")
+apply (rule_tac x = "O_proc p'" in exI, simp add:co2sobj_open_none)
+apply (rule_tac x = obj in exI, simp add:co2sobj_open_none alive_co2sobj_some_open_none)
+done
+
+lemma alive_co2sobj_some_open_some:
+ "\<lbrakk>alive s obj; valid (Open p f flag fd (Some i) # s)\<rbrakk>
+ \<Longrightarrow> alive (Open p f flag fd (Some i) # s) obj"
+by (case_tac obj, auto simp:alive_simps is_file_simps is_dir_simps)
+
+lemma alive_co2sobj_some_open_some':
+ "\<lbrakk>co2sobj (Open p f flag fd (Some i) # s) obj = Some sobj; alive (Open p f flag fd (Some i) # s) obj;
+ valid (Open p f flag fd (Some i) # s); obj \<noteq> O_file f\<rbrakk> \<Longrightarrow> alive s obj"
+by (erule co2sobj_some_caseD, auto simp:alive_simps is_file_simps is_dir_simps)
+
+lemma s2ss_open_some:
+ "valid (Open p f flag fd (Some i) # s) \<Longrightarrow> s2ss (Open p f flag fd (Some i) # s) = (
+ case (co2sobj s (O_proc p), co2sobj (Open p f flag fd (Some i) # s) (O_proc p),
+ co2sobj (Open p f flag fd (Some i) # s) (O_file f)) of
+ (Some sp, Some sp', Some sf) \<Rightarrow>
+ if (\<exists> p'. p' \<in> current_procs s \<and> p' \<noteq> p \<and> co2sobj s (O_proc p') = Some sp)
+ then s2ss s \<union> {sp', sf}
+ else s2ss s - {sp} \<union> {sp', sf}
+ | _ \<Rightarrow> {} )"
+unfolding s2ss_def
+apply (frule vt_grant_os, frule vd_cons)
+apply (case_tac "co2sobj s (O_proc p)", simp add:co2sobj.simps split:option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (case_tac "co2sobj (Open p f flag fd (Some i) # s) (O_proc p)")
+apply (simp add:co2sobj.simps split:option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (case_tac "co2sobj (Open p f flag fd (Some i) # s) (O_file f)")
+apply (simp add:co2sobj.simps split:option.splits)
+apply (clarsimp split del:if_splits)
+
+apply (rule set_eqI, rule iffI, erule CollectE, erule exE, erule conjE)
+apply (split if_splits, rule conjI, rule impI, erule exE, erule conjE, erule conjE)
+apply (case_tac "obj = O_proc p", simp, case_tac "obj = O_file f", simp)
+apply (rule UnI1, rule CollectI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in alive_co2sobj_some_open_some', simp+)
+apply (simp add:co2sobj_open split:t_object.splits)
+apply (rule impI, case_tac "obj = O_proc p", simp, case_tac "obj = O_file f", simp)
+apply (rule UnI1, rule DiffI, rule CollectI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in alive_co2sobj_some_open_some', simp+)
+apply (simp add:co2sobj_open split:t_object.splits)
+apply (frule_tac obj = obj in co2sobj_open_some, simp+)
+apply (simp add:alive_co2sobj_some_open_some', simp)
+apply (rule notI)
+apply (frule_tac obj = obj and p = p in co2sobj_proc_obj, simp+, erule exE)
+apply (erule_tac x = p' in allE, simp)
+
+apply (simp split:if_splits, erule disjE)
+apply (rule_tac x = "O_proc p" in exI, simp)
+apply (erule disjE, rule_tac x = "O_file f" in exI, simp add:is_file_simps)
+apply (erule exE, erule conjE)
+apply (case_tac "obj = O_proc p", simp)
+apply (rule_tac x = "O_proc p'" in exI, simp add:co2sobj_open_some)
+apply (case_tac "obj = O_file f", simp add:is_file_in_current)
+apply (rule_tac x = obj in exI, simp add:co2sobj_open_some alive_co2sobj_some_open_some)
+apply (erule disjE, rule_tac x = "O_proc p" in exI, simp)
+apply (erule disjE, rule_tac x = "O_file f" in exI, simp add:is_file_simps)
+apply (erule conjE, erule exE, erule conjE)
+apply (case_tac "obj = O_proc p", simp)
+apply (case_tac "obj = O_file f", simp add:is_file_in_current)
+apply (rule_tac x = obj in exI, simp add:co2sobj_open_some alive_co2sobj_some_open_some)
+done
+
+lemma s2ss_open:
+ "valid (Open p f flag fd opt # s) \<Longrightarrow> s2ss (Open p f flag fd opt # s) = (
+ if opt = None
+ then (case (co2sobj s (O_proc p), co2sobj (Open p f flag fd opt # s) (O_proc p)) of
+ (Some sp, Some sp') \<Rightarrow>
+ if (\<exists> p'. p' \<in> current_procs s \<and> p' \<noteq> p \<and> co2sobj s (O_proc p') = Some sp)
+ then s2ss s \<union> {sp'}
+ else s2ss s - {sp} \<union> {sp'}
+ | _ \<Rightarrow> {} )
+ else (case (co2sobj s (O_proc p), co2sobj (Open p f flag fd opt # s) (O_proc p),
+ co2sobj (Open p f flag fd opt # s) (O_file f)) of
+ (Some sp, Some sp', Some sf) \<Rightarrow>
+ if (\<exists> p'. p' \<in> current_procs s \<and> p' \<noteq> p \<and> co2sobj s (O_proc p') = Some sp)
+ then s2ss s \<union> {sp', sf}
+ else s2ss s - {sp} \<union> {sp', sf}
+ | _ \<Rightarrow> {} ) )"
+apply (case_tac opt)
+apply (simp add:s2ss_open_some s2ss_open_none)+
+done
+
+lemma s2ss_readfile:
+ "valid (ReadFile p fd # s) \<Longrightarrow> s2ss (ReadFile p fd # s) = (
+ case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> if (O_file f \<in> Tainted s)
+ then update_s2ss_shm s p
+ else s2ss s
+ | _ \<Rightarrow> {})"
+apply (frule vt_grant_os, frule vd_cons, clarsimp split:option.splits)
+apply (rule conjI, rule impI, rule impI, simp)
+unfolding update_s2ss_shm_def s2ss_def s2ss_shm_no_backup_def
+apply (rule set_eqI, rule iffI)
+apply (drule CollectD, (erule exE|erule conjE)+)
+apply (erule co2sobj_some_caseD)
+apply (rule DiffI)
+apply (case_tac "O_proc pa \<in> Tainted s")
+apply (rule UnI1, simp, rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:cp2sproc_other tainted_eq_Tainted co2sobj.simps split:option.splits)
+apply (case_tac "info_flow_shm s p pa")
+apply (rule UnI2, rule CollectI, simp only:co2sobj.simps split:option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (rule_tac x = ab in exI, rule_tac x = pa in exI)
+apply (clarsimp simp:cp2sproc_other tainted_eq_Tainted split:option.splits)
+apply (rule UnI1, simp)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:cp2sproc_other tainted_eq_Tainted co2sobj.simps split:option.splits)
+apply (rule notI, clarsimp simp:cp2sproc_other tainted_eq_Tainted co2sobj.simps split:option.splits)
+apply (erule_tac x = pa in allE, simp)
+
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI)
+apply (simp add:alive_simps co2sobj_readfile, rule notI, clarsimp simp:co2sobj.simps)
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI)
+apply (simp add:alive_simps co2sobj_readfile, rule notI, clarsimp simp:co2sobj.simps split:option.splits)
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI)
+apply (simp add:alive_simps co2sobj_readfile, rule notI, clarsimp simp:co2sobj.simps split:option.splits)
+apply (rule DiffI, rule UnI1, rule CollectI, rule_tac x = obj in exI)
+apply (simp add:alive_simps co2sobj_readfile, rule notI, clarsimp simp:co2sobj.simps split:option.splits)
+
+apply (erule DiffE, drule s2ss_ptrace1_aux, erule UnE)
+apply (erule CollectE, (erule exE|erule conjE)+, rule CollectI)
+apply (erule co2sobj_some_caseD)
+apply (case_tac "O_proc pa \<in> Tainted s")
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:tainted_eq_Tainted cp2sproc_other co2sobj.simps split:option.splits)
+apply (case_tac "info_flow_shm s p pa", simp only:co2sobj.simps split:option.splits)
+apply (drule current_proc_has_sp', simp, simp)
+apply (drule Meson.not_exD, erule_tac x = ab in allE, drule Meson.not_exD, erule_tac x = pa in allE)
+apply (simp add:tainted_eq_Tainted, (erule exE|erule conjE)+)
+apply (rule_tac x = "O_proc p'" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted co2sobj.simps split:option.splits)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted co2sobj.simps split:option.splits)
+
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_readfile alive_simps simp del:co2sobj.simps)[1]
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_readfile alive_simps simp del:co2sobj.simps)[1]
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_readfile alive_simps simp del:co2sobj.simps)[1]
+apply (rule_tac x = obj in exI,
+ auto split:option.splits simp:co2sobj_readfile alive_simps simp del:co2sobj.simps)[1]
+
+apply (erule CollectE, (erule exE|erule conjE)+, rule CollectI)
+apply (rule_tac x = "O_proc pa" in exI)
+apply (clarsimp simp:Tainted_ptrace' cp2sproc_other tainted_eq_Tainted info_shm_flow_in_procs co2sobj.simps)
+
+
+apply (rule impI, rule impI)
+apply (rule set_eqI, rule iffI, erule CollectE, erule exE, erule conjE, rule CollectI)
+apply (rule_tac x = obj in exI, simp add:alive_simps co2sobj_readfile split:t_object.splits)
+apply (erule CollectE, erule exE, erule conjE, rule CollectI)
+apply (rule_tac x = obj in exI, simp add:alive_simps)
+apply (drule_tac obj = obj in co2sobj_readfile, simp)
+apply (simp split:t_object.splits)
+done
+
+lemma same_inode_files_prop9:
+ "is_file s f \<Longrightarrow> f \<in> same_inode_files s f"
+by (simp add:same_inode_files_def)
+
+lemma cf2sfiles_prop:
+ "\<lbrakk>f \<in> same_inode_files s f'; valid s\<rbrakk> \<Longrightarrow> cf2sfiles s f = cf2sfiles s f'"
+apply (auto simp:cf2sfiles_def)
+apply (rule_tac x = f'a in bexI, simp)
+apply (erule same_inode_files_prop4, simp)
+apply (rule_tac x = f'a in bexI, simp)
+apply (drule same_inode_files_prop5)
+apply (erule same_inode_files_prop4, simp)
+done
+
+lemma co2sobj_writefile_unchange:
+ "\<lbrakk>valid (WriteFile p fd # s); alive s obj; file_of_proc_fd s p fd = Some f;
+ O_proc p \<in> Tainted s \<longrightarrow> O_file f \<in> Tainted s\<rbrakk>
+ \<Longrightarrow> co2sobj (WriteFile p fd # s) obj = co2sobj s obj"
+apply (frule vd_cons, frule co2sobj_writefile, simp, simp split:t_object.splits if_splits)
+apply (simp add:co2sobj.simps)
+apply (case_tac "O_proc p \<in> Tainted s")
+apply (simp add:tainted_eq_Tainted same_inodes_Tainted)+
+done
+
+lemma s2ss_writefile:
+ "valid (WriteFile p fd # s) \<Longrightarrow> s2ss (WriteFile p fd # s) = (
+ case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> if (O_proc p \<in> Tainted s \<and> O_file f \<notin> Tainted s)
+ then (if (\<exists> f'. f' \<notin> same_inode_files s f \<and> is_file s f' \<and>
+ co2sobj s (O_file f') = co2sobj s (O_file f))
+ then s2ss s \<union> {S_file (cf2sfiles s f) True}
+ else s2ss s - {S_file (cf2sfiles s f) False}
+ \<union> {S_file (cf2sfiles s f) True})
+ else s2ss s
+ | _ \<Rightarrow> {})"
+apply (frule vd_cons, frule vt_grant_os)
+apply (clarsimp split:option.splits)
+unfolding s2ss_def
+apply (rule conjI|rule impI|erule exE|erule conjE)+
+apply (rule set_eqI, rule iffI, erule CollectE, erule exE, erule conjE)
+apply (frule_tac obj = obj in co2sobj_writefile, simp add:alive_other)
+apply (simp split:t_object.splits if_splits)
+apply (rule disjI2, rule_tac x= "O_proc nat" in exI, simp)
+apply (rule disjI1, simp add:cf2sfiles_prop)
+apply (rule disjI2, rule_tac x = obj in exI, simp add:is_file_simps)
+apply (simp add:co2sobj.simps)
+apply (rule disjI2, rule_tac x = obj in exI, simp add:is_dir_simps)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (rule disjI2, rule_tac x = obj in exI, simp)
+apply (rule disjI2, rule_tac x = obj in exI, simp)
+apply (simp add:co2sobj.simps)
+
+apply (simp, erule disjE)
+apply (rule_tac x = "O_file aa" in exI, simp add:is_file_simps file_of_pfd_is_file)
+apply (frule_tac obj = "O_file aa" in co2sobj_writefile, simp add:file_of_pfd_is_file)
+apply (simp split:if_splits add:same_inode_files_def file_of_pfd_is_file)
+apply (erule exE, erule conjE)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI, simp add:co2sobj_writefile)
+apply (case_tac "fa \<in> same_inode_files s aa")
+apply (frule_tac f = fa and f' = aa in cf2sfiles_prop, simp)
+apply (rule_tac x = "O_file f'" in exI, simp add:co2sobj_writefile is_file_simps)
+apply (rule conjI, rule impI, simp add:same_inode_files_prop5)
+apply (rule impI, simp add:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted)
+apply (rule_tac x = "O_file fa" in exI, simp add:co2sobj_writefile is_file_simps)
+apply (rule impI, simp add:same_inode_files_prop5)
+apply (rule_tac x = obj in exI, simp add:co2sobj_writefile)
+apply (rule_tac x = obj in exI, simp add:co2sobj_writefile is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_writefile)
+
+apply (rule impI, rule impI, simp, rule set_eqI, rule iffI, erule CollectE, (erule conjE|erule exE)+)
+apply (rule CollectI, rule_tac x = obj in exI, simp add:alive_simps)
+apply (simp add:co2sobj_writefile split:t_object.splits if_splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted)
+apply (case_tac "O_proc p \<in> Tainted s", simp, simp)
+apply (erule CollectE, (erule conjE|erule exE)+, rule CollectI)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_writefile_unchange alive_simps)
+
+apply (rule impI| rule conjI|erule conjE)+
+apply (rule set_eqI, rule iffI, erule CollectE, erule exE, erule conjE)
+apply (simp add:alive_simps co2sobj_writefile split:t_object.splits)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI, simp,
+ rule notI, simp add:co2sobj.simps split:option.splits)
+apply (simp split:if_splits)
+apply (rule disjI1, simp add:cf2sfiles_prop)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI, simp)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (erule_tac x = list in allE, simp add:tainted_eq_Tainted same_inode_files_prop5)
+apply (simp add:co2sobj.simps)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI, simp,
+ rule notI, simp add:co2sobj.simps split:option.splits)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI, simp,
+ rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI, simp,
+ rule notI, simp add:co2sobj.simps split:option.splits)
+apply (simp add:co2sobj.simps)
+apply (simp)
+apply (erule disjE)
+apply (rule_tac x= "O_file aa" in exI, simp add:co2sobj_writefile alive_simps file_of_pfd_is_file)
+apply (rule impI, simp add:same_inode_files_def file_of_pfd_is_file)
+apply (erule exE|erule conjE)+
+apply (erule co2sobj_some_caseD)
+apply (rule_tac x = obj in exI, simp add:co2sobj_writefile)
+apply (case_tac "fa \<in> same_inode_files s aa")
+apply (frule cf2sfiles_prop, simp, simp add:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted)
+apply (rule_tac x = obj in exI, simp add:co2sobj_writefile is_file_simps)
+apply (rule impI, simp add:same_inode_files_prop5)
+apply (rule_tac x = obj in exI, simp add:co2sobj_writefile)
+apply (rule_tac x = obj in exI, simp add:co2sobj_writefile is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_writefile)
+apply (rule impI, rule impI)
+
+apply (rule set_eqI, rule iffI, erule CollectE,erule exE,erule conjE,rule CollectI)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_writefile_unchange alive_simps)
+apply (erule CollectE, erule exE, erule conjE)
+apply (rule CollectI, rule_tac x = obj in exI)
+apply (simp add:co2sobj_writefile_unchange alive_simps)
+done
+
+
+definition update_s2ss_obj :: "t_state \<Rightarrow> t_static_state \<Rightarrow> t_object \<Rightarrow> t_sobject \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
+where
+ "update_s2ss_obj s ss obj sobj sobj' =
+ (if (\<exists> obj'. alive s obj' \<and> obj' \<noteq> obj \<and> co2sobj s obj' = Some sobj)
+ then ss \<union> {sobj'}
+ else ss - {sobj} \<union> {sobj'})"
+
+definition update_s2ss_sfile_del :: "t_state \<Rightarrow> t_static_state \<Rightarrow> t_file \<Rightarrow> t_sfile \<Rightarrow> t_static_state"
+where
+ "update_s2ss_sfile_del s ss f sf \<equiv>
+ if (same_inode_files s f = {f})
+ then ss
+ else ss \<union> {S_file (cf2sfiles s f - {sf}) (O_file f \<in> Tainted s)}"
+
+definition del_s2ss_file:: "t_state \<Rightarrow> t_static_state \<Rightarrow> t_file \<Rightarrow> t_sfile \<Rightarrow> t_static_state"
+where
+ "del_s2ss_file s ss f sf =
+ (if (\<exists> f' \<in> same_inode_files s f. f' \<noteq> f \<and> cf2sfile s f' = Some sf)
+ then ss
+ else if (\<exists> f'. is_file s f' \<and> f' \<notin> same_inode_files s f \<and> co2sobj s (O_file f') = co2sobj s (O_file f))
+ then update_s2ss_sfile_del s ss f sf
+ else update_s2ss_sfile_del s (ss - {S_file (cf2sfiles s f) (O_file f \<in> Tainted s)}) f sf)"
+
+lemma alive_co2sobj_closefd1:
+ "\<lbrakk>alive s obj; valid (CloseFd p fd # s); co2sobj s obj = Some sobj;
+ file_of_proc_fd s p fd = Some f; \<not> (f \<in> files_hung_by_del s \<and> proc_fd_of_file s f = {(p, fd)})\<rbrakk>
+ \<Longrightarrow> alive (CloseFd p fd # s) obj"
+apply (erule co2sobj_some_caseD)
+by (auto simp:alive_simps is_file_simps is_dir_simps split:option.splits)
+
+lemma alive_co2sobj_closefd3:
+ "\<lbrakk>alive s obj; valid (CloseFd p fd # s); co2sobj s obj = Some sobj; obj \<noteq> O_file f;
+ file_of_proc_fd s p fd = Some f; f \<in> files_hung_by_del s; proc_fd_of_file s f = {(p, fd)}\<rbrakk>
+ \<Longrightarrow> alive (CloseFd p fd # s) obj"
+apply (erule co2sobj_some_caseD)
+by (auto simp:alive_simps is_file_simps is_dir_simps split:option.splits)
+
+lemma alive_co2sobj_closefd2:
+ "\<lbrakk>alive s obj; valid (CloseFd p fd # s); co2sobj s obj = Some sobj; file_of_proc_fd s p fd = None\<rbrakk>
+ \<Longrightarrow> alive (CloseFd p fd # s) obj"
+apply (erule co2sobj_some_caseD)
+by (auto simp:alive_simps is_file_simps is_dir_simps split:option.splits)
+
+lemma alive_co2sobj_closefd':
+ "\<lbrakk>co2sobj (CloseFd p fd # s) obj = Some sobj; alive (CloseFd p fd # s) obj;
+ valid (CloseFd p fd # s)\<rbrakk> \<Longrightarrow> alive s obj"
+apply (erule co2sobj_some_caseD)
+by (auto simp:alive_simps is_file_simps is_dir_simps split:option.splits if_splits)
+
+ML {*asm_full_simp_tac*}
+
+ML {*
+fun my_setiff_tac i =
+ (etac @{thm CollectE} i
+ ORELSE ( asm_full_simp_tac (HOL_ss addsimps @{thms Set.insert_iff}) i
+ THEN etac @{thm disjE} i)
+ ORELSE ( asm_full_simp_tac (HOL_ss addsimps @{thms Set.Diff_iff}) i
+ THEN etac @{thm conjE} i
+ THEN (REPEAT (etac @{thm CollectE} i))))
+THEN (REPEAT (( etac @{thm exE}
+ ORELSE' etac @{thm conjE}
+ ORELSE' etac @{thm bexE}) i))
+THEN (rtac @{thm CollectI} i
+ ORELSE ( asm_full_simp_tac (HOL_ss addsimps @{thms Set.insert_iff}) i))
+
+*}
+
+ML {*
+fun my_seteq_tac i =
+ (simp_tac (HOL_ss addsimps @{thms s2ss_def}) 1)
+THEN (rtac @{thm set_eqI} i)
+THEN (rtac @{thm iffI} i)
+THEN my_setiff_tac i
+*}
+
+lemma co2sobj_sproc_imp:
+ "co2sobj s obj = Some (S_proc sp tag) \<Longrightarrow> \<exists> p. obj = O_proc p"
+by (case_tac obj, auto simp:co2sobj.simps split:option.splits)
+
+lemma co2sobj_sfile_imp:
+ "co2sobj s obj = Some (S_file sfs tag) \<Longrightarrow> \<exists> f. obj = O_file f"
+by (case_tac obj, auto simp:co2sobj.simps split:option.splits)
+
+lemma co2sobj_sdir_imp:
+ "co2sobj s obj = Some (S_dir sf) \<Longrightarrow> \<exists> f. obj = O_dir f"
+by (case_tac obj, auto simp:co2sobj.simps split:option.splits)
+
+lemma co2sobj_sshm_imp:
+ "co2sobj s obj = Some (S_shm sh) \<Longrightarrow> \<exists> h. obj = O_shm h"
+by (case_tac obj, auto simp:co2sobj.simps split:option.splits)
+
+lemma co2sobj_smsgq_imp:
+ "co2sobj s obj = Some (S_msgq sq) \<Longrightarrow> \<exists> q. obj = O_msgq q"
+by (case_tac obj, auto simp:co2sobj.simps split:option.splits)
+
+lemma same_inode_files_prop10:
+ "\<lbrakk>same_inode_files s f \<noteq> {f}; is_file s f\<rbrakk> \<Longrightarrow> \<exists> f'. f' \<in> same_inode_files s f \<and> f' \<noteq> f"
+by (auto simp:same_inode_files_def split:if_splits)
+
+lemma same_inode_files_prop11:
+ "f \<in> same_inode_files s f' \<Longrightarrow> is_file s f"
+by (auto simp:same_inode_files_def is_file_def split:if_splits)
+
+lemma same_inode_files_prop11':
+ "f \<in> same_inode_files s f' \<Longrightarrow> is_file s f'"
+by (auto simp:same_inode_files_def is_file_def split:if_splits)
+
+lemma s2ss_closefd:
+ "valid (CloseFd p fd # s) \<Longrightarrow> s2ss (CloseFd p fd # s) = (
+ case (file_of_proc_fd s p fd) of
+ Some f \<Rightarrow> if (f \<in> files_hung_by_del s \<and> proc_fd_of_file s f = {(p, fd)})
+ then (case (cf2sfile s f, cp2sproc s p, cp2sproc (CloseFd p fd # s) p) of
+ (Some sf, Some sp, Some sp') \<Rightarrow>
+ (del_s2ss_file s (
+ update_s2ss_obj s (s2ss s) (O_proc p)
+ (S_proc sp (O_proc p \<in> Tainted s))
+ (S_proc sp' (O_proc p \<in> Tainted s))) f sf)
+ | _ \<Rightarrow> {})
+ else (case (cp2sproc s p, cp2sproc (CloseFd p fd # s) p) of
+ (Some sp, Some sp') \<Rightarrow>
+ (update_s2ss_obj s (s2ss s) (O_proc p)
+ (S_proc sp (O_proc p \<in> Tainted s))
+ (S_proc sp' (O_proc p \<in> Tainted s)))
+ | _ \<Rightarrow> {})
+ | _ \<Rightarrow> s2ss s)"
+apply (frule vd_cons, frule vt_grant_os)
+apply (clarsimp simp only:os_grant.simps)
+apply (frule current_proc_has_sp, simp, erule exE)
+apply (case_tac "file_of_proc_fd s p fd")
+
+apply (simp add:s2ss_def)
+apply (rule set_eqI, rule iffI, erule CollectE, erule exE, erule conjE, rule CollectI)
+apply (rule_tac x = obj in exI, simp add:alive_co2sobj_closefd')
+apply (frule co2sobj_closefd, simp)
+apply (frule cp2sproc_closefd, simp)
+apply (simp add:proc_file_fds_def split:t_object.splits)
+apply (simp split:if_splits add:co2sobj.simps tainted_eq_Tainted)
+apply (erule CollectE, erule exE, erule conjE, rule CollectI)
+apply (rule_tac x = obj in exI, simp add:alive_co2sobj_closefd2)
+apply (frule_tac obj = obj in co2sobj_closefd, simp add:alive_co2sobj_closefd2)
+apply (frule cp2sproc_closefd, simp)
+apply (auto simp add:proc_file_fds_def co2sobj.simps tainted_eq_Tainted
+ split:t_object.splits option.splits if_splits)[1]
+
+apply (case_tac "cp2sproc (CloseFd p fd # s) p")
+apply (drule current_proc_has_sp', simp, simp)
+apply (case_tac "cf2sfile s a")
+apply (drule current_file_has_sfile', simp, simp add:file_of_pfd_in_current)
+apply (simp)
+
+apply (rule conjI, rule impI, erule conjE)
+apply (simp add:del_s2ss_file_def)
+apply (rule conjI|rule impI|erule exE|erule conjE|erule bexE)+
+
+apply (simp add:update_s2ss_obj_def)
+apply (rule conjI|rule impI|erule exE|erule conjE|erule bexE)+
+apply (tactic {*my_seteq_tac 1*})
+apply simp
+apply (case_tac "obj = O_proc p")
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp add:alive_simps)
+apply (simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_proc p" in exI)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_proc p")
+apply (rule_tac x = obj' in exI)
+apply (frule co2sobj_sproc_imp, erule exE, simp add:co2sobj_closefd)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (case_tac "obj = O_file a")
+apply (rule_tac x = "O_file f'" in exI)
+apply (case_tac "f' = a", simp add:same_inode_files_prop9 file_of_pfd_is_file)
+apply (frule_tac obj = "O_file f'" in co2sobj_closefd, simp add:alive_simps)
+apply (simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in alive_co2sobj_closefd3, simp+)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+
+apply (rule impI)+
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_proc p", rule disjI1, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2)
+apply (case_tac "obj = O_file a", simp add:alive_simps)
+apply (rule DiffI, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in alive_co2sobj_closefd', simp+)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)
+apply (simp, rule notI, simp, frule co2sobj_sproc_imp, erule exE, simp add:co2sobj_closefd)
+apply (erule_tac x = "O_proc pa" in allE, simp)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_proc p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (case_tac "obj = O_file a", rule_tac x = "O_file f'" in exI)
+apply (case_tac "f' = a", simp add:same_inode_files_prop9 file_of_pfd_is_file)
+apply (frule_tac obj = "O_file f'" in co2sobj_closefd, simp add:alive_simps)
+apply (simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in alive_co2sobj_closefd3, simp+)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+
+apply (rule impI, tactic {*my_seteq_tac 1*})
+apply (simp add:update_s2ss_obj_def update_s2ss_sfile_del_def)
+apply (rule conjI| rule impI|erule exE|erule conjE)+
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa = p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (case_tac "f = a", simp add:alive_simps)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule conjI| rule impI|erule exE|erule conjE)+
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa = p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+apply (case_tac "f = a", simp add:alive_simps)
+apply (case_tac "f \<in> same_inode_files s a", rule disjI1)
+apply (simp add:co2sobj_simps split:if_splits option.splits t_sobject.splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted cf2sfiles_prop)
+apply (erule bexE, erule conjE)
+apply (erule_tac x = f'' in ballE, simp, simp)
+apply (rule disjI2, rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule disjI2, rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule disjI2, rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule disjI2, rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule impI, rule conjI, rule impI)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa = p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj_closefd, erule_tac x = obj in allE, simp add:is_file_simps)
+apply (case_tac "f = a", simp add:alive_simps)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj_closefd, erule_tac x = obj in allE, simp add:is_file_simps)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj_closefd, erule_tac x = obj in allE, simp)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj_closefd, erule_tac x = obj in allE, simp add:is_dir_simps)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj_closefd, erule_tac x = obj in allE, simp)
+apply (rule impI)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa = p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+apply (rule notI, simp add:co2sobj_closefd)
+apply (erule_tac x = obj in allE, simp)
+apply (case_tac "f = a", simp add:alive_simps)
+apply (case_tac "f \<in> same_inode_files s a", rule disjI1)
+apply (simp add:co2sobj_simps split:if_splits option.splits t_sobject.splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted cf2sfiles_prop)
+apply (erule bexE, erule conjE)
+apply (erule_tac x = f'' in ballE, simp, simp)
+apply (rule disjI2, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj_closefd, erule_tac x = obj in allE, simp add:is_file_simps)
+apply (rule disjI2, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj_closefd, erule_tac x = obj in allE, simp)
+apply (rule disjI2, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj_closefd, erule_tac x = obj in allE, simp add:is_dir_simps)
+apply (rule disjI2, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj_closefd, erule_tac x = obj in allE, simp)
+
+apply (simp add:update_s2ss_sfile_del_def update_s2ss_obj_def split:if_splits)
+apply (erule disjE, rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule exE, erule conjE)
+apply (case_tac "obj = O_proc p")
+apply (rule_tac x = obj' in exI)
+apply (frule co2sobj_sproc_imp, erule exE, simp add:co2sobj_closefd)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (case_tac "obj = O_file a")
+apply (rule_tac x = "O_file f'" in exI)
+apply (case_tac "f' = a", simp add:same_inode_files_prop9 file_of_pfd_is_file)
+apply (frule_tac obj = "O_file f'" in co2sobj_closefd, simp add:alive_simps)
+apply (simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in alive_co2sobj_closefd3, simp+)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+apply (erule disjE, rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule conjE, erule exE, erule conjE)
+apply (case_tac "obj = O_proc p")
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (case_tac "obj = O_file a")
+apply (rule_tac x = "O_file f'" in exI)
+apply (case_tac "f' = a", simp add:same_inode_files_prop9 file_of_pfd_is_file)
+apply (frule_tac obj = "O_file f'" in co2sobj_closefd, simp add:alive_simps)
+apply (simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in alive_co2sobj_closefd3, simp+)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+apply (erule disjE)
+apply (drule same_inode_files_prop10, simp add:file_of_pfd_is_file, erule exE, erule conjE)
+apply (rule_tac x = "O_file f'a" in exI)
+apply (frule same_inode_files_prop11)
+apply (frule_tac obj = "O_file f'a" in co2sobj_closefd)
+apply (simp add:alive_simps)+
+apply (frule_tac f = "f'a" in is_file_has_sfile', simp, erule exE)
+apply (simp add:co2sobj.simps cf2sfiles_prop tainted_eq_Tainted same_inodes_Tainted split:if_splits)
+apply (rule impI, erule bexE, erule conjE, erule_tac x = f'' in ballE, simp, simp)
+apply (erule bexE, erule conjE, erule_tac x = f'' in ballE, simp, simp)
+apply (erule disjE)
+apply (rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule exE, erule conjE)
+apply (case_tac "obj = O_proc p")
+apply (rule_tac x = obj' in exI)
+apply (frule co2sobj_sproc_imp, erule exE, simp add:co2sobj_closefd)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_closefd)
+apply (case_tac "f \<in> same_inode_files s a")
+apply (rule_tac x = "O_file f'" in exI)
+apply (simp add:co2sobj_simps is_file_simps split:if_splits option.splits t_sobject.splits)
+apply (rule conjI, rule notI, simp add:same_inode_files_prop9)
+apply (rule impI, simp add:co2sobj.simps cf2sfiles_prop tainted_eq_Tainted same_inodes_Tainted)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_closefd is_file_simps)
+apply (rule notI, simp add:same_inode_files_prop9)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+
+apply (erule disjE)
+apply (drule same_inode_files_prop10, simp add:file_of_pfd_is_file, erule exE, erule conjE)
+apply (rule_tac x = "O_file f'a" in exI)
+apply (frule same_inode_files_prop11)
+apply (frule_tac obj = "O_file f'a" in co2sobj_closefd)
+apply (simp add:alive_simps)+
+apply (frule_tac f = "f'a" in is_file_has_sfile', simp, erule exE)
+apply (simp add:co2sobj.simps cf2sfiles_prop tainted_eq_Tainted same_inodes_Tainted split:if_splits)
+apply (rule impI, erule bexE, erule conjE, erule_tac x = f'' in ballE, simp, simp)
+apply (erule bexE, erule conjE, erule_tac x = f'' in ballE, simp, simp)
+apply (erule disjE)
+apply (rule_tac x = "O_proc p" in exI)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule conjE, erule exE, erule conjE)
+apply (case_tac "obj = O_proc p")
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_closefd)
+apply (case_tac "f \<in> same_inode_files s a")
+apply (rule_tac x = "O_file f'" in exI)
+apply (simp add:co2sobj_simps is_file_simps split:if_splits option.splits t_sobject.splits)
+apply (rule conjI, rule notI, simp add:same_inode_files_prop9)
+apply (rule impI, simp add:co2sobj.simps cf2sfiles_prop tainted_eq_Tainted same_inodes_Tainted)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_closefd is_file_simps)
+apply (rule notI, simp add:same_inode_files_prop9)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+
+apply (rule impI, rule conjI, rule impI)
+apply (tactic {*my_seteq_tac 1*})
+apply (simp add:update_s2ss_obj_def update_s2ss_sfile_del_def)
+apply (rule conjI| rule impI|erule exE|erule conjE)+
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa = p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (case_tac "f = a", simp add:alive_simps)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, clarsimp simp:alive_simps split:if_splits)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule conjI| rule impI|erule exE|erule conjE)+
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa = p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+apply (rule notI, simp, erule_tac x = "O_proc pa" in allE, simp add:co2sobj_closefd)
+apply (case_tac "f = a", simp add:alive_simps)
+apply (case_tac "f \<in> same_inode_files s a", rule disjI2)
+apply (simp add:co2sobj_simps split:if_splits option.splits t_sobject.splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted cf2sfiles_prop)
+apply (erule bexE, erule conjE)
+apply (rule conjI, rule_tac x = "O_file f''" in exI)
+apply (simp add:same_inode_files_prop11 co2sobj.simps tainted_eq_Tainted cf2sfiles_prop same_inodes_Tainted)
+apply (rule notI, simp)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+
+apply (erule bexE, erule conjE)
+apply (simp add:update_s2ss_obj_def split:if_splits)
+apply (erule disjE, rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule exE, erule conjE)
+apply (case_tac "obj = O_proc p")
+apply (rule_tac x = obj' in exI)
+apply (frule co2sobj_sproc_imp, erule exE, simp add:co2sobj_closefd)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (case_tac "obj = O_file a")
+apply (rule_tac x = "O_file f'" in exI)
+apply (frule_tac obj = "O_file f'" in co2sobj_closefd, simp add:alive_simps same_inode_files_prop11)
+apply (simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)
+apply (rule conjI)
+apply (rule impI)
+apply (rule_tac x = f' in ballE, simp, simp, simp)
+apply (simp add:same_inode_files_prop11 co2sobj.simps cf2sfiles_prop same_inodes_Tainted)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in alive_co2sobj_closefd3, simp+)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+apply (erule disjE, rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule conjE, erule exE, erule conjE)
+apply (case_tac "obj = O_proc p")
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (case_tac "obj = O_file a")
+apply (rule_tac x = "O_file f'" in exI)
+apply (frule_tac obj = "O_file f'" in co2sobj_closefd, simp add:alive_simps same_inode_files_prop11)
+apply (simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)
+apply (rule conjI)
+apply (rule impI)
+apply (rule_tac x = f' in ballE, simp, simp, simp)
+apply (simp add:same_inode_files_prop11 co2sobj.simps cf2sfiles_prop same_inodes_Tainted)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in alive_co2sobj_closefd3, simp+)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:alive_simps co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+
+apply (rule impI)
+apply (tactic {*my_seteq_tac 1*})
+apply (simp add:update_s2ss_obj_def update_s2ss_sfile_del_def)
+apply (rule conjI| rule impI|erule exE|erule conjE)+
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa = p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (case_tac "f = a", simp add:alive_simps)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, clarsimp simp:alive_simps split:if_splits)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, rule notI, simp)
+apply (frule_tac obj = obj in co2sobj_sfile_imp, erule exE, simp add:is_file_simps split:if_splits)
+apply (erule_tac x= f in allE, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule conjI| rule impI|erule exE|erule conjE)+
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa = p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule notI, simp)
+apply (rule disjI2, rule conjI, rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (case_tac "f = a", simp add:alive_simps)
+apply (case_tac "f \<in> same_inode_files s a", rule disjI1)
+apply (simp add:co2sobj_simps split:if_splits option.splits t_sobject.splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted cf2sfiles_prop)
+apply (erule bexE, erule conjE)
+apply (erule_tac x = f'' in ballE, simp, simp)
+apply (rule disjI2, rule conjI, rule disjI2, rule_tac x = obj in exI)
+apply (simp add:is_file_simps co2sobj_closefd tainted_eq_Tainted)
+apply (rule notI, simp add:co2sobj_closefd)
+apply (erule_tac x = f in allE, simp add:is_file_simps co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule conjI, rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule disjI2, rule conjI, rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule disjI2, rule conjI, rule disjI2, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule impI, rule conjI, rule impI)
+
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa = p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule notI, simp)
+apply (rule conjI, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+apply (erule_tac x = obj in allE, simp add:co2sobj_closefd)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (case_tac "f = a", simp add:alive_simps)
+apply (case_tac "f \<in> same_inode_files s a", rule conjI, rule disjI2, rule conjI)
+apply (simp add:co2sobj_simps split:if_splits option.splits t_sobject.splits)
+apply (simp add:co2sobj.simps)
+apply (rule notI, simp add:co2sobj.simps)
+apply (rule conjI, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (simp add:is_file_simps co2sobj_closefd tainted_eq_Tainted)
+apply (rule notI, simp add:co2sobj.simps)
+apply (rule notI, simp add:co2sobj_closefd tainted_eq_Tainted)
+apply (erule_tac x = f in allE, simp add:is_file_simps co2sobj.simps tainted_eq_Tainted)
+apply (rule conjI, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)+
+apply (rule conjI, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)+
+apply (rule conjI, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)+
+apply (rule impI)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa = p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule notI, simp)
+apply (rule disjI2, rule conjI, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+apply (erule_tac x = obj in allE, simp add:co2sobj_closefd)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (case_tac "f = a", simp add:alive_simps)
+apply (case_tac "f \<in> same_inode_files s a", rule disjI1)
+apply (simp add:co2sobj_closefd split:if_splits option.splits t_sobject.splits)
+apply (simp add:co2sobj.simps cf2sfiles_prop same_inodes_Tainted tainted_eq_Tainted)
+apply (erule bexE, erule conjE, erule_tac x = "f''" in ballE, simp, simp)
+apply (rule disjI2, rule conjI, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (simp add:is_file_simps co2sobj_closefd tainted_eq_Tainted)
+apply (rule notI, simp add:co2sobj.simps)
+apply (rule notI, simp add:co2sobj_closefd tainted_eq_Tainted)
+apply (erule_tac x = f in allE, simp add:is_file_simps co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule conjI, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)+
+apply (rule disjI2, rule conjI, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)+
+apply (rule disjI2, rule conjI, rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp add:alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)+
+
+apply (simp add:update_s2ss_sfile_del_def update_s2ss_obj_def split:if_splits)
+apply (erule conjE, erule disjE)
+apply (rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule exE, erule conjE)
+apply (case_tac "obj = O_file a", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (case_tac "obj = O_proc p")
+apply (rule_tac x = obj' in exI, frule_tac obj = obj' in co2sobj_sproc_imp, erule exE)
+apply (frule_tac obj = obj' in alive_co2sobj_closefd3, simp+)
+apply (simp add:co2sobj_closefd)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (frule_tac obj = obj in alive_co2sobj_closefd3, simp+)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+apply (erule conjE|erule exE|erule disjE)+
+apply (rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule conjE, erule exE, erule conjE)
+apply (case_tac "obj = O_file a", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (case_tac "obj = O_proc p")
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (frule_tac obj = obj in alive_co2sobj_closefd3, simp+)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)[1]
+apply (erule conjE|erule exE|erule disjE)+
+apply (drule same_inode_files_prop10, simp add:file_of_pfd_is_file, erule exE, erule conjE)
+apply (rule_tac x = "O_file f'" in exI)
+apply (frule same_inode_files_prop11)
+apply (frule_tac obj = "O_file f'" in co2sobj_closefd)
+apply (simp add:alive_simps)+
+apply (frule_tac f = "f'" in is_file_has_sfile', simp, erule exE)
+apply (simp add:co2sobj.simps cf2sfiles_prop tainted_eq_Tainted same_inodes_Tainted split:if_splits)
+apply (rule impI, erule bexE, erule conjE, erule_tac x = f'' in ballE, simp, simp)
+apply (erule bexE, erule conjE, erule_tac x = f'' in ballE, simp, simp)
+apply (erule conjE, erule disjE)
+apply (rule_tac x = "O_proc p" in exI)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule exE, erule conjE)
+apply (case_tac "obj = O_proc p")
+apply (rule_tac x = "obj'" in exI, simp, frule_tac obj = obj' in co2sobj_sproc_imp, erule exE)
+apply (frule_tac obj = obj' in alive_co2sobj_closefd3, simp+)
+apply (simp add:co2sobj_closefd tainted_eq_Tainted)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (case_tac "obj = O_file a", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (frule_tac obj = obj in alive_co2sobj_closefd3, simp+)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_closefd)
+apply (case_tac "f = a", simp add:alive_simps)
+apply (case_tac "f \<in> same_inode_files s a")
+apply (simp add:co2sobj.simps cf2sfiles_prop tainted_eq_Tainted same_inodes_Tainted)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+apply (erule disjE)
+apply (drule same_inode_files_prop10, simp add:file_of_pfd_is_file, erule exE, erule conjE)
+apply (rule_tac x = "O_file f'" in exI)
+apply (frule same_inode_files_prop11)
+apply (frule_tac obj = "O_file f'" in co2sobj_closefd)
+apply (simp add:alive_simps)+
+apply (frule_tac f = "f'" in is_file_has_sfile', simp, erule exE)
+apply (simp add:co2sobj.simps cf2sfiles_prop tainted_eq_Tainted same_inodes_Tainted split:if_splits)
+apply (rule impI, erule bexE, erule conjE, erule_tac x = f'' in ballE, simp, simp)
+apply (erule bexE, erule conjE, erule_tac x = f'' in ballE, simp, simp)
+apply (erule conjE, erule disjE)
+apply (rule_tac x = "O_proc p" in exI)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule conjE, erule exE, erule conjE)
+apply (case_tac "obj = O_proc p", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (case_tac "obj = O_file a", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_closefd)
+apply (case_tac "f \<in> same_inode_files s a")
+apply (simp add:co2sobj.simps cf2sfiles_prop tainted_eq_Tainted same_inodes_Tainted)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd is_file_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_closefd)
+
+apply (rule impI)
+apply (simp add:update_s2ss_obj_def)
+apply (rule conjI, rule impI, erule exE, erule conjE)
+apply (simp add:s2ss_def)
+apply (rule set_eqI, rule iffI, erule CollectE, erule exE, erule conjE)
+apply (simp)
+apply (case_tac "obj = O_proc p")
+apply (simp add:co2sobj.simps tainted_eq_Tainted split:if_splits)
+apply (rule disjI2, rule_tac x = obj in exI, erule conjE)
+apply (simp add:alive_co2sobj_closefd')
+apply (frule_tac obj = obj in co2sobj_closefd, simp, simp split:t_object.splits if_splits)
+apply (simp, erule disjE)
+apply (rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule exE, erule conjE)
+apply (case_tac "obj = O_proc p")
+apply (rule_tac x = obj' in exI, simp add:alive_co2sobj_closefd1)
+apply (frule_tac obj = obj' in co2sobj_closefd, simp add:alive_co2sobj_closefd1)
+apply (clarsimp split:t_object.splits if_splits option.splits simp:tainted_eq_Tainted co2sobj.simps)
+apply (rule_tac x = obj in exI, simp add:alive_co2sobj_closefd1)
+apply (frule_tac obj = obj in co2sobj_closefd, simp add:alive_co2sobj_closefd1)
+apply (clarsimp split:t_object.splits if_splits option.splits simp:tainted_eq_Tainted co2sobj.simps)
+apply (rule impI)
+apply (simp add:s2ss_def)
+apply (rule set_eqI, rule iffI, erule CollectE, erule exE, erule conjE)
+apply (simp)
+apply (case_tac "obj = O_proc p")
+apply (rule disjI1, simp add:co2sobj.simps tainted_eq_Tainted split:if_splits)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI, simp add:alive_co2sobj_closefd')
+apply (frule_tac obj = obj in co2sobj_closefd, simp add:alive_co2sobj_closefd1)
+apply (clarsimp split:t_object.splits if_splits option.splits simp:tainted_eq_Tainted co2sobj.simps)
+apply (rule notI, erule_tac x = obj in allE, simp add:alive_co2sobj_closefd')
+apply (frule_tac obj = obj in co2sobj_closefd, simp add:alive_co2sobj_closefd1)
+apply (clarsimp split:t_object.splits if_splits option.splits)
+apply (simp)
+apply (erule disjE)
+apply (rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule exE|erule conjE)+
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_closefd, simp add:alive_co2sobj_closefd1)
+apply (clarsimp split:t_object.splits if_splits option.splits
+ simp:tainted_eq_Tainted co2sobj.simps alive_co2sobj_closefd1)
+done
+
+lemma alive_co2sobj_unlink:
+ "\<lbrakk>alive s obj; valid (UnLink p f # s); obj \<noteq> O_file f\<rbrakk>
+ \<Longrightarrow> alive (UnLink p f # s) obj"
+by (auto simp add:alive_simps split:t_object.splits)
+
+lemma s2ss_unlink:
+ "valid (UnLink p f # s) \<Longrightarrow> s2ss (UnLink p f # s) = (
+ if (proc_fd_of_file s f = {})
+ then (case (cf2sfile s f) of
+ Some sf \<Rightarrow> del_s2ss_file s (s2ss s) f sf
+ | _ \<Rightarrow> {})
+ else s2ss s)"
+apply (frule vd_cons, frule vt_grant_os, clarsimp split:if_splits)
+apply (frule is_file_has_sfile', simp, erule exE, simp)
+apply (rule conjI, rule impI)
+apply (simp add:update_s2ss_sfile_del_def del_s2ss_file_def)
+apply (rule impI|erule conjE|erule exE|rule conjI|erule bexE)+ defer
+apply (rule impI|erule conjE|erule exE|rule conjI|erule bexE)+
+
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_file f", simp add:is_file_simps)
+apply simp
+apply (rule conjI)
+apply (rule_tac x = obj in exI,simp add:co2sobj_unlink is_file_simps is_dir_simps split:t_object.splits)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (rule notI, simp, frule_tac obj = obj in co2sobj_sfile_imp, erule exE, simp)
+apply (frule_tac obj = obj in co2sobj_unlink, simp)
+apply (erule_tac x = fa in allE, simp add:is_file_simps)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_file f", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (frule_tac alive_co2sobj_unlink, simp, simp)
+apply (frule_tac obj = obj in co2sobj_unlink, simp)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj.simps tainted_eq_Tainted split:t_object.splits if_splits)
+
+apply (rule impI|erule conjE|erule exE|rule conjI|erule bexE)+ defer
+
+apply (rule impI|erule conjE|erule exE|rule conjI|erule bexE)+
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_file f", simp add:alive_simps)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_unlink)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (rule disjI1)
+apply (simp add:co2sobj_unlink)
+apply (simp add:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted cf2sfiles_prop split:if_splits)
+apply (erule bexE, erule_tac x = f'' in ballE, simp, simp)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_unlink is_file_simps)
+apply (rule disjI2, simp, rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (rule disjI2, simp, rule_tac x = obj in exI, simp add:co2sobj_unlink is_dir_simps)
+apply (rule disjI2, simp, rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (tactic {*my_setiff_tac 1*})
+apply (drule same_inode_files_prop10, simp, erule exE, erule conjE)
+apply (rule_tac x = "O_file f'a" in exI, simp add:is_file_simps)
+apply (frule_tac obj = "O_file f'a" in co2sobj_unlink, simp add:same_inode_files_prop11 is_file_simps)
+apply (simp add:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted cf2sfiles_prop
+ is_file_simps same_inode_files_prop11 split:if_splits)
+apply (rule impI, erule bexE, erule_tac x = f'' in ballE, simp, simp)
+apply (erule bexE, erule_tac x = f'' in ballE, simp, simp)
+
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "f' = f", simp add:same_inode_files_prop9)
+apply (case_tac "obj= O_file f")
+apply (rule_tac x = "O_file f'" in exI, simp add:is_file_simps)
+apply (frule_tac f' = f' in cf2sfiles_unlink, simp add:current_files_simps is_file_in_current)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (rule_tac x = "O_file f'" in exI)
+apply (frule_tac f' = f' in cf2sfiles_unlink, simp add:current_files_simps is_file_in_current)
+apply (simp add:co2sobj.simps tainted_eq_Tainted is_file_simps cf2sfiles_prop same_inodes_Tainted)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink is_file_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+
+apply (rule impI|erule conjE|erule exE|rule conjI|erule bexE)+ defer
+
+apply (rule impI|erule conjE|erule exE|rule conjI|erule bexE)+
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_file f", simp add:alive_simps, simp)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule disjI2, rule conjI, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_unlink)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (rule disjI1)
+apply (simp add:co2sobj_unlink)
+apply (simp add:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted cf2sfiles_prop split:if_splits)
+apply (erule bexE, erule_tac x = f'' in ballE, simp, simp)
+apply (rule disjI2, rule conjI, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_unlink is_file_simps)
+apply (rule notI, simp add:co2sobj_unlink tainted_eq_Tainted)
+apply (erule_tac x = fa in allE, simp add:co2sobj.simps is_file_simps tainted_eq_Tainted)
+apply (rule disjI2, rule conjI, simp, rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule disjI2, rule conjI, simp, rule_tac x = obj in exI, simp add:co2sobj_unlink is_dir_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule disjI2, rule conjI, simp, rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (drule same_inode_files_prop10, simp, erule exE, erule conjE)
+apply (rule_tac x = "O_file f'" in exI, simp add:is_file_simps)
+apply (frule_tac obj = "O_file f'" in co2sobj_unlink, simp add:same_inode_files_prop11 is_file_simps)
+apply (simp add:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted cf2sfiles_prop
+ is_file_simps same_inode_files_prop11 split:if_splits)
+apply (rule impI, erule bexE, erule_tac x = f'' in ballE, simp, simp)
+apply (erule bexE, erule_tac x = f'' in ballE, simp, simp)
+apply (tactic {*my_setiff_tac 1*}, simp)
+apply (case_tac "obj = O_file f", simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (simp add:co2sobj.simps cf2sfiles_prop same_inodes_Tainted tainted_eq_Tainted)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink is_file_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+
+apply (rule impI)
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_unlink is_file_simps is_dir_simps split:t_object.splits)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (subgoal_tac "alive (UnLink p f # s) obj")
+apply (auto simp add:co2sobj_unlink is_file_simps is_dir_simps split:t_object.splits)[1]
+apply (auto simp add:co2sobj_unlink alive_simps split:t_object.splits)[1]
+
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_file f", simp add:alive_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink is_file_simps is_dir_simps split:t_object.splits)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_file f")
+apply (rule_tac x = "O_file f'" in exI)
+apply (auto simp add:co2sobj_unlink is_file_simps is_dir_simps split:t_object.splits)[1]
+apply (rule_tac x =obj in exI)
+apply (subgoal_tac "alive (UnLink p f # s) obj")
+apply (auto simp add:co2sobj_unlink is_file_simps is_dir_simps split:t_object.splits)[1]
+apply (auto simp add:co2sobj_unlink alive_simps split:t_object.splits)[1]
+
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_file f", simp add:alive_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink is_file_simps is_dir_simps split:t_object.splits if_splits)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_file f")
+apply (rule_tac x = "O_file f'" in exI)
+apply (auto simp add:co2sobj_unlink is_file_simps is_dir_simps same_inode_files_prop9 split:t_object.splits)[1]
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (rule_tac x = "O_file f'" in exI)
+apply (simp add:alive_simps co2sobj.simps)
+apply (rule conjI, rule notI, simp add:same_inode_files_prop9)
+apply (rule impI, frule_tac f' = f' in cf2sfiles_unlink)
+apply (simp add:current_files_simps is_file_simps is_file_in_current)
+apply (simp add:tainted_eq_Tainted same_inodes_Tainted cf2sfiles_prop)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink is_file_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_file f", simp add:alive_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink is_file_simps is_dir_simps split:t_object.splits if_splits)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_file f")
+apply (rule_tac x = "O_file f'" in exI)
+apply (subgoal_tac "alive (UnLink p f # s) (O_file f')")
+apply (frule same_inode_files_prop11, frule_tac f = f' in is_file_has_sfile', simp add:vd_cons, erule exE)
+apply (frule_tac obj = "O_file f'" in co2sobj_unlink, simp)
+apply (simp split:if_splits option.splits add:is_file_simps)
+apply (simp add:co2sobj.simps cf2sfiles_prop same_inodes_Tainted tainted_eq_Tainted)
+apply (auto split:t_sobject.splits)[1]
+apply (simp add:is_file_simps same_inode_files_prop11)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (rule_tac x = "O_file f'" in exI)
+apply (subgoal_tac "alive (UnLink p f # s) (O_file f')")
+apply (frule same_inode_files_prop11, frule_tac f = f' in is_file_has_sfile', simp add:vd_cons, erule exE)
+apply (frule_tac obj = "O_file f'" in co2sobj_unlink, simp)
+apply (simp split:if_splits option.splits add:is_file_simps)
+apply (simp add:co2sobj.simps cf2sfiles_prop same_inodes_Tainted tainted_eq_Tainted)
+apply (auto split:t_sobject.splits)[1]
+apply (simp add:is_file_simps same_inode_files_prop11)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink is_file_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_unlink)
+done
+
+definition del_s2ss_obj :: "t_state \<Rightarrow> t_static_state \<Rightarrow> t_object \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
+where
+ "del_s2ss_obj s ss obj sobj \<equiv>
+ if (\<exists> obj'. alive s obj' \<and> obj' \<noteq> obj \<and> co2sobj s obj' = Some sobj)
+ then ss
+ else ss - {sobj}"
+
+lemma del_update_s2ss_obj:
+ "update_s2ss_obj s ss obj sobj sobj' = del_s2ss_obj s ss obj sobj \<union> {sobj'}"
+by (auto simp:update_s2ss_obj_def del_s2ss_obj_def split:if_splits)
+
+lemma s2ss_rmdir: "valid (Rmdir p f # s) \<Longrightarrow> s2ss (Rmdir p f # s) = (
+ case (co2sobj s (O_dir f)) of
+ Some sdir \<Rightarrow> del_s2ss_obj s (s2ss s) (O_dir f) sdir
+ | _ \<Rightarrow> {})"
+apply (frule vd_cons, frule vt_grant_os)
+apply (clarsimp simp:dir_is_empty_def)
+apply (frule is_dir_has_sdir', simp, erule exE)
+apply (simp split:option.splits, rule conjI, rule impI, simp add:co2sobj.simps)
+apply (rule allI, rule impI)
+
+apply (simp add:del_s2ss_obj_def)
+apply (rule conjI|rule impI|erule exE|erule conjE)+
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_rmdir is_file_simps is_dir_simps alive_simps split:t_object.splits if_splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_dir f")
+apply (rule_tac x = obj' in exI)
+apply (subgoal_tac "alive (Rmdir p f # s) obj'")
+apply (auto simp add:co2sobj_rmdir is_file_simps is_dir_simps split:t_object.splits)[1]
+apply (simp add:alive_rmdir)
+apply (rule_tac x = obj in exI)
+apply (subgoal_tac "alive (Rmdir p f # s) obj")
+apply (auto simp add:co2sobj_rmdir is_file_simps is_dir_simps split:t_object.splits)[1]
+apply (simp add:alive_rmdir)
+
+apply (rule conjI|rule impI|erule exE|erule conjE)+
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply simp
+apply (case_tac "obj = O_dir f", simp add:alive_rmdir)
+apply (rule conjI)
+apply (rule_tac x = obj in exI, simp add:co2sobj_rmdir alive_rmdir)
+apply (simp add:co2sobj_rmdir)
+apply (simp add:alive_rmdir, erule_tac x = obj in allE, simp)
+apply (tactic {*my_setiff_tac 1*}, simp)
+apply (case_tac "obj = O_dir f", simp)
+apply (rule_tac x = obj in exI, simp add:co2sobj_rmdir alive_rmdir)
+done
+
+lemma s2ss_mkdir: "valid (Mkdir p f inum # s) \<Longrightarrow> s2ss (Mkdir p f inum # s) = (
+ case (cf2sfile (Mkdir p f inum # s) f) of
+ Some sf \<Rightarrow> (s2ss s) \<union> {S_dir sf}
+ | _ \<Rightarrow> {})"
+apply (frule vt_grant_os, frule vd_cons, clarsimp)
+apply (case_tac "cf2sfile (Mkdir p f inum # s) f")
+apply (drule current_file_has_sfile', simp, simp add:current_files_simps, simp)
+
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*}, simp)
+apply (case_tac "obj = O_dir f")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2, rule_tac x = obj in exI, simp add:co2sobj_mkdir alive_simps)
+apply (tactic {*my_setiff_tac 1*}, simp)
+apply (rule_tac x = "O_dir f" in exI, simp add:alive_mkdir co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_dir f", simp add:is_dir_in_current)
+apply (rule_tac x = obj in exI, simp add:co2sobj_mkdir alive_mkdir)
+done
+
+definition update_s2ss_sfile:: "t_state \<Rightarrow> t_static_state \<Rightarrow> t_file \<Rightarrow> t_sfile \<Rightarrow> t_static_state"
+where
+ "update_s2ss_sfile s ss f sf \<equiv>
+ if (\<exists> f'. is_file s f' \<and> f' \<notin> same_inode_files s f \<and> co2sobj s (O_file f') = co2sobj s (O_file f))
+ then ss \<union> {S_file (cf2sfiles s f \<union> {sf}) (O_file f \<in> Tainted s)}
+ else ss - {S_file (cf2sfiles s f) (O_file f \<in> Tainted s)}
+ \<union> {S_file (cf2sfiles s f \<union> {sf}) (O_file f \<in> Tainted s)}"
+
+lemma s2ss_linkhard: "valid (LinkHard p f f' # s) \<Longrightarrow> s2ss (LinkHard p f f' # s) = (
+ case (cf2sfile (LinkHard p f f' # s) f') of
+ Some sf \<Rightarrow> update_s2ss_sfile s (s2ss s) f sf
+ | _ \<Rightarrow> {})"
+apply (frule vt_grant_os, frule vd_cons, clarsimp)
+apply (split option.splits)
+apply (rule conjI, rule impI, drule current_file_has_sfile', simp, simp add:current_files_simps)
+apply (rule allI, rule impI)
+
+apply (simp add:update_s2ss_sfile_def)
+apply (rule conjI, rule impI, erule exE, erule conjE, erule conjE)
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_file f'")
+apply (rule disjI1, simp add:co2sobj.simps cf2sfiles_linkhard tainted_eq_Tainted
+ same_inode_files_linkhard split:if_splits)
+apply (case_tac "O_file f' \<in> Tainted s")
+apply (drule Tainted_in_current, simp, simp add:is_file_in_current alive.simps, simp)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_linkhard alive_linkhard)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (rule disjI1, simp add:co2sobj.simps cf2sfiles_linkhard tainted_eq_Tainted
+ same_inodes_Tainted split:if_splits)
+apply (rule disjI2, simp, rule_tac x = obj in exI, simp add:co2sobj_linkhard is_file_simps)
+apply (rule disjI2, simp, rule_tac x = obj in exI, simp add:co2sobj_linkhard)
+apply (rule disjI2, simp, rule_tac x = obj in exI, simp add:co2sobj_linkhard is_dir_simps)
+apply (rule disjI2, simp, rule_tac x = obj in exI, simp add:co2sobj_linkhard)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_file f" in exI)
+apply (frule_tac obj = "O_file f" in co2sobj_linkhard)
+apply (simp add:alive_linkhard)
+apply (simp add:alive_linkhard same_inode_files_prop9 split:t_object.splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_file f'", simp add:alive_linkhard is_file_in_current)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI, simp add:co2sobj_linkhard alive_linkhard)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (rule_tac x = "O_file f'a" in exI, simp add:co2sobj_linkhard alive_linkhard)
+apply (rule conjI, rule impI, simp add:is_file_in_current)
+apply (rule impI, simp add:co2sobj.simps cf2sfiles_prop tainted_eq_Tainted same_inodes_Tainted)
+apply (rule_tac x = obj in exI, simp add:co2sobj_linkhard alive_linkhard)
+apply (rule_tac x = obj in exI, simp add:co2sobj_linkhard alive_linkhard)
+apply (rule_tac x = obj in exI, simp add:co2sobj_linkhard alive_linkhard)
+apply (rule_tac x = obj in exI, simp add:co2sobj_linkhard alive_linkhard)
+
+apply (rule impI)
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_file f'", simp)
+apply (rule disjI1, simp add:co2sobj.simps cf2sfiles_linkhard tainted_eq_Tainted
+ same_inode_files_linkhard split:if_splits)
+apply (case_tac "O_file f' \<in> Tainted s")
+apply (drule Tainted_in_current, simp, simp add:is_file_in_current alive.simps, simp)
+apply (erule_tac obj = obj in co2sobj_some_caseD, simp)
+apply (rule disjI2, rule conjI, rule_tac x = obj in exI)
+apply (simp add:co2sobj_linkhard alive_linkhard)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (rule disjI1, simp add:co2sobj.simps cf2sfiles_linkhard tainted_eq_Tainted
+ same_inodes_Tainted split:if_splits)
+apply (simp, rule disjI2, rule conjI, rule_tac x = obj in exI, simp add:co2sobj_linkhard is_file_simps)
+apply (erule_tac x = fa in allE, rule notI)
+apply (simp add:co2sobj_linkhard is_file_simps)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, simp, rule conjI, rule_tac x = obj in exI, simp add:co2sobj_linkhard)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule disjI2, simp, rule conjI, rule_tac x = obj in exI, simp add:co2sobj_linkhard is_dir_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule disjI2, simp, rule conjI, rule_tac x = obj in exI, simp add:co2sobj_linkhard)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_file f" in exI)
+apply (frule_tac obj = "O_file f" in co2sobj_linkhard)
+apply (simp add:alive_linkhard)
+apply (simp add:alive_linkhard same_inode_files_prop9 split:t_object.splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_file f'", simp add:alive_linkhard is_file_in_current)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI, simp add:co2sobj_linkhard alive_linkhard)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (simp add:co2sobj.simps cf2sfiles_prop tainted_eq_Tainted same_inodes_Tainted)
+apply (rule_tac x = obj in exI, simp add:co2sobj_linkhard alive_linkhard)
+apply (rule_tac x = obj in exI, simp add:co2sobj_linkhard alive_linkhard)
+apply (rule_tac x = obj in exI, simp add:co2sobj_linkhard alive_linkhard)
+apply (rule_tac x = obj in exI, simp add:co2sobj_linkhard alive_linkhard)
+done
+
+lemma same_inode_files_prop12:
+ "is_file s f \<Longrightarrow> f \<in> same_inode_files s f "
+by (auto simp:is_file_def same_inode_files_def split:option.splits)
+
+definition update_s2ss_sfile_tainted:: "t_state \<Rightarrow> t_static_state \<Rightarrow> t_file \<Rightarrow> bool \<Rightarrow> t_static_state"
+where
+ "update_s2ss_sfile_tainted s ss f tag \<equiv>
+ if (\<exists> f'. is_file s f' \<and> f' \<notin> same_inode_files s f \<and>
+ co2sobj s (O_file f') = Some (S_file (cf2sfiles s f) False))
+ then ss \<union> {S_file (cf2sfiles s f) True}
+ else ss - {S_file (cf2sfiles s f) False}
+ \<union> {S_file (cf2sfiles s f) True}"
+
+lemma s2ss_truncate: "valid (Truncate p f len # s) \<Longrightarrow> s2ss (Truncate p f len # s) = (
+ if (O_file f \<notin> Tainted s \<and> O_proc p \<in> Tainted s \<and> len > 0)
+ then update_s2ss_sfile_tainted s (s2ss s) f True
+ else s2ss s)"
+apply (frule vt_grant_os, frule vd_cons, simp split:if_splits)
+apply (rule conjI, rule impI, (erule conjE)+)
+
+apply (simp add:update_s2ss_sfile_tainted_def)
+apply (rule conjI|rule impI|erule exE|erule conjE)+
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_file f")
+apply (rule disjI1, simp add:co2sobj.simps tainted_eq_Tainted same_inode_files_prop12 cf2sfiles_other)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule disjI2, simp, rule_tac x = obj in exI, simp add:co2sobj_truncate alive_simps)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (rule disjI1, simp add:co2sobj.simps tainted_eq_Tainted cf2sfiles_prop cf2sfiles_other)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_truncate is_file_simps)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_truncate)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_truncate is_dir_simps)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_truncate)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_file f" in exI)
+apply (simp add:co2sobj.simps is_file_simps tainted_eq_Tainted cf2sfiles_other same_inode_files_prop12)
+apply (tactic {*my_setiff_tac 1*})
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI, simp add:co2sobj_truncate)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (rule_tac x = "O_file f'" in exI)
+apply (auto simp:co2sobj_truncate is_file_simps is_dir_simps split:t_object.splits)[1]
+apply (simp add:co2sobj.simps same_inodes_Tainted tainted_eq_Tainted cf2sfiles_prop)
+apply (rule_tac x = obj in exI, simp add:co2sobj_truncate is_file_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_truncate)
+apply (rule_tac x = obj in exI, simp add:co2sobj_truncate is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_truncate)
+
+apply (rule conjI|rule impI|erule exE|erule conjE)+
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_file f")
+apply (rule disjI1, simp add:co2sobj.simps tainted_eq_Tainted same_inode_files_prop12 cf2sfiles_other)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule disjI2, simp, rule conjI, rule_tac x = obj in exI, simp add:co2sobj_truncate alive_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (rule disjI1, simp add:co2sobj.simps tainted_eq_Tainted cf2sfiles_prop cf2sfiles_other)
+apply (rule disjI2, simp, rule conjI, rule_tac x = obj in exI)
+apply (simp add:co2sobj_truncate is_file_simps)
+apply (rule notI, simp add:co2sobj_truncate is_file_simps)
+apply (erule_tac x = fa in allE)
+apply (simp add:co2sobj.simps)
+apply (rule disjI2, simp, rule conjI, rule_tac x = obj in exI)
+apply (simp add:co2sobj_truncate)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule disjI2, simp, rule conjI, rule_tac x = obj in exI)
+apply (simp add:co2sobj_truncate is_dir_simps)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule disjI2, simp, rule conjI, rule_tac x = obj in exI)
+apply (simp add:co2sobj_truncate)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_file f" in exI)
+apply (simp add:co2sobj.simps is_file_simps tainted_eq_Tainted cf2sfiles_other same_inode_files_prop12)
+apply (tactic {*my_setiff_tac 1*})
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (rule_tac x = obj in exI, simp add:co2sobj_truncate)
+apply (case_tac "fa \<in> same_inode_files s f")
+apply (simp add:co2sobj.simps same_inodes_Tainted tainted_eq_Tainted cf2sfiles_prop)
+apply (rule_tac x = obj in exI, simp add:co2sobj_truncate is_file_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_truncate)
+apply (rule_tac x = obj in exI, simp add:co2sobj_truncate is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_truncate)
+
+apply (rule impI, simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:alive_simps co2sobj_truncate)
+apply (simp split:t_object.splits if_splits add:co2sobj.simps tainted_eq_Tainted)
+apply (case_tac "O_proc p \<in> Tainted s", simp add:same_inodes_Tainted)
+apply simp
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:alive_simps co2sobj_truncate)
+apply (auto split:t_object.splits if_splits simp:co2sobj.simps tainted_eq_Tainted same_inodes_Tainted)
+done
+
+lemma s2ss_createmsgq: "valid (CreateMsgq p q # s) \<Longrightarrow> s2ss (CreateMsgq p q # s) =
+ (case (cq2smsgq (CreateMsgq p q # s) q) of
+ Some sq \<Rightarrow> s2ss s \<union> {S_msgq sq}
+ | _ \<Rightarrow> {})"
+apply (frule vd_cons, frule vt_grant_os, clarsimp)
+apply (case_tac "cq2smsgq (CreateMsgq p q # s) q")
+apply (drule current_has_smsgq', simp+)
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_createmsgq is_file_simps is_dir_simps split:t_object.splits if_splits)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_msgq q" in exI, simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply simp
+apply (rule_tac x = obj in exI)
+apply (auto simp add:co2sobj_createmsgq alive_simps split:t_object.splits if_splits)
+done
+
+ML {*fun my_clarify_tac i =
+REPEAT (( rtac @{thm impI}
+ ORELSE' rtac @{thm allI}
+ ORELSE' rtac @{thm ballI}
+ ORELSE' rtac @{thm conjI}
+ ORELSE' etac @{thm conjE}
+ ORELSE' etac @{thm exE}
+ ORELSE' etac @{thm bexE}
+ ORELSE' etac @{thm disjE}) i)
+*}
+
+lemma s2ss_sendmsg: "valid (SendMsg p q m # s) \<Longrightarrow> s2ss (SendMsg p q m # s) = (
+ case (cq2smsgq s q, cq2smsgq (SendMsg p q m # s) q) of
+ (Some sq, Some sq') \<Rightarrow> update_s2ss_obj s (s2ss s) (O_msgq q) (S_msgq sq) (S_msgq sq')
+ | _ \<Rightarrow> {})"
+apply (frule vd_cons, frule vt_grant_os, clarsimp)
+apply (case_tac "cq2smsgq s q")
+apply (drule current_has_smsgq', simp+)
+apply (case_tac "cq2smsgq (SendMsg p q m # s) q")
+apply (drule current_has_smsgq', simp+)
+
+apply (simp add:update_s2ss_obj_def)
+apply (tactic {*my_clarify_tac 1*})
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_sendmsg is_file_simps is_dir_simps split:t_object.splits if_splits)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_msgq q" in exI, simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule_tac x = obj' in exI)
+apply (simp add:co2sobj_sendmsg alive_sendmsg split:t_object.splits if_splits)
+apply (auto simp:co2sobj.simps)[1]
+apply (rule_tac x = obj in exI, simp add:co2sobj_sendmsg alive_sendmsg split:t_object.splits)
+apply (auto simp:co2sobj.simps)[1]
+
+apply (rule impI)
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2, simp, rule conjI, rule_tac x = obj in exI)
+apply (simp add:co2sobj_sendmsg is_file_simps is_dir_simps split:t_object.splits if_splits)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (rule notI, simp)
+apply (frule_tac obj = obj in co2sobj_smsgq_imp, erule exE, simp)
+apply (erule_tac x = obj in allE, simp add:co2sobj_sendmsg)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_msgq q" in exI, simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (simp add:co2sobj.simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_sendmsg alive_sendmsg split:t_object.splits)
+apply (auto simp:co2sobj.simps)[1]
+done
+
+lemma alive_co2sobj_removemsgq:
+ "\<lbrakk>alive s obj; valid (RemoveMsgq p q # s); co2sobj s obj = Some sobj; obj \<noteq> O_msgq q\<rbrakk>
+ \<Longrightarrow> alive (RemoveMsgq p q # s) obj"
+apply (erule co2sobj_some_caseD)
+apply (auto simp:is_file_simps is_dir_simps)
+done
+
+lemma s2ss_removemsgq: "valid (RemoveMsgq p q # s) \<Longrightarrow> s2ss (RemoveMsgq p q # s) =
+ (case (cq2smsgq s q) of
+ Some sq \<Rightarrow> del_s2ss_obj s (s2ss s) (O_msgq q) (S_msgq sq)
+ | _ \<Rightarrow> {})"
+apply (frule vd_cons, frule vt_grant_os, clarsimp)
+apply (split option.splits, rule conjI, rule impI)
+apply (drule current_has_smsgq', simp, simp)
+apply (rule allI, rule impI)
+
+apply (simp add:del_s2ss_obj_def)
+apply (tactic {*my_clarify_tac 1*})
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q", simp)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_removemsgq alive_simps split:t_object.splits if_splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_msgq q", simp)
+apply (rule_tac x = obj' in exI)
+apply (frule_tac obj = obj' in co2sobj_smsgq_imp, erule exE)
+apply (simp add:co2sobj_removemsgq alive_simps split:t_object.splits if_splits)
+apply (simp add:co2sobj.simps)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_removemsgq alive_co2sobj_removemsgq)
+
+apply (rule impI)
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q", simp)
+apply (simp, rule conjI, rule_tac x = obj in exI)
+apply (simp add:co2sobj_removemsgq alive_simps split:t_object.splits if_splits)
+apply (rule notI, simp, frule_tac obj = obj in co2sobj_smsgq_imp, erule exE)
+apply (erule_tac x = obj in allE, simp add:co2sobj_removemsgq alive_co2sobj_removemsgq)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_msgq q", simp)
+apply (simp add:co2sobj.simps)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_removemsgq alive_co2sobj_removemsgq)
+done
+
+declare Product_Type.split_paired_Ex Product_Type.split_paired_All [simp del]
+
+lemma s2ss_shm_no_bk_elim':
+ "\<lbrakk>x \<notin> s2ss_shm_no_backup s pfrom; co2sobj s (O_proc p) = Some x; O_proc p \<notin> Tainted s;
+ valid s; info_flow_shm s pfrom p\<rbrakk>
+ \<Longrightarrow> \<exists> p'. \<not> info_flow_shm s pfrom p' \<and> p' \<in> current_procs s \<and> co2sobj s (O_proc p') = Some x"
+apply (auto simp:s2ss_shm_no_backup_def co2sobj.simps tainted_eq_Tainted split:option.splits)
+apply (erule_tac x = p in allE, auto)
+apply (rule_tac x = p' in exI, auto)
+done
+
+lemma s2ss_recvmsg: "valid (RecvMsg p q m # s) \<Longrightarrow> s2ss (RecvMsg p q m # s) = (
+ case (cq2smsgq s q, cq2smsgq (RecvMsg p q m # s) q) of
+ (Some sq, Some sq') \<Rightarrow> if (O_msg q m \<in> Tainted s \<and> O_proc p \<notin> Tainted s)
+ then update_s2ss_obj s (update_s2ss_shm s p) (O_msgq q) (S_msgq sq) (S_msgq sq')
+ else update_s2ss_obj s (s2ss s) (O_msgq q) (S_msgq sq) (S_msgq sq')
+ | _ \<Rightarrow> {})"
+apply (frule vt_grant_os, frule vd_cons)
+apply (case_tac "cq2smsgq s q")
+apply (drule current_has_smsgq', simp, simp)
+apply (case_tac "cq2smsgq (RecvMsg p q m # s) q")
+apply (drule current_has_smsgq', simp, simp)
+apply (simp split:if_splits add:update_s2ss_obj_def)
+
+apply (tactic {*my_clarify_tac 1*})
+unfolding s2ss_def update_s2ss_shm_def
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2)
+apply (rule DiffI)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "info_flow_shm s p pa")
+apply (rule UnI2)
+apply (simp add:co2sobj_recvmsg del:Product_Type.split_paired_Ex
+ split:t_object.splits option.splits)
+apply (rule_tac x = ab in exI, simp, rule_tac x = pa in exI, simp)
+apply (rule UnI1,simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_recvmsg split:t_object.splits option.splits)
+apply (rule UnI1,simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps split:t_object.splits option.splits)
+apply (rule UnI1, simp,rule_tac x = obj in exI)
+apply (simp add:co2sobj_recvmsg split:t_object.splits option.splits)
+apply (rule UnI1, simp,rule_tac x = obj in exI)
+apply (simp add:co2sobj_recvmsg is_dir_simps split:t_object.splits option.splits)
+apply (rule UnI1, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_recvmsg split:t_object.splits option.splits if_splits)
+apply (rule notI, simp add:s2ss_shm_no_backup_def del:Product_Type.split_paired_Ex)
+apply (erule exE|erule conjE)+
+apply (simp, frule co2sobj_sproc_imp, erule exE, simp)
+apply (simp add:co2sobj_recvmsg split:if_splits option.splits)
+apply (erule_tac x = paa in allE, simp)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_msgq q" in exI, simp add:co2sobj.simps alive_recvmsg)
+apply (tactic {*my_setiff_tac 1*}, erule UnE)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule_tac x = obj' in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps is_dir_simps alive_recvmsg
+ split:t_object.splits if_splits option.splits)
+apply (auto simp:co2sobj.simps is_file_simps is_dir_simps)[1]
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "info_flow_shm s p pa \<and> O_proc pa \<notin> Tainted s")
+apply (drule_tac p = pa in s2ss_shm_no_bk_elim', simp+)
+apply (erule exE|erule conjE)+
+apply (rule_tac x = "O_proc p'" in exI)
+apply (simp add:co2sobj_recvmsg alive_simps)
+apply (rule_tac x = obj in exI)
+apply (case_tac "info_flow_shm s p pa", simp)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp+)
+apply (case_tac "cp2sproc s pa", simp add:co2sobj.simps)
+apply (simp, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (simp, frule_tac obj = obj in co2sobj_recvmsg, simp+)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg alive_simps)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp split:t_object.splits)
+apply (simp split:t_object.splits)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_proc pa" in exI)
+apply (frule_tac obj = "O_proc pa" in co2sobj_recvmsg)
+apply (simp add:info_shm_flow_in_procs, simp add:info_shm_flow_in_procs)
+
+apply (tactic {*my_clarify_tac 1*})
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps is_dir_simps split:t_object.splits if_splits option.splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted info_flow_shm_Tainted)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (case_tac "msgs_of_queue s q", simp, simp)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_msgq q" in exI)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1 *})
+apply (case_tac "obj = O_msgq q")
+apply (rule_tac x = obj' in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps is_dir_simps alive_recvmsg
+ split:t_object.splits if_splits option.splits)
+apply (auto simp:co2sobj.simps is_file_simps is_dir_simps)[1]
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps is_dir_simps alive_recvmsg
+ split:t_object.splits if_splits option.splits)
+apply (auto simp:co2sobj.simps is_file_simps is_dir_simps tainted_eq_Tainted info_flow_shm_Tainted)[1]
+
+apply (tactic {*my_clarify_tac 1*})
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "info_flow_shm s p pa")
+apply (case_tac "O_proc pa \<in> Tainted s")
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule s2ss_shm_no_bk_intro2, simp, simp)
+apply (rule DiffI, rule UnI2, rule CollectI)
+apply (simp only:co2sobj.simps tainted_eq_Tainted split:option.splits, simp)
+apply (rule_tac x = ab in exI, simp)
+apply (rule_tac x = pa in exI, simp add:cp2sproc_other)
+apply (rule s2ss_shm_no_bk_intro2, simp, simp, simp)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (rule_tac p = pa in s2ss_shm_no_bk_intro3)
+apply (simp add:co2sobj_recvmsg, simp, simp)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp add:is_file_simps split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp add:is_dir_simps split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_msgq q" in exI)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (erule UnE, tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule_tac x = obj' in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps is_dir_simps alive_recvmsg
+ split:t_object.splits if_splits option.splits)
+apply (auto simp:co2sobj.simps is_file_simps is_dir_simps)[1]
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "info_flow_shm s p pa \<and> O_proc pa \<notin> Tainted s")
+apply (drule_tac p = pa in s2ss_shm_no_bk_elim', simp+)
+apply (erule exE|erule conjE)+
+apply (rule_tac x = "O_proc p'" in exI)
+apply (simp add:co2sobj_recvmsg alive_simps)
+apply (rule_tac x = obj in exI)
+apply (case_tac "info_flow_shm s p pa", simp)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp+)
+apply (case_tac "cp2sproc s pa", simp add:co2sobj.simps)
+apply (simp, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (simp, frule_tac obj = obj in co2sobj_recvmsg, simp+)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg alive_simps)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp split:t_object.splits)
+apply (simp split:t_object.splits)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_proc pa" in exI)
+apply (frule_tac obj = "O_proc pa" in co2sobj_recvmsg)
+apply (simp add:info_shm_flow_in_procs, simp add:info_shm_flow_in_procs)
+
+apply (rule impI, tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2, simp, rule conjI)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps is_dir_simps alive_recvmsg
+ split:t_object.splits if_splits option.splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted info_flow_shm_Tainted)
+apply (simp add:co2sobj_recvmsg)
+apply (simp add:alive_recvmsg split:if_splits)
+apply (erule_tac x = obj in allE, simp split:t_object.splits if_splits option.splits)
+apply (rule notI, simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_msgq q" in exI)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply simp
+apply (tactic {*my_clarify_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule_tac x = obj' in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps is_dir_simps alive_recvmsg
+ split:t_object.splits if_splits option.splits)
+apply (auto simp:co2sobj.simps is_file_simps is_dir_simps)[1]
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg)
+apply (simp add:alive_recvmsg, rule notI, simp add:co2sobj.simps)
+apply (rule conjI)
+apply (simp add:alive_recvmsg, rule notI, simp add:co2sobj.simps)
+apply (auto simp:co2sobj.simps tainted_eq_Tainted split:t_object.splits option.splits if_splits)[1]
+
+apply (tactic {*my_clarify_tac 1*})
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "info_flow_shm s p pa")
+apply (case_tac "O_proc pa \<in> Tainted s")
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule s2ss_shm_no_bk_intro2, simp, simp)
+apply (rule DiffI, rule UnI2, rule CollectI)
+apply (simp only:co2sobj.simps tainted_eq_Tainted split:option.splits, simp)
+apply (rule_tac x = ab in exI, simp)
+apply (rule_tac x = pa in exI, simp add:cp2sproc_other)
+apply (rule s2ss_shm_no_bk_intro2, simp, simp, simp)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (rule_tac p = pa in s2ss_shm_no_bk_intro3)
+apply (simp add:co2sobj_recvmsg, simp, simp)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp add:is_file_simps split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp add:is_dir_simps split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_msgq q" in exI)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (erule UnE, tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule_tac x = obj' in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps is_dir_simps alive_recvmsg
+ split:t_object.splits if_splits option.splits)
+apply (auto simp:co2sobj.simps is_file_simps is_dir_simps)[1]
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "info_flow_shm s p pa \<and> O_proc pa \<notin> Tainted s")
+apply (drule_tac p = pa in s2ss_shm_no_bk_elim', simp+)
+apply (erule exE|erule conjE)+
+apply (rule_tac x = "O_proc p'" in exI)
+apply (simp add:co2sobj_recvmsg alive_simps)
+apply (rule_tac x = obj in exI)
+apply (case_tac "info_flow_shm s p pa", simp)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp+)
+apply (case_tac "cp2sproc s pa", simp add:co2sobj.simps)
+apply (simp, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (simp, frule_tac obj = obj in co2sobj_recvmsg, simp+)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg alive_simps)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp split:t_object.splits)
+apply (simp split:t_object.splits)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_proc pa" in exI)
+apply (frule_tac obj = "O_proc pa" in co2sobj_recvmsg)
+apply (simp add:info_shm_flow_in_procs, simp add:info_shm_flow_in_procs)
+
+apply (rule impI, tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2, simp)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps is_dir_simps alive_recvmsg
+ split:t_object.splits if_splits option.splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted info_flow_shm_Tainted)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_msgq q" in exI)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule_tac x = obj' in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps is_dir_simps alive_recvmsg
+ split:t_object.splits if_splits option.splits)
+apply (auto simp:co2sobj.simps is_file_simps is_dir_simps)[1]
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg)
+apply (simp add:alive_recvmsg, rule notI, simp add:co2sobj.simps)
+apply (rule conjI)
+apply (simp add:alive_recvmsg, rule notI, simp add:co2sobj.simps)
+apply (auto simp:co2sobj.simps tainted_eq_Tainted split:t_object.splits option.splits if_splits)[1]
+
+apply (tactic {*my_clarify_tac 1*})
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "info_flow_shm s p pa")
+apply (case_tac "O_proc pa \<in> Tainted s")
+apply (rule DiffI, rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (erule s2ss_shm_no_bk_intro2, simp, simp)
+apply (simp, rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule DiffI, rule DiffI, rule UnI2, rule CollectI)
+apply (simp only:co2sobj.simps tainted_eq_Tainted split:option.splits, simp)
+apply (rule_tac x = ab in exI, simp)
+apply (rule_tac x = pa in exI, simp add:cp2sproc_other)
+apply (rule s2ss_shm_no_bk_intro2, simp, simp, simp)
+apply (simp, rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule DiffI, rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (rule_tac p = pa in s2ss_shm_no_bk_intro3)
+apply (simp add:co2sobj_recvmsg, simp, simp)
+apply (simp, rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule DiffI, rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp add:is_file_simps split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+apply (simp, rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule DiffI, rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+apply (simp, rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule DiffI, rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp add:is_dir_simps split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+apply (simp, rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule DiffI, rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp, simp split:option.splits)
+apply (simp add:co2sobj_recvmsg s2ss_shm_no_bk_intro1)
+apply (simp, rule notI, erule_tac x = obj in allE, simp)
+apply (simp add:co2sobj_recvmsg)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_msgq q" in exI)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (erule UnE, tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (simp add:co2sobj.simps)
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "info_flow_shm s p pa \<and> O_proc pa \<notin> Tainted s")
+apply (drule_tac p = pa in s2ss_shm_no_bk_elim', simp+)
+apply (erule exE|erule conjE)+
+apply (rule_tac x = "O_proc p'" in exI)
+apply (simp add:co2sobj_recvmsg alive_simps)
+apply (rule_tac x = obj in exI)
+apply (case_tac "info_flow_shm s p pa", simp)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp+)
+apply (case_tac "cp2sproc s pa", simp add:co2sobj.simps)
+apply (simp, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (simp, frule_tac obj = obj in co2sobj_recvmsg, simp+)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg alive_simps)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg, simp split:t_object.splits)
+apply (simp split:t_object.splits)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_recvmsg)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_proc pa" in exI)
+apply (frule_tac obj = "O_proc pa" in co2sobj_recvmsg)
+apply (simp add:info_shm_flow_in_procs, simp add:info_shm_flow_in_procs)
+
+apply (rule impI, tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_msgq q")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2, simp, rule conjI)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_recvmsg is_file_simps is_dir_simps alive_recvmsg
+ split:t_object.splits if_splits option.splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted info_flow_shm_Tainted)
+apply (rule notI, simp)
+apply (frule co2sobj_smsgq_imp, erule exE, simp)
+apply (simp add:co2sobj_recvmsg)
+apply (erule_tac x = obj in allE, simp)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_msgq q" in exI)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (simp, erule exE, erule conjE)
+apply (case_tac "obj = O_msgq q")
+apply (simp add:co2sobj.simps)
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_recvmsg)
+apply (simp add:alive_recvmsg, rule notI, simp add:co2sobj.simps)
+apply (rule conjI)
+apply (simp add:alive_recvmsg, rule notI, simp add:co2sobj.simps)
+apply (auto simp:co2sobj.simps tainted_eq_Tainted intro: info_flow_shm_Tainted
+split:t_object.splits option.splits if_splits)[1]
+done
+
+lemma s2ss_createshm:
+ "valid (CreateShM p h # s) \<Longrightarrow> s2ss (CreateShM p h # s) =
+ (case (ch2sshm (CreateShM p h # s) h) of
+ Some sh \<Rightarrow> s2ss s \<union> {S_shm sh}
+ | _ \<Rightarrow> {})"
+apply (frule vd_cons, frule vt_grant_os, clarsimp)
+apply (case_tac "ch2sshm (CreateShM p h # s) h")
+apply (drule current_shm_has_sh', simp+)
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_shm h")
+apply (rule disjI1, simp add:co2sobj.simps)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (simp add:co2sobj_createshm is_file_simps is_dir_simps split:t_object.splits if_splits)
+apply (simp add:co2sobj.simps)
+apply (simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_shm h" in exI, simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_shm h")
+apply simp
+apply (rule_tac x = obj in exI)
+apply (auto simp add:co2sobj_createshm alive_simps split:t_object.splits if_splits)
+done
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/S2ss_prop2.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,449 @@
+(*<*)
+theory S2ss_prop2
+imports Main Flask Flask_type Static Static_type Init_prop Tainted_prop Valid_prop Alive_prop Co2sobj_prop S2ss_prop
+begin
+(*>*)
+
+context tainting_s begin
+
+definition unbackuped_sprocs :: "t_state \<Rightarrow> t_event \<Rightarrow> t_process set \<Rightarrow> t_sobject set"
+where
+ "unbackuped_sprocs s e procs \<equiv>
+ {sp | p sp. p \<in> procs \<and> co2sobj s (O_proc p) = Some sp \<and>
+ (\<forall> p' \<in> procs. co2sobj (e # s) (O_proc p') \<noteq> Some sp) \<and>
+ (\<forall> p' \<in> (current_procs s - procs). co2sobj s (O_proc p') \<noteq> Some sp)}"
+
+definition update_s2ss_procs :: "t_state \<Rightarrow> t_static_state \<Rightarrow> t_event \<Rightarrow> t_process set \<Rightarrow> t_static_state"
+where
+ "update_s2ss_procs s ss e procs \<equiv>
+ ss \<union> {sp | p sp. p \<in> procs \<and> co2sobj (e # s) (O_proc p) = Some sp}
+ - unbackuped_sprocs s e procs"
+ (* new sp after event may exists as same before the event in procs *)
+
+lemma unbked_sps_D:
+ "\<lbrakk>x \<in> unbackuped_sprocs s e procs; p \<in> procs\<rbrakk> \<Longrightarrow> co2sobj (e # s) (O_proc p) \<noteq> Some x"
+by (auto simp add:unbackuped_sprocs_def)
+
+lemma unbked_sps_D':
+ "\<lbrakk>x \<in> unbackuped_sprocs s e procs; p \<notin> procs; p \<in> current_procs s;
+ co2sobj (e # s) (O_proc p) = co2sobj s (O_proc p)\<rbrakk>
+ \<Longrightarrow> co2sobj (e # s) (O_proc p) \<noteq> Some x"
+by (auto simp:unbackuped_sprocs_def)
+
+lemma not_unbked_sps_D:
+ "\<lbrakk>x \<notin> unbackuped_sprocs s e procs; p \<in> procs; co2sobj s (O_proc p) = Some x\<rbrakk>
+ \<Longrightarrow> (\<exists> p' \<in> procs. co2sobj (e # s) (O_proc p') = Some x) \<or>
+ (\<exists> p' \<in> current_procs s - procs. co2sobj s (O_proc p') = Some x)"
+by (auto simp:unbackuped_sprocs_def)
+
+lemma unbked_sps_I:
+ "\<lbrakk>co2sobj s obj = Some x; \<forall> p. obj \<noteq> O_proc p\<rbrakk> \<Longrightarrow> x \<notin> unbackuped_sprocs s' e procs"
+apply (case_tac obj)
+apply (auto simp add:unbackuped_sprocs_def co2sobj.simps split:option.splits)
+done
+
+lemma co2sobj_proc_deleteshm:
+ "\<lbrakk>valid (DeleteShM p h # s); \<forall>flag. (pa, flag) \<notin> procs_of_shm s h; pa \<in> current_procs s\<rbrakk>
+ \<Longrightarrow> co2sobj (DeleteShM p h # s) (O_proc pa) = co2sobj s (O_proc pa)"
+thm co2sobj_deleteshm
+apply (frule_tac obj = "O_proc pa" in co2sobj_deleteshm, simp)
+apply (frule vd_cons, frule_tac p = pa in current_proc_has_sp, simp, erule exE)
+apply (auto dest!:current_proc_has_sp' current_has_sec' current_shm_has_sh'
+ split:t_object.splits option.splits if_splits dest:flag_of_proc_shm_prop1
+ simp:co2sobj.simps tainted_eq_Tainted cp2sproc_deleteshm)
+done
+
+lemma s2ss_deleteshm:
+ "valid (DeleteShM p h # s) \<Longrightarrow> s2ss (DeleteShM p h # s) =
+ (case ch2sshm s h of
+ Some sh \<Rightarrow> del_s2ss_obj s
+ (update_s2ss_procs s (s2ss s) (DeleteShM p h) {p'| p' flag. (p', flag) \<in> procs_of_shm s h})
+ (O_shm h) (S_shm sh)
+ | _ \<Rightarrow> {})"
+apply (frule vt_grant_os, frule vd_cons)
+apply (case_tac "ch2sshm s h")
+apply (drule current_shm_has_sh', simp, simp)
+apply (simp add:del_s2ss_obj_def)
+apply (tactic {*my_clarify_tac 1*})
+
+unfolding update_s2ss_procs_def
+apply (tactic {*my_seteq_tac 1*})
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "\<exists> flag. (pa, flag) \<in> procs_of_shm s h")
+apply (erule exE, rule DiffI, rule UnI2, simp)
+apply (rule_tac x = pa in exI, simp, rule_tac x = flag in exI, simp)
+apply (rule notI, drule_tac p = pa in unbked_sps_D, simp)
+apply (rule_tac x = flag in exI, simp, simp)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI, simp add:co2sobj_deleteshm split:option.splits)
+apply (simp add:cp2sproc_deleteshm split:option.splits if_splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (drule current_has_sec', simp, simp)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (drule flag_of_proc_shm_prop1, simp, simp)
+apply (drule flag_of_proc_shm_prop1, simp, simp)
+apply (rule notI, drule_tac p = pa in unbked_sps_D', simp+)
+apply (simp add:co2sobj_proc_deleteshm)
+apply (simp add:co2sobj_proc_deleteshm)
+apply (rule DiffI, rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_deleteshm is_file_simps)
+apply (erule unbked_sps_I, simp)
+apply (rule DiffI,rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_deleteshm)
+apply (erule unbked_sps_I, simp)
+apply (rule DiffI, rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_deleteshm is_dir_simps)
+apply (erule unbked_sps_I, simp)
+apply (rule DiffI, rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_deleteshm)
+apply (erule unbked_sps_I, simp)
+
+apply (erule DiffE, erule UnE)
+apply (tactic {*my_setiff_tac 1*})
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "\<exists> flag. (pa, flag) \<in> procs_of_shm s h", erule exE)
+apply (drule_tac p = pa in not_unbked_sps_D, simp)
+apply (rule_tac x = flag in exI, simp)
+apply (simp, erule disjE, clarsimp)
+apply (rule_tac x = "O_proc p'" in exI, simp add:procs_of_shm_prop2)
+apply (erule bexE, simp, (erule conjE)+)
+apply (frule_tac pa = p' in co2sobj_proc_deleteshm, simp+)
+apply (rule_tac x = "O_proc p'" in exI, simp)
+apply (frule_tac pa = pa in co2sobj_proc_deleteshm, simp+)
+apply (rule_tac x = "O_proc pa" in exI, simp)
+apply (rule_tac x = obj in exI, simp add:co2sobj_deleteshm is_file_simps)
+apply (frule_tac co2sobj_sshm_imp, erule exE)
+apply (case_tac "ha = h")
+apply (rule_tac x = obj' in exI, simp add:co2sobj_deleteshm)
+apply (simp add:co2sobj.simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_deleteshm)
+apply (rule_tac x = obj in exI, simp add:co2sobj_deleteshm is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_deleteshm)
+apply (tactic {*my_setiff_tac 1*}, clarsimp)
+apply (rule_tac x = "O_proc pa" in exI, simp add:procs_of_shm_prop2)
+
+apply (tactic {*my_clarify_tac 1*})
+unfolding update_s2ss_procs_def
+apply (tactic {*my_seteq_tac 1*})
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "\<exists> flag. (pa, flag) \<in> procs_of_shm s h")
+apply (erule exE, rule DiffI, rule DiffI, rule UnI2, simp)
+apply (rule_tac x = pa in exI, simp, rule_tac x = flag in exI, simp)
+apply (rule notI, drule_tac p = pa in unbked_sps_D, simp)
+apply (rule_tac x = flag in exI, simp, simp)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule DiffI, rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI, simp add:co2sobj_deleteshm split:option.splits)
+apply (simp add:cp2sproc_deleteshm split:option.splits if_splits)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (drule current_has_sec', simp, simp)
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (drule flag_of_proc_shm_prop1, simp, simp)
+apply (drule flag_of_proc_shm_prop1, simp, simp)
+apply (rule notI, drule_tac p = pa in unbked_sps_D', simp+)
+apply (simp add:co2sobj_proc_deleteshm)
+apply (simp add:co2sobj_proc_deleteshm)
+apply (rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule DiffI, rule DiffI, rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_deleteshm is_file_simps)
+apply (erule unbked_sps_I, simp, rule notI, simp add:co2sobj.simps)
+apply (case_tac "ha = h", simp)
+apply (rule DiffI, rule DiffI, rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_deleteshm)
+apply (erule unbked_sps_I, simp)
+apply (rule notI, simp add:co2sobj_deleteshm, erule_tac x = "O_shm ha" in allE, simp)
+apply (rule DiffI, rule DiffI, rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_deleteshm is_dir_simps)
+apply (erule unbked_sps_I, simp, rule notI, simp add:co2sobj.simps split:option.splits)
+apply (rule DiffI, rule DiffI, rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_deleteshm)
+apply (erule unbked_sps_I, simp, rule notI, simp add:co2sobj.simps split:option.splits)
+
+apply (erule DiffE, erule DiffE, erule UnE)
+apply (tactic {*my_setiff_tac 1*})
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "\<exists> flag. (pa, flag) \<in> procs_of_shm s h", erule exE)
+apply (drule_tac p = pa in not_unbked_sps_D, simp)
+apply (rule_tac x = flag in exI, simp)
+apply (simp, erule disjE, clarsimp)
+apply (rule_tac x = "O_proc p'" in exI, simp add:procs_of_shm_prop2)
+apply (erule bexE, simp, (erule conjE)+)
+apply (frule_tac pa = p' in co2sobj_proc_deleteshm, simp+)
+apply (rule_tac x = "O_proc p'" in exI, simp)
+apply (frule_tac pa = pa in co2sobj_proc_deleteshm, simp+)
+apply (rule_tac x = "O_proc pa" in exI, simp)
+apply (rule_tac x = obj in exI, simp add:co2sobj_deleteshm is_file_simps)
+apply (case_tac "ha = h", simp add:co2sobj.simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_deleteshm)
+apply (rule_tac x = obj in exI, simp add:co2sobj_deleteshm is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_deleteshm)
+apply (tactic {*my_setiff_tac 1*}, clarsimp)
+apply (rule_tac x = "O_proc pa" in exI, simp add:procs_of_shm_prop2)
+done
+
+lemma s2ss_detach:
+ "valid (Detach p h # s) \<Longrightarrow> s2ss (Detach p h # s) = (
+ case (cp2sproc s p, cp2sproc (Detach p h # s) p) of
+ (Some sp, Some sp') \<Rightarrow> update_s2ss_obj s (s2ss s) (O_proc p)
+ (S_proc sp (O_proc p \<in> Tainted s)) (S_proc sp' (O_proc p \<in> Tainted s))
+ | _ \<Rightarrow> {} )"
+apply (frule vd_cons, frule vt_grant_os)
+apply (case_tac "cp2sproc s p")
+apply (drule current_proc_has_sp', simp+)
+apply (case_tac "cp2sproc (Detach p h # s) p")
+apply (drule current_proc_has_sp', simp+)
+apply (erule exE|erule conjE)+
+apply (simp add:update_s2ss_obj_def)
+apply (tactic {*my_clarify_tac 1*})
+
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_proc p")
+apply (rule disjI1, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, simp, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_detach, simp add:alive_simps)
+apply (simp add:is_file_simps is_dir_simps split:t_object.splits)
+apply (simp add:co2sobj.simps, simp add:co2sobj.simps)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (tactic {*my_setiff_tac 1*})
+apply (case_tac "obj = O_proc p")
+apply (rule_tac x = obj' in exI)
+apply (frule_tac obj = obj' in co2sobj_detach, simp)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted is_file_simps is_dir_simps split:t_object.splits)[1]
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_detach)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted is_file_simps is_dir_simps split:t_object.splits)[1]
+
+apply (tactic {*my_clarify_tac 1*})
+apply (tactic {*my_seteq_tac 1*})
+apply (case_tac "obj = O_proc p")
+apply (rule disjI1, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule disjI2, rule DiffI, simp, rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_detach, simp add:alive_simps)
+apply (simp add:is_file_simps is_dir_simps split:t_object.splits)
+apply (simp add:co2sobj.simps, simp add:co2sobj.simps)
+apply (rule notI, simp, erule_tac x = obj in allE, erule impE, simp add:alive_simps, simp)
+apply (frule_tac obj = obj in co2sobj_detach)
+apply (simp add:alive_simps)
+apply (simp split:t_object.splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = "O_proc p" in exI, simp add:co2sobj.simps tainted_eq_Tainted)
+apply (tactic {*my_setiff_tac 1*}, simp)
+apply (case_tac "obj = O_proc p")
+apply (simp add:co2sobj.simps tainted_eq_Tainted)
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_detach)
+apply (auto simp add:co2sobj.simps tainted_eq_Tainted is_file_simps is_dir_simps split:t_object.splits)[1]
+done
+
+definition attach_Tainted_procs :: "t_state \<Rightarrow> t_process \<Rightarrow> t_shm \<Rightarrow> t_shm_attach_flag \<Rightarrow> t_process set"
+where
+ "attach_Tainted_procs s p h flag \<equiv>
+ (if (O_proc p \<in> Tainted s \<and> flag = SHM_RDWR)
+ then {p''. \<exists> p' flag'. (p', flag') \<in> procs_of_shm s h \<and> info_flow_shm s p' p'' \<and>
+ O_proc p'' \<notin> Tainted s}
+ else if (\<exists> p'. O_proc p' \<in> Tainted s \<and> (p', SHM_RDWR) \<in> procs_of_shm s h)
+ then {p'. info_flow_shm s p p' \<and> O_proc p' \<notin> Tainted s}
+ else {})"
+
+lemma attach_Tainted_procs_prop1:
+ "valid s \<Longrightarrow> Tainted (Attach p h flag # s) = Tainted s \<union> {O_proc p' | p'. p' \<in> attach_Tainted_procs s p h flag}"
+apply (auto simp:attach_Tainted_procs_def intro:info_flow_shm_Tainted)
+done
+
+lemma attach_Tainted_procs_prop2:
+ "valid (Attach p h flag # s) \<Longrightarrow> {O_proc p'| p'. p' \<in> attach_Tainted_procs s p h flag} \<inter> Tainted s = {}"
+by (auto simp:attach_Tainted_procs_def)
+
+lemma attach_Tainted_procs_prop3:
+ "\<lbrakk>p' \<in> attach_Tainted_procs s p h flag; valid s\<rbrakk> \<Longrightarrow> p' \<in> current_procs s"
+by (auto simp:attach_Tainted_procs_def info_shm_flow_in_procs split:if_splits)
+
+lemma co2sobj_attach':
+ "\<lbrakk>valid (Attach p h flag # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (Attach p h flag # s) obj =
+ (case obj of
+ O_proc p' \<Rightarrow> if (p' = p)
+ then (case cp2sproc (Attach p h flag # s) p of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s \<or> p \<in> attach_Tainted_procs s p h flag))
+ | _ \<Rightarrow> None)
+ else if (p' \<in> attach_Tainted_procs s p h flag)
+ then case cp2sproc s p' of
+ None \<Rightarrow> None
+ | Some sp \<Rightarrow> Some (S_proc sp True)
+ else co2sobj s obj
+ | _ \<Rightarrow> co2sobj s obj)"
+apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted co2sobj.simps)
+
+apply (frule current_proc_has_sp, simp, erule exE, (erule conjE)+)
+apply (case_tac "cp2sproc (Attach p h flag # s) nat", drule current_proc_has_sp', simp+)
+apply (case_tac "cp2sproc (Attach p h flag # s) p", drule current_proc_has_sp', simp+)
+apply (simp add:tainted_eq_Tainted attach_Tainted_procs_prop1 del:Tainted.simps)
+apply (simp add:cp2sproc_attach split:if_splits)
+
+apply (auto split:if_splits option.splits dest!:current_file_has_sfile'
+ simp:current_files_simps cf2sfiles_simps cf2sfile_simps tainted_eq_Tainted
+ same_inode_files_prop6
+ dest:is_file_in_current is_dir_in_current)
+done
+
+lemma s2ss_attach:
+ "valid (Attach p h flag # s) \<Longrightarrow> s2ss (Attach p h flag # s) =
+ update_s2ss_procs s (s2ss s) (Attach p h flag) (attach_Tainted_procs s p h flag \<union> {p})"
+apply (frule vt_grant_os, frule vd_cons)
+apply (case_tac "cp2sproc s p", drule current_proc_has_sp', simp, simp)
+apply (case_tac "ch2sshm s h", drule current_shm_has_sh', simp, simp)
+apply (case_tac "cp2sproc (Attach p h flag # s) p", drule current_proc_has_sp', simp, simp)
+
+unfolding update_s2ss_procs_def
+apply (tactic {*my_seteq_tac 1*})
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa = p")
+apply (rule DiffI, rule UnI2, simp, rule_tac x = pa in exI, simp)
+apply (rule notI, drule_tac p = pa in unbked_sps_D, simp, simp)
+apply (case_tac "pa \<in> attach_Tainted_procs s p h flag")
+apply (rule DiffI, rule UnI2, simp)
+apply (rule_tac x = pa in exI, simp)
+apply (rule notI, drule_tac p = pa in unbked_sps_D, simp, simp)
+apply (rule DiffI, rule UnI1, simp)
+apply (rule_tac x = obj in exI, simp add:co2sobj_attach')
+apply (rule notI, drule_tac p = pa in unbked_sps_D', simp+)
+apply (simp add:co2sobj_attach')
+apply (simp add:co2sobj_attach')
+apply (rule DiffI, rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_attach' is_file_simps)
+apply (erule unbked_sps_I, simp)
+apply (rule DiffI,rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_attach')
+apply (erule unbked_sps_I, simp)
+apply (rule DiffI, rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_attach' is_dir_simps)
+apply (erule unbked_sps_I, simp)
+apply (rule DiffI, rule UnI1, simp, rule_tac x = obj in exI, simp add:co2sobj_attach')
+apply (erule unbked_sps_I, simp)
+
+apply (erule DiffE, erule UnE)
+apply (tactic {*my_setiff_tac 1*})
+apply (erule_tac obj = obj in co2sobj_some_caseD)
+apply (case_tac "pa \<in> attach_Tainted_procs s p h flag \<union> {p}")
+apply (drule_tac p = pa in not_unbked_sps_D, simp, simp)
+apply (erule disjE, erule bexE)
+apply (rule_tac x = "O_proc p'" in exI, simp)
+apply (erule disjE, simp, simp add:attach_Tainted_procs_prop3)
+apply (erule bexE, simp, (erule conjE)+)
+apply (rule_tac x = "O_proc p'" in exI, simp add:co2sobj_attach')
+apply (rule_tac x = "O_proc pa" in exI, simp add:co2sobj_attach')
+apply (rule_tac x = obj in exI, simp add:co2sobj_attach' is_file_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_attach')
+apply (rule_tac x = obj in exI, simp add:co2sobj_attach' is_dir_simps)
+apply (rule_tac x = obj in exI, simp add:co2sobj_attach')
+apply (tactic {*my_setiff_tac 1*}, clarsimp)
+apply (rule_tac x = "O_proc pa" in exI, simp)
+apply (erule disjE, simp, simp add:attach_Tainted_procs_prop3)
+done
+
+
+(* should be modified when socket is model in static *)
+lemma s2ss_createsock:
+ "valid (CreateSock p af st fd inum # s) \<Longrightarrow> s2ss (CreateSock p af st fd inum # s) = s2ss s"
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_other)
+apply (simp add:co2sobj.simps alive_simps split:t_object.splits if_splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_other)
+apply (auto simp:co2sobj.simps alive_simps split:t_object.splits if_splits)
+done
+
+lemma s2ss_bind:
+ "valid (Bind p fd addr # s) \<Longrightarrow> s2ss (Bind p fd addr # s) = s2ss s"
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_other)
+apply (simp add:co2sobj.simps alive_simps split:t_object.splits if_splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_other)
+apply (auto simp:co2sobj.simps alive_simps split:t_object.splits if_splits)
+done
+
+lemma s2ss_connect:
+ "valid (Connect p fd addr # s) \<Longrightarrow> s2ss (Connect p fd addr # s) = s2ss s"
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_other)
+apply (simp add:co2sobj.simps alive_simps split:t_object.splits if_splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_other)
+apply (auto simp:co2sobj.simps alive_simps split:t_object.splits if_splits)
+done
+
+lemma s2ss_listen:
+ "valid (Listen p fd # s) \<Longrightarrow> s2ss (Listen p fd # s) = s2ss s"
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_other)
+apply (simp add:co2sobj.simps alive_simps split:t_object.splits if_splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_other)
+apply (auto simp:co2sobj.simps alive_simps split:t_object.splits if_splits)
+done
+
+lemma s2ss_accept:
+ "valid (Accept p fd addr port fd' inum # s) \<Longrightarrow> s2ss (Accept p fd addr port fd' inum # s) = s2ss s"
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_other)
+apply (simp add:co2sobj.simps alive_simps split:t_object.splits if_splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_other)
+apply (auto simp:co2sobj.simps alive_simps split:t_object.splits if_splits)
+done
+
+lemma s2ss_send:
+ "valid (SendSock p fd # s) \<Longrightarrow> s2ss (SendSock p fd # s) = s2ss s"
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_other)
+apply (simp add:co2sobj.simps alive_simps split:t_object.splits if_splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_other)
+apply (auto simp:co2sobj.simps alive_simps split:t_object.splits if_splits)
+done
+
+lemma s2ss_recv:
+ "valid (RecvSock p fd # s) \<Longrightarrow> s2ss (RecvSock p fd # s) = s2ss s"
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_other)
+apply (simp add:co2sobj.simps alive_simps split:t_object.splits if_splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_other)
+apply (auto simp:co2sobj.simps alive_simps split:t_object.splits if_splits)
+done
+
+lemma s2ss_shutdown:
+ "valid (Shutdown p fd how # s) \<Longrightarrow> s2ss (Shutdown p fd how # s) = s2ss s"
+apply (simp add:s2ss_def)
+apply (tactic {*my_seteq_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (simp add:co2sobj_other)
+apply (simp add:co2sobj.simps alive_simps split:t_object.splits if_splits)
+apply (tactic {*my_setiff_tac 1*})
+apply (rule_tac x = obj in exI)
+apply (frule_tac obj = obj in co2sobj_other)
+apply (auto simp:co2sobj.simps alive_simps split:t_object.splits if_splits)
+done
+
+lemmas s2ss_simps = s2ss_execve s2ss_clone s2ss_ptrace s2ss_kill s2ss_exit s2ss_open
+ s2ss_readfile s2ss_writefile s2ss_closefd s2ss_unlink s2ss_rmdir s2ss_linkhard
+ s2ss_truncate s2ss_createmsgq s2ss_sendmsg s2ss_removemsgq s2ss_recvmsg
+ s2ss_createshm s2ss_detach s2ss_deleteshm s2ss_attach
+ s2ss_createsock s2ss_bind s2ss_connect s2ss_listen s2ss_accept s2ss_send
+ s2ss_recv s2ss_shutdown
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Sectxt_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,772 @@
+theory Sectxt_prop
+imports Main Flask Flask_type Current_files_prop Current_sockets_prop Alive_prop
+begin
+
+context flask begin
+
+lemma alive_obj_has_type:
+ assumes alive: "alive s obj" and vs: "valid s"
+ shows "\<exists> t. type_of_obj s obj = Some t"
+using alive vs
+proof (induct s arbitrary:obj)
+ case Nil
+ thus ?case
+ by (simp add:init_alive_prop[THEN sym] init_obj_has_type)
+next
+ case (Cons e s)
+ hence a2: "alive (e # s) obj" and a3: "valid (e # s)" by auto
+ hence a4: "valid s" and a5: "os_grant s e" and a6: "grant s e"
+ by (auto intro:vd_cons vt_grant_os vt_grant)
+ hence a1: "\<And> obj. alive s obj \<Longrightarrow> \<exists>t. type_of_obj s obj = Some t" using Cons by auto
+ have a1': "\<And> obj. type_of_obj s obj = None \<Longrightarrow> \<not> alive s obj" by (rule_tac notI, auto dest:a1)
+ have p1: "\<And> p. p \<in> current_procs s \<Longrightarrow> \<exists> t. type_of_obj s (O_proc p) = Some t" by (auto intro:a1)
+ have p2: "\<And> f. is_file s f \<Longrightarrow> \<exists> t. type_of_obj s (O_file f) = Some t"
+ by (auto intro!:a1 simp:is_file_in_current)
+ have p3: "\<And> f. is_dir s f \<Longrightarrow> \<exists> t. type_of_obj s (O_dir f) = Some t"
+ by (auto intro!:a1 simp:is_dir_in_current)
+ have p4: "\<And> sock. is_tcp_sock s sock \<Longrightarrow> \<exists> t. type_of_obj s (O_tcp_sock sock) = Some t"
+ by (auto intro!:a1 simp:is_tcp_in_current)
+ have p5: "\<And> sock. is_udp_sock s sock \<Longrightarrow> \<exists> t. type_of_obj s (O_udp_sock sock) = Some t"
+ by (auto intro!:a1 simp:is_udp_in_current)
+ have p6: "\<And> p fd. fd \<in> current_proc_fds s p \<Longrightarrow> \<exists> t. type_of_obj s (O_fd p fd) = Some t"
+ by (auto intro:a1)
+ have p7: "\<And> n. n \<in> init_nodes \<Longrightarrow> \<exists> t. type_of_obj s (O_node n) = Some t" by (auto intro:a1)
+ have p8: "\<And> h. h \<in> current_shms s \<Longrightarrow> \<exists> t. type_of_obj s (O_shm h) = Some t" by (auto intro:a1)
+ have p9: "\<And> q. q \<in> current_msgqs s \<Longrightarrow> \<exists> t. type_of_obj s (O_msgq q) = Some t" by (auto intro:a1)
+ have p10: "\<And> q m. \<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s\<rbrakk>
+ \<Longrightarrow> \<exists> t. type_of_obj s (O_msg q m) = Some t" by (auto intro:a1)
+ show ?case using a5 a2 a4 a3
+ apply (case_tac e)
+ apply (auto split:option.splits t_object.splits if_splits t_socket_type.splits
+ dest!:a1' intro:a1 intro:p1 p2 p3 p4 p5 p6 p7 p8 p9 p10 simp:alive_simps)
+ apply (frule_tac p = nat1 in file_fds_subset_pfds)
+ apply (rule a1, auto)
+ done
+qed
+
+lemma alive_obj_has_user:
+ assumes alive: "alive s obj" and vs: "valid s"
+ shows "\<exists> t. user_of_obj s obj = Some t"
+using alive vs
+proof (induct s arbitrary:obj)
+ case Nil
+ thus ?case
+ by (simp add:init_alive_prop[THEN sym] init_obj_has_user)
+next
+ case (Cons e s)
+ hence a2: "alive (e # s) obj" and a3: "valid (e # s)" by auto
+ hence a4: "valid s" and a5: "os_grant s e" and a6: "grant s e"
+ by (auto intro:vd_cons vt_grant_os vt_grant)
+ hence a1: "\<And> obj. alive s obj \<Longrightarrow> \<exists>t. user_of_obj s obj = Some t" using Cons by auto
+ have a1': "\<And> obj. user_of_obj s obj = None \<Longrightarrow> \<not> alive s obj" by (rule_tac notI, auto dest:a1)
+ have p1: "\<And> p. p \<in> current_procs s \<Longrightarrow> \<exists> t. user_of_obj s (O_proc p) = Some t" by (auto intro:a1)
+ have p2: "\<And> f. is_file s f \<Longrightarrow> \<exists> t. user_of_obj s (O_file f) = Some t"
+ by (auto intro!:a1 simp:is_file_in_current)
+ have p3: "\<And> f. is_dir s f \<Longrightarrow> \<exists> t. user_of_obj s (O_dir f) = Some t"
+ by (auto intro!:a1 simp:is_dir_in_current)
+ have p4: "\<And> sock. is_tcp_sock s sock \<Longrightarrow> \<exists> t. user_of_obj s (O_tcp_sock sock) = Some t"
+ by (auto intro!:a1 simp:is_tcp_in_current)
+ have p5: "\<And> sock. is_udp_sock s sock \<Longrightarrow> \<exists> t. user_of_obj s (O_udp_sock sock) = Some t"
+ by (auto intro!:a1 simp:is_udp_in_current)
+ have p6: "\<And> p fd. fd \<in> current_proc_fds s p \<Longrightarrow> \<exists> t. user_of_obj s (O_fd p fd) = Some t"
+ by (auto intro:a1)
+ have p7: "\<And> n. n \<in> init_nodes \<Longrightarrow> \<exists> t. user_of_obj s (O_node n) = Some t" by (auto intro:a1)
+ have p8: "\<And> h. h \<in> current_shms s \<Longrightarrow> \<exists> t. user_of_obj s (O_shm h) = Some t" by (auto intro:a1)
+ have p9: "\<And> q. q \<in> current_msgqs s \<Longrightarrow> \<exists> t. user_of_obj s (O_msgq q) = Some t" by (auto intro:a1)
+ have p10: "\<And> q m. \<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s\<rbrakk>
+ \<Longrightarrow> \<exists> t. user_of_obj s (O_msg q m) = Some t" by (auto intro:a1)
+ show ?case using a5 a2 a4 a3
+ apply (case_tac e)
+ apply (auto split:option.splits t_object.splits if_splits t_socket_type.splits
+ dest!:a1' intro:a1 intro:p1 p2 p3 p4 p5 p6 p7 p8 p9 p10 simp:alive_simps)
+ apply (frule_tac p = nat1 in file_fds_subset_pfds)
+ apply (rule a1, auto)
+ done
+qed
+
+lemma alive_obj_has_type': "\<lbrakk>type_of_obj s obj = None; valid s\<rbrakk> \<Longrightarrow> \<not> alive s obj"
+by (rule_tac notI, auto dest:alive_obj_has_type)
+
+lemma current_proc_has_type:
+ "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> t. type_of_obj s (O_proc p) = Some t"
+by (auto intro:alive_obj_has_type)
+
+lemma current_proc_has_type':
+ "\<lbrakk>type_of_obj s (O_proc p) = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
+by (rule notI, auto dest:current_proc_has_type)
+
+lemma is_file_has_type: "\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> t. type_of_obj s (O_file f) = Some t"
+by (auto intro:alive_obj_has_type simp:is_file_in_current)
+
+lemma is_file_has_type': "\<lbrakk>type_of_obj s (O_file f) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_file s f"
+by (auto intro:notI dest:is_file_has_type)
+
+lemma is_dir_has_type: "\<lbrakk>is_dir s f; valid s\<rbrakk> \<Longrightarrow> \<exists> t. type_of_obj s (O_dir f) = Some t"
+by (auto intro:alive_obj_has_type simp:is_dir_in_current)
+
+lemma is_dir_has_type': "\<lbrakk>type_of_obj s (O_dir f) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_dir s f"
+by (auto intro:notI dest:is_dir_has_type)
+
+lemma is_tcp_has_type: "\<lbrakk>is_tcp_sock s sock; valid s\<rbrakk> \<Longrightarrow> \<exists> t. type_of_obj s (O_tcp_sock sock) = Some t"
+by (auto intro:alive_obj_has_type simp:is_tcp_in_current)
+
+lemma is_tcp_has_type': "\<lbrakk>type_of_obj s (O_tcp_sock sock) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_tcp_sock s sock"
+by (auto intro:notI dest:is_tcp_has_type)
+
+lemma is_udp_has_type: "\<lbrakk>is_udp_sock s sock; valid s\<rbrakk> \<Longrightarrow> \<exists> t. type_of_obj s (O_udp_sock sock) = Some t"
+by (auto intro:alive_obj_has_type simp:is_udp_in_current)
+
+lemma is_udp_has_type': "\<lbrakk>type_of_obj s (O_udp_sock sock) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_udp_sock s sock"
+by (auto intro:notI dest:is_udp_has_type)
+
+lemma current_fd_has_type: "\<lbrakk>fd \<in> current_proc_fds s p; valid s\<rbrakk> \<Longrightarrow> \<exists> t. type_of_obj s (O_fd p fd) = Some t"
+by (auto intro:alive_obj_has_type)
+
+lemma current_fd_has_type': "\<lbrakk>type_of_obj s (O_fd p fd) = None; valid s\<rbrakk> \<Longrightarrow> fd \<notin> current_proc_fds s p"
+by (auto intro:notI dest:current_fd_has_type)
+
+lemma init_node_has_type: "\<lbrakk>n \<in> init_nodes; valid s\<rbrakk> \<Longrightarrow> \<exists> t. type_of_obj s (O_node n) = Some t"
+by (auto intro: alive_obj_has_type)
+
+lemma init_node_has_type': "\<lbrakk>type_of_obj s (O_node n) = None; valid s\<rbrakk> \<Longrightarrow> n \<notin> init_nodes"
+by (auto intro:notI dest:init_node_has_type)
+
+lemma current_shm_has_type: "\<lbrakk>h \<in> current_shms s; valid s\<rbrakk> \<Longrightarrow> \<exists> t. type_of_obj s (O_shm h) = Some t"
+by (auto intro:alive_obj_has_type)
+
+lemma current_shm_has_type': "\<lbrakk>type_of_obj s (O_shm h) = None; valid s\<rbrakk> \<Longrightarrow> h \<notin> current_shms s"
+by (auto intro:notI dest:current_shm_has_type)
+
+lemma current_msgq_has_type: "\<lbrakk>q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> \<exists> t. type_of_obj s (O_msgq q) = Some t"
+by (auto intro:alive_obj_has_type)
+
+lemma current_msgq_has_type': "\<lbrakk>type_of_obj s (O_msgq q) = None; valid s\<rbrakk> \<Longrightarrow> q \<notin> current_msgqs s"
+by (auto intro:notI dest:current_msgq_has_type)
+
+lemma current_msg_has_type: "\<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s; valid s\<rbrakk>
+ \<Longrightarrow> \<exists> t. type_of_obj s (O_msg q m) = Some t"
+by (auto intro:alive_obj_has_type)
+
+lemma current_msg_has_type': "\<lbrakk>type_of_obj s (O_msg q m) = None; valid s\<rbrakk>
+ \<Longrightarrow> m \<notin> set (msgs_of_queue s q) \<or> q \<notin> current_msgqs s"
+by (auto dest!:current_msg_has_type)
+
+lemmas current_has_type = alive_obj_has_type current_proc_has_type is_file_has_type is_dir_has_type
+ is_tcp_has_type is_udp_has_type current_fd_has_type init_node_has_type current_shm_has_type
+ current_msgq_has_type current_msg_has_type
+
+lemmas current_has_type' = alive_obj_has_type' current_proc_has_type' is_file_has_type' is_dir_has_type'
+ is_tcp_has_type' is_udp_has_type' current_fd_has_type' init_node_has_type' current_shm_has_type'
+ current_msgq_has_type' current_msg_has_type'
+
+lemma alive_obj_has_user': "\<lbrakk>user_of_obj s obj = None; valid s\<rbrakk> \<Longrightarrow> \<not> alive s obj"
+by (rule_tac notI, auto dest:alive_obj_has_user)
+
+lemma current_proc_has_user:
+ "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> t. user_of_obj s (O_proc p) = Some t"
+by (auto intro:alive_obj_has_user)
+
+lemma current_proc_has_user':
+ "\<lbrakk>user_of_obj s (O_proc p) = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
+by (rule notI, auto dest:current_proc_has_user)
+
+lemma is_file_has_user: "\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> t. user_of_obj s (O_file f) = Some t"
+by (auto intro:alive_obj_has_user simp:is_file_in_current)
+
+lemma is_file_has_user': "\<lbrakk>user_of_obj s (O_file f) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_file s f"
+by (auto intro:notI dest:is_file_has_user)
+
+lemma is_dir_has_user: "\<lbrakk>is_dir s f; valid s\<rbrakk> \<Longrightarrow> \<exists> t. user_of_obj s (O_dir f) = Some t"
+by (auto intro:alive_obj_has_user simp:is_dir_in_current)
+
+lemma is_dir_has_user': "\<lbrakk>user_of_obj s (O_dir f) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_dir s f"
+by (auto intro:notI dest:is_dir_has_user)
+
+lemma is_tcp_has_user: "\<lbrakk>is_tcp_sock s sock; valid s\<rbrakk> \<Longrightarrow> \<exists> t. user_of_obj s (O_tcp_sock sock) = Some t"
+by (auto intro:alive_obj_has_user simp:is_tcp_in_current)
+
+lemma is_tcp_has_user': "\<lbrakk>user_of_obj s (O_tcp_sock sock) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_tcp_sock s sock"
+by (auto intro:notI dest:is_tcp_has_user)
+
+lemma is_udp_has_user: "\<lbrakk>is_udp_sock s sock; valid s\<rbrakk> \<Longrightarrow> \<exists> t. user_of_obj s (O_udp_sock sock) = Some t"
+by (auto intro:alive_obj_has_user simp:is_udp_in_current)
+
+lemma is_udp_has_user': "\<lbrakk>user_of_obj s (O_udp_sock sock) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_udp_sock s sock"
+by (auto intro:notI dest:is_udp_has_user)
+
+lemma current_fd_has_user: "\<lbrakk>fd \<in> current_proc_fds s p; valid s\<rbrakk> \<Longrightarrow> \<exists> t. user_of_obj s (O_fd p fd) = Some t"
+by (auto intro:alive_obj_has_user)
+
+lemma current_fd_has_user': "\<lbrakk>user_of_obj s (O_fd p fd) = None; valid s\<rbrakk> \<Longrightarrow> fd \<notin> current_proc_fds s p"
+by (auto intro:notI dest:current_fd_has_user)
+
+lemma init_node_has_user: "\<lbrakk>n \<in> init_nodes; valid s\<rbrakk> \<Longrightarrow> \<exists> t. user_of_obj s (O_node n) = Some t"
+by (auto intro: alive_obj_has_user)
+
+lemma init_node_has_user': "\<lbrakk>user_of_obj s (O_node n) = None; valid s\<rbrakk> \<Longrightarrow> n \<notin> init_nodes"
+by (auto intro:notI dest:init_node_has_user)
+
+lemma current_shm_has_user: "\<lbrakk>h \<in> current_shms s; valid s\<rbrakk> \<Longrightarrow> \<exists> t. user_of_obj s (O_shm h) = Some t"
+by (auto intro:alive_obj_has_user)
+
+lemma current_shm_has_user': "\<lbrakk>user_of_obj s (O_shm h) = None; valid s\<rbrakk> \<Longrightarrow> h \<notin> current_shms s"
+by (auto intro:notI dest:current_shm_has_user)
+
+lemma current_msgq_has_user: "\<lbrakk>q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> \<exists> t. user_of_obj s (O_msgq q) = Some t"
+by (auto intro:alive_obj_has_user)
+
+lemma current_msgq_has_user': "\<lbrakk>user_of_obj s (O_msgq q) = None; valid s\<rbrakk> \<Longrightarrow> q \<notin> current_msgqs s"
+by (auto intro:notI dest:current_msgq_has_user)
+
+lemma current_msg_has_user: "\<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s; valid s\<rbrakk>
+ \<Longrightarrow> \<exists> t. user_of_obj s (O_msg q m) = Some t"
+by (auto intro:alive_obj_has_user)
+
+lemma current_msg_has_user': "\<lbrakk>user_of_obj s (O_msg q m) = None; valid s\<rbrakk>
+ \<Longrightarrow> m \<notin> set (msgs_of_queue s q) \<or> q \<notin> current_msgqs s"
+by (auto dest!:current_msg_has_user)
+
+lemmas current_has_user = alive_obj_has_user current_proc_has_user is_file_has_user is_dir_has_user
+ is_tcp_has_user is_udp_has_user current_fd_has_user init_node_has_user current_shm_has_user
+ current_msgq_has_user current_msg_has_user
+
+lemmas current_has_user' = alive_obj_has_user' current_proc_has_user' is_file_has_user' is_dir_has_user'
+ is_tcp_has_user' is_udp_has_user' current_fd_has_user' init_node_has_user' current_shm_has_user'
+ current_msgq_has_user' current_msg_has_user'
+
+lemma current_proc_has_role:
+ "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> r. role_of_proc s p = Some r"
+apply (induct s arbitrary:p, simp add:init_procrole_prop2)
+apply (frule vd_cons, frule vt_grant_os, frule vt_grant, case_tac a)
+apply (auto split:option.splits dest:current_has_type' simp:sectxt_of_obj_def)
+done
+
+lemma current_proc_has_role':
+ "\<lbrakk>role_of_proc s p = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
+by (auto dest:current_proc_has_role)
+
+lemma alive_obj_has_role:
+ "\<lbrakk>alive s obj; valid s\<rbrakk> \<Longrightarrow> \<exists> r. role_of_obj s obj = Some r"
+by (case_tac obj, auto intro:current_proc_has_role)
+
+lemma alive_obj_has_sec:
+ "\<lbrakk>alive s obj; valid s\<rbrakk> \<Longrightarrow> \<exists> sec. sectxt_of_obj s obj = Some sec"
+apply (frule alive_obj_has_type, simp)
+apply (frule alive_obj_has_role, simp)
+apply (frule alive_obj_has_user, simp)
+apply (auto split:option.splits simp:sectxt_of_obj_def)
+done
+
+lemma alive_obj_has_sec':
+ "\<lbrakk>sectxt_of_obj s obj = None; valid s\<rbrakk> \<Longrightarrow> \<not> alive s obj"
+by (auto dest:alive_obj_has_sec)
+
+lemma alive_obj_has_sec'':
+ "\<lbrakk>alive s obj; valid s\<rbrakk> \<Longrightarrow> \<exists> u r t. sectxt_of_obj s obj = Some (u,r,t)"
+by (auto dest:alive_obj_has_sec)
+
+lemma current_proc_has_sec:
+ "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sec. sectxt_of_obj s (O_proc p) = Some sec"
+by (rule alive_obj_has_sec, simp+)
+
+lemma current_proc_has_sec':
+ "\<lbrakk>sectxt_of_obj s (O_proc p) = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
+by (rule notI, auto dest:current_proc_has_sec)
+
+lemma current_proc_has_sec'':
+ "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> u r t. sectxt_of_obj s (O_proc p) = Some (u,r,t)"
+by (drule current_proc_has_sec, auto)
+
+lemma is_file_has_sec: "\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> sec. sectxt_of_obj s (O_file f) = Some sec"
+by (rule alive_obj_has_sec, simp_all add:is_file_in_current)
+
+lemma is_file_has_sec': "\<lbrakk>sectxt_of_obj s (O_file f) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_file s f"
+by (auto intro:notI dest:is_file_has_sec)
+
+lemma is_file_has_sec'': "\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> u r t. sectxt_of_obj s (O_file f) = Some (u,r,t)"
+by (drule is_file_has_sec, auto)
+
+lemma is_dir_has_sec: "\<lbrakk>is_dir s f; valid s\<rbrakk> \<Longrightarrow> \<exists> sec. sectxt_of_obj s (O_dir f) = Some sec"
+by (rule alive_obj_has_sec, simp_all add:is_dir_in_current)
+
+lemma is_dir_has_sec': "\<lbrakk>sectxt_of_obj s (O_dir f) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_dir s f"
+by (auto intro:notI dest:is_dir_has_sec)
+
+lemma is_dir_has_sec'': "\<lbrakk>is_dir s f; valid s\<rbrakk> \<Longrightarrow> \<exists> u r t. sectxt_of_obj s (O_dir f) = Some (u,r,t)"
+by (drule is_dir_has_sec, auto)
+
+lemma is_tcp_has_sec: "\<lbrakk>is_tcp_sock s sock; valid s\<rbrakk> \<Longrightarrow> \<exists> sec. sectxt_of_obj s (O_tcp_sock sock) = Some sec"
+by (rule alive_obj_has_sec, simp_all add:is_tcp_in_current)
+
+lemma is_tcp_has_sec': "\<lbrakk>sectxt_of_obj s (O_tcp_sock sock) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_tcp_sock s sock"
+by (auto intro:notI dest:is_tcp_has_sec)
+
+lemma is_tcp_has_sec'': "\<lbrakk>is_tcp_sock s sock; valid s\<rbrakk>
+ \<Longrightarrow> \<exists> u r t. sectxt_of_obj s (O_tcp_sock sock) = Some (u,r,t)"
+by (drule is_tcp_has_sec, simp_all add:is_tcp_in_current)
+
+lemma is_udp_has_sec: "\<lbrakk>is_udp_sock s sock; valid s\<rbrakk> \<Longrightarrow> \<exists> sec. sectxt_of_obj s (O_udp_sock sock) = Some sec"
+by (rule alive_obj_has_sec, simp_all add:is_udp_in_current)
+
+lemma is_udp_has_sec': "\<lbrakk>sectxt_of_obj s (O_udp_sock sock) = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_udp_sock s sock"
+by (auto intro:notI dest:is_udp_has_sec)
+
+lemma is_udp_has_sec'': "\<lbrakk>is_udp_sock s sock; valid s\<rbrakk>
+ \<Longrightarrow> \<exists> u r t. sectxt_of_obj s (O_udp_sock sock) = Some (u,r,t)"
+by (drule is_udp_has_sec, simp_all add:is_udp_in_current)
+
+lemma current_fd_has_sec: "\<lbrakk>fd \<in> current_proc_fds s p; valid s\<rbrakk> \<Longrightarrow> \<exists> sec. sectxt_of_obj s (O_fd p fd) = Some sec"
+by (rule alive_obj_has_sec, simp+)
+
+lemma current_fd_has_sec': "\<lbrakk>sectxt_of_obj s (O_fd p fd) = None; valid s\<rbrakk> \<Longrightarrow> fd \<notin> current_proc_fds s p"
+by (auto intro:notI dest:current_fd_has_sec)
+
+lemma current_fd_has_sec'': "\<lbrakk>fd \<in> current_proc_fds s p; valid s\<rbrakk>
+ \<Longrightarrow> \<exists> u r t. sectxt_of_obj s (O_fd p fd) = Some (u, r, t)"
+by (drule current_fd_has_sec, simp+)
+
+lemma init_node_has_sec: "\<lbrakk>n \<in> init_nodes; valid s\<rbrakk> \<Longrightarrow> \<exists> sec. sectxt_of_obj s (O_node n) = Some sec"
+by (rule alive_obj_has_sec, simp+)
+
+lemma init_node_has_sec': "\<lbrakk>sectxt_of_obj s (O_node n) = None; valid s\<rbrakk> \<Longrightarrow> n \<notin> init_nodes"
+by (auto intro:notI dest:init_node_has_sec)
+
+lemma init_node_has_sec'': "\<lbrakk>n \<in> init_nodes; valid s\<rbrakk> \<Longrightarrow> \<exists> u r t. sectxt_of_obj s (O_node n) = Some (u,r,t)"
+by (drule init_node_has_sec, simp+)
+
+lemma current_shm_has_sec: "\<lbrakk>h \<in> current_shms s; valid s\<rbrakk> \<Longrightarrow> \<exists> sec. sectxt_of_obj s (O_shm h) = Some sec"
+by (rule alive_obj_has_sec, simp+)
+
+lemma current_shm_has_sec': "\<lbrakk>sectxt_of_obj s (O_shm h) = None; valid s\<rbrakk> \<Longrightarrow> h \<notin> current_shms s"
+by (auto intro:notI dest:current_shm_has_sec)
+
+lemma current_shm_has_sec'': "\<lbrakk>h \<in> current_shms s; valid s\<rbrakk> \<Longrightarrow> \<exists> u r t. sectxt_of_obj s (O_shm h) = Some (u,r,t)"
+by (drule current_shm_has_sec, simp+)
+
+lemma current_msgq_has_sec: "\<lbrakk>q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sec. sectxt_of_obj s (O_msgq q) = Some sec"
+by (rule alive_obj_has_sec, simp+)
+
+lemma current_msgq_has_sec': "\<lbrakk>sectxt_of_obj s (O_msgq q) = None; valid s\<rbrakk> \<Longrightarrow> q \<notin> current_msgqs s"
+by (auto intro:notI dest:current_msgq_has_sec)
+
+lemma current_msgq_has_sec'': "\<lbrakk>q \<in> current_msgqs s; valid s\<rbrakk>
+ \<Longrightarrow> \<exists> u r t. sectxt_of_obj s (O_msgq q) = Some (u,r,t)"
+by (drule current_msgq_has_sec, simp+)
+
+lemma current_msg_has_sec: "\<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s; valid s\<rbrakk>
+ \<Longrightarrow> \<exists> sec. sectxt_of_obj s (O_msg q m) = Some sec"
+by (rule alive_obj_has_sec, simp+)
+
+lemma current_msg_has_sec': "\<lbrakk>sectxt_of_obj s (O_msg q m) = None; valid s\<rbrakk>
+ \<Longrightarrow> m \<notin> set (msgs_of_queue s q) \<or> q \<notin> current_msgqs s"
+by (auto dest!:current_msg_has_sec)
+
+lemma current_msg_has_sec'': "\<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s; valid s\<rbrakk>
+ \<Longrightarrow> \<exists> u r t. sectxt_of_obj s (O_msg q m) = Some (u, r, t)"
+by (drule current_msg_has_sec, simp+)
+
+lemmas current_has_sec = alive_obj_has_sec current_proc_has_sec is_file_has_sec is_dir_has_sec
+ is_tcp_has_sec is_udp_has_sec current_fd_has_sec init_node_has_sec current_shm_has_sec
+ current_msgq_has_sec current_msg_has_sec
+
+lemmas current_has_sec' = alive_obj_has_sec' current_proc_has_sec' is_file_has_sec' is_dir_has_sec'
+ is_tcp_has_sec' is_udp_has_sec' current_fd_has_sec' init_node_has_sec' current_shm_has_sec'
+ current_msgq_has_sec' current_msg_has_sec'
+
+lemmas current_has_sec'' = alive_obj_has_sec'' current_proc_has_sec'' is_file_has_sec'' is_dir_has_sec''
+ is_tcp_has_sec'' is_udp_has_sec'' current_fd_has_sec'' init_node_has_sec'' current_shm_has_sec''
+ current_msgq_has_sec'' current_msg_has_sec''
+
+(************ root sec remains ************)
+
+lemma root_type_remains:
+ "valid s \<Longrightarrow> type_of_obj s (O_dir []) = init_type_of_obj (O_dir [])"
+apply (induct s)
+apply (simp)
+apply (frule vd_cons, frule vt_grant_os, case_tac a) prefer 6
+by (case_tac option, auto)
+
+lemma root_user_remains:
+ "valid s \<Longrightarrow> user_of_obj s (O_dir []) = init_user_of_obj (O_dir [])"
+apply (induct s)
+apply (simp)
+apply (frule vd_cons, frule vt_grant_os, case_tac a) prefer 6
+by (case_tac option, auto)
+
+lemma root_has_type':
+ "\<lbrakk>type_of_obj s (O_dir []) = None; valid s\<rbrakk> \<Longrightarrow> False"
+apply (drule alive_obj_has_type', simp)
+by (drule root_is_dir, simp)
+
+lemma root_has_user':
+ "\<lbrakk>user_of_obj s (O_dir []) = None; valid s\<rbrakk> \<Longrightarrow> False"
+apply (drule alive_obj_has_user', simp)
+by (drule root_is_dir, simp)
+
+lemma root_has_init_type':
+ "init_type_of_obj (O_dir []) = None \<Longrightarrow> False"
+using init_obj_has_type[where obj = "O_dir []"] root_is_init_dir
+by auto
+
+lemma root_has_init_user':
+ "init_user_of_obj (O_dir []) = None \<Longrightarrow> False"
+using init_obj_has_user[where obj = "O_dir []"] root_is_init_dir
+by auto
+
+lemma root_sec_remains:
+ "valid s \<Longrightarrow> sectxt_of_obj s (O_dir []) = init_sectxt_of_obj (O_dir [])"
+by (auto simp:root_user_remains root_type_remains sectxt_of_obj_def init_sectxt_of_obj_def
+ split:option.splits)
+
+lemma root_sec_set:
+ "\<exists> u t. sec_of_root = (u, R_object, t)"
+by (auto simp:sec_of_root_def split:option.splits
+ dest!: root_has_init_type' root_has_init_user')
+
+lemma sec_of_root_set:
+ "init_sectxt_of_obj (O_dir []) = Some sec_of_root"
+using root_has_init_type' root_has_init_user'
+apply (auto simp:init_sectxt_of_obj_def sec_of_root_def split:option.splits)
+done
+
+(*************** sectxt_of_obj simpset ****************)
+
+lemma sec_execve:
+ "valid (Execve p f fds # s) \<Longrightarrow> sectxt_of_obj (Execve p f fds # s) =
+ (sectxt_of_obj s) (O_proc p := (
+ case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_file f)) of
+ (Some psec, Some fsec) \<Rightarrow> npctxt_execve psec fsec
+ | _ \<Rightarrow> None))"
+apply (rule ext, frule vd_cons, frule vt_grant_os, frule vt_grant,case_tac x)
+apply (auto simp:sectxt_of_obj_def split:option.splits
+ dest!:current_has_type' current_proc_has_role' current_has_user')
+done
+
+lemma sec_clone:
+ "valid (Clone p p' fds shms # s) \<Longrightarrow> sectxt_of_obj (Clone p p' fds shms # s) = (\<lambda> obj.
+ case obj of
+ O_proc p'' \<Rightarrow> if (p'' = p') then sectxt_of_obj s (O_proc p)
+ else sectxt_of_obj s obj
+ | O_fd p'' fd \<Rightarrow> if (p'' = p' \<and> fd \<in> fds) then sectxt_of_obj s (O_fd p fd)
+ else if (p'' = p') then None
+ else sectxt_of_obj s obj
+ | O_tcp_sock (p'', fd) \<Rightarrow> if (p'' = p' \<and> fd \<in> fds) then sectxt_of_obj s (O_tcp_sock (p, fd))
+ else if (p'' = p') then None
+ else sectxt_of_obj s obj
+ | O_udp_sock (p'', fd) \<Rightarrow> if (p'' = p' \<and> fd \<in> fds) then sectxt_of_obj s (O_udp_sock (p, fd))
+ else if (p'' = p') then None
+ else sectxt_of_obj s obj
+ | _ \<Rightarrow> sectxt_of_obj s obj )"
+apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac obj)
+apply (auto simp:sectxt_of_obj_def split:option.splits if_splits
+ dest!:current_has_type' current_proc_has_role' current_has_user')
+done
+
+lemma sec_open:
+ "valid (Open p f flags fd opt # s) \<Longrightarrow> sectxt_of_obj (Open p f flags fd opt # s) = (\<lambda> obj.
+ case obj of
+ O_file f' \<Rightarrow> if (f' = f \<and> opt \<noteq> None) then
+ (case (parent f) of
+ None \<Rightarrow> None
+ | Some pf \<Rightarrow> (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of
+ (Some psec, Some pfsec) \<Rightarrow> Some (fst psec, R_object,
+ type_transition ((snd o snd) psec) ((snd o snd) pfsec) C_file True)
+ | _ \<Rightarrow> None))
+ else sectxt_of_obj s obj
+ | O_fd p' fd' \<Rightarrow> if (p' = p \<and> fd' = fd) then
+ (case (sectxt_of_obj s (O_proc p)) of
+ Some psec \<Rightarrow> Some (fst psec, R_object, (snd o snd) psec)
+ | _ \<Rightarrow> None)
+ else sectxt_of_obj s obj
+ | _ \<Rightarrow> sectxt_of_obj s obj)"
+apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac obj)
+apply (auto simp:sectxt_of_obj_def split:option.splits if_splits
+ dest!:current_has_type' current_proc_has_role' current_has_user')
+done
+
+lemma sec_mkdir:
+ "valid (Mkdir p f inum # s) \<Longrightarrow> sectxt_of_obj (Mkdir p f inum # s) =
+ (sectxt_of_obj s) (O_dir f :=
+ (case parent f of
+ None \<Rightarrow> None
+ | Some pf \<Rightarrow> (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of
+ (Some psec, Some pfsec) \<Rightarrow> Some (fst psec, R_object,
+ type_transition ((snd o snd) psec) ((snd o snd) pfsec) C_dir True)
+ | _ \<Rightarrow> None)))"
+apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac x)
+apply (auto simp:sectxt_of_obj_def split:option.splits if_splits
+ dest!:current_has_type' current_proc_has_role' current_has_user')
+done
+
+lemma sec_linkhard:
+ "valid (LinkHard p f f' # s) \<Longrightarrow> sectxt_of_obj (LinkHard p f f' # s) =
+ (sectxt_of_obj s) (O_file f' :=
+ (case parent f' of
+ None \<Rightarrow> None
+ | Some pf \<Rightarrow> (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of
+ (Some psec, Some pfsec) \<Rightarrow> Some (fst psec, R_object,
+ type_transition ((snd o snd) psec) ((snd o snd) pfsec) C_file True)
+ | _ \<Rightarrow> None)))"
+apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac x)
+apply (auto simp:sectxt_of_obj_def split:option.splits if_splits
+ dest!:current_has_type' current_proc_has_role' current_has_user')
+done
+
+lemma sec_createmsgq:
+ "valid (CreateMsgq p q # s) \<Longrightarrow> sectxt_of_obj (CreateMsgq p q # s) = (sectxt_of_obj s) (O_msgq q :=
+ (case (sectxt_of_obj s (O_proc p)) of
+ Some psec \<Rightarrow> Some (fst psec, R_object, (snd o snd) psec)
+ | _ \<Rightarrow> None))"
+apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac x)
+apply (auto simp:sectxt_of_obj_def split:option.splits if_splits
+ dest!:current_has_type' current_proc_has_role' current_has_user')
+done
+
+lemma sec_sendmsg:
+ "valid (SendMsg p q m # s) \<Longrightarrow> sectxt_of_obj (SendMsg p q m # s) = (sectxt_of_obj s) (O_msg q m :=
+ (case (sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_msgq q)) of
+ (Some psec, Some qsec) \<Rightarrow> Some (fst psec, R_object,
+ type_transition ((snd o snd) psec) ((snd o snd) qsec) C_msg False)
+ | _ \<Rightarrow> None))"
+apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac x)
+apply (auto simp:sectxt_of_obj_def split:option.splits if_splits
+ dest!:current_has_type' current_proc_has_role' current_has_user')
+done
+
+lemma sec_createshm:
+ "valid (CreateShM p h # s) \<Longrightarrow> sectxt_of_obj (CreateShM p h # s) = (sectxt_of_obj s) (O_shm h :=
+ case (sectxt_of_obj s (O_proc p)) of
+ Some psec \<Rightarrow> Some (fst psec, R_object, (snd o snd) psec)
+ | _ \<Rightarrow> None)"
+apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac x)
+apply (auto simp:sectxt_of_obj_def split:option.splits if_splits
+ dest!:current_has_type' current_proc_has_role' current_has_user')
+done
+
+lemma sec_createsock:
+ "valid (CreateSock p af st fd inum # s) \<Longrightarrow> sectxt_of_obj (CreateSock p af st fd inum # s) = (\<lambda> obj.
+ case obj of
+ O_fd p' fd' \<Rightarrow> if (p' = p \<and> fd' = fd) then
+ (case (sectxt_of_obj s (O_proc p)) of
+ Some psec \<Rightarrow> Some (fst psec, R_object, (snd o snd) psec)
+ | _ \<Rightarrow> None)
+ else sectxt_of_obj s obj
+ | O_tcp_sock (p', fd') \<Rightarrow> if (p' = p \<and> fd' = fd \<and> st = STREAM) then
+ (case (sectxt_of_obj s (O_proc p)) of
+ Some psec \<Rightarrow> Some (fst psec, R_object, (snd o snd) psec)
+ | _ \<Rightarrow> None)
+ else sectxt_of_obj s obj
+ | O_udp_sock (p', fd') \<Rightarrow> if (p' = p \<and> fd' = fd \<and> st = DGRAM) then
+ (case (sectxt_of_obj s (O_proc p)) of
+ Some psec \<Rightarrow> Some (fst psec, R_object, (snd o snd) psec)
+ | _ \<Rightarrow> None)
+ else sectxt_of_obj s obj
+ | _ \<Rightarrow> sectxt_of_obj s obj )"
+apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac obj)
+apply (auto simp:sectxt_of_obj_def split:option.splits if_splits
+ dest!:current_has_type' current_proc_has_role' current_has_user')
+done
+
+lemma sec_accept:
+ "valid (Accept p fd addr port fd' inum # s) \<Longrightarrow> sectxt_of_obj (Accept p fd addr port fd' inum # s) = (\<lambda> obj.
+ case obj of
+ O_fd p' fd'' \<Rightarrow> if (p' = p \<and> fd'' = fd') then
+ (case (sectxt_of_obj s (O_proc p)) of
+ Some psec \<Rightarrow> Some (fst psec, R_object, (snd o snd) psec)
+ | _ \<Rightarrow> None)
+ else sectxt_of_obj s obj
+ | O_tcp_sock (p', fd'') \<Rightarrow> if (p' = p \<and> fd'' = fd') then
+ (case (sectxt_of_obj s (O_proc p)) of
+ Some psec \<Rightarrow> Some (fst psec, R_object, (snd o snd) psec)
+ | _ \<Rightarrow> None)
+ else sectxt_of_obj s obj
+ | _ \<Rightarrow> sectxt_of_obj s obj )"
+apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac obj)
+apply (auto simp:sectxt_of_obj_def split:option.splits if_splits
+ dest!:current_has_type' current_proc_has_role' current_has_user')
+done
+
+lemma sec_others :
+ "\<lbrakk>valid (e # s);
+ \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms;
+ \<forall> p f fds. e \<noteq> Execve p f fds;
+ \<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ \<forall> p f inum. e \<noteq> Mkdir p f inum;
+ \<forall> p f f'. e \<noteq> LinkHard p f f';
+ \<forall> p q. e \<noteq> CreateMsgq p q;
+ \<forall> p q m. e \<noteq> SendMsg p q m;
+ \<forall> p h. e \<noteq> CreateShM p h;
+ \<forall> p af st fd inum. e \<noteq> CreateSock p af st fd inum;
+ \<forall> p fd addr port fd' inum. e \<noteq> Accept p fd addr port fd' inum
+ \<rbrakk> \<Longrightarrow> sectxt_of_obj (e # s) = sectxt_of_obj s"
+apply (rule ext, frule vd_cons, frule vt_grant_os, case_tac e, case_tac[!] x)
+apply (auto simp:sectxt_of_obj_def split:option.splits if_splits
+ dest!:current_has_type' current_proc_has_role' current_has_user')
+done
+
+lemma sec_kill:
+ "valid (Kill p p' # s) \<Longrightarrow> sectxt_of_obj (Kill p p' # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_ptrace:
+ "valid (Ptrace p p' # s) \<Longrightarrow> sectxt_of_obj (Ptrace p p' # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_exit:
+ "valid (Exit p # s) \<Longrightarrow> sectxt_of_obj (Exit p # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_readfile:
+ "valid (ReadFile p fd # s) \<Longrightarrow> sectxt_of_obj (ReadFile p fd # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_writefile:
+ "valid (WriteFile p fd # s) \<Longrightarrow> sectxt_of_obj (WriteFile p fd # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_closefd:
+ "valid (CloseFd p fd # s) \<Longrightarrow> sectxt_of_obj (CloseFd p fd # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_unlink:
+ "valid (UnLink p f # s) \<Longrightarrow> sectxt_of_obj (UnLink p f # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_Rmdir:
+ "valid (Rmdir p f # s) \<Longrightarrow> sectxt_of_obj (Rmdir p f # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_truncate:
+ "valid (Truncate p f len # s) \<Longrightarrow> sectxt_of_obj (Truncate p f len # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_recvmsg:
+ "valid (RecvMsg p q m # s) \<Longrightarrow> sectxt_of_obj (RecvMsg p q m # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_removemsgq:
+ "valid (RemoveMsgq p q # s) \<Longrightarrow> sectxt_of_obj (RemoveMsgq p q # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_attach:
+ "valid (Attach p h flag # s) \<Longrightarrow> sectxt_of_obj (Attach p h flag # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_detach:
+ "valid (Detach p h # s) \<Longrightarrow> sectxt_of_obj (Detach p h # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_deleteshm:
+ "valid (DeleteShM p h # s) \<Longrightarrow> sectxt_of_obj (DeleteShM p h # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_bind:
+ "valid (Bind p fd addr # s) \<Longrightarrow> sectxt_of_obj (Bind p fd addr # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_connect:
+ "valid (Connect p fd addr # s) \<Longrightarrow> sectxt_of_obj (Connect p fd addr # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_listen:
+ "valid (Listen p fd # s) \<Longrightarrow> sectxt_of_obj (Listen p fd # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_sendsock:
+ "valid (SendSock p fd # s) \<Longrightarrow> sectxt_of_obj (SendSock p fd # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_recvsock:
+ "valid (RecvSock p fd # s) \<Longrightarrow> sectxt_of_obj (RecvSock p fd # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemma sec_shutdown:
+ "valid (Shutdown p fd how # s) \<Longrightarrow> sectxt_of_obj (Shutdown p fd how # s) = sectxt_of_obj s"
+by (auto dest!:sec_others)
+
+lemmas sectxt_of_obj_simps = sec_execve sec_open sec_mkdir sec_linkhard sec_createmsgq sec_sendmsg
+ sec_createshm sec_createsock sec_accept sec_clone sec_kill sec_ptrace sec_exit sec_readfile
+ sec_writefile sec_closefd sec_unlink sec_Rmdir sec_truncate sec_recvmsg sec_removemsgq
+ sec_attach sec_detach sec_deleteshm sec_bind sec_connect sec_listen sec_sendsock
+ sec_recvsock sec_shutdown
+(* init_sectxt_prop *)
+
+(**************** get_parentfs_ctxts simpset **************)
+
+lemma parentf_is_dir_prop1: "\<lbrakk>is_dir s (x # pf); valid s\<rbrakk> \<Longrightarrow> is_dir s pf"
+apply (rule_tac f = "x # pf" in parentf_is_dir)
+by (auto simp:is_dir_def current_files_def split:option.splits t_inode_tag.splits)
+
+lemma parentf_is_dir_prop2: "\<lbrakk>is_file s (x # pf); valid s\<rbrakk> \<Longrightarrow> is_dir s pf"
+apply (rule_tac f = "x # pf" in parentf_is_dir)
+by (auto simp:is_dir_def is_file_def current_files_def split:option.splits t_inode_tag.splits)
+
+lemma parentf_is_dir_prop3: "\<lbrakk>(x # pf) \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> is_dir s pf"
+apply (rule_tac f = "x # pf" in parentf_is_dir)
+by (auto simp:is_dir_def current_files_def split:option.splits t_inode_tag.splits)
+
+lemma get_pfs_secs_prop':
+ "\<lbrakk>get_parentfs_ctxts s f = None; valid s\<rbrakk> \<Longrightarrow> \<not> is_dir s f"
+apply (induct f)
+by (auto split:option.splits dest:current_has_sec' parentf_is_dir_prop1)
+
+lemma get_pfs_secs_prop:
+ "\<lbrakk>is_dir s f; valid s\<rbrakk> \<Longrightarrow> \<exists> asecs. get_parentfs_ctxts s f = Some asecs"
+using get_pfs_secs_prop'
+by (auto)
+
+lemma get_pfs_secs_open:
+ "valid (Open p f flags fd opt # s) \<Longrightarrow> get_parentfs_ctxts (Open p f flags fd opt # s) = get_parentfs_ctxts s"
+apply (frule noroot_events, frule vd_cons, frule vt_grant_os)
+apply (rule ext, induct_tac x)
+by (auto split:option.splits simp:sectxt_of_obj_simps)
+
+lemma get_pfs_secs_other:
+ "\<lbrakk>valid (e # s); \<forall> p f inum. e \<noteq> Mkdir p f inum\<rbrakk>
+ \<Longrightarrow> get_parentfs_ctxts (e # s) = get_parentfs_ctxts s"
+apply (frule vd_cons, frule vt_grant_os, rule ext, induct_tac x)
+apply (case_tac [!] e)
+apply (auto split:option.splits if_splits simp:sectxt_of_obj_simps)
+done
+
+lemma get_pfs_secs_mkdir1:
+ assumes mkdir: "valid (Mkdir p f inum # s)" and noteq: "is_dir s f'"
+ shows "get_parentfs_ctxts (Mkdir p f inum # s) f' = get_parentfs_ctxts s f'"
+proof-
+ from mkdir have vd: "valid s" and os: "os_grant s (Mkdir p f inum)"
+ by (frule_tac vd_cons, simp, frule_tac vt_grant_os, simp)
+ from mkdir have notroot: "f \<noteq> []" by (auto intro!:noroot_mkdir)
+ show ?thesis using noteq
+ proof (induct f')
+ case Nil
+ show ?case using notroot mkdir by (simp add:sectxt_of_obj_simps)
+ next
+ case (Cons a f')
+ hence p1: "is_dir s f'" using vd by (simp add:parentf_is_dir_prop1)
+ from Cons have p2: "is_dir s (a # f')" by simp
+ from Cons p1 have p3: "get_parentfs_ctxts (Mkdir p f inum # s) f' = get_parentfs_ctxts s f'" by simp
+ from p2 os have p4: "a # f' \<noteq> f" by (auto simp:is_dir_in_current)
+ from p1 os have p5: "f' \<noteq> f" by (auto simp:is_dir_in_current)
+ show ?case using mkdir vd os p4 p5 p1
+ by (auto simp:sectxt_of_obj_simps is_dir_simps p3
+ split:option.splits dest:current_has_sec' get_pfs_secs_prop')
+ qed
+qed
+
+lemma get_pfs_secs_mkdir2:
+ "valid (Mkdir p f inum # s) \<Longrightarrow> get_parentfs_ctxts (Mkdir p f inum # s) f = (
+ case f of
+ [] \<Rightarrow> get_parentfs_ctxts s []
+ | x#pf \<Rightarrow> (case (get_parentfs_ctxts s pf, sectxt_of_obj s (O_proc p), sectxt_of_obj s (O_dir pf)) of
+ (Some pfsecs, Some psec, Some pfsec) \<Rightarrow> Some ((fst psec, R_object, type_transition ((snd o snd) psec) ((snd o snd) pfsec) C_dir True) # pfsecs)
+ | _ \<Rightarrow> None))"
+apply (frule vd_cons, frule vt_grant_os)
+apply (frule noroot_events, case_tac f)
+by (auto split:option.splits dest:current_has_sec' get_pfs_secs_prop' get_pfs_secs_mkdir1
+ simp:sectxt_of_obj_simps is_dir_simps)
+
+lemmas get_parentfs_ctxts_simps = get_pfs_secs_other get_pfs_secs_mkdir1 get_pfs_secs_mkdir2
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Static.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,654 @@
+theory Static
+imports Static_type OS_type_def Flask_type Flask
+begin
+
+locale tainting_s = tainting
+
+begin
+
+(*
+definition init_sectxt_of_file:: "t_file \<Rightarrow> security_context_t option"
+where
+ "init_sectxt_of_file f \<equiv>
+ if (is_init_file f) then init_sectxt_of_obj (O_file f)
+ else if (is_init_dir f) then init_sectxt_of_obj (O_dir f)
+ else None"
+*)
+
+definition sroot :: "t_sfile"
+where
+ "sroot \<equiv> (Init [], sec_of_root, None, {})"
+
+definition init_cf2sfile :: "t_file \<Rightarrow> t_sfile option"
+where
+ "init_cf2sfile f \<equiv>
+ case (parent f) of
+ None \<Rightarrow> Some sroot
+ | Some pf \<Rightarrow> if (is_init_file f) then
+ (case (init_sectxt_of_obj (O_file f), init_sectxt_of_obj (O_dir pf), get_parentfs_ctxts [] pf) of
+ (Some sec, Some psec, Some aseclist) \<Rightarrow> Some (Init f, sec, Some psec, set aseclist)
+ | _ \<Rightarrow> None) else
+ (case (init_sectxt_of_obj (O_dir f), init_sectxt_of_obj (O_dir pf), get_parentfs_ctxts [] pf) of
+ (Some sec, Some psec, Some aseclist) \<Rightarrow> Some (Init f, sec, Some psec, set aseclist)
+ | _ \<Rightarrow> None)"
+
+definition init_cf2sfiles :: "t_file \<Rightarrow> t_sfile set"
+where
+ "init_cf2sfiles f \<equiv> {sf. \<exists>f' \<in> init_same_inode_files f. init_cf2sfile f' = Some sf}"
+
+definition init_cfd2sfd :: "t_process \<Rightarrow> t_fd \<Rightarrow> (security_context_t \<times> t_open_flags \<times> t_sfile) option"
+where
+ "init_cfd2sfd p fd =
+ (case (init_file_of_proc_fd p fd, init_oflags_of_proc_fd p fd, init_sectxt_of_obj (O_fd p fd)) of
+ (Some f, Some flags, Some sec) \<Rightarrow> (case (init_cf2sfile f) of
+ Some sf \<Rightarrow> Some (sec, flags, sf)
+ | _ \<Rightarrow> None)
+ | _ \<Rightarrow> None)"
+
+definition init_cfds2sfds :: "t_process \<Rightarrow> (security_context_t \<times> t_open_flags \<times> t_sfile) set"
+where
+ "init_cfds2sfds p \<equiv> { sfd. \<exists> fd \<in> init_proc_file_fds p. init_cfd2sfd p fd = Some sfd}"
+
+(*
+definition init_ch2sshm :: "t_shm \<Rightarrow> t_sshm option"
+where
+ "init_ch2sshm h \<equiv> (case (init_sectxt_of_obj (O_shm h)) of
+ Some sec \<Rightarrow> Some (Init h, sec)
+ | _ \<Rightarrow> None)"
+
+definition init_cph2spshs
+ :: "t_process \<Rightarrow> (t_sshm \<times> t_shm_attach_flag) set"
+where
+ "init_cph2spshs p \<equiv> {(sh, flag)| sh flag h. (p, flag) \<in> init_procs_of_shm h \<and>
+ init_ch2sshm h = Some sh}"
+*)
+
+definition init_cp2sproc :: "t_process \<Rightarrow> t_sproc option"
+where
+ "init_cp2sproc p \<equiv> (case (init_sectxt_of_obj (O_proc p)) of
+ Some sec \<Rightarrow> Some (Init p, sec, (init_cfds2sfds p))
+ | None \<Rightarrow> None)"
+
+(* acient files of init-file
+definition init_o2s_afs :: "t_file \<Rightarrow> security_context_t set"
+where
+ "init_o2s_afs f \<equiv> {sec. \<exists> f'. f' \<preceq> f \<and> init_sectxt_of_obj (O_dir f') = Some sec}" *)
+
+definition init_cm2smsg :: "t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsg option"
+where
+ "init_cm2smsg q m \<equiv> (case (init_sectxt_of_obj (O_msg q m)) of
+ Some sec \<Rightarrow> Some (Init m, sec, (O_msg q m) \<in> seeds)
+ | _ \<Rightarrow> None)"
+
+fun init_cqm2sms :: "t_msgq \<Rightarrow> t_msg list \<Rightarrow> (t_smsg list) option"
+where
+ "init_cqm2sms q [] = Some []"
+| "init_cqm2sms q (m#ms) =
+ (case (init_cqm2sms q ms, init_cm2smsg q m) of
+ (Some sms, Some sm) \<Rightarrow> Some (sm # sms)
+ | _ \<Rightarrow> None)"
+
+definition init_cq2smsgq :: "t_msgq \<Rightarrow> t_smsgq option"
+where
+ "init_cq2smsgq q \<equiv> (case (init_sectxt_of_obj (O_msgq q), init_cqm2sms q (init_msgs_of_queue q)) of
+ (Some sec, Some sms) \<Rightarrow> Some (Init q, sec, sms)
+ | _ \<Rightarrow> None)"
+
+fun init_obj2sobj :: "t_object \<Rightarrow> t_sobject option"
+where
+ "init_obj2sobj (O_proc p) =
+ (case (init_cp2sproc p) of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> seeds))
+ | _ \<Rightarrow> None)"
+| "init_obj2sobj (O_file f) =
+ (case (init_cf2sfile f) of
+ Some sf \<Rightarrow> Some (S_file sf (O_file f \<in> seeds))
+ | _ \<Rightarrow> None)"
+| "init_obj2sobj (O_dir f) =
+ (case (init_cf2sfile f) of
+ Some sf \<Rightarrow> Some (S_dir sf)
+ | _ \<Rightarrow> None)"
+| "init_obj2sobj (O_msgq q) =
+ (case (init_cq2smsgq q) of
+ Some sq \<Rightarrow> Some (S_msgq sq)
+ | _ \<Rightarrow> None)"
+(*
+| "init_obj2sobj (O_shm h) =
+ (case (init_ch2sshm h) of
+ Some sh \<Rightarrow> Some (S_shm sh)
+ | _ \<Rightarrow> None)"
+| "init_obj2sobj (O_msg q m) =
+ (case (init_cq2smsgq q, init_cm2smsg q m) of
+ (Some sq, Some sm) \<Rightarrow> Some (S_msg sq sm)
+ | _ \<Rightarrow> None)" *)
+| "init_obj2sobj _ = None"
+
+definition
+ "init_static_state \<equiv> {sobj. \<exists> obj. init_alive obj \<and> init_obj2sobj obj = Some sobj}"
+
+(**************** translation from dynamic to static *******************)
+
+definition cf2sfile :: "t_state \<Rightarrow> t_file \<Rightarrow> t_sfile option"
+where
+ "cf2sfile s f \<equiv>
+ case (parent f) of
+ None \<Rightarrow> Some sroot
+ | Some pf \<Rightarrow> if (is_file s f)
+ then (case (sectxt_of_obj s (O_file f), sectxt_of_obj s (O_dir pf), get_parentfs_ctxts s pf) of
+ (Some sec, Some psec, Some asecs) \<Rightarrow>
+ Some (if (\<not> deleted (O_file f) s \<and> is_init_file f) then Init f else Created, sec, Some psec, set asecs)
+ | _ \<Rightarrow> None)
+ else (case (sectxt_of_obj s (O_dir f), sectxt_of_obj s (O_dir pf), get_parentfs_ctxts s pf) of
+ (Some sec, Some psec, Some asecs) \<Rightarrow>
+ Some (if (\<not> deleted (O_dir f) s \<and> is_init_dir f) then Init f else Created, sec, Some psec, set asecs)
+ | _ \<Rightarrow> None)"
+
+(*
+definition cf2sfiles :: "t_state \<Rightarrow> t_file \<Rightarrow> t_sfile set"
+where
+ "cf2sfiles s f \<equiv> {sf. \<exists> f' \<in> (same_inode_files s f). cf2sfile s f' = Some sf}"
+*)
+
+(* here cf2sfile is passed with True, because, process' fds are only for files not dirs *)
+definition cfd2sfd :: "t_state \<Rightarrow> t_process \<Rightarrow> t_fd \<Rightarrow> t_sfd option"
+where
+ "cfd2sfd s p fd \<equiv>
+ (case (file_of_proc_fd s p fd, flags_of_proc_fd s p fd, sectxt_of_obj s (O_fd p fd)) of
+ (Some f, Some flags, Some sec) \<Rightarrow> (case (cf2sfile s f) of
+ Some sf \<Rightarrow> Some (sec, flags, sf)
+ | _ \<Rightarrow> None)
+ | _ \<Rightarrow> None)"
+
+
+definition cpfd2sfds :: "t_state \<Rightarrow> t_process \<Rightarrow> t_sfd set"
+where
+ "cpfd2sfds s p \<equiv> {sfd. \<exists> fd \<in> proc_file_fds s p. cfd2sfd s p fd = Some sfd}"
+
+(*
+definition ch2sshm :: "t_state \<Rightarrow> t_shm \<Rightarrow> t_sshm option"
+where
+ "ch2sshm s h \<equiv> (case (sectxt_of_obj s (O_shm h)) of
+ Some sec \<Rightarrow>
+ Some (if (\<not> deleted (O_shm h) s \<and> h \<in> init_shms) then Init h else Created, sec)
+ | _ \<Rightarrow> None)"
+
+definition cph2spshs :: "t_state \<Rightarrow> t_process \<Rightarrow> t_sproc_sshm set"
+where
+ "cph2spshs s p \<equiv> {(sh, flag)| sh flag h. (p, flag) \<in> procs_of_shm s h \<and> ch2sshm s h = Some sh}"
+*)
+
+definition cp2sproc :: "t_state \<Rightarrow> t_process \<Rightarrow> t_sproc option"
+where
+ "cp2sproc s p \<equiv> (case (sectxt_of_obj s (O_proc p)) of
+ Some sec \<Rightarrow>
+ Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs) then Init p else Created, sec,
+ cpfd2sfds s p)
+ | _ \<Rightarrow> None)"
+
+definition cm2smsg :: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsg option"
+where
+ "cm2smsg s q m \<equiv> (case (sectxt_of_obj s (O_msg q m)) of
+ Some sec \<Rightarrow>
+ Some (if (\<not> deleted (O_msg q m) s \<and> m \<in> set (init_msgs_of_queue q)) then Init m else Created,
+ sec, O_msg q m \<in> tainted s)
+ | _ \<Rightarrow> None)"
+
+fun cqm2sms:: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_msg list \<Rightarrow> (t_smsg list) option"
+where
+ "cqm2sms s q [] = Some []"
+| "cqm2sms s q (m#ms) =
+ (case (cqm2sms s q ms, cm2smsg s q m) of
+ (Some sms, Some sm) \<Rightarrow> Some (sm#sms)
+ | _ \<Rightarrow> None)"
+
+definition cq2smsgq :: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_smsgq option"
+where
+ "cq2smsgq s q \<equiv> (case (sectxt_of_obj s (O_msgq q), cqm2sms s q (msgs_of_queue s q)) of
+ (Some sec, Some sms) \<Rightarrow>
+ Some (if (\<not> deleted (O_msgq q) s \<and> q \<in> init_msgqs) then Init q else Created,
+ sec, sms)
+ | _ \<Rightarrow> None)"
+
+fun co2sobj :: "t_state \<Rightarrow> t_object \<Rightarrow> t_sobject option"
+where
+ "co2sobj s (O_proc p) =
+ (case (cp2sproc s p) of
+ Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> tainted s))
+ | _ \<Rightarrow> None)"
+| "co2sobj s (O_file f) =
+ (case (cf2sfile s f) of
+ Some sf \<Rightarrow> Some (S_file sf (O_file f \<in> tainted s))
+ | _ \<Rightarrow> None)"
+| "co2sobj s (O_dir f) =
+ (case (cf2sfile s f) of
+ Some sf \<Rightarrow> Some (S_dir sf)
+ | _ \<Rightarrow> None)"
+| "co2sobj s (O_msgq q) =
+ (case (cq2smsgq s q) of
+ Some sq \<Rightarrow> Some (S_msgq sq)
+ | _ \<Rightarrow> None)"
+(*
+| "co2sobj s (O_shm h) =
+ (case (ch2sshm s h) of
+ Some sh \<Rightarrow> Some (S_shm sh)
+ | _ \<Rightarrow> None)"
+| "co2sobj s (O_msg q m) =
+ (case (cq2smsgq s q, cm2smsg s q m) of
+ (Some sq, Some sm) \<Rightarrow> Some (S_msg sq sm)
+ | _ \<Rightarrow> None)"
+*)
+| "co2sobj s _ = None"
+
+
+definition s2ss :: "t_state \<Rightarrow> t_static_state"
+where
+ "s2ss s \<equiv> {sobj. \<exists> obj. alive s obj \<and> co2sobj s obj = Some sobj}"
+
+
+(* ******************************************************** *)
+
+
+fun is_init_sfile :: "t_sfile \<Rightarrow> bool"
+where
+ "is_init_sfile (Init _, sec, psec,asec) = True"
+| "is_init_sfile _ = False"
+
+fun is_many_sfile :: "t_sfile \<Rightarrow> bool"
+where
+ "is_many_sfile (Created, sec, psec, asec) = True"
+| "is_many_sfile _ = False"
+
+fun is_init_sproc :: "t_sproc \<Rightarrow> bool"
+where
+ "is_init_sproc (Init p, sec, fds) = True"
+| "is_init_sproc _ = False"
+
+fun is_many_sproc :: "t_sproc \<Rightarrow> bool"
+where
+ "is_many_sproc (Created, sec,fds) = True"
+| "is_many_sproc _ = False"
+
+fun is_many_smsg :: "t_smsg \<Rightarrow> bool"
+where
+ "is_many_smsg (Created,sec,tag) = True"
+| "is_many_smsg _ = False"
+
+(* wrong def
+fun is_many_smsgq :: "t_smsgq \<Rightarrow> bool"
+where
+ "is_many_smsgq (Created,sec,sms) = (True \<and> (\<forall> sm \<in> (set sms). is_many_smsg sm))"
+| "is_many_smsgq _ = False"
+*)
+
+fun is_many_smsgq :: "t_smsgq \<Rightarrow> bool"
+where
+ "is_many_smsgq (Created,sec,sms) = True"
+| "is_many_smsgq _ = False"
+
+fun is_many :: "t_sobject \<Rightarrow> bool"
+where
+ "is_many (S_proc sp tag) = is_many_sproc sp"
+| "is_many (S_file sf tag) = is_many_sfile sf"
+| "is_many (S_dir sf ) = is_many_sfile sf"
+| "is_many (S_msgq sq ) = is_many_smsgq sq"
+
+fun is_init_smsgq :: "t_smsgq \<Rightarrow> bool"
+where
+ "is_init_smsgq (Init q,sec,sms) = True"
+| "is_init_smsgq _ = False"
+
+fun is_init_sobj :: "t_sobject \<Rightarrow> bool"
+where
+ "is_init_sobj (S_proc sp tag ) = is_init_sproc sp"
+| "is_init_sobj (S_file sf tag ) = is_init_sfile sf"
+| "is_init_sobj (S_dir sf ) = is_init_sfile sf"
+| "is_init_sobj (S_msgq sq ) = is_init_smsgq sq"
+
+(*
+fun update_ss_sp:: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
+where
+ "update_ss_sp ss (S_proc sp tag) (S_proc sp' tag') =
+ (if (is_many_sproc sp) then ss \<union> {S_proc sp' tag'}
+ else (ss - {S_proc sp tag}) \<union> {S_proc sp' tag'})"
+
+fun update_ss_sd:: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
+where
+ "update_ss_sd ss (S_dir sf tag) (S_dir sf' tag') =
+ (if (is_many_sfile sf) then ss \<union> {S_dir sf' tag'}
+ else (ss - {S_dir sf tag}) \<union> {S_dir sf' tag'})"
+*)
+
+(*
+fun sparent :: "t_sfile \<Rightarrow> t_sfile option"
+where
+ "sparent (Sroot si sec) = None"
+| "sparent (Sfile si sec spf) = Some spf"
+
+inductive is_ancesf :: "t_sfile \<Rightarrow> t_sfile \<Rightarrow> bool"
+where
+ "is_ancesf sf sf"
+| "sparent sf = Some spf \<Longrightarrow> is_ancesf spf sf"
+| "\<lbrakk>sparent sf = Some spf; is_ancesf saf spf\<rbrakk> \<Longrightarrow> is_ancesf saf sf"
+
+definition sfile_reparent :: "t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile"
+where
+ "sfile_reparent (Sroot)"
+*)
+
+(*
+(* sfds rename aux definitions *)
+definition sfds_rename_notrelated
+ :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
+where
+ "sfds_rename_notrelated sfds from to sfds' \<equiv>
+ (\<forall> sec flag sf. (sec,flag,sf) \<in> sfds \<and> (\<not> from \<preceq> sf) \<longrightarrow> (sec,flag,sf) \<in> sfds')"
+
+definition sfds_rename_renamed
+ :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
+where
+ "sfds_rename_renamed sfds sf from to sfds' \<equiv>
+ (\<forall> sec flag sf'. (sec,flag,sf) \<in> sfds \<and> (sf \<preceq> sf') \<longrightarrow>
+ (sec, flag, file_after_rename sf' from to) \<in> sfds' \<and> (sec,flag,sf) \<notin> sfds')"
+
+definition sfds_rename_remain
+ :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
+where
+ "sfds_rename_remain sfds sf from to sfds' \<equiv>
+ (\<forall> sec flag sf'. (sec,flag,sf') \<in> sfds \<and> (sf \<preceq> sf') \<longrightarrow>
+ (sec, flag, sf') \<in> sfds' \<and> (sec,flag, file_after_rename sf' from to) \<notin> sfds')"
+
+ (* for not many, choose on renamed or not *)
+definition sfds_rename_choices
+ :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
+where
+ "sfds_rename_choices sfds sf from to sfds' \<equiv>
+ sfds_rename_remain sfds sf from to sfds' \<or> sfds_rename_renamed sfds sf from to sfds'"
+
+(* for many, merge renamed with not renamed *)
+definition sfds_rename_both
+ :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
+where
+ "sfds_rename_both sfds sf from to sfds' \<equiv>
+ (\<forall> sec flag sf'. (sec,flag,sf') \<in> sfds \<and> (sf \<preceq> sf') \<longrightarrow>
+ (sec, flag, sf') \<in> sfds' \<or> (sec,flag, file_after_rename sf' from to) \<in> sfds')"
+
+(* added to the new sfds, are those only under the new sfile *)
+definition sfds_rename_added
+ :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
+where
+ "sfds_rename_added sfds from to sfds' \<equiv>
+ (\<forall> sec' flag' sf' sec flag. (sec',flag',sf') \<in> sfds' \<and> (sec,flag,sf') \<notin> sfds \<longrightarrow>
+ (\<exists> sf. (sec,flag,sf) \<in> sfds \<and> sf' = file_after_rename from to sf))"
+
+definition sproc_sfds_renamed
+ :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "sproc_sfds_renamed ss sf from to ss' \<equiv>
+ (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow>
+ (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_renamed sfds sf from to sfds'))"
+
+definition sproc_sfds_remain
+ :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "sproc_sfds_remain ss sf from to ss' \<equiv>
+ (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow>
+ (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_remain sfds sf from to sfds'))"
+
+(* for not many, choose on renamed or not *)
+definition sproc_sfds_choices
+ :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "sproc_sfds_choices ss sf from to ss' \<equiv>
+ (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow>
+ (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_choices sfds sf from to sfds'))"
+
+(* for many, merge renamed with not renamed *)
+definition sproc_sfds_both
+ :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "sproc_sfds_both ss sf from to ss' \<equiv>
+ (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow>
+ (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_both sfds sf from to sfds'))"
+
+(* remove (\<forall> sp tag. S_proc sp tag \<in> ss \<longrightarrow> S_proc sp tag \<in> ss'),
+ * cause sfds contains sfs informations *)
+definition ss_rename_notrelated
+ :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "ss_rename_notrelated ss sf sf' ss' \<equiv>
+ (\<forall> sq. S_msgq sq \<in> ss \<longrightarrow> S_msgq sq \<in> ss') \<and>
+ (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow> (\<exists> sfds'.
+ S_proc (pi,sec,sfds',sshms) tagp \<in> ss'\<and> sfds_rename_notrelated sfds sf sf' sfds')) \<and>
+ (\<forall> sfs sf'' tag. S_file sfs tag \<in> ss \<and> sf'' \<in> sfs \<and> (\<not> sf \<preceq> sf'') \<longrightarrow> (\<exists> sfs'.
+ S_file sfs tag \<in> ss' \<and> sf'' \<in> sfs')) \<and>
+ (\<forall> sf'' tag. S_dir sf'' tag \<in> ss \<and> (\<not> sf \<preceq> sf'') \<longrightarrow> S_dir sf'' tag \<in> ss')"
+
+(* rename from to, from should definited renamed if not many *)
+definition all_descendant_sf_renamed
+ :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "all_descendant_sf_renamed ss sf from to ss' \<equiv>
+ (\<forall> sfs sf' tagf. sf \<preceq> sf' \<and> S_file sfs tagf \<in> ss \<and> sf' \<in> sfs \<longrightarrow> (\<exists> sfs'.
+ S_file sfs' tagf \<in> ss' \<and> file_after_rename from to sf' \<in> sfs' \<and> sf' \<notin> sfs')) \<and>
+ (\<forall> sf' tagf. sf \<preceq> sf' \<and> S_dir sf' tagf \<in> ss \<longrightarrow> S_dir (file_after_rename from to sf') tagf \<in> ss' \<and> S_dir sf' tagf \<notin> ss') \<and>
+ sproc_sfds_renamed ss sf from to ss'"
+
+(* not renamed *)
+definition all_descendant_sf_remain
+ :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "all_descendant_sf_remain ss sf from to ss' \<equiv>
+ (\<forall> sfs sf' tag'. sf \<preceq> sf' \<and> S_file sfs tag' \<in> ss \<and> sf' \<in> sfs \<longrightarrow> (\<exists> sfs'.
+ S_file sfs' tag' \<in> ss \<and> file_after_rename from to sf' \<notin> sfs' \<and> sf' \<in> sfs')) \<and>
+ (\<forall> sf' tag'. sf \<preceq> sf' \<and> S_dir sf' tag' \<in> ss \<longrightarrow> S_dir (file_after_rename from to sf') tag' \<notin> ss' \<and> S_dir sf' tag' \<in> ss') \<and>
+ sproc_sfds_remain ss sf from to ss'"
+
+definition all_descendant_sf_choices
+ :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "all_descendant_sf_choices ss sf from to ss' \<equiv>
+ all_descendant_sf_renamed ss sf from to ss' \<or> all_descendant_sf_remain ss sf from to ss'"
+
+definition all_descendant_sf_both
+ :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "all_descendant_sf_both ss sf from to ss' \<equiv>
+ (\<forall> sfs sf' tag'. sf \<preceq> sf' \<and> S_file sfs tag' \<in> ss \<and> sf' \<in> sfs \<longrightarrow> (\<exists> sfs'.
+ S_file sfs' tag' \<in> ss \<and> file_after_rename from to sf' \<in> sfs' \<or> sf' \<in> sfs')) \<and>
+ (\<forall> sf' tag'. sf \<preceq> sf' \<and> S_dir sf' tag' \<in> ss \<longrightarrow>
+ S_dir (file_after_rename from to sf') tag' \<in> ss' \<or> S_dir sf' tag' \<in> ss') \<and>
+ sproc_sfds_both ss sf from to ss'"
+
+definition ss_renamed_file
+ :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "ss_renamed_file ss sf sf' ss' \<equiv>
+ (if (is_many_sfile sf)
+ then all_descendant_sf_choices ss sf sf sf' ss'
+ else all_descendant_sf_renamed ss sf sf sf' ss')"
+
+(* added to the new sfs, are those only under the new sfile *)
+definition sfs_rename_added
+ :: "t_sfile set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile set \<Rightarrow> bool"
+where
+ "sfs_rename_added sfs from to sfs' \<equiv>
+ (\<forall> sf'. sf' \<in> sfs' \<and> sf' \<notin> sfs \<longrightarrow> (\<exists> sf. sf \<in> sfs \<and> sf' = file_after_rename from to sf))"
+
+(* added to the new sfile, are those only under the new sfile *)
+definition ss_rename_added
+ :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "ss_rename_added ss from to ss' \<equiv>
+ (\<forall> pi sec fds fds' shms tagp. S_proc (pi, sec, fds,shms) tagp \<in> ss \<and>
+ S_proc (pi,sec,fds',shms) tagp \<in> ss' \<longrightarrow> sfds_rename_added fds from to fds') \<and>
+ (\<forall> sq. S_msgq sq \<in> ss' \<longrightarrow> S_msgq sq \<in> ss) \<and>
+ (\<forall> sfs sfs' tagf. S_file sfs' tagf \<in> ss' \<and> S_file sfs tagf \<in> ss \<longrightarrow>
+ sfs_rename_added sfs from to sfs') \<and>
+ (\<forall> sf' tagf. S_dir sf' tagf \<in> ss' \<and> S_dir sf' tagf \<notin> ss \<longrightarrow>
+ (\<exists> sf. S_dir sf tagf \<in> ss \<and> sf' = file_after_rename from to sf))"
+
+definition sfile_alive :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> bool"
+where
+ "sfile_alive ss sf \<equiv> (\<exists> sfs tagf. sf \<in> sfs \<and> S_file sfs tagf \<in> ss)"
+
+definition sf_alive :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> bool"
+where
+ "sf_alive ss sf \<equiv> (\<exists> tagd. S_dir sf tagd \<in> ss) \<or> sfile_alive ss sf"
+
+(* constrains that the new static state should satisfy *)
+definition ss_rename:: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "ss_rename ss sf sf' ss' \<equiv>
+ ss_rename_notrelated ss sf sf' ss' \<and>
+ ss_renamed_file ss sf sf' ss' \<and> ss_rename_added ss sf sf' ss' \<and>
+ (\<forall> sf''. sf_alive ss sf'' \<and> sf \<prec> sf'' \<longrightarrow>
+ (if (is_many_sfile sf'')
+ then all_descendant_sf_choices ss sf'' sf sf' ss'
+ else all_descendant_sf_both ss sf'' sf sf' ss'))"
+
+(* two sfile, the last fname should not be equal *)
+fun sfile_same_fname:: "t_sfile \<Rightarrow> t_sfile \<Rightarrow> bool"
+where
+ "sfile_same_fname ((Init n, sec)#spf) ((Init n', sec')#spf') = (n = n')"
+| "sfile_same_fname _ _ = False"
+
+(* no same init sfile/only sfile in the target-to spf, should be in static_state addmissble check *)
+definition ss_rename_no_same_fname:: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> bool"
+where
+ "ss_rename_no_same_fname ss from spf \<equiv>
+ \<not> (\<exists> to. sfile_same_fname from to \<and> parent to = Some spf \<and> sf_alive ss to)"
+
+(* is not a function, is a relation (one 2 many)
+definition update_ss_rename :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state"
+where
+ "update_ss_rename ss sf sf' \<equiv> if (is_many_sfile sf)
+ then (ss \<union> {S_file (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_file sf'' tag \<in> ss}
+ \<union> {S_dir (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_dir sf'' tag \<in> ss})
+ else (ss - {S_file sf'' tag | sf'' tag. sf \<preceq> sf'' \<and> S_file sf'' tag \<in> ss}
+ - {S_dir sf'' tag | sf'' tag. sf \<preceq> sf'' \<and> S_dir sf'' tag \<in> ss})
+ \<union> {S_file (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_file sf'' tag \<in> ss}
+ \<union> {S_dir (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_dir sf'' tag \<in> ss}"
+*)
+*)
+
+fun sectxt_of_sproc :: "t_sproc \<Rightarrow> security_context_t"
+where
+ "sectxt_of_sproc (pi,sec,fds) = sec"
+
+fun sectxt_of_sfile :: "t_sfile \<Rightarrow> security_context_t"
+where
+ "sectxt_of_sfile (fi,sec,psec,asecs) = sec"
+
+fun asecs_of_sfile :: "t_sfile \<Rightarrow> security_context_t set"
+where
+ "asecs_of_sfile (fi,sec,psec,asecs) = asecs"
+
+definition search_check_s :: "security_context_t \<Rightarrow> t_sfile \<Rightarrow> bool \<Rightarrow> bool"
+where
+ "search_check_s pctxt sf if_file =
+ (if if_file
+ then search_check_ctxt pctxt (sectxt_of_sfile sf) (asecs_of_sfile sf) True
+ else search_check_ctxt pctxt (sectxt_of_sfile sf) (asecs_of_sfile sf) False)"
+
+definition sectxts_of_sfds :: "t_sfd set \<Rightarrow> security_context_t set"
+where
+ "sectxts_of_sfds sfds \<equiv> {ctxt. \<exists> flag sf. (ctxt, flag, sf) \<in> sfds}"
+
+definition inherit_fds_check_s :: "security_context_t \<Rightarrow> t_sfd set \<Rightarrow> bool"
+where
+ "inherit_fds_check_s pctxt sfds \<equiv> inherit_fds_check_ctxt pctxt (sectxts_of_sfds sfds)"
+
+(*
+fun smsg_related :: "t_msg \<Rightarrow> t_smsg list \<Rightarrow> bool"
+where
+ "smsg_related m [] = False"
+| "smsg_related m ((mi, sec, tag)#sms) =
+ (if (mi = Init m) then True else smsg_related m sms)"
+
+fun smsgq_smsg_related :: "t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsgq \<Rightarrow> bool"
+where
+ "smsgq_smsg_related q m (qi, sec, smsgslist) = ((qi = Init q) \<and> (smsg_related m smsgslist))"
+
+fun smsg_relatainted :: "t_msg \<Rightarrow> t_smsg list \<Rightarrow> bool"
+where
+ "smsg_relatainted m [] = False"
+| "smsg_relatainted m ((mi, sec, tag)#sms) =
+ (if (mi = Init m \<and> tag = True) then True else smsg_relatainted m sms)"
+
+fun smsgq_smsg_relatainted :: "t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsgq \<Rightarrow> bool"
+where
+ "smsgq_smsg_relatainted q m (qi, sec, smsgslist) =
+ ((qi = Init q) \<and> (smsg_relatainted m smsgslist))"
+*)
+
+fun sfile_related :: "t_sfile \<Rightarrow> t_file \<Rightarrow> bool"
+where
+ "sfile_related (fi,sec,psec,asecs) f = (fi = Init f)"
+(*
+fun sproc_related :: "t_process \<Rightarrow> t_sproc \<Rightarrow> bool"
+where
+ "sproc_related p (pi, sec, fds, shms) = (pi = Init p)"
+*)
+fun init_obj_related :: "t_sobject \<Rightarrow> t_object \<Rightarrow> bool"
+where
+ "init_obj_related (S_proc (pi, sec, fds) tag) (O_proc p') = (pi = Init p')"
+| "init_obj_related (S_file sf tag) (O_file f) = (sfile_related sf f)"
+| "init_obj_related (S_dir sf) (O_dir f) = (sfile_related sf f)"
+| "init_obj_related (S_msgq (qi, sec, sms)) (O_msgq q') = (qi = Init q')"
+(*
+| "init_obj_related (S_msg (qi, sec, sms) (mi, secm, tagm)) (O_msg q' m') = (qi = Init q' \<and> mi = Init m')"
+*)
+| "init_obj_related _ _ = False"
+
+
+
+(***************** for backward proof when Instancing static objects ******************)
+
+definition next_nat :: "nat set \<Rightarrow> nat"
+where
+ "next_nat nset = (Max nset) + 1"
+
+definition new_proc :: "t_state \<Rightarrow> t_process"
+where
+ "new_proc \<tau> = next_nat (current_procs \<tau>)"
+
+definition new_inode_num :: "t_state \<Rightarrow> nat"
+where
+ "new_inode_num \<tau> = next_nat (current_inode_nums \<tau>)"
+
+definition new_msgq :: "t_state \<Rightarrow> t_msgq"
+where
+ "new_msgq s = next_nat (current_msgqs s)"
+
+definition new_msg :: "t_state \<Rightarrow> t_msgq \<Rightarrow> nat"
+where
+ "new_msg s q = next_nat (set (msgs_of_queue s q))"
+
+definition new_proc_fd :: "t_state \<Rightarrow> t_process \<Rightarrow> t_fd"
+where
+ "new_proc_fd \<tau> p = next_nat (current_proc_fds \<tau> p)"
+
+definition all_fname_under_dir:: "t_file \<Rightarrow> t_state \<Rightarrow> t_fname set"
+where
+ "all_fname_under_dir d \<tau> = {fn. \<exists> f. fn # d = f \<and> f \<in> current_files \<tau>}"
+
+fun fname_all_a:: "nat \<Rightarrow> t_fname"
+where
+ "fname_all_a 0 = []" |
+ "fname_all_a (Suc n) = ''a''@(fname_all_a n)"
+
+definition fname_length_set :: "t_fname set \<Rightarrow> nat set"
+where
+ "fname_length_set fns = length`fns"
+
+definition next_fname:: "t_file \<Rightarrow> t_state \<Rightarrow> t_fname"
+where
+ "next_fname pf \<tau> = fname_all_a ((Max (fname_length_set (all_fname_under_dir pf \<tau>))) + 1)"
+
+definition new_childf:: "t_file \<Rightarrow> t_state \<Rightarrow> t_file"
+where
+ "new_childf pf \<tau> = next_fname pf \<tau> # pf"
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Static_type.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,98 @@
+theory Static_type
+imports Main OS_type_def Flask_type OS_type_def
+begin
+
+(* option: if some \<rightarrow> initial else none \<rightarrow> new created
+ * the constructor each is one-to-one according to security_class_t
+ * but dynamic-special objects are not included, such as:
+ * 1. file-descriptor: C_fd
+ * 2. network node : C_node
+ * 3.
+ *)
+(* Init \<rightarrow> exists in the initial state; Only: the only static object(cannot be cloned, e.g.)
+ * Many \<rightarrow> Many instantiation of this static object in the dynamic world
+ * tainted under:
+ * Init \<rightarrow> this init obj is tainted
+ * Only \<rightarrow> this new created is tainted
+ * Many \<rightarrow> at least one of them is tainted or all of them can be tainted
+ *)
+datatype 'a Instance =
+ Init 'a
+| Created (* new created *)
+(*
+| Only (* for process, cannot clone(that means no childprocesses in ss!);
+ for file, renamed init-file(still keeps Init fname!)) *)
+| Many
+*)
+
+(* (fi, file sectxt, parent file's sectxt, ancestral files' sectxts) *)
+type_synonym t_sfile = "(t_file Instance) \<times> security_context_t \<times> (security_context_t option) \<times> (security_context_t set)"
+
+(*
+type_synonym t_sfn = "(t_fname Instance) \<times> security_context_t"
+type_synonym t_sfile = "t_sfn list" (* apparently, "sfile set" is not a finite set! *)
+*)
+(*
+(* Sroot si sectxt : root sfile
+ * Sfile si sectxt Sdir : normal sfile *)
+datatype t_sfile =
+ Sroot "t_file Instance" "security_context_t"
+| Sfile "t_file Instance" "security_context_t" "t_sfile"
+*)
+
+
+(* This is not finite set! think it as nature numbers
+lemma finite_sfile_set: "finite (UNIV::(t_sfile set))"
+apply (rule finite.intros)
+unfolding UNIV_def
+apply (induct x rule:t_sfile.induct)
+apply (simp add:UNIV_def)
+
+done *)
+
+(* the flags tells the sfile is a file or dir: NO!
+ * here sfile means sfd is opened to a sfile that has that "type", it is associated with all dynamic
+ * files that according to this sfile. When we need to give such a file instance, we can give any dynamic
+ * file in that state which according to this sfd's "creating static_state", so we might record ss in fds??
+ * that means our static is a "(static-state list) set"?? *)
+type_synonym t_sfd = "(security_context_t \<times> t_open_flags \<times> t_sfile)"
+
+(*
+type_synonym t_sshm = "(t_shm Instance \<times> security_context_t)"
+
+type_synonym t_sproc_sshm = "(t_sshm \<times> t_shm_attach_flag)"
+*)
+
+(* (si, sectxt, sectxts of fds, sectxt of attached shms) *)
+type_synonym t_sproc =
+ "t_process Instance \<times> security_context_t \<times> (t_sfd set)"
+
+(* here is a exception, the tainted-flag is in this type, case, msgq will not be-tainted *)
+type_synonym t_smsg = "t_msg Instance \<times> security_context_t \<times> bool"
+
+(* (qmi, sectxt, sectxt of queue) *)
+type_synonym t_smsgq = "t_msgq Instance \<times> security_context_t \<times> (t_smsg list)"
+
+(* O_fd, O_*_sock, O_shm, O_msgq are all that can not be tainted *)
+datatype t_sobject =
+ S_proc "t_sproc" "bool"
+| S_file "t_sfile" "bool" (* here sfile set is for "linking", tainted is the inode *)
+| S_dir "t_sfile"
+| S_msgq "t_smsgq"
+(*
+| S_shm "t_sshm"
+| S_msg "t_smsgq" "t_smsg"
+*)
+(*
+datatype t_tainted_sobj =
+ TS_proc "t_sproc" "bool"
+| TS_file "t_sfile set" "bool"
+| TS_msg "t_smsgq" "t_smsg"
+*)
+
+(* the static state is the current static objects
+ * But, how to record the static fds ??? fd is a dynamic concept !!!
+ *)
+type_synonym t_static_state = "t_sobject set"
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Tainted_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,223 @@
+theory Tainted_prop
+imports Main Flask Flask_type Init_prop Current_files_prop Current_sockets_prop Delete_prop Proc_fd_of_file_prop Current_prop Alive_prop
+begin
+
+ML {*quick_and_dirty := true*}
+
+context tainting begin
+
+lemma valid_Tainted_obj:
+ "\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> (\<forall> f. obj \<noteq> O_dir f) \<and> (\<forall> q. obj \<noteq> O_msgq q) \<and> (\<forall> h. obj \<noteq> O_shm h) \<and> (\<forall> p fd. obj \<noteq> O_fd p fd) \<and> (\<forall> s. obj \<noteq> O_tcp_sock s) \<and> (\<forall> s. obj \<noteq> O_udp_sock s)"
+apply (induct s, simp, drule seeds_in_init, case_tac obj, simp+)
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto split:if_splits option.splits)
+done
+
+lemma Tainted_in_current:
+ "\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> alive s obj"
+apply (induct s, simp)
+apply (drule seeds_in_init, case_tac obj, simp_all add:is_file_nil)
+apply (frule vd_cons, frule valid_Tainted_obj, simp, frule vt_grant_os, case_tac a)
+apply (auto simp:alive_simps split:if_splits option.splits t_object.splits
+ intro:same_inode_files_prop1 procs_of_shm_prop2
+ dest:info_shm_flow_in_procs)
+apply (auto simp:same_inode_files_def is_file_def split:if_splits)
+done
+
+lemma Tainted_proc_in_current:
+ "\<lbrakk>O_proc p \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s"
+by (drule Tainted_in_current, simp+)
+
+
+lemma info_flow_shm_Tainted:
+ "\<lbrakk>O_proc p \<in> Tainted s; info_flow_shm s p p'; valid s\<rbrakk> \<Longrightarrow> O_proc p' \<in> Tainted s"
+proof (induct s arbitrary:p p')
+ case Nil
+ thus ?case by (simp add:flow_shm_in_seeds)
+next
+ case (Cons e s)
+ hence p1: "O_proc p \<in> Tainted (e # s)" and p2: "info_flow_shm (e # s) p p'" and p3: "valid (e # s)"
+ and p4: "\<And> p p'. \<lbrakk>O_proc p \<in> Tainted s; info_flow_shm s p p'\<rbrakk> \<Longrightarrow> O_proc p' \<in> Tainted s"
+ and p5: "valid s" and p6: "os_grant s e"
+ by (auto dest:vd_cons intro:vd_cons vt_grant_os)
+ have p4':
+ "\<And> p p' h flag. \<lbrakk>O_proc p \<in> Tainted s; (p, SHM_RDWR) \<in> procs_of_shm s h; (p', flag) \<in> procs_of_shm s h\<rbrakk>
+ \<Longrightarrow> O_proc p' \<in> Tainted s"
+ by (rule p4, auto simp:info_flow_shm_def one_flow_shm_def procs_of_shm_prop2 p5)
+ from p2 p3 have p7: "p \<in> current_procs (e # s)" and p8: "p' \<in> current_procs (e # s)"
+ by (auto dest:info_shm_flow_in_procs)
+ show ?case
+ proof (cases "self_shm s p p'")
+ case True with p1 show ?thesis by simp
+ next
+ case False
+ with p1 p2 p5 p6 p7 p8 p3 show ?thesis
+ apply (case_tac e)(*
+ prefer 7
+ apply (simp add:info_flow_shm_simps split:if_splits option.splits)
+ apply (rule allI|rule impI|rule conjI)+
+ apply simp
+ apply (case_tac "O_proc p \<in> Tainted s", drule_tac p'=p' in p4, simp+)
+ apply simp
+
+
+
+
+ apply (auto simp:info_flow_shm_simps one_flow_shm_def dest:Tainted_in_current
+ intro:p4 p4' split:if_splits option.splits)
+ apply (auto simp:info_flow_shm_def one_flow_shm_def)
+
+
+
+ apply (auto simp:one_flow_shm_def intro:p4 p4' split:if_splits option.splits)
+
+
+
+ prefer 7
+ apply (simp split:if_splits option.splits)
+ apply (rule allI|rule impI|rule conjI)+
+
+
+ apply (auto dest:p4' procs_of_shm_prop2 Tainted_in_current split:if_splits option.splits)[1]
+
+ apply (erule disjE, drule_tac p = p and p' = p' in p4', simp+)
+ apply (erule disjE, rule disjI2, rule disjI2, rule_tac x = h in exI, simp, rule_tac x= toflag in exI, simp)
+ apply ((erule exE|erule conjE)+)
+
+
+ apply (auto simp:info_flow_shm_def dest:p4'
+ procs_of_shm_prop2 Tainted_in_current split:if_splits option.splits)[1]
+ apply (drule_tac p = p and p' = p' in p4')
+ apply (erule_tac x = ha in allE, simp)
+ apply (drule_tac p = "nat1" and p' = p' in p4')
+ apply (auto dest:p4'[where p = nat1 and p' = p'])
+
+apply (induct s)
+apply simp defer
+apply (frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:info_flow_shm_def elim!:disjE)
+sorry *)
+ sorry
+qed
+qed
+
+lemma has_same_inode_comm:
+ "has_same_inode s f f' = has_same_inode s f' f"
+by (auto simp add:has_same_inode_def same_inode_files_def is_file_def)
+
+lemma tainted_imp_Tainted:
+ "obj \<in> tainted s \<Longrightarrow> obj \<in> Tainted s"
+apply (induct rule:tainted.induct) (*
+apply (simp_all add:vd_cons)
+apply (rule impI)
+
+apply (rule disjI2, rule_tac x = flag' in exI, simp)
+apply ((rule impI)+, erule_tac x = p' in allE, simp)
+*)
+apply (auto intro:info_flow_shm_Tainted simp:has_same_inode_def dest:vd_cons)
+apply (case_tac e, auto split:option.splits if_splits simp:alive_simps)
+done
+
+lemma Tainted_imp_tainted_aux_remain:
+ "\<lbrakk>obj \<in> Tainted s; valid (e # s); alive (e # s) obj; \<And> obj. obj \<in> Tainted s \<Longrightarrow> obj \<in> tainted s\<rbrakk>
+ \<Longrightarrow> obj \<in> tainted (e # s)"
+by (rule t_remain, auto)
+
+lemma Tainted_imp_tainted:
+ "\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> obj \<in> tainted s"
+apply (induct s arbitrary:obj, simp add:t_init)
+apply (frule Tainted_in_current, simp+)
+apply (frule vt_grant_os, frule vd_cons, case_tac a)
+apply (auto split:if_splits option.splits elim:Tainted_imp_tainted_aux_remain intro:tainted.intros
+ simp:has_same_inode_def)
+done
+
+lemma tainted_eq_Tainted:
+ "valid s \<Longrightarrow> (obj \<in> tainted s) = (obj \<in> Tainted s)"
+by (rule iffI, auto intro:Tainted_imp_tainted tainted_imp_Tainted)
+
+lemma info_flow_shm_tainted:
+ "\<lbrakk>O_proc p \<in> tainted s; info_flow_shm s p p'; valid s\<rbrakk> \<Longrightarrow> O_proc p' \<in> tainted s"
+by (simp only:tainted_eq_Tainted info_flow_shm_Tainted)
+
+
+lemma same_inode_files_Tainted:
+ "\<lbrakk>O_file f \<in> Tainted s; f' \<in> same_inode_files s f; valid s\<rbrakk> \<Longrightarrow> O_file f' \<in> Tainted s"
+apply (induct s arbitrary:f f', simp add:same_inode_in_seeds has_same_inode_def)
+apply (frule vt_grant_os, frule vd_cons, case_tac a)
+prefer 6
+apply (simp split:if_splits option.splits add:same_inode_files_open current_files_simps)
+prefer 8
+apply (frule Tainted_in_current, simp, simp add:alive.simps, drule is_file_in_current)
+apply (auto simp add:same_inode_files_closefd split:option.splits if_splits)[1]
+prefer 8
+apply (frule Tainted_in_current, simp, simp add:alive.simps, drule is_file_in_current)
+apply (auto simp add:same_inode_files_unlink split:option.splits if_splits)[1]
+prefer 10
+apply (auto split:if_splits option.splits simp:same_inode_files_linkhard current_files_simps)[1]
+apply (drule Tainted_in_current, simp, simp add:alive.simps is_file_in_current)
+apply (drule same_inode_files_prop5, simp)
+apply (drule same_inode_files_prop5, drule_tac f' = list1 and f'' = f' in same_inode_files_prop4, simp, simp)
+
+apply (auto simp:same_inode_files_other split:if_splits)
+apply (drule_tac f'' = f' and f' = f and f = fa in same_inode_files_prop4, simp+)
+apply (drule_tac f'' = f' and f' = f and f = list in same_inode_files_prop4, simp+)
+done
+
+lemma has_same_inode_Tainted:
+ "\<lbrakk>O_file f \<in> Tainted s; has_same_inode s f f'; valid s\<rbrakk> \<Longrightarrow> O_file f' \<in> Tainted s"
+by (simp add:has_same_inode_def same_inode_files_Tainted)
+
+lemma has_same_inode_tainted:
+ "\<lbrakk>O_file f \<in> tainted s; has_same_inode s f f'; valid s\<rbrakk> \<Longrightarrow> O_file f' \<in> tainted s"
+by (simp add:has_same_inode_Tainted tainted_eq_Tainted)
+
+lemma same_inodes_Tainted:
+ "\<lbrakk>f \<in> same_inode_files s f'; valid s\<rbrakk> \<Longrightarrow> (O_file f \<in> Tainted s) = (O_file f' \<in> Tainted s)"
+apply (frule same_inode_files_prop8, frule same_inode_files_prop7)
+apply (auto intro:has_same_inode_Tainted)
+done
+
+
+
+lemma tainted_in_current:
+ "obj \<in> tainted s \<Longrightarrow> alive s obj"
+apply (erule tainted.induct)
+apply (auto dest:vt_grant_os vd_cons info_shm_flow_in_procs procs_of_shm_prop2 simp:is_file_simps)
+apply (drule seeds_in_init, simp add:tobj_in_alive)
+apply (erule has_same_inode_prop2, simp, simp add:vd_cons)
+apply (frule vt_grant_os, simp)
+apply (erule has_same_inode_prop1, simp, simp add:vd_cons)
+done
+
+lemma tainted_is_valid:
+ "obj \<in> tainted s \<Longrightarrow> valid s"
+by (erule tainted.induct, auto intro:valid.intros)
+
+lemma t_remain_app:
+ "\<lbrakk>obj \<in> tainted s; \<not> deleted obj (s' @ s); valid (s' @ s)\<rbrakk>
+ \<Longrightarrow> obj \<in> tainted (s' @ s)"
+apply (induct s', simp)
+apply (simp (no_asm) only:cons_app_simp_aux, rule t_remain)
+apply (simp_all add:not_deleted_cons_D vd_cons)
+apply (drule tainted_in_current, simp add:not_deleted_imp_alive_cons)
+done
+
+lemma valid_tainted_obj:
+ "obj \<in> tainted s \<Longrightarrow> (\<forall> f. obj \<noteq> O_dir f) \<and> (\<forall> q. obj \<noteq> O_msgq q) \<and> (\<forall> h. obj \<noteq> O_shm h) \<and> (\<forall> p fd. obj \<noteq> O_fd p fd) \<and> (\<forall> s. obj \<noteq> O_tcp_sock s) \<and> (\<forall> s. obj \<noteq> O_udp_sock s)"
+apply (erule tainted.induct)
+apply (drule seeds_in_init)
+by auto
+
+lemma dir_not_tainted: "O_dir f \<in> tainted s \<Longrightarrow> False"
+by (auto dest:valid_tainted_obj)
+
+lemma msgq_not_tainted: "O_msgq q \<in> tainted s \<Longrightarrow> False"
+by (auto dest:valid_tainted_obj)
+
+lemma shm_not_tainted: "O_shm h \<in> tainted s \<Longrightarrow> False"
+by (auto dest:valid_tainted_obj)
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Temp.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,233 @@
+theory Temp
+imports Static_type OS_type_def Flask_type Flask Static
+begin
+
+context tainting_s
+
+begin
+
+definition update_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
+where
+ "update_ss ss so so' \<equiv> if (is_many so) then ss \<union> {so'} else (ss - {so}) \<union> {so'}"
+
+definition add_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
+where
+ "add_ss ss so \<equiv> ss \<union> {so}"
+
+definition del_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
+where
+ "del_ss ss so \<equiv> if (is_many so) then ss else ss - {so}"
+
+
+definition update_ss_shms :: "t_static_state \<Rightarrow> t_sproc \<Rightarrow> bool \<Rightarrow> t_static_state"
+where
+ "update_ss_shms ss spfrom tag \<equiv> {sobj. \<exists> sobj' \<in> ss.
+ (case sobj' of
+ S_proc sp tagp \<Rightarrow> if (info_flow_sshm spfrom sp)
+ then if (is_many_sproc sp)
+ then (sobj = S_proc sp tagp \<or> sobj = S_proc sp (tagp \<or> tag))
+ else (sobj = S_proc sp (tagp \<or> tag))
+ else (sobj = S_proc sp tag)
+ | _ \<Rightarrow> sobj = sobj')}"
+
+
+
+(* all reachable static states(sobjects set) *)
+inductive_set static :: "t_static_state set"
+where
+ s_init: "init_static_state \<in> static"
+| s_execve: "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; S_file sfs tagf \<in> ss;
+ (fi,fsec,pfsec,asecs) \<in> sfs; npctxt_execve pctxt fsec = Some pctxt';
+ grant_execve pctxt fsec pctxt'; search_check_s pctxt (fi,fsec,pfsec,asecs) True;
+ inherit_fds_check_s pctxt' fds'; fds' \<subseteq> fds\<rbrakk>
+ \<Longrightarrow> (update_ss ss (S_proc (pi, pctxt, fds, shms) tagp)
+ (S_proc (pi, pctxt', fds', {}) (tagp \<or> tagf))) \<in> static"
+| s_clone: "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss;
+ permission_check pctxt pctxt C_process P_fork;
+ inherit_fds_check_s pctxt fds'; fds' \<subseteq> fds;
+ inherit_shms_check_s pctxt shms'; shms' \<subseteq> shms\<rbrakk>
+ \<Longrightarrow> (add_ss ss (S_proc (Created, pctxt, fds', shms') tagp)) \<in> static"
+| s_kill: "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss;
+ S_proc (pi', pctxt', fds', shms') tagp' \<in> ss;
+ permission_check pctxt pctxt' C_process P_sigkill\<rbrakk>
+ \<Longrightarrow> (del_ss ss (S_proc (pi', pctxt', fds', shms') tagp')) \<in> static"
+| s_ptrace: "\<lbrakk>ss \<in> static; S_proc sp tagp \<in> ss; S_proc sp' tagp' \<in> ss;
+ permission_check (sextxt_of_sproc sp) (sectxt_of_sproc sp') C_process P_ptrace\<rbrakk>
+ \<Longrightarrow> (update_ss_shms (update_ss_shms ss sp tagp) sp' tagp') \<in> static"
+| s_exit: "\<lbrakk>ss \<in> static; S_proc sp tagp \<in> ss\<rbrakk> \<Longrightarrow> (del_ss ss (S_proc sp tagp)) \<in> static"
+| s_open: "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; S_file sfs tagf \<in> ss; sf \<in> sfs;
+ search_check_s pctxt sf True; \<not> is_creat_excl_flag flags;
+ oflags_check flags pctxt (sectxt_of_sfile sf); permission_check pctxt pctxt C_fd P_setattr\<rbrakk>
+ \<Longrightarrow> (update_ss ss (S_proc (pi, pctxt, fds, shms) tagp)
+ (S_proc (pi, pctxt, fds \<union> {(pctxt,flags,sf)}, shms) tagp)) \<in> static"
+| s_open': "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; is_creat_excl_flag flags;
+ S_dir (pfi,fsec,pfsec,asecs) \<in> ss; search_check_s pctxt (pfi,fsec,pfsec,asecs) False;
+ nfsec = nfctxt_create pctxt fsec C_file; oflags_check flags pctxt nfsec;
+ permission_check pctxt fsec C_dir P_add_rename; permission_check pctxt pctxt C_fd P_setattr\<rbrakk>
+ \<Longrightarrow> (update_ss (add_ss ss (S_file {(Created, nfsec, Some fsec, asecs \<union> {fsec})} tagp))
+ (S_proc (pi, pctxt, fds, shms) tagp)
+ (S_proc (pi, pctxt, fds \<union> {(pctxt, flags, (Created, nfsec, Some fsec, asecs \<union> {fsec}))}, shms) tagp)
+ ) \<in> static"
+| S_readf: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; (fdctxt,flags,sf) \<in> fds;
+ permission_check pctxt fdctxt C_fd P_setattr; S_file sfs tagf \<in> ss; sf \<in> sfs;
+ permission_check pctxt (sectxt_of_sfile sf) C_file P_read; is_read_flag flags\<rbrakk>
+ \<Longrightarrow> (update_ss_shms ss (pi, pctxt,fds,shms) (tagp \<or> tagf)) \<in> static"
+| S_writef: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; (fdctxt,flags,sf) \<in> fds;
+ permission_check pctxt fdctxt C_fd P_setattr; sf \<in> sfs; S_file sfs tagf \<in> ss;
+ permission_check pctxt (sectxt_of_sfile sf) C_file P_write; is_write_flag flags\<rbrakk>
+ \<Longrightarrow> (update_ss ss (S_file sfs tagf) (S_file sfs (tagp \<or> tagf))) \<in> static"
+| S_unlink: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_file sfs tagf \<in> ss;
+ (Init f,fsec,Some pfsec,asecs) \<in> sfs;
+ search_check_s pctxt (Init f,fsec,Some pfsec,asecs) True;
+ permission_check pctxt fsec C_file P_unlink;
+ permission_check pctxt pfsec C_dir P_remove_name\<rbrakk>
+ \<Longrightarrow> ((ss - {S_file sfs tagf}) \<union> {S_file (sfs - {(Init f,fsec,Some pfsec,asecs)}) tagf}) \<in> static"
+| S_rmdir: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss;
+ S_dir (fi,fsec,Some pfsec,asecs) \<in> ss;
+ search_check_s pctxt (fi,fsec,Some pfsec,asecs) False;
+ permission_check pctxt fsec C_dir P_rmdir;
+ permission_check pctxt pfsec C_dir P_remove_name\<rbrakk>
+ \<Longrightarrow> (del_ss ss (S_dir (fi,fsec,Some pfsec,asecs))) \<in> static"
+| S_mkdir: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_dir (fi,fsec,pfsec,asecs) \<in> ss;
+ search_check_s pctxt (fi,fsec,pfsec,asecs) False;
+ permission_check pctxt (nfctxt_create pctxt fsec C_dir) C_dir P_create;
+ permission_check pctxt fsec C_dir P_add_name\<rbrakk>
+ \<Longrightarrow> (add_ss ss (S_dir (Created,nfctxt_create pctxt fsec C_dir,Some fsec,asecs \<union> {fsec}))) \<in> static"
+| s_link: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_dir (pfi,pfsec,ppfsec,asecs) \<in> ss;
+ S_file sfs tagf \<in> ss; sf \<in> sfs; nfsec = nfctxt_create pctxt pfsec C_file;
+ search_check_s pctxt (pfi,pfsec,ppfsec,asecs) False; search_check_s pctxt sf True;
+ permission_check pctxt (sectxt_of_sfile sf) C_file P_link;
+ permission_check pctxt pfsec C_dir P_add_name\<rbrakk>
+ \<Longrightarrow> (update_ss ss (S_file sfs tagf)
+ (S_file (sfs \<union> {(Created,nfsec,Some pfsec, asecs \<union> {pfsec})}) tagf)) \<in> static"
+| s_trunc: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_file sfs tagf \<in> ss; sf \<in> sfs;
+ search_check_s pctxt sf True; permission_check pctxt (sectxt_of_sfile sf) C_file P_setattr\<rbrakk>
+ \<Longrightarrow> (update_ss ss (S_file sfs tagf) (S_file sfs (tagf \<or> tagp))) \<in> static"
+(*
+| s_rename: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_file sfs tagf \<in> ss;
+ (sf#spf') \<in> sfs; S_dir spf tagpf \<in> ss; \<not>((sf#spf') \<preceq> (sf#spf));
+ search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True;
+ sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt;
+ permission_check pctxt fctxt C_file P_rename;
+ permission_check pctxt pfctxt C_dir P_add_name;
+ ss_rename ss (sf#spf') (sf#spf) ss';
+ ss_rename_no_same_fname ss (sf#spf') (sf#spf)\<rbrakk>
+ \<Longrightarrow> ss' \<in> static"
+| s_rename': "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_dir (sf#spf') tagf \<in> ss;
+ S_dir spf tagpf \<in> ss; \<not>((sf#spf') \<preceq> (sf#spf));
+ search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True;
+ sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt;
+ permission_check pctxt fctxt C_dir P_reparent;
+ permission_check pctxt pfctxt C_dir P_add_name;
+ ss_rename ss (sf#spf') (sf#spf) ss';
+ ss_rename_no_same_fname ss (sf#spf') (sf#spf)\<rbrakk>
+ \<Longrightarrow> ss' \<in> static"
+*)
+| s_createq: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss;
+ permission_check pctxt pctxt C_msgq P_associate;
+ permission_check pctxt pctxt C_msgq P_create\<rbrakk>
+ \<Longrightarrow> (add_ss ss (S_msgq (Created,pctxt,[]))) \<in> static"
+| s_sendmsg: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_msgq (qi,qctxt,sms) \<in> ss;
+ permission_check pctxt qctxt C_msgq P_enqueue;
+ permission_check pctxt qctxt C_msgq P_write;
+ permission_check pctxt pctxt C_msg P_create\<rbrakk>
+ \<Longrightarrow> (update_ss ss (S_msgq (qi,qctxt,sms))
+ (S_msgq (qi,qctxt,sms @ [(Created, pctxt, tagp)]))) \<in> static"
+| s_recvmsg: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss;
+ S_msgq (qi,qctxt,(mi,mctxt,tagm)#sms) \<in> ss;
+ permission_check pctxt qctxt C_msgq P_read;
+ permission_check pctxt mctxt C_msg P_receive\<rbrakk>
+ \<Longrightarrow> (update_ss (update_ss_shms ss (pi,pctxt,fds,shms) (tagp \<or> tagm))
+ (S_msgq (qi, qctxt, (mi, mctxt, tagm)#sms))
+ (S_msgq (qi, qctxt, sms))) \<in> static"
+| s_removeq: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_msgq (qi,qctxt,sms) \<in> ss;
+ permission_check pctxt qctxt C_msgq P_destroy\<rbrakk>
+ \<Longrightarrow> (del_ss ss (S_msgq (qi,qctxt,sms))) \<in> static"
+| s_createh: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss;
+ permission_check pctxt pctxt C_shm P_associate;
+ permission_check pctxt pctxt C_shm P_create\<rbrakk>
+ \<Longrightarrow> (add_ss ss (S_shm (Created, pctxt))) \<in> static"
+| s_attach: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_shm (hi,hctxt) \<in> ss;
+ if flag = SHM_RDONLY then permission_check pctxt hctxt C_shm P_read
+ else (permission_check pctxt hctxt C_shm P_read \<and>
+ permission_check pctxt hctxt C_shm P_write)\<rbrakk>
+ \<Longrightarrow> (update_ss ss (S_proc (pi,pctxt,fds,shms) tagp)
+ (S_proc (pi,pctxt,fds,shms \<union> {((hi,hctxt),flag)}) tagp)) \<in> static"
+| s_detach: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_shm sh \<in> ss;
+ (sh,flag) \<in> shms; \<not> is_many_sshm sh\<rbrakk>
+ \<Longrightarrow> (update_ss ss (S_proc (pi,pctxt,fds,shms) tagp)
+ (S_proc (pi,pctxt,fds,shms - {(sh,flag)}) tagp)) \<in> static"
+| s_deleteh: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_shm (hi,hctxt) \<in> ss;
+ permission_check pctxt hctxt C_shm P_destroy; \<not> is_many_sshm sh\<rbrakk>
+ \<Longrightarrow> (remove_sproc_sshm (del_ss ss (S_shm (hi,hctxt))) (hi,hctxt)) \<in> static"
+
+
+fun tainted_s :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> bool"
+where
+ "tainted_s ss (S_proc sp tag) = (S_proc sp tag \<in> ss \<and> tag = True)"
+| "tainted_s ss (S_file sfs tag) = (S_file sfs tag \<in> ss \<and> tag = True)"
+(*
+| "tainted_s ss (S_msg (qi, sec, sms) (mi, secm, tag)) =
+ (S_msgq (qi, sec, sms) \<in> ss \<and> (mi,secm,tag) \<in> set sms \<and> tag = True)"
+*)
+| "tainted_s ss _ = False"
+
+(*
+fun tainted_s :: "t_object \<Rightarrow> t_static_state \<Rightarrow> bool"
+where
+ "tainted_s (O_proc p) ss = (\<exists> sp. S_proc sp True \<in> ss \<and> sproc_related p sp)"
+| "tainted_s (O_file f) ss = (\<exists> sfs sf. S_file sfs True \<in> ss \<and> sf \<in> sfs \<and> sfile_related f sf)"
+| "tainted_s (O_msg q m) ss = (\<exists> sq. S_msgq sq \<in> ss \<and> smsgq_smsg_relatainted q m sq)"
+| "tainted_s _ ss = False"
+*)
+
+definition taintable_s :: "t_object \<Rightarrow> bool"
+where
+ "taintable_s obj \<equiv> \<exists> ss \<in> static. \<exists> sobj. tainted_s ss sobj \<and> init_obj_related sobj obj \<and> init_alive obj"
+
+(* this definition is wrong, cause process can exited from static-world
+definition deletable_s :: "t_object \<Rightarrow> bool"
+where
+ "deletable_s obj \<equiv> init_alive obj \<and> (\<exists> ss \<in> static. \<forall> sobj \<in> ss. \<not> init_obj_related sobj obj)"
+*)
+
+fun deleted_s :: "t_sproc \<Rightarrow> t_object \<Rightarrow> bool"
+where
+ "deleted_s (pi, pctxt, fds, shms) (O_proc p) = (pi \<noteq> Init p \<and> permission_check pctxt "
+
+definition deletable_s :: "t_object \<Rightarrow> bool"
+ "deletable_s obj \<equiv> init_alive obj \<and> (\<exists> ss \<in> static. \<exists> sp tagp. S_proc sp tagp \<in> ss \<and> deleted_s sp obj)"
+
+definition undeletable_s :: "t_object \<Rightarrow> bool"
+where
+ "undeletable_s obj \<equiv> init_alive obj \<and> (\<forall> ss \<in> static. \<exists> sobj \<in> ss. init_obj_related sobj obj)"
+
+definition init_ss_eq:: "t_static_state \<Rightarrow> t_static_state \<Rightarrow> bool" (infix "\<doteq>" 100)
+where
+ "ss \<doteq> ss' \<equiv> ss \<subseteq> ss' \<and> {sobj. is_init_sobj sobj \<and> sobj \<in> ss'} \<subseteq> ss"
+
+lemma [simp]: "ss \<doteq> ss"
+by (auto simp:init_ss_eq_def)
+
+definition init_ss_in:: "t_static_state \<Rightarrow> t_static_state set \<Rightarrow> bool" (infix "\<propto>" 101)
+where
+ "ss \<propto> sss \<equiv> \<exists> ss' \<in> sss. ss \<doteq> ss'"
+
+lemma s2ss_included_sobj:
+ "\<lbrakk>alive s obj; co2sobj s obj= Some sobj\<rbrakk> \<Longrightarrow> sobj \<in> (s2ss s)"
+by (simp add:s2ss_def, rule_tac x = obj in exI, simp)
+
+lemma init_ss_in_prop:
+ "\<lbrakk>s2ss s \<propto> static; co2sobj s obj = Some sobj; alive s obj; init_obj_related sobj obj\<rbrakk>
+ \<Longrightarrow> \<exists> ss \<in> static. sobj \<in> ss"
+apply (simp add:init_ss_in_def init_ss_eq_def)
+apply (erule bexE, erule conjE)
+apply (rule_tac x = ss' in bexI, auto dest!:s2ss_included_sobj)
+done
+
+
+
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/simple_selinux/Valid_prop.thy Mon Dec 02 10:52:40 2013 +0800
@@ -0,0 +1,33 @@
+theory Valid_prop
+imports Main Flask Flask_type My_list_prefix
+begin
+
+context flask begin
+
+lemma vd_cons'[rule_format]: "valid s' \<Longrightarrow> e # s = s' \<longrightarrow> valid s"
+by (erule valid.induct, auto)
+
+lemma vd_cons: "valid (e # s) \<Longrightarrow> valid s"
+by (simp only:vd_cons')
+
+lemma vd_appd: " valid (\<tau>' @ \<tau>) \<Longrightarrow> valid \<tau>"
+apply (induct \<tau>')
+by (auto simp:vd_cons)
+
+lemma vd_preceq: "\<lbrakk>\<tau>' \<preceq> \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> valid \<tau>'"
+apply (erule no_juniorE)
+by (simp only:vd_appd)
+
+lemma vd_prec: "\<lbrakk>\<tau>' \<prec> \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> valid \<tau>'"
+apply (drule is_ancestor_no_junior)
+by (simp only:vd_preceq)
+
+lemma vt_grant_os: "valid (e # \<tau>) \<Longrightarrow> os_grant \<tau> e"
+by (erule valid.cases, simp+)
+
+lemma vt_grant: "valid (e # \<tau>) \<Longrightarrow> grant \<tau> e"
+by (erule valid.cases, simp+)
+
+end
+
+end
\ No newline at end of file