--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Alive_prop.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,299 @@
+theory Alive_prop
+imports Main Flask_type Flask Current_files_prop Current_sockets_prop Init_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_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)
+
+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
+
+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_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 shms. e \<noteq> Clone p p' fds shms;
+ \<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 h. e \<noteq> CreateShM p h;
+ \<forall> p h. e \<noteq> DeleteShM 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> 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_linkhard alive_createmsgq alive_removemsgq alive_createshm alive_deleteshm
+ 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/Co2sobj_prop.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,494 @@
+(*<*)
+theory Co2sobj_prop
+imports Main Flask Flask_type Static Static_type Sectxt_prop Init_prop Current_files_prop Current_sockets_prop
+begin
+(*<*)
+
+context tainting_s begin
+
+(************ simpset for cf2sfile ***************)
+
+declare get_parentfs_ctxts.simps [simp del]
+
+lemma cf2sfile_open_none:
+ "valid (Open p f flag fd None # s) \<Longrightarrow> cf2sfile (Open p f flag fd None # s) f b = cf2sfile s f b"
+apply (frule vd_cons, frule vt_grant_os)
+apply (frule noroot_events)
+by (auto split:if_splits option.splits dest!:current_has_sec'
+ simp:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ get_parentfs_ctxts_simps)
+
+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, rule ext)
+apply (induct f')
+apply (auto split:if_splits option.splits dest!:current_has_sec'
+ simp: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' True = cf2sfile s f' True"
+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' False = cf2sfile s f' False"
+apply (frule vd_cons, drule is_dir_in_current)
+by (simp add:cf2sfile_open_some1)
+
+lemma cf2sfile_open_some:
+ "valid (Open p f flag fd (Some inum) # s) \<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f True = (
+ 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 (auto split:if_splits option.splits dest!:current_has_sec'
+ simp:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ get_parentfs_ctxts_simps)
+
+
+
+lemma cf2sfile_open:
+ "valid (Open p f flag fd opt # s) \<Longrightarrow> cf2sfile (Open p f flag fd opt # s) = (\<lambda> f' b.
+ if (opt = None) then cf2sfile s f' b
+ else if (f' = f \<and> b = True)
+ then (case (parent f) of
+ Some pf \<Rightarrow> (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> f \<in> init_files)
+ then Init f else Created, sec, Some psec, set asecs)
+ | _ \<Rightarrow> None)
+ | None \<Rightarrow> None)
+ else cf2sfile s f' b)"
+apply (frule vd_cons, frule vt_grant_os, rule ext, rule ext)
+apply (auto split:if_splits option.splits
+ simp:sectxt_of_obj_simps)
+
+
+
+
+
+
+
+
+
+
+
+
+lemma cf2sfile_keep_path: "\<lbrakk>f \<preceq> f'; \<tau> \<in> vt rc_cs\<rbrakk> \<Longrightarrow> cf2sfile \<tau> f \<preceq> cf2sfile \<tau> f'"
+apply (induct f', simp add:no_junior_def)
+apply (case_tac "f = a # f'", simp add:cf2sfile.simps)
+apply (drule no_junior_noteq, simp, simp add:cf2sfile.simps)
+done
+
+lemma ckp'_aux: "\<not> f \<preceq> a # f' \<Longrightarrow> \<not> f \<preceq> f'"
+by (auto simp:no_junior_def)
+
+lemma cf2sfile_keep_path': "\<lbrakk>\<not> f \<preceq> f'; \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>; f' \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> \<not> cf2sfile \<tau> f \<preceq> cf2sfile \<tau> f'"
+apply (induct f', simp add:no_junior_def cf2sfile.simps, rule notI, drule cf2sfile_root_file, simp)
+apply (case_tac "f = a # f'", simp add:cf2sfile.simps)
+apply (rule notI)
+apply (frule ckp'_aux, simp, frule parentf_in_current', simp+)
+apply (case_tac "cf2sfile \<tau> f = cf2sfile \<tau> (a # f')", drule cf2sfile_inj, simp+)
+apply (simp add:cf2sfile.simps)
+by (drule no_junior_noteq, simp+)
+
+lemma cf2sfile_keep_path'': "\<lbrakk>cf2sfile \<tau> f \<preceq> cf2sfile \<tau> f'; \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>; f' \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> f \<preceq> f'"
+using cf2sfile_keep_path'
+by (auto)
+
+lemma cf2sfile_open_some': "\<lbrakk>f \<in> current_files \<tau>; Open p f' flags fd (Some inum) # \<tau> \<in> vt rc_cs\<rbrakk> \<Longrightarrow> cf2sfile (Open p f' flags fd (Some inum) # \<tau>) f = cf2sfile \<tau> f"
+apply (frule vt_cons')
+apply (drule vt_grant_os)
+apply (induct f)
+apply (simp add:cf2sfile.simps)+
+apply (frule parentf_in_current', simp)
+apply (auto simp:os_grant.simps index_of_file.simps split:option.splits)
+done
+
+lemma cf2sfile_open_some: "\<lbrakk>Open p f flags fd (Some inum) # \<tau> \<in> vt rc_cs; parent f = Some pf\<rbrakk>
+ \<Longrightarrow> cf2sfile (Open p f flags fd (Some inum) # \<tau>) f = (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf))"
+apply (case_tac f, simp)
+apply (frule vt_grant_os)
+apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_open_some')
+done
+
+lemma cf2sfile_open_none: "cf2sfile (Open p f flags fd None # \<tau>) f' = cf2sfile \<tau> f'"
+apply (induct f')
+by (simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits)+
+
+lemma cf2sfile_open: "\<lbrakk>Open p f flags fd opt # \<tau> \<in> vt rc_cs; f' \<in> current_files (Open p f flags fd opt # \<tau>)\<rbrakk> \<Longrightarrow>
+ cf2sfile (Open p f flags fd opt # \<tau>) f' = (
+ if (opt = None) then cf2sfile \<tau> f'
+ else if (f' = f) then (case (parent f) of
+ Some pf \<Rightarrow> (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf))
+ | _ \<Rightarrow> [])
+ else cf2sfile \<tau> f' )"
+apply (frule vt_grant_os)
+by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_open_none cf2sfile_open_some cf2sfile_open_some' split:if_splits option.splits)
+
+lemma cf2sfile_mkdir_some': "\<lbrakk>Mkdir p f' inum # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> cf2sfile (Mkdir p f' inum # \<tau>) f = cf2sfile \<tau> f"
+apply (frule vt_cons', drule vt_grant_os)
+apply (induct f, (simp add:cf2sfile.simps)+)
+apply (frule parentf_in_current', simp)
+apply (auto simp:os_grant.simps index_of_file.simps split:option.splits)
+done
+
+lemma cf2sfile_mkdir_some: "\<lbrakk>Mkdir p f inum # \<tau> \<in> vt rc_cs; parent f = Some pf\<rbrakk>
+ \<Longrightarrow> cf2sfile (Mkdir p f inum # \<tau>) f = (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf))"
+apply (case_tac f, simp)
+apply (frule vt_grant_os)
+apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_mkdir_some')
+done
+
+lemma cf2sfile_mkdir: "\<lbrakk>Mkdir p f inum # \<tau> \<in> vt rc_cs; f' \<in> current_files (Mkdir p f inum # \<tau>)\<rbrakk> \<Longrightarrow>
+ cf2sfile (Mkdir p f inum # \<tau>) f' = (
+ if (f' = f) then (case (parent f) of
+ Some pf \<Rightarrow> (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf))
+ | _ \<Rightarrow> [])
+ else cf2sfile \<tau> f' ) "
+apply (frule vt_grant_os)
+by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_mkdir_some cf2sfile_mkdir_some' split:if_splits option.splits)
+
+lemma cf2sfile_linkhard_none: "\<lbrakk>LinkHard p f\<^isub>1 f\<^isub>2 # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>) f = cf2sfile \<tau> f"
+apply (frule vt_cons')
+apply (drule vt_grant_os)
+apply (induct f)
+apply (simp add:cf2sfile.simps)+
+apply (frule parentf_in_current', simp)
+apply (auto simp:os_grant.simps index_of_file.simps split:option.splits)
+done
+
+lemma cf2sfile_linkhard_some:
+ "\<lbrakk>LinkHard p f\<^isub>1 f\<^isub>2 # \<tau> \<in> vt rc_cs; parent f\<^isub>2 = Some pf\<^isub>2\<rbrakk> \<Longrightarrow> cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>) f\<^isub>2 = (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>2))"
+apply (case_tac f\<^isub>2, simp)
+apply (frule vt_grant_os)
+apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_linkhard_none)
+done
+
+lemma cf2sfile_linkhard: "\<lbrakk>LinkHard p f\<^isub>1 f\<^isub>2 # \<tau> \<in> vt rc_cs; f \<in> current_files (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>)\<rbrakk> \<Longrightarrow>
+ cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>) f = (
+ if (f = f\<^isub>2) then (case (parent f\<^isub>2) of
+ Some pf\<^isub>2 \<Rightarrow> SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>2)
+ | _ \<Rightarrow> [])
+ else cf2sfile \<tau> f )"
+apply (frule vt_grant_os)
+by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_linkhard_none cf2sfile_linkhard_some split:if_splits option.splits)
+
+
+lemma no_junior_aux: "\<not> f\<^isub>2 \<preceq> a # f \<Longrightarrow> \<not> f\<^isub>2 \<preceq> f"
+by (auto simp:no_junior_def)
+
+lemma cf2sfile_rename_aux: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> f \<noteq> f\<^isub>3 \<and> \<not> f\<^isub>3 \<preceq> f"
+apply (frule vt_grant_os, simp add:os_grant.simps, (erule exE|erule conjE)+)
+apply (rule conjI, rule notI, simp)
+apply (rule notI, drule vt_cons', simp add:ancenf_in_current)
+done
+
+lemma cf2sfile_rename'1:
+ "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>; \<not> (f\<^isub>2 \<preceq> f)\<rbrakk> \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = cf2sfile \<tau> f"
+apply (frule vt_cons',frule vt_grant_os, induct f)
+apply (simp add:cf2sfile.simps)
+apply (frule parentf_in_current', simp, simp)
+apply (frule no_junior_aux, simp)
+apply (simp add:os_grant.simps, (erule exE|erule conjE)+)
+apply (drule cf2sfile_rename_aux, simp, erule conjE)
+apply (auto simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits)
+done
+
+lemma cf2sfile_rename'2:
+ "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; \<not> (f\<^isub>3 \<preceq> f)\<rbrakk> \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = cf2sfile \<tau> f"
+apply (frule vt_cons', induct f)
+by (auto simp add:index_of_file.simps cf2sfile.simps no_junior_def split:option.splits nat.splits)
+
+lemma cf2sfile_rename1:
+ "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; parent f\<^isub>3 = Some pf\<^isub>3\<rbrakk> \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f\<^isub>3 = SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)"
+apply (case_tac f\<^isub>3, simp add:cf2sfile.simps)
+apply (auto dest!:cf2sfile_rename'2 simp add:no_junior_def cf2sfile.simps index_of_file.simps split:option.splits)
+done
+
+lemma index_of_file_rename1: "\<lbrakk>f\<^isub>2 \<preceq> f; f \<noteq> f\<^isub>2\<rbrakk> \<Longrightarrow> index_of_file (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) (file_after_rename f\<^isub>2 f\<^isub>3 f) = index_of_file \<tau> f"
+apply (clarsimp simp add:index_of_file.simps split:option.splits)
+by (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop3, simp, erule conjE, simp add:file_renaming_prop5)
+
+lemma cf2sfile_rename2: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; (file_before_rename f\<^isub>2 f\<^isub>3 f) \<in> current_files \<tau>; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>3 \<preceq> f\<rbrakk>
+ \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> (file_before_rename f\<^isub>2 f\<^isub>3 f))"
+apply (induct f, simp add:no_junior_def)
+apply (case_tac "a # f = f\<^isub>3", simp)
+apply (drule cf2sfile_rename1, simp, simp add:file_renaming_prop0' file_renaming_prop0)
+
+apply (frule no_junior_noteq, simp, simp)
+apply (frule_tac file_renaming_prop1')
+apply (frule_tac f = f\<^isub>2 and \<tau> = \<tau> in cf2sfile_keep_path, simp add:vt_cons')
+apply (frule_tac f\<^isub>2 = f\<^isub>2 and a = a in file_renaming_prop8')
+apply (frule_tac f\<^isub>2 = f\<^isub>2 and a = a in file_renaming_prop6')
+apply (frule_tac f = "file_before_rename f\<^isub>2 f\<^isub>3 (a # f)" and \<tau> = \<tau> and f\<^isub>2 = f\<^isub>2 and f\<^isub>3 = f\<^isub>3 and p = p in index_of_file_rename1, simp add:file_before_rename_def)
+apply (drule_tac f\<^isub>3 = f\<^isub>3 and f\<^isub>1 = f and a = a and f\<^isub>2 = f\<^isub>2 in file_renaming_prop9', simp)
+apply (drule parentf_in_current', simp add:vt_cons')
+apply (simp add:cf2sfile.simps)
+apply (erule file_renaming_prop6[THEN sym])
+done
+
+lemma cf2sfile_rename2': "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>2 \<preceq> f\<rbrakk>
+ \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> f)"
+apply (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop5)
+apply (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop1)
+apply (drule_tac f = "file_after_rename f\<^isub>2 f\<^isub>3 f" in cf2sfile_rename2, simp+)
+done
+
+lemma cf2sfile_rename3: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>; parent f\<^isub>3 = Some pf\<^isub>3\<rbrakk>
+ \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> f)"
+apply (case_tac "f\<^isub>2 \<preceq> f")
+apply (rule cf2sfile_rename2', simp+)
+apply (frule vt_grant_os)
+apply (frule_tac \<tau> = \<tau> in cf2sfile_keep_path', simp add:vt_cons', simp add:os_grant.simps, simp)
+apply (simp add:file_after_rename_def)
+by (rule cf2sfile_rename'1, simp+)
+
+lemma cf2sfile_rename: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; f \<in> current_files (Rename p f\<^isub>2 f\<^isub>3 # \<tau>)\<rbrakk> \<Longrightarrow>
+ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = (
+ if (f\<^isub>3 \<preceq> f) then (case (parent f\<^isub>3) of
+ Some pf\<^isub>3 \<Rightarrow> file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> (file_before_rename f\<^isub>2 f\<^isub>3 f))
+ | _ \<Rightarrow> [])
+ else cf2sfile \<tau> f )"
+apply (frule vt_grant_os)
+apply (case_tac "f = f\<^isub>3", clarsimp simp:os_grant.simps, drule cf2sfile_rename1, simp+, simp add:file_renaming_prop0' file_renaming_prop0)
+
+apply (auto simp:os_grant.simps current_files_simps intro:cf2sfile_rename'2 cf2sfile_rename1 split:if_splits option.splits)
+apply (rule cf2sfile_rename2, simp, drule rename_renaming_decom, simp+)
+apply (drule_tac f\<^isub>2 = f\<^isub>2 and f\<^isub>1 = f\<^isub>1 and f\<^isub>3 = f\<^isub>3 in file_renaming_prop5, simp+)
+done
+
+lemma cf2sfile_other: "\<lbrakk>
+ \<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ \<forall> p f im. e \<noteq> Mkdir p f im;
+ \<forall> p f\<^isub>1 f\<^isub>2. e \<noteq> LinkHard p f\<^isub>1 f\<^isub>2;
+ \<forall> p f\<^isub>2 f\<^isub>3. e \<noteq> Rename p f\<^isub>2 f\<^isub>3\<rbrakk> \<Longrightarrow> cf2sfile (e # \<tau>) f' = cf2sfile \<tau> f'"
+apply (induct f', simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits)
+apply (case_tac e, auto simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits)
+done
+
+lemmas cf2sfile_nil = init_cf2sfile
+
+lemma cf2sfile_nil': "f \<in> current_files [] \<Longrightarrow> cf2sfile [] f = map SInit f"
+by (simp add:cf2sfile_nil current_files_simps)
+
+lemmas cf2sfile_simps = cf2sfile_nil cf2sfile_nil' cf2sfile_open cf2sfile_mkdir cf2sfile_linkhard cf2sfile_rename cf2sfile_other
+
+lemmas cf2sfile_simpss = cf2sfile_nil cf2sfile_nil' cf2sfile_open_some' cf2sfile_open_some cf2sfile_open_none cf2sfile_open cf2sfile_mkdir_some' cf2sfile_mkdir_some
+ cf2sfile_mkdir cf2sfile_linkhard_none cf2sfile_linkhard_some cf2sfile_linkhard cf2sfile_rename'1 cf2sfile_rename'2 cf2sfile_rename1
+ cf2sfile_rename2 cf2sfile_rename2' cf2sfile_rename3 cf2sfile_rename cf2sfile_other
+
+lemma cf2sfile_open_some'_inum: "\<lbrakk>Open p f' flags fd (Some inum) # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im\<rbrakk> \<Longrightarrow> cf2sfile (Open p f' flags fd (Some inum) # \<tau>) f = cf2sfile \<tau> f"
+by (simp add:cf2sfile_open_some' current_files_def)
+
+lemma cf2sfile_mkdir_some'_inum: "\<lbrakk>Mkdir p f' inum # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im\<rbrakk> \<Longrightarrow> cf2sfile (Mkdir p f' inum # \<tau>) f = cf2sfile \<tau> f"
+by (simp add:cf2sfile_mkdir_some' current_files_def)
+
+lemma cf2sfile_linkhard_none_inum: "\<lbrakk>LinkHard p f\<^isub>1 f\<^isub>2 # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im\<rbrakk> \<Longrightarrow> cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>) f = cf2sfile \<tau> f"
+by (simp add:cf2sfile_linkhard_none current_files_def)
+
+lemma cf2sfile_rename'1_inum:
+ "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im ; \<not> (f\<^isub>2 \<preceq> f)\<rbrakk> \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = cf2sfile \<tau> f"
+by (simp add:cf2sfile_rename'1 current_files_def)
+
+lemma cf2sfile_rename2_inum: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; inum_of_file \<tau> (file_before_rename f\<^isub>2 f\<^isub>3 f) = Some im; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>3 \<preceq> f\<rbrakk>
+ \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> (file_before_rename f\<^isub>2 f\<^isub>3 f))"
+by (simp add:cf2sfile_rename2 current_files_def)
+
+lemma cf2sfile_rename2'_inum: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>2 \<preceq> f\<rbrakk>
+ \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> f)"
+by (simp add:cf2sfile_rename2' current_files_def)
+
+lemma cf2sfile_rename3_inum: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im; parent f\<^isub>3 = Some pf\<^isub>3\<rbrakk>
+ \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> f)"
+by (simp add:cf2sfile_rename3 current_files_def)
+
+lemma cf2sfile_nil'_inum: "inum_of_file [] f = Some im \<Longrightarrow> cf2sfile [] f = map SInit f"
+by (simp add:cf2sfile_nil' current_files_def)
+
+lemmas cf2sfile_simps_inum = cf2sfile_nil cf2sfile_nil'_inum cf2sfile_open_some'_inum cf2sfile_open_some cf2sfile_open_none cf2sfile_open cf2sfile_mkdir_some'_inum cf2sfile_mkdir_some
+ cf2sfile_mkdir cf2sfile_linkhard_none_inum cf2sfile_linkhard_some cf2sfile_linkhard cf2sfile_rename'1_inum cf2sfile_rename'2 cf2sfile_rename1
+ cf2sfile_rename2_inum cf2sfile_rename2'_inum cf2sfile_rename3_inum cf2sfile_rename cf2sfile_other
+
+(*************** cp2sproc simpset *********************)
+
+lemma cp2sproc_nil: "p \<in> init_processes \<Longrightarrow> cp2sproc [] p = SInit p"
+by (simp add:cp2sproc_def index_of_proc.simps)
+
+lemma cp2sproc_nil': "p \<in> current_procs [] \<Longrightarrow> cp2sproc [] p = SInit p"
+by (simp add:cp2sproc_nil current_procs.simps)
+
+lemma cp2sproc_clone: "cp2sproc (Clone p p' # \<tau>) p'' = (
+ if (p'' = p') then SCrea (Suc (length \<tau>))
+ else cp2sproc \<tau> p'' )"
+by (auto simp:cp2sproc_def index_of_proc.simps d2s_aux.simps)
+
+lemma cp2sproc_other: "\<forall> p p'. e \<noteq> Clone p p' \<Longrightarrow> cp2sproc (e # \<tau>) p'' = cp2sproc \<tau> p''"
+apply (case_tac e)
+by (auto simp:cp2sproc_def index_of_proc.simps d2s_aux.simps)
+
+lemmas cp2sproc_simps = cp2sproc_nil cp2sproc_nil' cp2sproc_clone cp2sproc_other
+
+(******************** ch2sshm simpset **************************)
+
+lemma ch2sshm_nil: "h \<in> init_shms \<Longrightarrow> ch2sshm [] h = SInit h"
+by (simp add:ch2sshm_def index_of_shm.simps)
+
+lemma ch2sshm_nil': "h \<in> current_shms [] \<Longrightarrow> ch2sshm [] h = SInit h"
+by (simp add:ch2sshm_nil current_shms.simps)
+
+lemma ch2sshm_createshm: "ch2sshm (CreateShM p h # \<tau>) h' = (if (h' = h) then SCrea (Suc (length \<tau>)) else ch2sshm \<tau> h')"
+by (simp add:ch2sshm_def index_of_shm.simps d2s_aux.simps)
+
+lemma ch2sshm_other: "\<forall> p h. e \<noteq> CreateShM p h \<Longrightarrow> ch2sshm (e # \<tau>) h' = ch2sshm \<tau> h'"
+apply (case_tac e)
+by (auto simp add:ch2sshm_def index_of_shm.simps d2s_aux.simps)
+
+lemmas ch2sshm_simps = ch2sshm_nil ch2sshm_nil' ch2sshm_createshm ch2sshm_other
+
+(********************* cm2smsg simpset ***********************)
+
+lemma cm2smsg_nil: "m \<in> init_msgs \<Longrightarrow> cm2smsg [] m = SInit m"
+by (simp add:cm2smsg_def index_of_msg.simps)
+
+lemma cm2smsg_nil': "m \<in> current_msgs [] \<Longrightarrow> cm2smsg [] m = SInit m"
+by (simp add:cm2smsg_nil current_msgs.simps)
+
+lemma cm2smsg_createmsg: "cm2smsg (CreateMsg p m # \<tau>) m' = (if (m' = m) then SCrea (Suc (length \<tau>)) else cm2smsg \<tau> m')"
+by (simp add:cm2smsg_def index_of_msg.simps d2s_aux.simps)
+
+lemma cm2smsg_other: "\<forall> p m. e \<noteq> CreateMsg p m \<Longrightarrow> cm2smsg (e # \<tau>) m' = cm2smsg \<tau> m'"
+apply (case_tac e)
+by (auto simp:cm2smsg_def index_of_msg.simps d2s_aux.simps)
+
+lemmas cm2smsg_simps = cm2smsg_nil cm2smsg_nil' cm2smsg_createmsg cm2smsg_other
+
+(********************** cfd2fd_s simpset ******************************)
+
+lemma cfd2fd_s_nil: "fd \<in> init_fds_of_proc p \<Longrightarrow> cfd2fd_s [] p fd = SInit fd"
+by (simp add:cfd2fd_s_def index_of_fd.simps)
+
+lemma cfd2fd_s_nil': "fd \<in> current_proc_fds [] p \<Longrightarrow> cfd2fd_s [] p fd = SInit fd"
+by (simp add:cfd2fd_s_nil current_proc_fds.simps)
+
+lemma cfd2fd_s_open: "cfd2fd_s (Open p f flags fd opt # \<tau>) p' fd' = (
+ if (p = p') then (if (fd = fd') then SCrea (Suc (length \<tau>))
+ else cfd2fd_s \<tau> p' fd')
+ else cfd2fd_s \<tau> p' fd' )"
+by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps)
+
+lemma cfd2fd_s_createsock: "cfd2fd_s (CreateSock p af st fd im # \<tau>) p' fd' = (
+ if (p = p') then (if (fd = fd') then SCrea (Suc (length \<tau>))
+ else cfd2fd_s \<tau> p' fd')
+ else cfd2fd_s \<tau> p' fd' )"
+by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps)
+
+lemma cfd2fd_s_accept: "cfd2fd_s (Accept p fd addr port fd' im # \<tau>) p' fd'' = (
+ if (p' = p) then (if (fd'' = fd') then SCrea (Suc (length \<tau>))
+ else cfd2fd_s \<tau> p' fd'')
+ else cfd2fd_s \<tau> p' fd'' )"
+by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps)
+
+lemma cfd2fd_s_clone: "cfd2fd_s (Clone p p' # \<tau>) p'' fd = (if (p'' = p') then cfd2fd_s \<tau> p fd else cfd2fd_s \<tau> p'' fd)"
+by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps)
+
+lemma cfd2fd_s_other: "\<lbrakk>\<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ \<forall> p af st fd im. e \<noteq> CreateSock p af st fd im;
+ \<forall> p fd addr port fd' im. e \<noteq> Accept p fd addr port fd' im;
+ \<forall> p p'. e \<noteq> Clone p p'\<rbrakk> \<Longrightarrow> cfd2fd_s (e # \<tau>) p'' fd'' = cfd2fd_s \<tau> p'' fd''"
+by (case_tac e, auto simp:cfd2fd_s_def index_of_fd.simps d2s_aux.simps)
+
+lemmas cfd2fd_s_simps = cfd2fd_s_nil cfd2fd_s_nil' cfd2fd_s_open cfd2fd_s_createsock cfd2fd_s_accept cfd2fd_s_clone cfd2fd_s_other
+
+(************* cim2im_s simpset **************************)
+
+(* no such lemma
+lemma cim2im_s_nil: "init_itag_of_inum im = Some tag \<Longrightarrow> cim2im_s [] im = SInit im"
+by (simp add:cim2im_s_def)
+*)
+
+lemma cim2im_s_open: "cim2im_s (Open p f flags fd (Some im) # \<tau>) im' = (if (im' = im) then SCrea (Suc (length \<tau>)) else cim2im_s \<tau> im')"
+by (simp add:cim2im_s_def)
+
+lemma cim2im_s_open': "cim2im_s (Open p f flags fd None # \<tau>) im = cim2im_s \<tau> im"
+by (simp add:cim2im_s_def)
+
+lemma cim2im_s_mkdir: "cim2im_s (Mkdir p f im # \<tau>) im' = (if (im' = im) then SCrea (Suc (length \<tau>)) else cim2im_s \<tau> im')"
+by (simp add:cim2im_s_def)
+
+lemma cim2im_s_createsock: "cim2im_s (CreateSock p sf st fd im # \<tau>) im' = (if (im' = im) then SCrea (Suc (length \<tau>)) else cim2im_s \<tau> im')"
+by (simp add:cim2im_s_def)
+
+lemma cim2im_s_accept: "cim2im_s (Accept p fd addr port fd' im # \<tau>) im' = (if (im' = im) then SCrea (Suc (length \<tau>)) else cim2im_s \<tau> im')"
+by (simp add:cim2im_s_def)
+
+lemma cim2im_s_other: "\<lbrakk>\<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ \<forall> p f im. e \<noteq> Mkdir p f im;
+ \<forall> p sf st fd im. e \<noteq> CreateSock p sf st fd im;
+ \<forall> p fd addr port fd' im. e \<noteq> Accept p fd addr port fd' im\<rbrakk> \<Longrightarrow> cim2im_s (e # \<tau>) im = cim2im_s \<tau> im"
+by (case_tac e, auto simp:cim2im_s_def)
+
+lemmas cim2im_s_simps = cim2im_s_open cim2im_s_open' cim2im_s_mkdir cim2im_s_createsock cim2im_s_accept cim2im_s_other
+
+
+lemma cig2ig_s_simp: "cig2ig_s (e # \<tau>) tag = cig2ig_s \<tau> tag"
+apply (case_tac tag)
+by auto
+
+(******************* cobj2sobj no Suc (length \<tau>) ***********************)
+
+lemma cf2sfile_le_len: "\<lbrakk>cf2sfile \<tau> f = SCrea (Suc (length \<tau>)) # spf; f \<in> current_files \<tau>; \<tau> \<in> vt rc_cs\<rbrakk> \<Longrightarrow> False"
+apply (case_tac f, (simp add:cf2sfile.simps d2s_aux.simps)+)
+apply (case_tac "index_of_file \<tau> (a # list)", (simp add:d2s_aux.simps)+)
+by (drule index_of_file_le_length', simp+)
+
+lemma cf2sfile_le_len': "\<lbrakk>SCrea (Suc (length \<tau>)) # spf \<preceq> cf2sfile \<tau> f; f \<in> current_files \<tau>; \<tau> \<in> vt rc_cs\<rbrakk> \<Longrightarrow> False"
+apply (induct f)
+apply (simp add:no_junior_def cf2sfile.simps d2s_aux.simps)
+apply (case_tac "cf2sfile \<tau> (a # f) = SCrea (Suc (length \<tau>)) # spf")
+apply (drule_tac f = "a # f" in cf2sfile_le_len, simp+)
+apply (simp only:cf2sfile.simps d2s_aux.simps)
+apply (drule_tac no_junior_noteq, simp+)
+apply (rule impI, erule impE, simp+)
+apply (drule parentf_in_current', simp+)
+done
+
+lemma cp2sproc_le_len: "cp2sproc \<tau> p = SCrea (Suc (length \<tau>)) \<Longrightarrow> False"
+apply (simp add:cp2sproc_def, case_tac "index_of_proc \<tau> p")
+apply (simp add:d2s_aux.simps)+
+by (drule index_of_proc_le_length', simp)
+
+lemma ch2sshm_le_len: "ch2sshm \<tau> h = SCrea (Suc (length \<tau>)) \<Longrightarrow> False"
+apply (simp add:ch2sshm_def, case_tac "index_of_shm \<tau> h")
+apply (simp add:d2s_aux.simps)+
+by (drule index_of_shm_le_length', simp)
+
+lemma cm2smsg_le_len: "cm2smsg \<tau> m = SCrea (Suc (length \<tau>)) \<Longrightarrow> False"
+apply (simp add:cm2smsg_def, case_tac "index_of_msg \<tau> m")
+apply (simp add:d2s_aux.simps)+
+by (drule index_of_msg_le_length', simp)
+
+lemma cim2im_s_le_len: "cim2im_s \<tau> im = SCrea (Suc (length \<tau>)) \<Longrightarrow> False"
+apply (simp add:cim2im_s_def, case_tac "inum2ind \<tau> im")
+apply (simp add:d2s_aux.simps)+
+by (drule inum2ind_le_length', simp)
+
+lemma cfd2fd_s_le_len: "cfd2fd_s \<tau> p fd = SCrea (Suc (length \<tau>)) \<Longrightarrow> False"
+apply (simp add:cfd2fd_s_def, case_tac "index_of_fd \<tau> p fd")
+apply (simp add:d2s_aux.simps)+
+by (drule index_of_fd_le_length', simp)
+
+end
+
+(*<*)
+end
+(*>*)
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Current_files_prop.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,736 @@
+(*<*)
+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
+ 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 split:if_splits option.splits)
+ apply (drule_tac f = pf and f' = f in hns', simp, simp, simp)
+ apply (drule_tac f = list and f' = f in fns', simp, simp, simp)
+ apply (drule_tac f = list and f' = f in dns', simp, 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''
+ 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 (drule_tac f = f' and pf = list in pin', simp, simp)
+ apply (drule_tac f = f' and pf = list2 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)
+
+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)
+
+
+
+(************** file_of_proc_fd in current ********************)
+
+lemma file_of_pfd_imp_inode_aux: "\<forall> p f. file_of_proc_fd \<tau> p fd = Some f \<and> valid \<tau> \<longrightarrow> inum_of_file \<tau> f \<noteq> None"
+apply (induct \<tau>)
+apply (clarsimp simp add:file_of_proc_fd.simps inum_of_file.simps init_filefd_prop1 init_file_has_inum)
+apply ((rule_tac allI|rule_tac impI|erule_tac conjE)+, frule vd_cons, frule vt_grant_os, case_tac a)
+apply (auto simp:inum_of_file.simps file_of_proc_fd.simps os_grant.simps current_files_def
+ split:if_splits option.splits)
+apply (simp add:pfdof_simp3)+
+apply (simp add:proc_fd_of_file_def)+
+done
+
+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 (rule file_of_pfd_imp_inode_aux[rule_format], 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_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"
+apply (induct \<tau> arbitrary:p, simp)
+apply (drule init_filefd_prop5, simp add:is_init_file_def split:option.splits t_inode_tag.splits)
+apply (frule vd_cons, frule vt_grant_os, simp, case_tac a)
+by (auto split:option.splits t_inode_tag.splits if_splits t_socket_type.splits
+ dest:file_of_pfd_imp_inode' iof's_im_in_cim
+ simp:is_file_def is_dir_def dir_is_empty_def current_files_def)
+
+lemma file_of_pfd_is_file:
+ "\<lbrakk>file_of_proc_fd \<tau> p fd = Some f; valid \<tau>\<rbrakk> \<Longrightarrow> is_file \<tau> f"
+apply (frule file_of_pfd_imp_inode, simp, erule exE)
+apply (drule file_of_pfd_is_file_tag, simp+)
+by (simp add: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)
+
+(************** 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 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>) = (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_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 current_files_linkhard:
+ "valid (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>) \<Longrightarrow> current_files (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>) = insert f\<^isub>2 (current_files \<tau>)"
+apply (frule vt_grant_os, frule vd_cons)
+by (auto simp:current_files_def inum_of_file.simps os_grant.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;
+ \<forall> p f f'. e \<noteq> LinkHard p f f'\<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_linkhard current_files_other
+
+
+(******************** is_file simpset *********************)
+
+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_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_linkhard:
+ "valid (LinkHard p f f' # s) \<Longrightarrow> is_file (LinkHard p f f' # s) = (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)
+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;
+ \<forall> p f f'. e \<noteq> LinkHard p f 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_linkhard is_file_other
+
+(********* is_dir simpset **********)
+
+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_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)
+apply (drule pfdof_simp2)
+apply (drule file_of_pfd_is_file, simp)
+apply (simp add:is_file_def split:t_inode_tag.splits option.splits)
+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_linkhard:
+ "valid (LinkHard p f f' # s) \<Longrightarrow> f \<noteq> [] \<and> f' \<noteq> []"
+by (frule vd_cons, drule vt_grant_os, auto dest!:root_is_dir')
+
+lemma noroot_linkhard':
+ "valid (LinkHard p [] f # s) \<Longrightarrow> False"
+by (drule noroot_linkhard, simp)
+
+lemma noroot_linkhard'':
+ "valid (LinkHard p f [] # s) \<Longrightarrow> False"
+by (drule noroot_linkhard, 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_linkhard noroot_truncate
+
+lemmas noroot_events' = noroot_execve' noroot_open' noroot_filefd' noroot_unlink' noroot_rmdir'
+ noroot_mkdir' noroot_linkhard' noroot_linkhard'' noroot_truncate'
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Current_sockets_prop.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,365 @@
+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 shms # \<tau>)
+ \<Longrightarrow> current_sockets (Clone p p' fds shms # \<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 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> 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 shms # s) \<Longrightarrow> is_tcp_sock (Clone p p' fds shms # 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 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> 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 shms # s) \<Longrightarrow> is_udp_sock (Clone p p' fds shms # 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 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> 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
+
+(***** 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/Delete_prop.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,53 @@
+theory Deleted_prop
+imports Main Flask Flask_type Init_prop Alive_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 not_deleted_imp_exists:
+ "\<lbrakk>\<not> deleted obj s; valid s; init_alive obj\<rbrakk> \<Longrightarrow> alive s obj"
+apply (induct s, simp add:init_alive_prop)
+apply (frule vd_cons)
+apply (case_tac a, auto simp:alive_simps split:t_object.splits)
+pr 19
+done
+
+lemma cons_app_simp_aux:
+ "(a # b) @ c = a # (b @ c)" by auto
+
+lemma not_deleted_imp_exists':
+ "\<lbrakk>\<not> deleted obj (s'@s); exists s obj\<rbrakk> \<Longrightarrow> exists (s'@s) obj"
+apply (induct s', simp, simp only:cons_app_simp_aux)
+apply (frule not_deleted_cons_D)
+apply (case_tac a, case_tac [!] obj, auto)
+done
+
+(*
+
+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/Dynamic2static.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,205 @@
+theory Dynamic2static
+imports Main Flask Static Init_prop Valid_prop
+begin
+
+context tainting_s begin
+
+lemma d2s_main:
+ "valid s \<Longrightarrow> s2ss s \<in> static"
+apply (induct s, simp add:s2ss_nil_prop s_init)
+apply (frule vd_cons, simp)
+apply (case_tac a, simp_all)
+(*
+apply
+induct s, case tac e, every event analysis
+*)
+sorry
+
+lemma is_file_has_sfile: "is_file s f \<Longrightarrow> \<exists> sf. cf2sfile s f True = Some sf"
+sorry
+
+lemma is_dir_has_sfile: "is_dir s f \<Longrightarrow> \<exists> sf. cf2sfile s f False = Some sf"
+sorry
+
+lemma is_file_imp_alive: "is_file s f \<Longrightarrow> alive s (O_file f)"
+sorry
+
+
+lemma d2s_main':
+ "\<lbrakk>alive s obj; co2sobj s obj= Some sobj\<rbrakk> \<Longrightarrow> sobj \<in> (s2ss s)"
+apply (induct s)
+apply (simp add:s2ss_def)
+apply (rule_tac x = obj in exI, simp)
+sorry
+
+lemma tainted_prop1:
+ "obj \<in> tainted s \<Longrightarrow> alive s obj"
+sorry
+
+lemma tainted_prop2:
+ "obj \<in> tainted s \<Longrightarrow> valid s"
+sorry
+
+lemma alive_has_sobj:
+ "\<lbrakk>alive s obj; valid s\<rbrakk> \<Longrightarrow> \<exists> sobj. co2sobj s obj = Some sobj"
+sorry
+
+lemma t2ts:
+ "obj \<in> tainted s \<Longrightarrow> co2sobj s obj = Some sobj \<Longrightarrow> tainted_s (s2ss s) sobj"
+apply (frule tainted_prop1, frule tainted_prop2)
+apply (simp add:s2ss_def)
+apply (case_tac sobj, simp_all)
+apply (case_tac [!] obj, simp_all split:option.splits)
+apply (rule_tac x = "O_proc nat" in exI, simp)
+apply (rule_tac x = "O_file list" in exI, simp)
+defer defer defer
+apply (case_tac prod1, simp, case_tac prod2, clarsimp)
+apply (rule conjI)
+apply (rule_tac x = "O_msgq nat1" in exI, simp)
+sorry (* doable, need properties about cm2smsg and cq2smsgq *)
+
+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 undel_init_file_remains:
+ "\<lbrakk>is_init_file f; \<not> deleted (O_file f) s\<rbrakk> \<Longrightarrow> is_file s f"
+sorry
+
+
+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_prop2)
+ hence static: "s2ss s \<in> static" using d2s_main by auto
+ from tainted have alive: "alive s obj"
+ using tainted_prop1 by auto
+ then obtain sobj where sobj: "co2sobj s obj = Some sobj"
+ using vs alive_has_sobj by blast
+ 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:cp2sproc_def ch2sshm_def cq2smsgq_def cm2smsg_def)
+ apply (frule undel_init_file_remains, simp, drule is_file_has_sfile, erule exE)
+ apply (rule_tac x = sf in bexI)
+ apply (case_tac list, auto split:option.splits simp:is_init_file_props)[1]
+ apply (simp add:same_inode_files_def cfs2sfiles_def)
+ apply (rule_tac x = list in exI, simp)
+ apply (case_tac list, auto split:option.splits simp:is_init_dir_props delq_imp_delqm)
+ done
+ with tainted t2ts init_alive sobj static
+ show ?thesis unfolding taintable_s_def
+ apply (rule_tac x = "s2ss s" in bexI, simp)
+ apply (rule_tac x = "sobj" in exI, auto)
+ done
+qed
+
+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 (induct s, simp)
+apply (frule vd_cons)
+apply (case_tac a, auto)
+(* need simpset for s2ss *)
+sorry
+
+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 (rule_tac x = "s2ss s" in bexI)
+apply (clarify, simp add:init_deled_imp_deled_s)
+apply (erule d2s_main)
+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 \<in> static" by (rule d2s_main)
+ with alive_s obtain sobj where in_ss: "sobj \<in> (s2ss s)"
+ and related: "init_obj_related sobj obj" by auto
+ 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 created_can_have_many:
+ "\<lbrakk>valid s; alive s obj; \<not> init_alive obj\<rbrakk> \<Longrightarrow> \<exists> s'. valid s' \<and> alive s' obj \<and> alive s' obj' \<and> s2ss s = s2ss s'"
+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 (erule exE, erule conjE)+
+
+sorry
+
+
+
+lemma tainted_s_imp_tainted:
+ "\<lbrakk>tainted_s ss sobj; ss \<in> static\<rbrakk> \<Longrightarrow> \<exists> obj s. s2ss s = ss \<and> valid s \<and> co2sobj s obj = Some sobj \<and> obj \<in> tainted s"
+sorry
+
+
+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 s2ss: "s2ss s = ss" and co2sobj: "co2sobj s obj' = Some sobj"
+ and tainted: "obj' \<in> tainted s" and vs: "valid s" by blast
+
+ from co2sobj related
+ have eq:"obj = obj'"
+ apply (case_tac obj', case_tac [!] obj, case_tac [!] sobj)
+ apply auto
+ apply (auto split:option.splits if_splits)
+ apply (case_tac a, simp+)
+ apply (simp add:cp2sproc_def split:option.splits if_splits)
+ apply simp
+ sorry
+ with tainted vs init_alive
+ show ?thesis by (auto simp:taintable_def)
+qed
+
+
+
+lemma ts2t:
+ "obj \<in> tainted_s ss \<Longrightarrow> \<exists> s. obj \<in> tainted s"
+ "obj \<in> tainted_s ss \<Longrightarrow> \<exists> so. so True \<in> ss \<Longrightarrow> so True \<in> ss \<Longrightarrow> \<exists> s. valid s \<and> s2ss s = ss \<Longrightarrow> so True \<in> s2ss s \<Longrightarrow> tainted s obj. "
+
+
+
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/File_renaming.thy Fri May 03 08:20:21 2013 +0100
@@ -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/Finite_current.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,109 @@
+theory Finite_current
+imports Main Valid_prop Flask Flask_type
+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)
+apply (rule_tac B = "insert list2 {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 (metis double_diff equalityE finite_Diff)
+by (metis double_diff finite_Diff subset_refl)
+
+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> set1}" 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_ch: "finite (current_shms \<tau>)"
+apply (induct \<tau>) defer
+apply (case_tac a, auto simp:current_shms.simps init_finite_sets)
+done
+
+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/Flask.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,1410 @@
+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"
+
+ 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 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_msgs_distinct: "\<And> q. distinct (init_msgs_of_queue q)"
+
+ 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 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_shms \<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)"
+
+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 shms # \<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}"
+
+
+(****** 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>) = (
+ if (proc_fd_of_file \<tau> f = {})
+ then (inum_of_file \<tau>) (f := None)
+ else inum_of_file \<tau> )"
+| "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}"
+
+definition has_same_inode :: "t_state \<Rightarrow> t_file \<Rightarrow> t_file \<Rightarrow> bool"
+where
+ "has_same_inode \<tau> fa fb =
+ (\<exists>i. inum_of_file \<tau> fa = Some i \<and> inum_of_file \<tau> fb = Some i)"
+
+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 shms # \<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}"
+
+(* 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>" (* may be sth wrong with nettemp representing addr \<times> netdev in statical analysis ??? *)
+
+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 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 shms # \<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 shms # \<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 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). \<exists> 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). \<exists> 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). \<exists> flag. (p, flag) \<in> procs_of_shm \<tau> h})"
+| "procs_of_shm (e # \<tau>) = procs_of_shm \<tau>"
+
+definition info_flow_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
+where
+ "info_flow_shm \<tau> from to \<equiv> (from = to) \<or> (\<exists> h \<in> current_shms \<tau>. \<exists> toflag.
+ (((from, SHM_RDWR) \<in> procs_of_shm \<tau> h) \<and> ((to, toflag) \<in> procs_of_shm \<tau> h)))"
+
+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 shms # \<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 shms # \<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> pf \<in> current_files \<tau> \<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> f \<in> current_files \<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> f \<in> current_files \<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> f \<in> current_files \<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> pf \<in> current_files \<tau> \<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> f\<^isub>1 \<in> current_files \<tau> \<and> f\<^isub>2 \<notin> current_files \<tau> \<and>
+ (\<exists> pf\<^isub>2. parent f\<^isub>2 = Some pf\<^isub>2 \<and> pf\<^isub>2 \<in> current_files \<tau> \<and> is_dir \<tau> pf\<^isub>2 \<and>
+ pf\<^isub>2 \<notin> files_hung_by_del \<tau>) \<and> is_file \<tau> f\<^isub>1)"
+| "os_grant \<tau> (Truncate p f len) =
+ (p \<in> current_procs \<tau> \<and> f \<in> current_files \<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> f \<in> current_files \<tau> \<and> is_file \<tau> f \<and> fds \<subseteq> current_proc_fds \<tau> p)" (* file_at_writing_by \<tau> f = {} *)
+(*
+| "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))"
+| "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> (p, fd) \<in> current_sockets \<tau> \<and>
+ socket_state \<tau> (p, fd) = Some SS_bind \<and> is_tcp_sock \<tau> (p, fd))" (* 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 shms) =
+ (p \<in> current_procs \<tau> \<and> p' \<notin> (current_procs \<tau>) \<and> fds \<subseteq> current_proc_fds \<tau> p \<and>
+ (\<forall> h \<in> shms. \<exists> flag. (p, flag) \<in> procs_of_shm \<tau> h))"
+| "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) = (f \<in> current_files s \<and> is_file s f)"
+| "alive s (O_dir f) = (f \<in> current_files s \<and> 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) = (k \<in> current_sockets s \<and> is_tcp_sock s k)"
+| "alive s (O_udp_sock k) = (k \<in> current_sockets s \<and> 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> 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 (UnLink p f # s) = ((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 shms # 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 shms # 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 shms # 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)"
+
+(******* 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)"
+
+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_dir pctxt fctxt
+ | _ \<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_file pctxt fctxt \<and> search_check_allp pctxt (set pfctxts))
+ | _ \<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_dir pctxt fctxt \<and> search_check_allp pctxt (set pfctxts))
+ | _ \<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 :: "t_state \<Rightarrow> security_context_t \<Rightarrow> t_process \<Rightarrow> t_fd set \<Rightarrow> bool"
+where
+ "inherit_fds_check s npsectxt p fds \<equiv>
+ (\<forall> ctxt \<in> sectxts_of_fds s p fds. permission_check npsectxt ctxt C_fd P_inherit)"
+
+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 :: "t_state \<Rightarrow> security_context_t \<Rightarrow> t_shm set \<Rightarrow> bool"
+where
+ "inherit_shms_check s npsectxt shms \<equiv>
+ (\<forall> ctxt \<in> sectxts_of_shms s shms. permission_check npsectxt ctxt C_shm P_inherit)"
+
+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 shms) =
+ (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 \<and>
+ inherit_shms_check s (up,rp,tp) shms)
+ | _ \<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) = (f \<in> init_files)"
+(* 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) = (member (init_msgs_of_queue q) ((op =) m))"
+| "tobj_in_init _ = False" (* other kind of obj cannot be tainted *)
+
+(* no use
+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 (_ # \<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"
+
+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_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)"
+
+definition taintable:: "t_object \<Rightarrow> bool"
+where
+ "taintable obj \<equiv> init_alive 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/Flask_type.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,239 @@
+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/Init_prop.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,735 @@
+(*<*)
+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
+
+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 = (f \<in> init_files \<and> is_file [] f)"
+by (auto simp add:is_init_file_def is_file_def init_inum_of_file_props split:option.splits)
+
+lemma is_init_file_prop2: "is_init_file f = (init_alive (O_file f))"
+by (auto simp add:is_init_file_def is_file_def init_inum_of_file_props split:option.splits)
+
+lemmas is_init_file_props = is_init_file_prop1 is_init_file_prop2
+
+lemma is_init_dir_prop1: "is_init_dir f = (f \<in> init_files \<and> is_dir [] f)"
+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 = (init_alive (O_dir f))"
+by (auto simp add:is_init_dir_def is_dir_def init_inum_of_file_props split:option.splits)
+
+lemmas is_init_dir_props = is_init_dir_prop1 is_init_dir_prop2
+
+lemma is_init_udp_sock_prop1: "is_init_udp_sock s = (s \<in> init_sockets \<and> is_udp_sock [] s)"
+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 = (init_alive (O_udp_sock s))"
+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
+
+lemmas is_init_udp_sock_props = is_init_udp_sock_prop1 is_init_udp_sock_prop2
+
+lemma is_init_tcp_sock_prop1: "is_init_tcp_sock s = (s \<in> init_sockets \<and> is_tcp_sock [] s)"
+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 = (init_alive (O_tcp_sock s))"
+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
+
+lemmas is_init_tcp_sock_props = is_init_tcp_sock_prop1 is_init_tcp_sock_prop2
+
+
+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)
+
+lemmas init_file_of_proc_fd_props = init_filefd_prop1 init_filefd_prop2 init_filefd_prop3 init_filefd_prop4 init_filefd_prop5
+
+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)
+
+
+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)
+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_shm: "h \<in> init_shms \<Longrightarrow> init_alive (O_shm h)" 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_shm 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 (simp only: is_init_file_prop2)
+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 (simp only: is_init_dir_prop2)
+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
+
+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_shm_has_ctxt:
+ "h \<in> init_shms \<Longrightarrow> \<exists> sec. init_sectxt_of_obj (O_shm h) = 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_shm_has_ctxt':
+ "init_sectxt_of_obj (O_shm h) = None \<Longrightarrow> h \<notin> init_shms"
+by (rule notI, drule init_shm_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_shm_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_shm_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 [] False = 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_shm_has_sshm:
+ "h \<in> init_shms \<Longrightarrow> \<exists> sh. init_ch2sshm h = Some sh"
+apply (drule init_shm_has_ctxt)
+by (auto simp add:init_ch2sshm_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_prop1:
+ "f \<in> init_files \<Longrightarrow> cf2sfile [] f (is_init_file f) = init_cf2sfile f"
+apply (case_tac f)
+apply (simp add:init_sectxt_prop cf2sfile_def init_cf2sfile_def)
+apply (rule notI, drule root_is_init_dir', simp)
+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 split:option.splits t_inode_tag.splits
+ dest:init_file_has_inum inof_has_file_tag)
+done
+
+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)
+
+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_prop2:
+ "f \<in> init_files \<Longrightarrow> cf2sfile [] f (\<not> is_init_file f) = None"
+apply (case_tac f)
+apply (simp add:init_sectxt_prop cf2sfile_def init_cf2sfile_def)
+apply (rule notI, drule root_is_init_dir', simp)
+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 split:option.splits t_inode_tag.splits
+ dest:init_file_has_inum inof_has_file_tag init_sec_file_dir)
+done
+
+lemma cf2sfile_nil_prop:
+ "f \<in> init_files \<Longrightarrow> cf2sfile [] f = (\<lambda> b. if (b = is_init_file f) then init_cf2sfile f else None)"
+apply (frule cf2sfile_nil_prop1, frule cf2sfile_nil_prop2)
+by (rule ext, auto split:if_splits)
+
+lemma cf2sfile_nil_prop3:
+ "is_init_file f \<Longrightarrow> cf2sfile [] f True = init_cf2sfile f"
+by (simp add:is_init_file_prop1 cf2sfile_nil_prop)
+
+lemma cf2sfile_nil_prop4:
+ "is_init_dir f \<Longrightarrow> cf2sfile [] f False = init_cf2sfile f"
+apply (frule init_file_dir_conflict2)
+by (simp add:is_init_file_prop1 is_init_dir_prop1 cf2sfile_nil_prop)
+
+lemma cfs2sfiles_nil_prop:
+ "\<forall> f \<in> fs. f \<in> init_files \<Longrightarrow> cfs2sfiles [] fs = init_cfs2sfiles fs"
+apply (simp add:cfs2sfiles_def init_cfs2sfiles_def)
+using cf2sfile_nil_prop apply auto
+
+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 (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)
+apply (rule set_eqI, rule iffI)
+apply (drule CollectD, rule CollectI, (erule exE)+)
+apply (rule_tac x = fd in exI, rule_tac x = sfd in exI, rule_tac x = f in exI) defer
+apply (drule CollectD, rule CollectI, (erule exE)+)
+apply (rule_tac x = fd in exI, rule_tac x = sfd in exI, rule_tac x = f in exI)
+using cfd2sfd_nil_prop by auto
+
+lemma ch2sshm_nil_prop:
+ "h \<in> init_shms \<Longrightarrow> ch2sshm [] h = init_ch2sshm h"
+by (simp add:ch2sshm_def init_sectxt_prop init_ch2sshm_def)
+
+lemma cph2spshs_nil_prop:
+ "cph2spshs [] p = init_cph2spshs p"
+apply (auto simp add:init_cph2spshs_def cph2spshs_def init_sectxt_prop)
+apply (rule_tac x = h in exI, simp) defer
+apply (rule_tac x = h in exI, simp)
+by (auto simp:ch2sshm_nil_prop dest:init_procs_has_shm)
+
+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 cph2spshs_nil_prop cpfd2sfds_nil_prop
+ split:option.splits)
+
+lemma tainted_nil_prop:
+ "(x \<in> tainted []) = (x \<in> seeds)"
+apply (rule iffI)
+apply (erule tainted.cases, simp+)
+apply (erule t_init)
+done
+
+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 intro:t_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 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)
+
+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 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 cfs2sfiles_nil_prop is_init_dir_prop1 is_init_file_prop1
+ is_init_udp_sock_prop1 is_init_tcp_sock_prop1 ch2sshm_nil_prop
+ same_inode_nil_prop cm2smsg_nil_prop dest:init_same_inode_prop1
+ split:option.splits)
+apply (rule_tac x = list in exI, simp add:init_same_inode_files_def)
+by (simp add:init_files_props)
+
+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/List_Prefix.thy Fri May 03 08:20:21 2013 +0100
@@ -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/My_list_prefix.thy Fri May 03 08:20:21 2013 +0100
@@ -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/New_obj_prop.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,92 @@
+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 nh_notin_curh: "valid \<tau> \<Longrightarrow> new_shm \<tau> \<notin> current_shms \<tau>" using finite_ch
+by (simp add:new_shm_def nn_notin)
+
+lemma nh_notin_curh': "\<lbrakk>new_shm \<tau> \<in> current_shms \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> False"
+by (drule nh_notin_curh, 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/OS_type_def.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,276 @@
+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" *)
+
+
+(* need not consideration, for rdonly transfer virus too
+datatype t_sharememo_flag =
+ SHM_RDONLY
+| SHM_OTHER
+
+type_synonym t_sharememo_flags = "t_sharememo_flag set"
+*)
+
+
+
+(****************************** 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/Proc_fd_of_file_prop.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,115 @@
+theory Proc_fd_of_file_prop
+imports Main Flask Flask_type Valid_prop Current_files_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')
+
+
+(******************* 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/ROOT Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,22 @@
+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
+ Valid_prop
+ Init_prop
+ Current_files_prop
+ Current_sockets_prop
+ Alive_prop
+ Proc_fd_of_file_prop
+ Finite_current
+ New_obj_prop
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Sectxt_prop.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,627 @@
+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)
+ 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)
+ 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''
+
+(*************** 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
+
+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_others (* 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_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/Static.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,871 @@
+theory Static
+imports Static_type OS_type_def Flask_type Flask
+begin
+
+locale tainting_s = tainting
+
+begin
+
+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 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 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))"
+
+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_cfs2sfiles :: "t_file set \<Rightarrow> t_sfile set"
+where
+ "init_cfs2sfiles fs \<equiv> {sf. \<exists>f \<in> fs. 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 sfd f. init_file_of_proc_fd p fd = Some f \<and> 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), (init_cph2spshs 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)"
+
+definition init_same_inode_files :: "t_file \<Rightarrow> t_file set"
+where
+ "init_same_inode_files f \<equiv> {f'. init_inum_of_file f = init_inum_of_file f'}"
+
+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) =
+ (if (f \<in> init_files \<and> is_init_file f)
+ then Some (S_file (init_cfs2sfiles (init_same_inode_files f))
+ (\<exists> f'. f' \<in> (init_same_inode_files f) \<and> O_file f \<in> seeds))
+ else None)"
+| "init_obj2sobj (O_dir f) =
+ (if (is_init_dir f) then
+ (case (init_cf2sfile f) of
+ Some sf \<Rightarrow> Some (S_dir sf)
+ | _ \<Rightarrow> None)
+ else 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}"
+
+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, shms) = True"
+| "is_init_sproc _ = False"
+
+fun is_many_sproc :: "t_sproc \<Rightarrow> bool"
+where
+ "is_many_sproc (Created, sec,fds,shms) = 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"
+
+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_sshm :: "t_sshm \<Rightarrow> bool"
+where
+ "is_many_sshm (Created, sec) = True"
+| "is_many_sshm _ = False"
+
+(*
+fun is_init_sobj :: "t_sobject \<Rightarrow> bool"
+where
+ "is_init_sobj (S_proc (popt, sec, fds, shms) tag) = (popt \<noteq> None)"
+| "is_init_sobj (S_file sf tag) = (is_init_sfile sf)"
+| "is_init_sobj (S_dir sf tag) = (is_init_sfile sf)"
+| "is_init_sobj (S_msg" *)
+
+fun is_many :: "t_sobject \<Rightarrow> bool"
+where
+ "is_many (S_proc (Many, sec, fds, shms) tag) = True"
+| "is_many (S_file sfs tag) = (\<forall> sf \<in> sfs. is_many_sfile sf)"
+| "is_many (S_dir sf ) = is_many_sfile sf"
+| "is_many (S_msgq sq ) = is_many_smsgq sq"
+| "is_many (S_shm sh ) = is_many_sshm sh"
+
+(*
+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'})"
+*)
+
+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}"
+
+(*
+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,shms) = 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_file pctxt (sectxt_of_sfile sf) \<and> search_check_allp pctxt (asecs_of_sfile sf)
+ else search_check_dir pctxt (sectxt_of_sfile sf) \<and> search_check_allp pctxt (asecs_of_sfile sf))"
+
+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>
+ (\<forall> ctxt \<in> sectxts_of_sfds sfds. permission_check pctxt ctxt C_fd P_inherit)"
+
+definition sectxts_of_sproc_sshms :: "t_sproc_sshm set \<Rightarrow> security_context_t set"
+where
+ "sectxts_of_sproc_sshms sshms \<equiv> {ctxt. \<exists> hi flag. ((hi, ctxt),flag) \<in> sshms}"
+
+definition inherit_shms_check_s :: "security_context_t \<Rightarrow> t_sproc_sshm set \<Rightarrow> bool"
+where
+ "inherit_shms_check_s pctxt sshms \<equiv>
+ (\<forall> ctxt \<in> sectxts_of_sproc_sshms sshms. permission_check pctxt ctxt C_shm P_inherit)"
+
+(*
+fun info_flow_sshm :: "t_sproc \<Rightarrow> t_sproc \<Rightarrow> bool"
+where
+ "info_flow_sshm (pi,sec,fds,shms) (pi',sec',fds',shms') =
+ (\<exists> sh flag'. (sh, SHM_RDWR) \<in> shms \<and> (sh, flag') \<in> shms')"
+*)
+definition info_flow_sproc_sshms :: "t_sproc_sshm set \<Rightarrow> t_sproc_sshm set \<Rightarrow> bool"
+where
+ "info_flow_sproc_sshms shms shms' \<equiv> (\<exists> sh flag'. (sh, SHM_RDWR) \<in> shms \<and> (sh, flag') \<in> shms')"
+
+fun info_flow_sshm :: "t_sproc \<Rightarrow> t_sproc \<Rightarrow> bool"
+where
+ "info_flow_sshm (pi,sec,fds,shms) (pi',sec',fds',shms') = info_flow_sproc_sshms shms shms'"
+
+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_shms ss (pi,pctxt,fds,shms) (tagp \<or> tagm)) \<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 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 (Init p, sec, fds, shms) tag) (O_proc p') = (p = p')"
+| "init_obj_related (S_file sfs tag) (O_file f) = (\<exists> sf \<in> sfs. sfile_related sf f)"
+| "init_obj_related (S_dir sf) (O_dir f) = (sfile_related sf f)"
+| "init_obj_related (S_msgq (Init q, sec, sms)) (O_msgq q') = (q = q')"
+| "init_obj_related (S_shm (Init h, sec)) (O_shm h') = (h = h')"
+| "init_obj_related (S_msg (Init q, sec, sms) (Init m, secm, tagm)) (O_msg q' m') = (q = q' \<and> m = m')"
+| "init_obj_related _ _ = False"
+
+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,sec,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"
+
+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)"
+
+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)"
+
+
+(**************** translation from dynamic to static *******************)
+
+definition cf2sfile :: "t_state \<Rightarrow> t_file \<Rightarrow> bool \<Rightarrow> t_sfile option"
+where
+ "cf2sfile s f Is_file \<equiv>
+ case (parent f) of
+ None \<Rightarrow> if Is_file then None else Some sroot
+ | Some pf \<Rightarrow> if Is_file
+ 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 cfs2sfiles :: "t_state \<Rightarrow> t_file set \<Rightarrow> t_sfile set"
+where
+ "cfs2sfiles s fs \<equiv> {sf. \<exists> f \<in> fs. cf2sfile s f True = Some sf}"
+
+definition same_inode_files :: "t_state \<Rightarrow> t_file \<Rightarrow> t_file set"
+where
+ "same_inode_files s f \<equiv> {f'. inum_of_file s f = inum_of_file s f'}"
+
+(* 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 True) 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 sfd f. file_of_proc_fd s p fd = Some f \<and> 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, cph2spshs 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) =
+ (if (f \<in> current_files s) then
+ Some (S_file (cfs2sfiles s (same_inode_files s f)) (O_file f \<in> tainted s))
+ else None)"
+| "co2sobj s (O_dir f) =
+ (case (cf2sfile s f False) 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"
+
+
+(***************** 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_shm :: "t_state \<Rightarrow> nat"
+where
+ "new_shm \<tau> = next_nat (current_shms \<tau>)"
+
+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"
+
+definition s2ss :: "t_state \<Rightarrow> t_static_state"
+where
+ "s2ss s \<equiv> {sobj. \<exists> obj. alive s obj \<and> co2sobj s obj = Some sobj}"
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Static_type.thy Fri May 03 08:20:21 2013 +0100
@@ -0,0 +1,95 @@
+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) \<times> (t_sproc_sshm 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 set" "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/Valid_prop.thy Fri May 03 08:20:21 2013 +0100
@@ -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