|
1 theory Enrich |
|
2 imports Main Flask Static Init_prop Valid_prop Tainted_prop Delete_prop Co2sobj_prop S2ss_prop S2ss_prop2 |
|
3 Temp |
|
4 begin |
|
5 |
|
6 context tainting_s begin |
|
7 |
|
8 (* enrich s target_proc duplicated_pro *) |
|
9 fun enrich_proc :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> t_state" |
|
10 where |
|
11 "enrich_proc [] tp dp = []" |
|
12 | "enrich_proc (Execve p f fds # s) tp dp = ( |
|
13 if (tp = p) |
|
14 then Execve dp f (fds \<inter> proc_file_fds s p) # Execve p f fds # (enrich_proc s tp dp) |
|
15 else Execve p f fds # (enrich_proc s tp dp))" |
|
16 | "enrich_proc (Clone p p' fds # s) tp dp = ( |
|
17 if (tp = p') |
|
18 then Clone p dp (fds \<inter> proc_file_fds s p) # Clone p p' fds # s |
|
19 else Clone p p' fds # (enrich_proc s tp dp))" |
|
20 | "enrich_proc (Open p f flags fd opt # s) tp dp = ( |
|
21 if (tp = p) |
|
22 then Open dp f (remove_create_flag flags) fd opt # Open p f flags fd opt # (enrich_proc s tp dp) |
|
23 else Open p f flags fd opt # (enrich_proc s tp dp))" |
|
24 | "enrich_proc (CloseFd p fd # s) tp dp = ( |
|
25 if (tp = p) |
|
26 then CloseFd dp fd # CloseFd p fd # (enrich_proc s tp dp) |
|
27 else CloseFd p fd # (enrich_proc s tp dp))" |
|
28 (* |
|
29 | "enrich_proc (Attach p h flag # s) tp dp = ( |
|
30 if (tp = p) |
|
31 then Attach dp h flag # Attach p h flag # (enrich_proc s tp dp) |
|
32 else Attach p h flag # (enrich_proc s tp dp))" |
|
33 | "enrich_proc (Detach p h # s) tp dp = ( |
|
34 if (tp = p) |
|
35 then Detach dp h # Detach p h # (enrich_proc s tp dp) |
|
36 else Detach p h # (enrich_proc s tp dp))" |
|
37 *) |
|
38 | "enrich_proc (Kill p p' # s) tp dp = ( |
|
39 if (tp = p) then Kill p p' # s |
|
40 else Kill p p' # (enrich_proc s tp dp))" |
|
41 | "enrich_proc (Exit p # s) tp dp = ( |
|
42 if (tp = p) then Exit p # s |
|
43 else Exit p # (enrich_proc s tp dp))" |
|
44 | "enrich_proc (e # s) tp dp = e # (enrich_proc s tp dp)" |
|
45 |
|
46 definition is_created_proc:: "t_state \<Rightarrow> t_process \<Rightarrow> bool" |
|
47 where |
|
48 "is_created_proc s p \<equiv> p \<in> init_procs \<longrightarrow> died (O_proc p) s" |
|
49 |
|
50 lemma enrich_search_check: |
|
51 assumes grant: "search_check s (up, rp, tp) f" |
|
52 and cf2sf: "\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile s' f = cf2sfile s f" |
|
53 and vd: "valid s" and f_in: "is_file s f" and f_in': "is_file s' f" |
|
54 and sec: "sectxt_of_obj s' (O_file f) = sectxt_of_obj s (O_file f)" |
|
55 shows "search_check s' (up, rp, tp) f" |
|
56 proof (cases f) |
|
57 case Nil |
|
58 with f_in vd have "False" |
|
59 by (auto dest:root_is_dir') |
|
60 thus ?thesis by simp |
|
61 next |
|
62 case (Cons n pf) |
|
63 from vd f_in obtain sf where sf: "cf2sfile s f = Some sf" |
|
64 apply (drule_tac is_file_in_current, drule_tac current_file_has_sfile, simp) |
|
65 apply (erule exE, simp) |
|
66 done |
|
67 then obtain psfs where psfs: "get_parentfs_ctxts s pf = Some psfs" using Cons |
|
68 by (auto simp:cf2sfile_def split:option.splits if_splits) |
|
69 from sf cf2sf f_in have sf': "cf2sfile s' f = Some sf" by (auto dest:is_file_in_current) |
|
70 then obtain psfs' where psfs': "get_parentfs_ctxts s' pf = Some psfs'"using Cons |
|
71 by (auto simp:cf2sfile_def split:option.splits if_splits) |
|
72 with sf sf' psfs have psfs_eq: "set psfs' = set psfs" using Cons f_in f_in' |
|
73 apply (simp add:cf2sfile_def split:option.splits) |
|
74 apply (case_tac sf, simp) |
|
75 done |
|
76 show ?thesis using grant f_in f_in' psfs psfs' psfs_eq sec |
|
77 apply (simp add:Cons split:option.splits) |
|
78 by (case_tac a, simp) |
|
79 qed |
|
80 |
|
81 lemma proc_filefd_has_sfd: "\<lbrakk>fd \<in> proc_file_fds s p; valid s\<rbrakk> \<Longrightarrow> \<exists> sfd. cfd2sfd s p fd = Some sfd" |
|
82 apply (simp add:proc_file_fds_def) |
|
83 apply (auto dest: current_filefd_has_sfd) |
|
84 done |
|
85 |
|
86 lemma enrich_inherit_fds_check: |
|
87 assumes grant: "inherit_fds_check s (up, nr, nt) p fds" and vd: "valid s" |
|
88 and cfd2sfd: "\<forall> p fd. fd \<in> proc_file_fds s p\<longrightarrow> cfd2sfd s' p fd = cfd2sfd s p fd" |
|
89 and fd_in: "fds \<subseteq> proc_file_fds s p" and fd_in': "fds \<subseteq> proc_file_fds s' p" |
|
90 shows "inherit_fds_check s' (up, nr, nt) p fds" |
|
91 proof- |
|
92 have "\<And> fd. fd \<in> fds \<Longrightarrow> sectxt_of_obj s' (O_fd p fd) = sectxt_of_obj s (O_fd p fd)" |
|
93 proof- |
|
94 fix fd |
|
95 assume fd_in_fds: "fd \<in> fds" |
|
96 hence fd_in_cfds: "fd \<in> proc_file_fds s p" |
|
97 and fd_in_cfds': "fd \<in> proc_file_fds s' p" |
|
98 using fd_in fd_in' by auto |
|
99 with cfd2sfd |
|
100 have cfd_eq: "cfd2sfd s' p fd = cfd2sfd s p fd" by auto |
|
101 from fd_in_cfds obtain f where ffd: "file_of_proc_fd s p fd = Some f" |
|
102 by (auto simp:proc_file_fds_def) |
|
103 moreover have "flags_of_proc_fd s p fd \<noteq> None" |
|
104 using ffd vd by (auto dest:current_filefd_has_flags) |
|
105 moreover have "sectxt_of_obj s (O_fd p fd) \<noteq> None" |
|
106 using fd_in_cfds vd |
|
107 apply (rule_tac notI) |
|
108 by (auto dest!:current_has_sec' file_fds_subset_pfds[where p = p] intro:vd) |
|
109 moreover have "cf2sfile s f \<noteq> None" |
|
110 apply (rule notI) |
|
111 apply (drule current_file_has_sfile') |
|
112 using ffd |
|
113 by (auto simp:vd is_file_in_current dest:file_of_pfd_is_file) |
|
114 ultimately show "sectxt_of_obj s' (O_fd p fd) = sectxt_of_obj s (O_fd p fd)" |
|
115 using cfd_eq |
|
116 by (auto simp:cfd2sfd_def split:option.splits) |
|
117 qed |
|
118 hence "sectxts_of_fds s' p fds = sectxts_of_fds s p fds" |
|
119 by (simp add:sectxts_of_fds_def) |
|
120 thus ?thesis using grant |
|
121 by (simp add:inherit_fds_check_def) |
|
122 qed |
|
123 |
|
124 lemma not_all_procs_cons: |
|
125 "p \<notin> all_procs (e # s) \<Longrightarrow> p \<notin> all_procs s" |
|
126 by (case_tac e, auto) |
|
127 |
|
128 lemma not_all_procs_prop: |
|
129 "\<lbrakk>p' \<notin> all_procs s; p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> p' \<noteq> p" |
|
130 apply (induct s, rule notI, simp) |
|
131 apply (frule vt_grant_os, frule vd_cons, frule not_all_procs_cons, simp, rule notI) |
|
132 apply (case_tac a, auto) |
|
133 done |
|
134 |
|
135 fun enrich_not_alive :: "t_state \<Rightarrow> t_object \<Rightarrow> bool" |
|
136 where |
|
137 "enrich_not_alive s (O_file f) = (f \<notin> current_files s)" |
|
138 | "enrich_not_alive s (O_dir f) = (f \<notin> current_files s)" |
|
139 | "enrich_not_alive s (O_proc p) = (p \<notin> current_procs s)" |
|
140 | "enrich_not_alive s (O_fd p fd) = (fd \<notin> current_proc_fds s p)" |
|
141 | "enrich_not_alive s (O_msgq q) = (q \<notin> current_msgqs s)" |
|
142 | "enrich_not_alive s (O_msg q m) = (m \<notin> set (msgs_of_queue s q) \<or> q \<notin> current_msgqs s)" |
|
143 | "enrich_not_alive s _ = True" |
|
144 |
|
145 lemma enrich_valid_intro_cons: |
|
146 assumes vs': "valid s'" |
|
147 and os: "os_grant s e" and grant: "grant s e" and vd: "valid s" |
|
148 and alive: "\<forall> obj. alive s obj \<longrightarrow> alive s' obj" |
|
149 and alive': "\<forall> obj. enrich_not_alive s obj \<longrightarrow> enrich_not_alive s' obj" |
|
150 and cp2sp: "\<forall> p. p \<in> current_procs s \<longrightarrow> cp2sproc s' p = cp2sproc s p" |
|
151 and cf2sf: "\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile s' f = cf2sfile s f" |
|
152 and ffd_remain: "\<forall> p fd f. file_of_proc_fd s p fd = Some f \<longrightarrow> file_of_proc_fd s' p fd = Some f" |
|
153 and cfd2sfd: "\<forall> p fd. fd \<in> proc_file_fds s p \<longrightarrow> cfd2sfd s' p fd = cfd2sfd s p fd" |
|
154 shows "valid (e # s')" |
|
155 proof (cases e) |
|
156 case (Execve p f fds) |
|
157 have p_in: "p \<in> current_procs s'" using os alive |
|
158 apply (erule_tac x = "O_proc p" in allE) |
|
159 by (auto simp:Execve) |
|
160 have f_in: "is_file s' f" using os alive |
|
161 apply (erule_tac x = "O_file f" in allE) |
|
162 by (auto simp:Execve) |
|
163 have fd_in: "fds \<subseteq> proc_file_fds s' p" using os alive ffd_remain |
|
164 by (auto simp:Execve proc_file_fds_def) |
|
165 have "os_grant s' e" using p_in f_in fd_in by (simp add:Execve) |
|
166 moreover have "grant s' e" |
|
167 proof- |
|
168 from grant obtain up rp tp uf rf tf |
|
169 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
170 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
171 by (simp add:Execve split:option.splits, blast) |
|
172 with grant obtain pu nr nt where p3: "npctxt_execve (up, rp, tp) (uf, rf, tf) = Some (pu, nr, nt)" |
|
173 by (simp add:Execve split:option.splits del:npctxt_execve.simps, blast) |
|
174 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
175 using os cp2sp |
|
176 apply (erule_tac x = p in allE) |
|
177 by (auto simp:Execve co2sobj.simps cp2sproc_def split:option.splits) |
|
178 from os have f_in': "is_file s f" by (simp add:Execve) |
|
179 from vd os have "\<exists> sf. cf2sfile s f = Some sf" |
|
180 by (auto dest!:is_file_in_current current_file_has_sfile simp:Execve) |
|
181 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf |
|
182 apply (erule_tac x = f in allE) |
|
183 apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) |
|
184 apply (case_tac f, simp) |
|
185 apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
|
186 done |
|
187 have "inherit_fds_check s' (pu, nr, nt) p fds" |
|
188 proof- |
|
189 have "fds \<subseteq> proc_file_fds s' p" using os ffd_remain Execve |
|
190 by (auto simp:proc_file_fds_def) |
|
191 thus ?thesis using Execve grant vd cfd2sfd p1 p2 p3 os |
|
192 apply (rule_tac s = s in enrich_inherit_fds_check) |
|
193 by (simp_all split:option.splits) |
|
194 qed |
|
195 moreover have "search_check s' (pu, rp, tp) f" |
|
196 using p1 p2 p2' vd cf2sf f_in' grant Execve p3 f_in |
|
197 apply (rule_tac s = s in enrich_search_check) |
|
198 by (simp_all split:option.splits) |
|
199 ultimately show ?thesis using p1' p2' p3 |
|
200 apply (simp add:Execve split:option.splits) |
|
201 using grant Execve p1 p2 |
|
202 by (simp add:Execve grant p1 p2) |
|
203 qed |
|
204 ultimately show ?thesis using vs' |
|
205 by (erule_tac valid.intros(2), simp+) |
|
206 next |
|
207 case (Clone p p' fds) |
|
208 have p_in: "p \<in> current_procs s'" using os alive |
|
209 apply (erule_tac x = "O_proc p" in allE) |
|
210 by (auto simp:Clone) |
|
211 have p'_not_in: "p' \<notin> current_procs s'" using os alive' |
|
212 apply (erule_tac x = "O_proc p'" in allE) |
|
213 by (auto simp:Clone) |
|
214 have fd_in: "fds \<subseteq> proc_file_fds s' p" using os alive ffd_remain |
|
215 by (auto simp:Clone proc_file_fds_def) |
|
216 have "os_grant s' e" using p_in p'_not_in fd_in by (simp add:Clone) |
|
217 moreover have "grant s' e" |
|
218 proof- |
|
219 from grant obtain up rp tp |
|
220 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
221 apply (simp add:Clone split:option.splits) |
|
222 by (case_tac a, auto) |
|
223 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
224 using os cp2sp |
|
225 apply (erule_tac x = p in allE) |
|
226 by (auto simp:Clone co2sobj.simps cp2sproc_def split:option.splits) |
|
227 have p2: "inherit_fds_check s' (up, rp, tp) p fds" |
|
228 proof- |
|
229 have "fds \<subseteq> proc_file_fds s' p" using os ffd_remain Clone |
|
230 by (auto simp:proc_file_fds_def) |
|
231 thus ?thesis using Clone grant vd cfd2sfd p1 os |
|
232 apply (rule_tac s = s in enrich_inherit_fds_check) |
|
233 by (simp_all split:option.splits) |
|
234 qed |
|
235 show ?thesis using p1 p2 p1' grant |
|
236 by (simp add:Clone) |
|
237 qed |
|
238 ultimately show ?thesis using vs' |
|
239 by (erule_tac valid.intros(2), simp+) |
|
240 next |
|
241 |
|
242 |
|
243 |
|
244 |
|
245 |
|
246 |
|
247 |
|
248 |
|
249 lemma enrich_proc_prop: |
|
250 "\<lbrakk>valid s; is_created_proc s p; p' \<notin> all_procs s\<rbrakk> |
|
251 \<Longrightarrow> valid (enrich_proc s p p') \<and> |
|
252 (p \<in> current_procs s \<longrightarrow> co2sobj (enrich_proc s p p') (O_proc p') = co2sobj (enrich_proc s p p') (O_proc p)) \<and> |
|
253 (\<forall> obj. alive s obj \<longrightarrow> alive (enrich_proc s p p') obj) \<and> |
|
254 (\<forall> p'. p' \<in> current_procs s \<longrightarrow> cp2sproc (enrich_proc s p p') p' = cp2sproc s p) \<and> |
|
255 (\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile (enrich_proc s p p') f = cf2sfile s f) \<and> |
|
256 (Tainted (enrich_proc s p p') = (Tainted s \<union> (if (O_proc p \<in> Tainted s) then {O_proc p'} else {})))" |
|
257 proof (induct s) |
|
258 case Nil |
|
259 thus ?case by (auto simp:is_created_proc_def) |
|
260 next |
|
261 case (Cons e s) |
|
262 hence p1: "\<lbrakk>valid s; is_created_proc s p; p' \<notin> all_procs s\<rbrakk> |
|
263 \<Longrightarrow> valid (enrich_proc s p p') \<and> |
|
264 (p \<in> current_procs s \<longrightarrow> co2sobj (enrich_proc s p p') (O_proc p') = co2sobj (enrich_proc s p p') (O_proc p)) \<and> |
|
265 (alive s obj \<longrightarrow> alive (enrich_proc s p p') obj \<and> co2sobj (enrich_proc s p p') obj = co2sobj s obj)" |
|
266 and p2: "valid (e # s)" and p3: "is_created_proc (e # s) p" and p4: "p' \<notin> all_procs (e # s)" |
|
267 by auto |
|
268 from p2 have vd: "valid s" and os: "os_grant s e" and grant: "grant s e" |
|
269 by (auto dest:vd_cons vt_grant vt_grant_os) |
|
270 from p4 have p4': "p' \<notin> all_procs s" by (case_tac e, auto) |
|
271 from p1 p4' have a1: "is_created_proc s p \<Longrightarrow> valid (enrich_proc s p p')" by (auto simp:vd) |
|
272 have c1: "valid (enrich_proc (e # s) p p')" |
|
273 apply (case_tac e) |
|
274 using a1 os p3 |
|
275 apply (auto simp:is_created_proc_def) |
|
276 sorry |
|
277 moreover have c2: "p' \<in> current_procs (enrich_proc (e # s) p p')" |
|
278 sorry |
|
279 moreover have c3: "co2sobj (enrich_proc (e # s) p p') (O_proc p') = co2sobj (enrich_proc (e # s) p p') (O_proc p)" |
|
280 sorry |
|
281 moreover have c4: "alive (e # s) obj \<longrightarrow> |
|
282 alive (enrich_proc (e # s) p p') obj \<and> co2sobj (enrich_proc (e # s) p p') obj = co2sobj (e # s) obj" |
|
283 sorry |
|
284 ultimately show ?case by auto |
|
285 qed |
|
286 |