1
|
1 |
theory Proc_fd_of_file_prop
|
70
|
2 |
imports Main Flask Flask_type Valid_prop Current_files_prop Current_sockets_prop
|
1
|
3 |
begin
|
|
4 |
|
|
5 |
context flask begin
|
|
6 |
|
|
7 |
lemma proc_fd_in_procs: "\<lbrakk>file_of_proc_fd \<tau> p fd = Some f; valid \<tau>\<rbrakk> \<Longrightarrow> p \<in> current_procs \<tau>"
|
|
8 |
apply (induct \<tau> arbitrary: f) defer
|
|
9 |
apply (frule vd_cons, frule vt_grant_os, case_tac a)
|
|
10 |
apply (auto simp:file_of_proc_fd.simps current_procs.simps os_grant.simps split:if_splits option.splits)
|
|
11 |
by (drule init_filefd_valid, simp)
|
|
12 |
|
|
13 |
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"
|
|
14 |
apply (induct \<tau>)
|
|
15 |
apply (simp add:file_of_proc_fd.simps current_proc_fds.simps)
|
|
16 |
apply (clarify, drule init_filefd_valid, simp)
|
|
17 |
apply (clarify, frule vd_cons, frule vt_grant_os, case_tac a)
|
|
18 |
apply (auto simp:file_of_proc_fd.simps current_proc_fds.simps split:if_splits option.splits t_sock_addr.splits)
|
|
19 |
done
|
|
20 |
|
|
21 |
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"
|
|
22 |
by (rule proc_fd_in_fds_aux[rule_format], simp+)
|
|
23 |
|
|
24 |
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>"
|
|
25 |
by (auto simp:proc_fd_of_file_def intro:file_of_pfd_in_current)
|
|
26 |
|
|
27 |
lemma proc_fd_file_in_cur': "\<lbrakk>proc_fd_of_file \<tau> f \<noteq> {}; valid \<tau>\<rbrakk> \<Longrightarrow> f \<in> current_files \<tau>"
|
|
28 |
by (auto simp:proc_fd_file_in_cur)
|
|
29 |
|
|
30 |
lemma proc_fd_file_in_cur'': "\<lbrakk>proc_fd_of_file \<tau> f = {(p,fd)}; valid \<tau>\<rbrakk> \<Longrightarrow> f \<in> current_files \<tau>"
|
|
31 |
by (auto simp:proc_fd_file_in_cur')
|
|
32 |
|
|
33 |
lemma procfd_of_file_imp_fpfd: "proc_fd_of_file \<tau> f = {(p, fd)} \<Longrightarrow> file_of_proc_fd \<tau> p fd = Some f"
|
|
34 |
by (auto simp:proc_fd_of_file_def)
|
|
35 |
|
|
36 |
lemma procfd_of_file_imp_fpfd': "proc_fd_of_file \<tau> f = {(p, fd)} \<Longrightarrow> file_of_proc_fd \<tau> p fd \<noteq> None"
|
|
37 |
by (auto simp:proc_fd_of_file_def)
|
|
38 |
|
|
39 |
lemma procfd_of_file_eq_fpfd'': "(p, fd) \<in> proc_fd_of_file \<tau> f = (file_of_proc_fd \<tau> p fd = Some f)"
|
|
40 |
by (auto simp:proc_fd_of_file_def)
|
|
41 |
|
|
42 |
lemma procfd_of_file_non_empty: "file_of_proc_fd \<tau> p fd = Some f \<Longrightarrow> proc_fd_of_file \<tau> f \<noteq> {}"
|
|
43 |
by (auto simp:proc_fd_of_file_def)
|
|
44 |
|
|
45 |
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>"
|
|
46 |
by (drule procfd_of_file_non_empty, simp add:proc_fd_file_in_cur')
|
|
47 |
|
70
|
48 |
lemma file_fds_subset_pfds:
|
|
49 |
"valid s \<Longrightarrow> proc_file_fds s p \<subseteq> current_proc_fds s p"
|
|
50 |
by (auto simp add:proc_file_fds_def intro:proc_fd_in_fds)
|
|
51 |
|
|
52 |
lemma tcp_filefd_conflict:
|
|
53 |
"\<lbrakk>file_of_proc_fd s p fd = Some f; is_tcp_sock s (p, fd); valid s\<rbrakk> \<Longrightarrow> False"
|
|
54 |
apply (induct s, simp)
|
|
55 |
apply (auto simp:is_tcp_sock_simps)
|
|
56 |
|
|
57 |
|
|
58 |
lemma tcp_not_file_fd:
|
|
59 |
"\<lbrakk>is_tcp_sock s (p, fd); valid s\<rbrakk> \<Longrightarrow> fd \<notin> proc_file_fds s p"
|
|
60 |
apply (auto simp:proc_file_fds_def)
|
|
61 |
|
|
62 |
lemma udp_not_file_fd:
|
|
63 |
"\<lbrakk>is_udp_sock s (p, fd); valid s\<rbrakk> \<Longrightarrow> fd \<notin> proc_file_fds s p"
|
|
64 |
|
|
65 |
lemma
|
1
|
66 |
|
|
67 |
(******************* rebuild proc_fd_of_file simpset ***********************)
|
|
68 |
(*
|
|
69 |
lemma proc_fd_of_file_open: "Open p f flags fd iopt # valid \<tau> \<Longrightarrow>
|
|
70 |
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')"
|
|
71 |
apply (auto simp:proc_fd_of_file_def file_of_proc_fd.simps split:if_splits)
|
|
72 |
apply (frule vd_cons, drule vt_grant_os, case_tac iopt)
|
|
73 |
apply (drule proc_fd_in_fds, simp, simp add:os_grant.simps nfd_notin_curfd)+
|
|
74 |
done
|
|
75 |
|
|
76 |
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) "
|
|
77 |
by (auto simp:proc_fd_of_file_def file_of_proc_fd.simps split:if_splits)
|
|
78 |
|
|
79 |
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>
|
|
80 |
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)"
|
|
81 |
apply (frule vt_grant_os, frule vd_cons)
|
|
82 |
apply (case_tac "f\<^isub>3 \<preceq> f")
|
|
83 |
apply (subgoal_tac "f \<notin> current_files \<tau>") prefer 2 apply (rule notI)
|
|
84 |
apply (clarsimp simp:os_grant.simps, drule_tac f = f\<^isub>3 and f' = f in ancenf_in_current, simp, simp, simp)
|
|
85 |
apply (auto simp add:proc_fd_of_file_def)[1]
|
|
86 |
|
|
87 |
apply (simp add:file_of_proc_fd.simps split:option.splits if_splits)
|
|
88 |
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)
|
|
89 |
apply (drule file_of_pfd_in_current, simp+)
|
|
90 |
apply (simp add:file_of_proc_fd.simps)
|
|
91 |
apply (rule conjI, rule impI, simp add:file_renaming_prop5')
|
|
92 |
apply (rule impI, simp add:file_before_rename_def)
|
|
93 |
|
|
94 |
apply (simp add:proc_fd_of_file_def split:if_splits)
|
|
95 |
apply auto
|
|
96 |
apply (simp add:file_of_proc_fd.simps split:option.splits if_splits)
|
|
97 |
apply (drule_tac f\<^isub>3 = f\<^isub>3 and f\<^isub>2 = f\<^isub>2 and f = aa in file_renaming_prop1, simp)
|
|
98 |
apply (simp add:current_files_simps)
|
|
99 |
apply (erule exE| erule conjE)+
|
|
100 |
apply (simp add:file_of_proc_fd.simps split:option.splits if_splits)
|
|
101 |
apply (drule_tac f = f\<^isub>1 in rename_renaming_decom', simp+)
|
|
102 |
apply (simp add:file_after_rename_def)
|
|
103 |
done
|
|
104 |
|
|
105 |
|
|
106 |
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}"
|
|
107 |
by (auto simp:proc_fd_of_file_def file_of_proc_fd.simps)
|
|
108 |
|
|
109 |
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'}"
|
|
110 |
by (auto simp:proc_fd_of_file_def file_of_proc_fd.simps)
|
|
111 |
|
|
112 |
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}"
|
|
113 |
apply (auto simp:proc_fd_of_file_def file_of_proc_fd.simps)
|
|
114 |
apply (frule vd_cons, drule vt_grant_os)
|
|
115 |
apply (drule proc_fd_in_procs, (simp add:os_grant.simps np_notin_curp)+)
|
|
116 |
done
|
|
117 |
|
|
118 |
lemma proc_fd_of_file_other: "\<lbrakk>e # valid \<tau>;
|
|
119 |
\<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
|
|
120 |
\<forall> p fd. e \<noteq> CloseFd p fd;
|
|
121 |
\<forall> p f f'. e \<noteq> Rename p f f';
|
|
122 |
\<forall> p p'. e \<noteq> Kill p p';
|
|
123 |
\<forall> p. e \<noteq> Exit p;
|
|
124 |
\<forall> p p'. e \<noteq> Clone p p'\<rbrakk> \<Longrightarrow> proc_fd_of_file (e # \<tau>) f = proc_fd_of_file \<tau> f"
|
|
125 |
apply (case_tac e, auto simp:proc_fd_of_file_def file_of_proc_fd.simps)
|
|
126 |
done
|
|
127 |
|
|
128 |
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
|
|
129 |
*)
|
|
130 |
end
|
|
131 |
|
|
132 |
|
|
133 |
end |