--- a/no_shm_selinux/Enrich2.thy Thu Jan 09 14:39:00 2014 +0800
+++ b/no_shm_selinux/Enrich2.thy Thu Jan 09 19:09:09 2014 +0800
@@ -5,965 +5,11 @@
context tainting_s begin
-fun enrich_msgq :: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_msgq \<Rightarrow> t_state"
+fun enrich_file :: "t_state \<Rightarrow> t_file \<Rightarrow> t_state \<Rightarrow> (t_file set \<times> t_state)"
where
- "enrich_msgq [] tq dq = []"
-| "enrich_msgq (CreateMsgq p q # s) tq dq =
- (if (tq = q)
- then (CreateMsgq p dq # CreateMsgq p q # s)
- else CreateMsgq p q # (enrich_msgq s tq dq))"
-| "enrich_msgq (SendMsg p q m # s) tq dq =
- (if (tq = q)
- then (SendMsg p dq m # SendMsg p q m # (enrich_msgq s tq dq))
- else SendMsg p q m # (enrich_msgq s tq dq))"
-| "enrich_msgq (RecvMsg p q m # s) tq dq =
- (if (tq = q)
- then (RecvMsg p dq m # RecvMsg p q m # (enrich_msgq s tq dq))
- else RecvMsg p q m # (enrich_msgq s tq dq))"
-| "enrich_msgq (e # s) tq dq = e # (enrich_msgq s tq dq)"
-
-lemma enrich_msgq_duq_in:
- "\<lbrakk>q' \<notin> current_msgqs s; q \<in> current_msgqs s; no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> q' \<in> current_msgqs (enrich_msgq s q q')"
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons, case_tac a)
-apply auto
-done
-
-lemma enrich_msgq_duq_sms:
- "\<lbrakk>q' \<notin> current_msgqs s; q \<in> current_msgqs s; no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> msgs_of_queue (enrich_msgq s q q') q' = msgs_of_queue s q"
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons, case_tac a)
-apply auto
-done
-
-lemma enrich_msgq_cur_inof:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> inum_of_file (enrich_msgq s q q') f = inum_of_file s f"
-apply (induct s arbitrary:f, simp)
-apply (frule vt_grant_os, frule vd_cons, case_tac a)
-apply (auto split:option.splits)
-done
-
-lemma enrich_msgq_cur_inos:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> inum_of_socket (enrich_msgq s q q') = inum_of_socket s"
-apply (rule ext)
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons, case_tac a)
-apply (auto split:option.splits)
-done
-
-lemma enrich_msgq_cur_inos':
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> inum_of_socket (enrich_msgq s q q') sock = inum_of_socket s sock"
-apply (simp add:enrich_msgq_cur_inos)
-done
-
-lemma enrich_msgq_cur_inums:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> current_inode_nums (enrich_msgq s q q') = current_inode_nums s"
-apply (auto simp:current_inode_nums_def current_file_inums_def
- current_sock_inums_def enrich_msgq_cur_inof enrich_msgq_cur_inos)
-done
-
-lemma enrich_msgq_cur_itag:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> itag_of_inum (enrich_msgq s q q') = itag_of_inum s"
-apply (rule ext)
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons, case_tac a)
-apply (auto split:option.splits t_socket_type.splits)
-done
-
-lemma enrich_msgq_cur_tcps:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> is_tcp_sock (enrich_msgq s q q') = is_tcp_sock s"
-apply (rule ext)
-apply (auto simp:is_tcp_sock_def enrich_msgq_cur_itag enrich_msgq_cur_inos
- split:option.splits t_inode_tag.splits)
-done
-
-lemma enrich_msgq_cur_udps:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> is_udp_sock (enrich_msgq s q q') = is_udp_sock s"
-apply (rule ext)
-apply (auto simp:is_udp_sock_def enrich_msgq_cur_itag enrich_msgq_cur_inos
- split:option.splits t_inode_tag.splits)
-done
-
-lemma enrich_msgq_cur_msgqs:
- "\<lbrakk>q \<in> current_msgqs s; no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> current_msgqs (enrich_msgq s q q') = current_msgqs s \<union> {q'}"
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons)
-apply (case_tac a, auto)
-done
-
-lemma enrich_msgq_cur_msgs:
- "\<lbrakk>q' \<notin> current_msgqs s; q \<in> current_msgqs s; no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> msgs_of_queue (enrich_msgq s q q') = (msgs_of_queue s) (q' := msgs_of_queue s q)"
-apply (rule ext, simp, rule conjI, rule impI)
-apply (simp add:enrich_msgq_duq_sms)
-apply (rule impI)
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons)
-apply (case_tac a, auto)
-done
-
-lemma enrich_msgq_cur_procs:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> current_procs (enrich_msgq s q q') = current_procs s"
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons)
-apply (case_tac a, auto)
-done
-
-lemma enrich_msgq_cur_files:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> current_files (enrich_msgq s q q') = current_files s"
-apply (auto simp:current_files_def)
-apply (simp add:enrich_msgq_cur_inof)+
-done
-
-lemma enrich_msgq_cur_fds:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> current_proc_fds (enrich_msgq s q q') = current_proc_fds s"
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons, case_tac a)
-apply auto
-done
-
-lemma enrich_msgq_filefd:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> file_of_proc_fd (enrich_msgq s q q') = file_of_proc_fd s"
-apply (rule ext, rule ext)
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons, case_tac a)
-apply auto
-done
-
-lemma enrich_msgq_flagfd:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> flags_of_proc_fd (enrich_msgq s q q') = flags_of_proc_fd s"
-apply (rule ext, rule ext)
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons, case_tac a)
-apply auto
-done
-
-lemma enrich_msgq_proc_fds:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> proc_file_fds (enrich_msgq s q q') = proc_file_fds s"
-apply (auto simp:proc_file_fds_def enrich_msgq_filefd)
-done
-
-lemma enrich_msgq_hungs:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> files_hung_by_del (enrich_msgq s q q') = files_hung_by_del s"
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons, case_tac a)
-apply (auto simp:files_hung_by_del.simps)
-done
-
-lemma enrich_msgq_is_file:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> is_file (enrich_msgq s q q') = is_file s"
-apply (rule ext)
-apply (auto simp add:is_file_def enrich_msgq_cur_itag enrich_msgq_cur_inof
- split:option.splits t_inode_tag.splits)
-done
-
-lemma enrich_msgq_is_dir:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> is_dir (enrich_msgq s q q') = is_dir s"
-apply (rule ext)
-apply (auto simp add:is_dir_def enrich_msgq_cur_itag enrich_msgq_cur_inof
- split:option.splits t_inode_tag.splits)
-done
-
-lemma enrich_msgq_sameinode:
- "\<lbrakk>no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> (f \<in> same_inode_files (enrich_msgq s q q') f') = (f \<in> same_inode_files s f')"
-apply (induct s, simp)
-apply (simp add:same_inode_files_def)
-apply (auto simp:enrich_msgq_is_file enrich_msgq_cur_inof)
-done
-
-lemma enrich_msgq_tainted_aux1:
- "\<lbrakk>no_del_event s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; valid s\<rbrakk>
- \<Longrightarrow> (tainted s \<union> {O_msg q' m| m. O_msg q m \<in> tainted s}) \<subseteq> tainted (enrich_msgq s q q')"
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons)
-apply (case_tac a)
-apply (auto split:option.splits if_splits dest:tainted_in_current
- simp:enrich_msgq_filefd enrich_msgq_sameinode)
-done
-
-lemma enrich_msgq_tainted_aux2:
- "\<lbrakk>no_del_event s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; valid s; valid (enrich_msgq s q q')\<rbrakk>
- \<Longrightarrow> tainted (enrich_msgq s q q') \<subseteq> (tainted s \<union> {O_msg q' m| m. O_msg q m \<in> tainted s})"
-apply (induct s, simp)
-apply (frule vt_grant_os, frule vd_cons)
-apply (case_tac a)
-apply (auto split:option.splits if_splits simp:enrich_msgq_filefd enrich_msgq_sameinode
- dest:tainted_in_current vd_cons)
-apply (drule_tac vd_cons)+
-apply (simp)
-apply (drule set_mp)
-apply simp
-apply simp
-apply (drule_tac s = s in tainted_in_current)
-apply simp+
-done
-
-lemma enrich_msgq_alive:
- "\<lbrakk>alive s obj; valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk>
- \<Longrightarrow> alive (enrich_msgq s q q') obj"
-apply (case_tac obj)
-apply (simp_all add:enrich_msgq_is_file enrich_msgq_is_dir enrich_msgq_cur_msgqs
- enrich_msgq_cur_msgs enrich_msgq_cur_procs enrich_msgq_cur_fds
- enrich_msgq_cur_tcps enrich_msgq_cur_udps)
-apply (rule impI, simp)
-done
-
-lemma enrich_msgq_alive':
- "\<lbrakk>alive (enrich_msgq s q q') obj; valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk>
- \<Longrightarrow> alive s obj \<or> obj = O_msgq q' \<or> (\<exists> m. obj = O_msg q' m \<and> alive s (O_msg q m))"
-apply (case_tac obj)
-apply (simp_all add:enrich_msgq_is_file enrich_msgq_is_dir enrich_msgq_cur_msgqs
- enrich_msgq_cur_msgs enrich_msgq_cur_procs enrich_msgq_cur_fds
- enrich_msgq_cur_tcps enrich_msgq_cur_udps)
-apply (auto split:if_splits)
-done
-
-lemma enrich_msgq_not_alive:
- "\<lbrakk>enrich_not_alive s (E_msgq q') obj; q' \<notin> current_msgqs s; q \<in> current_msgqs s; no_del_event s; valid s\<rbrakk>
- \<Longrightarrow> enrich_not_alive (enrich_msgq s q q') (E_msgq q') obj"
-apply (case_tac obj)
-apply (auto simp:enrich_msgq_cur_fds enrich_msgq_cur_files
- enrich_msgq_cur_procs enrich_msgq_cur_inums enrich_msgq_cur_msgqs enrich_msgq_cur_msgs)
-done
-
-lemma enrich_msgq_nodel:
- "no_del_event (enrich_msgq s q q') = no_del_event s"
-apply (induct s, simp)
-by (case_tac a, auto)
-
-lemma enrich_msgq_died_proc:
- "died (O_proc p) (enrich_msgq s q q') = died (O_proc p) s"
-apply (induct s, simp)
-by (case_tac a, auto)
-
-lemma cf2sfile_execve:
- "\<lbrakk>ff \<in> current_files (Execve p f fds # s); valid (Execve p f fds # s)\<rbrakk>
- \<Longrightarrow> cf2sfile (Execve p f fds # s) ff= cf2sfile s ff"
-by (auto dest:cf2sfile_other' simp:current_files_simps)
-lemma cf2sfile_clone:
- "\<lbrakk>ff \<in> current_files (Clone p p' fds # s); valid (Clone p p' fds # s)\<rbrakk>
- \<Longrightarrow> cf2sfile (Clone p p' fds # s) ff= cf2sfile s ff"
-by (auto dest:cf2sfile_other' simp:current_files_simps)
-lemma cf2sfile_ptrace:
- "\<lbrakk>ff \<in> current_files (Ptrace p p' # s); valid (Ptrace p p' # s)\<rbrakk>
- \<Longrightarrow> cf2sfile (Ptrace p p' # s) ff= cf2sfile s ff"
-by (auto dest:cf2sfile_other' simp:current_files_simps)
-lemma cf2sfile_readfile:
- "\<lbrakk>ff \<in> current_files (ReadFile p fd # s); valid (ReadFile p fd # s)\<rbrakk>
- \<Longrightarrow> cf2sfile (ReadFile p fd # s) ff= cf2sfile s ff"
-by (auto dest:cf2sfile_other' simp:current_files_simps)
-lemma cf2sfile_writefile:
- "\<lbrakk>ff \<in> current_files (WriteFile p fd # s); valid (WriteFile p fd # s)\<rbrakk>
- \<Longrightarrow> cf2sfile (WriteFile p fd # s) ff= cf2sfile s ff"
-by (auto dest:cf2sfile_other' simp:current_files_simps)
-lemma cf2sfile_truncate:
- "\<lbrakk>ff \<in> current_files (Truncate p f len # s); valid (Truncate p f len # s)\<rbrakk>
- \<Longrightarrow> cf2sfile (Truncate p f len # s) ff= cf2sfile s ff"
-by (auto dest:cf2sfile_other' simp:current_files_simps)
-lemma cf2sfile_createmsgq:
- "\<lbrakk>ff \<in> current_files (CreateMsgq p q # s); valid (CreateMsgq p q # s)\<rbrakk>
- \<Longrightarrow> cf2sfile (CreateMsgq p q # s) ff= cf2sfile s ff"
-by (auto dest:cf2sfile_other' simp:current_files_simps)
-lemma cf2sfile_sendmsg:
- "\<lbrakk>ff \<in> current_files (SendMsg p q m # s); valid (SendMsg p q m # s)\<rbrakk>
- \<Longrightarrow> cf2sfile (SendMsg p q m # s) ff = cf2sfile s ff"
-by (auto dest:cf2sfile_other' simp:current_files_simps)
-lemma cf2sfile_recvmsg:
- "\<lbrakk>ff \<in> current_files (RecvMsg p q m # s); valid (RecvMsg p q m # s)\<rbrakk>
- \<Longrightarrow> cf2sfile (RecvMsg p q m # s) ff = cf2sfile s ff"
-by (auto dest:cf2sfile_other' simp:current_files_simps)
-lemmas cf2sfile_other'' = cf2sfile_recvmsg cf2sfile_sendmsg cf2sfile_createmsgq cf2sfile_truncate
- cf2sfile_writefile cf2sfile_readfile cf2sfile_ptrace cf2sfile_clone cf2sfile_execve
-
-lemma enrich_msgq_prop:
- "\<lbrakk>valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk>
- \<Longrightarrow> valid (enrich_msgq s q q') \<and>
- (\<forall> tp. tp \<in> current_procs s \<longrightarrow> cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp) \<and>
- (\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f) \<and>
- (\<forall> tq. tq \<in> current_msgqs s \<longrightarrow> cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq) \<and>
- (\<forall> tp fd. fd \<in> proc_file_fds s tp \<longrightarrow> cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd) \<and>
- (cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q) \<and>
- (tainted (enrich_msgq s q q') = (tainted s \<union> {O_msg q' m| m. O_msg q m \<in> tainted s}))"
-proof (induct s)
- case Nil
- thus ?case by (auto)
-next
- case (Cons e s)
- hence vd_cons': "valid (e # s)" and curq_cons: "q \<in> current_msgqs (e # s)"
- and curq'_cons: "q' \<notin> current_msgqs (e # s)" and nodel_cons: "no_del_event (e # s)"
- and os: "os_grant s e" and grant: "grant s e" and vd: "valid s"
- by (auto dest:vd_cons vt_grant_os vt_grant)
- from curq'_cons nodel_cons have curq': "q' \<notin> current_msgqs s" by (case_tac e, auto)
- from nodel_cons have nodel: "no_del_event s" by (case_tac e, auto)
- from nodel curq' vd Cons
- have pre: "q \<in> current_msgqs s \<Longrightarrow> valid (enrich_msgq s q q') \<and>
- (\<forall>tp. tp \<in> current_procs s \<longrightarrow> cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp) \<and>
- (\<forall>f. f \<in> current_files s \<longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f) \<and>
- (\<forall>tq. tq \<in> current_msgqs s \<longrightarrow> cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq) \<and>
- (\<forall>tp fd. fd \<in> proc_file_fds s tp \<longrightarrow> cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd) \<and>
- (cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q) \<and>
- (tainted (enrich_msgq s q q') = (tainted s \<union> {O_msg q' m| m. O_msg q m \<in> tainted s}))"
- by auto
-
- from pre have pre_vd: "q \<in> current_msgqs s \<Longrightarrow> valid (enrich_msgq s q q')" by simp
- from pre have pre_sp: "\<And> tp. \<lbrakk>tp \<in> current_procs s; q \<in> current_msgqs s\<rbrakk>
- \<Longrightarrow> cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp" by auto
- from pre have pre_sf: "\<And> f. \<lbrakk>f \<in> current_files s; q \<in> current_msgqs s\<rbrakk>
- \<Longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f" by auto
- from pre have pre_sq: "\<And> tq. \<lbrakk>tq \<in> current_msgqs s; q \<in> current_msgqs s\<rbrakk>
- \<Longrightarrow> cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq" by auto
- from pre have pre_sfd: "\<And> tp fd. \<lbrakk>fd \<in> proc_file_fds s tp; q \<in> current_msgqs s\<rbrakk>
- \<Longrightarrow> cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" by auto
- hence pre_sfd': "\<And> tp fd f. \<lbrakk>file_of_proc_fd s tp fd = Some f; q \<in> current_msgqs s\<rbrakk>
- \<Longrightarrow> cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" by (auto simp:proc_file_fds_def)
- from pre have pre_duq: "q \<in> current_msgqs s \<Longrightarrow> cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q"
- by auto
- have vd_enrich:"q \<in> current_msgqs s \<Longrightarrow> valid (e # enrich_msgq s q q')"
- apply (frule pre_vd)
- apply (erule_tac s = s and obj' = "E_msgq q'" in enrich_valid_intro_cons)
- apply (simp_all add:pre nodel nodel_cons curq_cons vd_cons' vd enrich_msgq_hungs)
- apply (clarsimp simp:nodel vd curq' enrich_msgq_alive)
- apply (rule allI, rule impI, erule enrich_msgq_not_alive)
- apply (simp_all add:curq' curq'_cons nodel vd enrich_msgq_cur_msgs enrich_msgq_filefd enrich_msgq_flagfd)
- done
-
- have q_neq_q': "q' \<noteq> q" using curq'_cons curq_cons by auto
-
- have vd_enrich_cons: "valid (enrich_msgq (e # s) q q')"
- proof-
- have "\<And> p q''. e = CreateMsgq p q'' \<Longrightarrow> valid (enrich_msgq (e # s) q q')"
- proof-
- fix p q'' assume ev: "e = CreateMsgq p q''"
- show "valid (enrich_msgq (e # s) q q')"
- proof (cases "q'' = q")
- case False with ev vd_enrich curq_cons
- show ?thesis by simp
- next
- case True
- have "os_grant (CreateMsgq p q # s) (CreateMsgq p q')"
- using os ev
- by (simp add:q_neq_q' curq')
- moreover have "grant (CreateMsgq p q # s) (CreateMsgq p q')"
- using grant ev
- by (auto simp add:sectxt_of_obj_def split:option.splits)
- ultimately
- show ?thesis using ev vd_cons' True
- by (auto intro: valid.intros(2))
- qed
- qed
- moreover have "\<And> p q'' m. \<lbrakk>e = SendMsg p q'' m; q \<in> current_msgqs s\<rbrakk>
- \<Longrightarrow> valid (enrich_msgq (e # s) q q')"
- proof-
- fix p q'' m assume ev: "e = SendMsg p q'' m" and q_in: "q \<in> current_msgqs s"
- show "valid (enrich_msgq (e # s) q q')"
- proof (cases "q'' = q")
- case False with ev vd_enrich q_in
- show ?thesis by simp
- next
- case True
- from grant os ev True obtain psec qsec
- where psec: "sectxt_of_obj s (O_proc p) = Some psec"
- and qsec: "sectxt_of_obj s (O_msgq q) = Some qsec"
- by (auto split:option.splits)
- from psec q_in os ev
- have psec':"sectxt_of_obj (enrich_msgq s q q') (O_proc p) = Some psec"
- by (auto dest!:pre_sp simp:cp2sproc_def split:option.splits)
- from qsec q_in pre_duq vd
- have qsec': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = Some qsec"
- by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms')
- show ?thesis using ev True vd_cons' q_in vd_enrich nodel vd
- curq' psec psec' qsec qsec' grant os q_neq_q'
- apply (simp, erule_tac valid.intros(2))
- apply (auto simp:q_neq_q' enrich_msgq_cur_msgqs enrich_msgq_cur_procs
- enrich_msgq_cur_msgs sectxt_of_obj_simps)
- done
- qed
- qed
- moreover have "\<And> p q'' m. \<lbrakk>e = RecvMsg p q'' m; q \<in> current_msgqs s\<rbrakk>
- \<Longrightarrow> valid (enrich_msgq (e # s) q q')"
- proof-
- fix p q'' m assume ev: "e = RecvMsg p q'' m" and q_in: "q \<in> current_msgqs s"
- show "valid (enrich_msgq (e # s) q q')"
- proof (cases "q'' = q")
- case False with ev vd_enrich q_in
- show ?thesis by simp
- next
- case True
- from grant os ev True obtain psec qsec msec
- where psec: "sectxt_of_obj s (O_proc p) = Some psec"
- and qsec: "sectxt_of_obj s (O_msgq q) = Some qsec"
- and msec: "sectxt_of_obj s (O_msg q (hd (msgs_of_queue s q))) = Some msec"
- by (auto split:option.splits)
- from psec q_in os ev
- have psec':"sectxt_of_obj (enrich_msgq s q q') (O_proc p) = Some psec"
- by (auto dest!:pre_sp simp:cp2sproc_def split:option.splits)
- from qsec q_in pre_duq vd
- have qsec': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = Some qsec"
- by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms')
- from qsec q_in vd
- have qsec'': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q) = Some qsec"
- apply (frule_tac pre_sq, simp_all)
- by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms')
- from msec q_in pre_duq vd qsec qsec' qsec'' curq' nodel
- have msec': "sectxt_of_obj (enrich_msgq s q q') (O_msg q' (hd (msgs_of_queue s q))) = Some msec"
- apply (auto simp:cq2smsgq_def enrich_msgq_cur_msgs
- split:option.splits dest!:current_has_sms')
- apply (case_tac "msgs_of_queue s q")
- using os ev True apply simp
- apply (simp add:cqm2sms.simps split:option.splits)
- apply (auto simp:cm2smsg_def split:option.splits)
- done
- show ?thesis using ev True vd_cons' q_in vd_enrich nodel vd
- curq' grant os q_neq_q' psec psec' msec msec' qsec qsec'
- apply (simp, erule_tac valid.intros(2))
- apply (auto simp:enrich_msgq_cur_msgqs enrich_msgq_cur_procs
- enrich_msgq_cur_msgs sectxt_of_obj_simps)
- done
- qed
- qed
- ultimately
- show ?thesis using vd_enrich curq_cons vd_cons'
- apply (case_tac e)
- apply (auto simp del:enrich_msgq.simps)
- apply (auto split:if_splits )
- done
- qed
-
- have curpsec: "\<And> tp. \<lbrakk>tp \<in> current_procs s; q \<in> current_msgqs s\<rbrakk> \<Longrightarrow>
- sectxt_of_obj (enrich_msgq s q q') (O_proc tp) = sectxt_of_obj s (O_proc tp)"
- using pre_vd vd
- apply (frule_tac pre_sp, simp)
- by (auto simp:cp2sproc_def split:option.splits if_splits dest!:current_has_sec')
- have curfsec: "\<And> f. \<lbrakk>is_file s f; q \<in> current_msgqs s\<rbrakk> \<Longrightarrow>
- sectxt_of_obj (enrich_msgq s q q') (O_file f) = sectxt_of_obj s (O_file f)"
- proof-
- fix f
- assume a1: "is_file s f" and a2: "q \<in> current_msgqs s"
- from a2 pre_sf pre_vd
- have pre': "\<And>f. f \<in> current_files s \<Longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f"
- and vd_enrich: "valid (enrich_msgq s q q')"
- by auto
- hence csf: "cf2sfile (enrich_msgq s q q') f = cf2sfile s f"
- using a1 by (auto simp:is_file_in_current)
- from a1 obtain sf where csf_some: "cf2sfile s f = Some sf"
- apply (case_tac "cf2sfile s f")
- apply (drule current_file_has_sfile')
- apply (simp add:vd, simp add:is_file_in_current, simp)
- done
- from a1 have a1': "is_file (enrich_msgq s q q') f"
- using vd nodel by (simp add:enrich_msgq_is_file)
- show "sectxt_of_obj (enrich_msgq s q q') (O_file f) = sectxt_of_obj s (O_file f)"
- using csf csf_some vd_enrich vd a1 a1'
- apply (auto simp:cf2sfile_def split:option.splits if_splits)
- apply (case_tac f, simp_all)
- apply (drule root_is_dir', simp+)
- done
- qed
- have curdsec: "\<And> tf. \<lbrakk>is_dir s tf; q \<in> current_msgqs s\<rbrakk>
- \<Longrightarrow> sectxt_of_obj (enrich_msgq s q q') (O_dir tf) = sectxt_of_obj s (O_dir tf)"
- proof-
- fix tf
- assume a1: "is_dir s tf" and a2: "q \<in> current_msgqs s"
- from a2 pre_sf pre_vd
- have pre': "\<And>f. f \<in> current_files s \<Longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f"
- and vd_enrich: "valid (enrich_msgq s q q')"
- by auto
- hence csf: "cf2sfile (enrich_msgq s q q') tf = cf2sfile s tf"
- using a1 by (auto simp:is_dir_in_current)
- from a1 obtain sf where csf_some: "cf2sfile s tf = Some sf"
- apply (case_tac "cf2sfile s tf")
- apply (drule current_file_has_sfile')
- apply (simp add:vd, simp add:is_dir_in_current, simp)
- done
- from a1 have a1': "is_dir (enrich_msgq s q q') tf"
- using enrich_msgq_is_dir vd nodel by simp
- from a1 have a3: "\<not> is_file s tf" using vd by (simp add:is_dir_not_file)
- from a1' vd have a3': "\<not> is_file (enrich_msgq s q q') tf" by (simp add:is_dir_not_file)
- show "sectxt_of_obj (enrich_msgq s q q') (O_dir tf) = sectxt_of_obj s (O_dir tf)"
- using csf csf_some a3 a3' vd_enrich vd
- apply (auto simp:cf2sfile_def split:option.splits)
- apply (case_tac tf)
- apply (simp add:root_sec_remains, simp)
- done
- qed
- have curpsecs: "\<And> tf ctxts'. \<lbrakk>is_dir s tf; q \<in> current_msgqs s; get_parentfs_ctxts s tf = Some ctxts'\<rbrakk>
- \<Longrightarrow> \<exists> ctxts. get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts \<and> set ctxts = set ctxts'"
- proof-
- fix tf ctxts'
- assume a1: "is_dir s tf" and a2: "q \<in> current_msgqs s"
- and a4: "get_parentfs_ctxts s tf = Some ctxts'"
- from a2 pre
- have pre': "\<And>f. f \<in> current_files s \<Longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f"
- and vd_enrich': "valid (enrich_msgq s q q')"
- by auto
- hence csf: "cf2sfile (enrich_msgq s q q') tf = cf2sfile s tf"
- using a1 by (auto simp:is_dir_in_current)
- from a1 obtain sf where csf_some: "cf2sfile s tf = Some sf"
- apply (case_tac "cf2sfile s tf")
- apply (drule current_file_has_sfile')
- apply (simp add:vd, simp add:is_dir_in_current, simp)
- done
- from a1 have a1': "is_dir (enrich_msgq s q q') tf"
- using enrich_msgq_is_dir vd nodel by simp
- from a1 have a5: "\<not> is_file s tf" using vd by (simp add:is_dir_not_file)
- from a1' vd have a5': "\<not> is_file (enrich_msgq s q q') tf" by (simp add:is_dir_not_file)
-
- from a1' pre_vd a2 obtain ctxts
- where a3: "get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts"
- apply (case_tac "get_parentfs_ctxts (enrich_msgq s q q') tf")
- apply (drule get_pfs_secs_prop', simp+)
- done
- moreover have "set ctxts = set ctxts'"
- proof (cases tf)
- case Nil
- with a3 a4 vd vd_enrich'
- show ?thesis
- by (simp add:get_parentfs_ctxts.simps root_sec_remains split:option.splits)
- next
- case (Cons a ff)
- with csf csf_some a5 a5' vd_enrich' vd a3 a4
- show ?thesis
- apply (auto simp:cf2sfile_def split:option.splits if_splits)
- done
- qed
- ultimately
- show "\<exists> ctxts. get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts \<and> set ctxts = set ctxts'"
- by auto
- qed
-
- have psec_cons: "\<And> tp. tp \<in> current_procs (e # s) \<Longrightarrow>
- sectxt_of_obj (enrich_msgq (e # s) q q') (O_proc tp) = sectxt_of_obj (e # s) (O_proc tp)"
- using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd
- apply (case_tac e)
- apply (auto intro:curpsec simp:sectxt_of_obj_simps)
- apply (frule curpsec, simp, frule curfsec, simp)
- apply (auto split:option.splits)[1]
- apply (frule vd_cons) defer apply (frule vd_cons)
- apply (auto intro:curpsec simp:sectxt_of_obj_simps)
- done
-
-
- have sf_cons: "\<And> f. f \<in> current_files (e # s) \<Longrightarrow> cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f"
- proof-
- fix f
- assume a1: "f \<in> current_files (e # s)"
- hence a1': "f \<in> current_files (enrich_msgq (e # s) q q')"
- using nodel_cons os vd vd_cons' vd_enrich_cons
- apply (case_tac e)
- apply (auto simp:current_files_simps enrich_msgq_cur_files dest:is_file_in_current split:option.splits)
- done
- have b1: "\<And> p f' flags fd opt. e = Open p f' flags fd opt \<Longrightarrow>
- cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f"
- proof-
- fix p f' flags fd opt
- assume ev: "e = Open p f' flags fd opt"
- show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f"
- proof (cases opt)
- case None
- with a1 a1' ev vd_cons' vd_enrich_cons curq_cons
- show ?thesis
- apply (simp add:cf2sfile_open_none)
- apply (simp add:pre_sf current_files_simps)
- done
- next
- case (Some inum)
- show ?thesis
- proof (cases "f = f'")
- case False
- with a1 a1' ev vd_cons' vd_enrich_cons curq_cons Some
- show ?thesis
- apply (simp add:cf2sfile_open)
- apply (simp add:pre_sf current_files_simps)
- done
- next
- case True
- with vd_cons' ev os Some
- obtain pf where pf: "parent f = Some pf" by auto
- then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts"
- using os vd ev Some True
- apply (case_tac "get_parentfs_ctxts s pf")
- apply (drule get_pfs_secs_prop', simp, simp)
- apply auto
- done
-
- have "sectxt_of_obj (Open p f' flags fd (Some inum) # enrich_msgq s q q') (O_file f') =
- sectxt_of_obj (Open p f' flags fd (Some inum) # s) (O_file f')"
- using Some vd_enrich_cons vd_cons' ev pf True os curq_cons
- by (simp add:sectxt_of_obj_simps curpsec curdsec)
- moreover
- have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)"
- using curq_cons ev pf Some True os
- by (simp add:curdsec)
- moreover
- have "\<exists> ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \<and> set ctxts' = set ctxts"
- using curq_cons ev pf Some True os vd psecs
- apply (case_tac "get_parentfs_ctxts s pf")
- apply (drule get_pfs_secs_prop', simp+)
- apply (rule curpsecs, simp+)
- done
- then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'"
- and psecs_eq: "set ctxts' = set ctxts" by auto
- ultimately show ?thesis
- using a1 a1' ev vd_cons' vd_enrich_cons Some True pf psecs
- by (simp add:cf2sfile_open split:option.splits)
- qed
- qed
- qed
- have b2: "\<And> p f' inum. e = Mkdir p f' inum \<Longrightarrow>
- cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f"
- proof-
- fix p f' inum
- assume ev: "e = Mkdir p f' inum"
- show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f"
- proof (cases "f' = f")
- case False
- with a1 a1' ev vd_cons' vd_enrich_cons curq_cons
- show ?thesis
- apply (simp add:cf2sfile_mkdir)
- apply (simp add:pre_sf current_files_simps)
- done
- next
- case True
- with vd_cons' ev os
- obtain pf where pf: "parent f = Some pf" by auto
- then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts"
- using os vd ev True
- apply (case_tac "get_parentfs_ctxts s pf")
- apply (drule get_pfs_secs_prop', simp, simp)
- apply auto
- done
-
- have "sectxt_of_obj (Mkdir p f' inum # enrich_msgq s q q') (O_dir f') =
- sectxt_of_obj (Mkdir p f' inum # s) (O_dir f')"
- using vd_enrich_cons vd_cons' ev pf True os curq_cons
- by (simp add:sectxt_of_obj_simps curpsec curdsec)
- moreover
- have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)"
- using curq_cons ev pf True os
- by (simp add:curdsec)
- moreover
- have "\<exists> ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \<and> set ctxts' = set ctxts"
- using curq_cons ev pf True os vd psecs
- apply (case_tac "get_parentfs_ctxts s pf")
- apply (drule get_pfs_secs_prop', simp+)
- apply (rule curpsecs, simp+)
- done
- then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'"
- and psecs_eq: "set ctxts' = set ctxts" by auto
- ultimately show ?thesis
- using a1 a1' ev vd_cons' vd_enrich_cons True pf psecs
- apply (simp add:cf2sfile_mkdir split:option.splits)
- done
- qed
- qed
- have b3: "\<And> p f' f''. e = LinkHard p f' f'' \<Longrightarrow>
- cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f"
- proof-
- fix p f' f''
- assume ev: "e = LinkHard p f' f''"
- show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f"
- proof (cases "f'' = f")
- case False
- with a1 a1' ev vd_cons' vd_enrich_cons curq_cons
- show ?thesis
- apply (simp add:cf2sfile_linkhard)
- apply (simp add:pre_sf current_files_simps)
- done
- next
- case True
- with vd_cons' ev os
- obtain pf where pf: "parent f = Some pf" by auto
- then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts"
- using os vd ev True
- apply (case_tac "get_parentfs_ctxts s pf")
- apply (drule get_pfs_secs_prop', simp, simp)
- apply auto
- done
-
- have "sectxt_of_obj (LinkHard p f' f'' # enrich_msgq s q q') (O_file f) =
- sectxt_of_obj (LinkHard p f' f'' # s) (O_file f)"
- using vd_enrich_cons vd_cons' ev pf True os curq_cons
- by (simp add:sectxt_of_obj_simps curpsec curdsec)
- moreover
- have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)"
- using curq_cons ev pf True os
- by (simp add:curdsec)
- moreover
- have "\<exists> ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \<and> set ctxts' = set ctxts"
- using curq_cons ev pf True os vd psecs
- apply (case_tac "get_parentfs_ctxts s pf")
- apply (drule get_pfs_secs_prop', simp+)
- apply (rule curpsecs, simp+)
- done
- then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'"
- and psecs_eq: "set ctxts' = set ctxts" by auto
- ultimately show ?thesis
- using a1 a1' ev vd_cons' vd_enrich_cons True pf psecs
- apply (simp add:cf2sfile_linkhard split:option.splits)
- done
- qed
- qed
- show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f"
- apply (case_tac e)
- prefer 6 apply (erule b1)
- prefer 11 apply (erule b2)
- prefer 11 apply (erule b3)
- apply (simp_all only:b1 b2 b3)
- using a1' a1 vd_enrich_cons vd_cons' curq_cons nodel_cons
- apply (simp_all add:cf2sfile_other'' cf2sfile_simps enrich_msgq.simps no_del_event.simps split:if_splits)
- apply (simp_all add:pre_sf cf2sfile_other' current_files_simps split:if_splits)
- apply (drule vd_cons, simp add:cf2sfile_other', drule pre_sf, simp+)+
- done
- qed
-
- have sfd_cons:"\<And> tp fd f. file_of_proc_fd (e # s) tp fd = Some f \<Longrightarrow>
- cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd"
- proof-
- fix tp fd f
- assume a1: "file_of_proc_fd (e # s) tp fd = Some f"
- hence a1': "file_of_proc_fd (enrich_msgq (e # s) q q') tp fd = Some f"
- using nodel_cons vd_enrich os vd_cons'
- apply (case_tac e, auto simp:enrich_msgq_filefd simp del:enrich_msgq.simps)
- done
- have b1: "\<And> p f' flags fd' opt. e = Open p f' flags fd' opt \<Longrightarrow>
- cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd"
- proof-
- fix p f' flags fd' opt
- assume ev: "e = Open p f' flags fd' opt"
- have c1': "file_of_proc_fd (Open p f' flags fd' opt # s) tp fd = Some f"
- using a1' ev a1 by (simp split:if_splits)
- show "cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" thm cfd2sfd_open
- proof (cases "tp = p \<and> fd = fd'")
- case False
- show ?thesis using ev vd_enrich_cons vd_cons' a1' a1 False curq_cons
- apply (simp add:cfd2sfd_open split:if_splits del:file_of_proc_fd.simps)
- apply (rule conjI, rule impI, simp)
- apply (rule conjI, rule impI, simp)
- apply (auto simp: False intro!:pre_sfd' split:if_splits)
- done
- next
- case True
- have "f' \<in> current_files (Open p f' flags fd' opt # s)" using ev vd_cons' os
- by (auto simp:current_files_simps is_file_in_current split:option.splits)
- hence "cf2sfile (Open p f' flags fd' opt # enrich_msgq s q q') f'
- = cf2sfile (Open p f' flags fd' opt # s) f'"
- using sf_cons ev by auto
- moreover have "sectxt_of_obj (enrich_msgq s q q') (O_proc p) = sectxt_of_obj s (O_proc p)"
- apply (rule curpsec)
- using os ev curq_cons
- by (auto split:option.splits)
- ultimately show ?thesis
- using ev True a1 a1' vd_enrich_cons vd_cons'
- apply (simp add:cfd2sfd_open sectxt_of_obj_simps del:file_of_proc_fd.simps)
- done
- qed
- qed
- show "cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd"
- apply (case_tac e)
- prefer 6 apply (erule b1)
- using a1' a1 vd_enrich_cons vd_cons' curq_cons
- apply (simp_all only:cfd2sfd_simps enrich_msgq.simps)
- apply (auto simp:cfd2sfd_simps pre_sfd' dest:vd_cons cfd2sfd_other split:if_splits)
- done
- qed
-
- have pfds_cons: "\<And> tp. tp \<in> current_procs (e # s) \<Longrightarrow>
- cpfd2sfds (enrich_msgq (e # s) q q') tp = cpfd2sfds (e # s) tp"
- apply (auto simp add:cpfd2sfds_def proc_file_fds_def)
- apply (rule_tac x = fd in exI, rule conjI, rule_tac x = f in exI)
- prefer 3
- apply (rule_tac x = fd in exI, rule conjI, rule_tac x = f in exI)
- apply (auto simp:sfd_cons enrich_msgq_filefd nodel_cons vd_cons')
- done
-
- have tainted_cons: "tainted (enrich_msgq (e # s) q q') =
- (tainted (e # s) \<union> {O_msg q' m | m. O_msg q m \<in> tainted (e # s)})"
- apply (rule equalityI)
- using nodel_cons curq_cons curq'_cons vd_cons' vd_enrich_cons
- apply (rule enrich_msgq_tainted_aux2)
- using nodel_cons curq_cons curq'_cons vd_cons'
- apply (rule enrich_msgq_tainted_aux1)
- done
- have pre_tainted: "q \<in> current_msgqs s \<Longrightarrow> tainted (enrich_msgq s q q') =
- (tainted s \<union> {O_msg q' m| m. O_msg q m \<in> tainted s})" by (simp add:pre)
-
- have "\<forall>tp fd. fd \<in> proc_file_fds (e # s) tp \<longrightarrow> cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd"
- by (auto simp:proc_file_fds_def elim!:sfd_cons)
- moreover
- have "\<forall>tp. tp \<in> current_procs (e # s) \<longrightarrow> cp2sproc (enrich_msgq (e # s) q q') tp = cp2sproc (e # s) tp"
- by (auto simp:cp2sproc_def pfds_cons psec_cons enrich_msgq_died_proc split:option.splits)
- moreover
- have "\<forall>tq. tq \<in> current_msgqs (e # s) \<longrightarrow> cq2smsgq (enrich_msgq (e # s) q q') tq = cq2smsgq (e # s) tq"
- proof clarify
- fix tq assume a1: "tq \<in> current_msgqs (e # s)"
-
- have curqsec: "\<And> tq. \<lbrakk>tq \<in> current_msgqs s; q \<in> current_msgqs s\<rbrakk> \<Longrightarrow>
- sectxt_of_obj (enrich_msgq s q q') (O_msgq tq) = sectxt_of_obj s (O_msgq tq)"
- using pre_vd vd
- apply (frule_tac pre_sq, simp)
- by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms')
- have cursms: "\<And> q''. \<lbrakk>q'' \<in> current_msgqs s; q \<in> current_msgqs s\<rbrakk> \<Longrightarrow>
- cqm2sms (enrich_msgq s q q') q'' (msgs_of_queue (enrich_msgq s q q') q'') =
- cqm2sms s q'' (msgs_of_queue s q'')"
- using pre_vd vd
- apply (frule_tac pre_sq, simp)
- by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms')
- have qsec_cons: "sectxt_of_obj (enrich_msgq (e # s) q q') (O_msgq tq) = sectxt_of_obj (e # s) (O_msgq tq)"
- using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons curq_cons a1
- apply (case_tac e)
- apply (auto intro:curqsec simp:sectxt_of_obj_simps curpsec split:option.splits dest!:current_has_sec')
- apply (frule vd_cons) defer apply (frule vd_cons)
- apply (auto intro:curqsec simp:sectxt_of_obj_simps)
- done
- have sms_cons: "cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) =
- cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)"
- proof-
- have b1: "\<And> p q'' m. e = SendMsg p q'' m \<Longrightarrow>
- cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) =
- cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)"
- apply (case_tac e)
- using a1 curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons
- apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms pre_tainted
- split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps)
- apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*})
- apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms pre_tainted
- split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps)
- done
- have b2: "\<And> p q''. e = CreateMsgq p q'' \<Longrightarrow>
- cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) =
- cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)"
- using a1 curq_cons curq'_cons vd_enrich_cons vd_cons'
- apply (auto simp:cqm2sms_simps intro:cursms)
- apply (auto simp:cqm2sms.simps)
- done
- have b3: "\<And> p q'' m. e = RecvMsg p q'' m \<Longrightarrow>
- cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) =
- cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)"
- using a1 curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons
- apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms
- split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps)
- apply (frule vd_cons) defer apply (frule vd_cons)
- apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms
- split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps)
- done
- show "cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) =
- cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)"
- apply (case_tac e)
- prefer 15 apply (erule b2)
- prefer 15 apply (erule b1)
- prefer 15 apply (erule b3)
- using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons a1
- apply (auto intro:cursms simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec
- split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps)
- done
- qed
-
- show "cq2smsgq (enrich_msgq (e # s) q q') tq = cq2smsgq (e # s) tq"
- using a1 curq_cons
- apply (simp add:cq2smsgq_def qsec_cons sms_cons)
- done
- qed
- moreover
- have "cq2smsgq (enrich_msgq (e # s) q q') q' = cq2smsgq (e # s) q"
- proof-
- have duqsec: "q \<in> current_msgqs s \<Longrightarrow>
- sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = sectxt_of_obj s (O_msgq q)"
- apply (frule pre_duq) using vd
- by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms')
- have duqsms: "q \<in> current_msgqs s \<Longrightarrow>
- cqm2sms (enrich_msgq s q q') q' (msgs_of_queue (enrich_msgq s q q') q') =
- cqm2sms s q (msgs_of_queue s q)"
- apply (frule pre_duq) using vd
- by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms')
- have qsec_cons: "sectxt_of_obj (enrich_msgq (e # s) q q') (O_msgq q') = sectxt_of_obj (e # s) (O_msgq q)"
- using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons curq_cons
- apply (case_tac e)
- apply (auto simp:duqsec sectxt_of_obj_simps curpsec split:option.splits dest!:current_has_sec')
- apply (frule vd_cons) defer apply (frule vd_cons)
- apply (auto intro:duqsec simp:sectxt_of_obj_simps)
- done
- have sms_cons: "cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') =
- cqm2sms (e # s) q (msgs_of_queue (e # s) q)"
- proof-
- have b1: "\<And> p q'' m. e = SendMsg p q'' m \<Longrightarrow>
- cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') =
- cqm2sms (e # s) q (msgs_of_queue (e # s) q)"
- apply (case_tac e)
- using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons
- apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted
- enrich_msgq_cur_procs enrich_msgq_cur_msgqs
- split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps)
- apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*})
- apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted
- enrich_msgq_cur_procs enrich_msgq_cur_msgqs dest:tainted_in_current
- split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps)
- done
- have b2: "\<And> p q''. e = CreateMsgq p q'' \<Longrightarrow>
- cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') =
- cqm2sms (e # s) q (msgs_of_queue (e # s) q)"
- using curq_cons curq'_cons vd_enrich_cons vd_cons'
- apply (auto simp:cqm2sms_simps intro:duqsms)
- apply (auto simp:cqm2sms.simps)
- done
- have b3: "\<And> p q'' m. e = RecvMsg p q'' m \<Longrightarrow>
- cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') =
- cqm2sms (e # s) q (msgs_of_queue (e # s) q)"
- using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons
- apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted
- enrich_msgq_cur_procs enrich_msgq_cur_msgqs
- split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps)
- apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*})
- apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted
- enrich_msgq_cur_procs enrich_msgq_cur_msgqs dest:tainted_in_current
- split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps)
- done
- show ?thesis
- apply (case_tac e)
- prefer 15 apply (erule b2)
- prefer 15 apply (erule b1)
- prefer 15 apply (erule b3)
- using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons
- apply (auto intro:duqsms simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsms
- split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps)
- done
- qed
- show ?thesis
- using curq_cons
- apply (simp add:cq2smsgq_def qsec_cons sms_cons)
- done
- qed
- ultimately show ?case using vd_enrich_cons sf_cons tainted_cons
- by auto
-qed
-
-fun enrich_file_link :: "t_state \<Rightarrow> t_file \<Rightarrow> t_state \<Rightarrow> (t_file \<times> t_state)"
-where
- "enrich_file_link [] tf df = []"
-
-fun enrich_file_set :: "t_state \<Rightarrow> t_file \<Rightarrow> t_file \<Rightarrow> t_state"
-where
- "enrich_file_set [] tf df = []"
-| "enrich_file_set [] (Open p f flags fd (Some inum) # s) =
+ "enrich_file [] tf s = ({}, [])"
+| "enrich_file (Open p f flags fd (Some inum) # s) tf s =
+ "