212 qed |
346 qed |
213 thus ?thesis using grant |
347 thus ?thesis using grant |
214 by (auto simp:inherit_fds_check_def inherit_fds_check_ctxt_def) |
348 by (auto simp:inherit_fds_check_def inherit_fds_check_ctxt_def) |
215 qed |
349 qed |
216 |
350 |
217 lemma not_all_procs_cons: |
|
218 "p \<notin> all_procs (e # s) \<Longrightarrow> p \<notin> all_procs s" |
|
219 by (case_tac e, auto) |
|
220 |
|
221 lemma not_all_procs_prop: |
|
222 "\<lbrakk>p' \<notin> all_procs s; p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> p' \<noteq> p" |
|
223 apply (induct s, rule notI, simp) |
|
224 apply (frule vt_grant_os, frule vd_cons, frule not_all_procs_cons, simp, rule notI) |
|
225 apply (case_tac a, auto) |
|
226 done |
|
227 |
|
228 fun enrich_not_alive :: "t_state \<Rightarrow> t_enrich_obj \<Rightarrow> bool" |
|
229 where |
|
230 "enrich_not_alive s (E_file f) = (f \<notin> current_files s)" |
|
231 | "enrich_not_alive s (E_proc p) = (p \<notin> current_procs s)" |
|
232 | "enrich_not_alive s (E_fd p fd) = (p \<in> current_procs s \<longrightarrow> fd \<notin> current_proc_fds s p)" |
|
233 | "enrich_not_alive s (E_msgq q) = (q \<notin> current_msgqs s)" |
|
234 | "enrich_not_alive s (E_inum i) = (i \<notin> current_inode_nums s)" |
|
235 | "enrich_not_alive s (E_msg q m) = (q \<in> current_msgqs s \<longrightarrow> m \<notin> set (msgs_of_queue s q))" |
|
236 |
|
237 lemma file_has_parent: "\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> pf. is_dir s pf \<and> parent f = Some pf" |
|
238 apply (case_tac f) |
|
239 apply (simp, drule root_is_dir', simp+) |
|
240 apply (simp add:parentf_is_dir_prop2) |
|
241 done |
|
242 |
351 |
243 lemma enrich_valid_intro_cons: |
352 lemma enrich_valid_intro_cons: |
244 assumes vs': "valid s'" |
353 assumes vs': "valid s'" and vd': "valid (e # s)" |
245 and os: "os_grant s e" and grant: "grant s e" and vd: "valid s" |
|
246 and alive: "\<forall> obj. alive s obj \<longrightarrow> alive s' obj" |
354 and alive: "\<forall> obj. alive s obj \<longrightarrow> alive s' obj" |
247 and alive': "\<forall> obj. enrich_not_alive s obj \<longrightarrow> enrich_not_alive s' obj" |
355 and alive': "\<forall> obj. enrich_not_alive s obj' obj \<longrightarrow> enrich_not_alive s' obj' obj" |
248 and hungs: "files_hung_by_del s' = files_hung_by_del s" |
356 and hungs: "files_hung_by_del s' = files_hung_by_del s" |
249 and cp2sp: "\<forall> p. p \<in> current_procs s \<longrightarrow> cp2sproc s' p = cp2sproc s p" |
357 and cp2sp: "\<forall> p. p \<in> current_procs s \<longrightarrow> cp2sproc s' p = cp2sproc s p" |
250 and cf2sf: "\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile s' f = cf2sfile s f" |
358 and cf2sf: "\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile s' f = cf2sfile s f" |
251 and cq2sq: "\<forall> q. q \<in> current_msgqs s \<longrightarrow> cq2smsgq s' q = cq2smsgq s q" |
359 and cq2sq: "\<forall> q. q \<in> current_msgqs s \<longrightarrow> cq2smsgq s' q = cq2smsgq s q" |
252 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" |
360 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" |
253 and fflags_remain: "\<forall> p fd flags. flags_of_proc_fd s p fd = Some flags \<longrightarrow> flags_of_proc_fd s' p fd = Some flags" |
361 and fflags_remain: "\<forall> p fd flags. flags_of_proc_fd s p fd = Some flags \<longrightarrow> flags_of_proc_fd s' p fd = Some flags" |
254 and sms_remain: "\<forall> q. msgs_of_queue s' q = msgs_of_queue s q" |
362 and sms_remain: "\<forall> q. msgs_of_queue s' q = msgs_of_queue s q" |
255 (* and empty_remain: "\<forall> f. dir_is_empty s f \<longrightarrow> dir_is_empty s' f" *) |
363 (* and empty_remain: "\<forall> f. dir_is_empty s f \<longrightarrow> dir_is_empty s' f" *) |
256 and cfd2sfd: "\<forall> p fd. fd \<in> proc_file_fds s p \<longrightarrow> cfd2sfd s' p fd = cfd2sfd s p fd" |
364 and cfd2sfd: "\<forall> p fd. fd \<in> proc_file_fds s p \<longrightarrow> cfd2sfd s' p fd = cfd2sfd s p fd" |
|
365 and nodel: "no_del_event (e # s)" |
|
366 and notin_all: "notin_all (e # s) obj'" |
257 shows "valid (e # s')" |
367 shows "valid (e # s')" |
258 proof (cases e) |
368 proof- |
259 case (Execve p f fds) |
369 from vd' have os: "os_grant s e" and grant: "grant s e" and vd: "valid s" |
260 have p_in: "p \<in> current_procs s'" using os alive |
370 by (auto dest:vt_grant_os vt_grant vd_cons) |
261 apply (erule_tac x = "O_proc p" in allE) |
|
262 by (auto simp:Execve) |
|
263 have f_in: "is_file s' f" using os alive |
|
264 apply (erule_tac x = "O_file f" in allE) |
|
265 by (auto simp:Execve) |
|
266 have fd_in: "fds \<subseteq> proc_file_fds s' p" using os alive ffd_remain |
|
267 by (auto simp:Execve proc_file_fds_def) |
|
268 have "os_grant s' e" using p_in f_in fd_in by (simp add:Execve) |
|
269 moreover have "grant s' e" |
|
270 proof- |
|
271 from grant obtain up rp tp uf rf tf |
|
272 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
273 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
274 by (simp add:Execve split:option.splits, blast) |
|
275 with grant obtain pu nr nt where p3: "npctxt_execve (up, rp, tp) (uf, rf, tf) = Some (pu, nr, nt)" |
|
276 by (simp add:Execve split:option.splits del:npctxt_execve.simps, blast) |
|
277 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
278 using os cp2sp |
|
279 apply (erule_tac x = p in allE) |
|
280 by (auto simp:Execve co2sobj.simps cp2sproc_def split:option.splits) |
|
281 from os have f_in': "is_file s f" by (simp add:Execve) |
|
282 from vd os have "\<exists> sf. cf2sfile s f = Some sf" |
|
283 by (auto dest!:is_file_in_current current_file_has_sfile simp:Execve) |
|
284 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf |
|
285 apply (erule_tac x = f in allE) |
|
286 apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) |
|
287 apply (case_tac f, simp) |
|
288 apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
|
289 done |
|
290 have "inherit_fds_check s' (pu, nr, nt) p fds" |
|
291 proof- |
|
292 have "fds \<subseteq> proc_file_fds s' p" using os ffd_remain Execve |
|
293 by (auto simp:proc_file_fds_def) |
|
294 thus ?thesis using Execve grant vd cfd2sfd p1 p2 p3 os |
|
295 apply (rule_tac s = s in enrich_inherit_fds_check) |
|
296 by (simp_all split:option.splits) |
|
297 qed |
|
298 moreover have "search_check s' (pu, rp, tp) f" |
|
299 using p1 p2 p2' vd cf2sf f_in' grant Execve p3 f_in |
|
300 apply (rule_tac s = s in enrich_search_check) |
|
301 by (simp_all split:option.splits) |
|
302 ultimately show ?thesis using p1' p2' p3 |
|
303 apply (simp add:Execve split:option.splits) |
|
304 using grant Execve p1 p2 |
|
305 by (simp add:Execve grant p1 p2) |
|
306 qed |
|
307 ultimately show ?thesis using vs' |
|
308 by (erule_tac valid.intros(2), simp+) |
|
309 next |
|
310 case (Clone p p' fds) |
|
311 have p_in: "p \<in> current_procs s'" using os alive |
|
312 apply (erule_tac x = "O_proc p" in allE) |
|
313 by (auto simp:Clone) |
|
314 have p'_not_in: "p' \<notin> current_procs s'" using os alive' |
|
315 apply (erule_tac x = "E_proc p'" in allE) |
|
316 by (auto simp:Clone) |
|
317 have fd_in: "fds \<subseteq> proc_file_fds s' p" using os alive ffd_remain |
|
318 by (auto simp:Clone proc_file_fds_def) |
|
319 have "os_grant s' e" using p_in p'_not_in fd_in by (simp add:Clone) |
|
320 moreover have "grant s' e" |
|
321 proof- |
|
322 from grant obtain up rp tp |
|
323 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
324 apply (simp add:Clone split:option.splits) |
|
325 by (case_tac a, auto) |
|
326 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
327 using os cp2sp |
|
328 apply (erule_tac x = p in allE) |
|
329 by (auto simp:Clone co2sobj.simps cp2sproc_def split:option.splits) |
|
330 have p2: "inherit_fds_check s' (up, rp, tp) p fds" |
|
331 proof- |
|
332 have "fds \<subseteq> proc_file_fds s' p" using os ffd_remain Clone |
|
333 by (auto simp:proc_file_fds_def) |
|
334 thus ?thesis using Clone grant vd cfd2sfd p1 os |
|
335 apply (rule_tac s = s in enrich_inherit_fds_check) |
|
336 by (simp_all split:option.splits) |
|
337 qed |
|
338 show ?thesis using p1 p2 p1' grant |
|
339 by (simp add:Clone) |
|
340 qed |
|
341 ultimately show ?thesis using vs' |
|
342 by (erule_tac valid.intros(2), simp+) |
|
343 next |
|
344 case (Kill p p') |
|
345 have p_in: "p \<in> current_procs s'" using os alive |
|
346 apply (erule_tac x = "O_proc p" in allE) |
|
347 by (auto simp:Kill) |
|
348 have p'_in: "p' \<in> current_procs s'" using os alive |
|
349 apply (erule_tac x = "O_proc p'" in allE) |
|
350 by (auto simp:Kill) |
|
351 have "os_grant s' e" using p_in p'_in by (simp add:Kill) |
|
352 moreover have "grant s' e" |
|
353 proof- |
|
354 from grant obtain up rp tp up' rp' tp' |
|
355 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
356 and p'1: "sectxt_of_obj s (O_proc p') = Some (up', rp', tp')" |
|
357 apply (simp add:Kill split:option.splits) |
|
358 by (case_tac a, case_tac aa, blast) |
|
359 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
360 using os cp2sp |
|
361 apply (erule_tac x = p in allE) |
|
362 by (auto simp:Kill co2sobj.simps cp2sproc_def split:option.splits) |
|
363 from p'1 have p'1': "sectxt_of_obj s' (O_proc p') = Some (up', rp', tp')" |
|
364 using os cp2sp |
|
365 apply (erule_tac x = p' in allE) |
|
366 by (auto simp:Kill co2sobj.simps cp2sproc_def split:option.splits) |
|
367 show ?thesis using p1 p'1 p1' p'1' grant |
|
368 by (simp add:Kill) |
|
369 qed |
|
370 ultimately show ?thesis using vs' |
|
371 by (erule_tac valid.intros(2), simp+) |
|
372 next |
|
373 case (Ptrace p p') |
|
374 have p_in: "p \<in> current_procs s'" using os alive |
|
375 apply (erule_tac x = "O_proc p" in allE) |
|
376 by (auto simp:Ptrace) |
|
377 have p'_in: "p' \<in> current_procs s'" using os alive |
|
378 apply (erule_tac x = "O_proc p'" in allE) |
|
379 by (auto simp:Ptrace) |
|
380 have "os_grant s' e" using p_in p'_in by (simp add:Ptrace) |
|
381 moreover have "grant s' e" |
|
382 proof- |
|
383 from grant obtain up rp tp up' rp' tp' |
|
384 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
385 and p'1: "sectxt_of_obj s (O_proc p') = Some (up', rp', tp')" |
|
386 apply (simp add:Ptrace split:option.splits) |
|
387 by (case_tac a, case_tac aa, blast) |
|
388 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
389 using os cp2sp |
|
390 apply (erule_tac x = p in allE) |
|
391 by (auto simp:Ptrace co2sobj.simps cp2sproc_def split:option.splits) |
|
392 from p'1 have p'1': "sectxt_of_obj s' (O_proc p') = Some (up', rp', tp')" |
|
393 using os cp2sp |
|
394 apply (erule_tac x = p' in allE) |
|
395 by (auto simp:Ptrace co2sobj.simps cp2sproc_def split:option.splits) |
|
396 show ?thesis using p1 p'1 p1' p'1' grant |
|
397 by (simp add:Ptrace) |
|
398 qed |
|
399 ultimately show ?thesis using vs' |
|
400 by (erule_tac valid.intros(2), simp+) |
|
401 next |
|
402 case (Exit p) |
|
403 have p_in: "p \<in> current_procs s'" using os alive |
|
404 apply (erule_tac x = "O_proc p" in allE) |
|
405 by (auto simp:Exit) |
|
406 have "os_grant s' e" using p_in by (simp add:Exit) |
|
407 moreover have "grant s' e" |
|
408 by (simp add:Exit) |
|
409 ultimately show ?thesis using vs' |
|
410 by (erule_tac valid.intros(2), simp+) |
|
411 next |
|
412 case (Open p f flags fd opt) |
|
413 show ?thesis |
371 show ?thesis |
414 proof (cases opt) |
372 proof (cases e) |
415 case None |
373 case (Execve p f fds) |
416 have p_in: "p \<in> current_procs s'" using os alive |
374 have p_in: "p \<in> current_procs s'" using os alive |
417 apply (erule_tac x = "O_proc p" in allE) |
375 apply (erule_tac x = "O_proc p" in allE) |
418 by (auto simp:Open None) |
376 by (auto simp:Execve) |
419 have f_in: "is_file s' f" using os alive |
377 have f_in: "is_file s' f" using os alive |
420 apply (erule_tac x = "O_file f" in allE) |
378 apply (erule_tac x = "O_file f" in allE) |
421 by (auto simp:Open None) |
379 by (auto simp:Execve) |
422 have fd_not_in: "fd \<notin> current_proc_fds s' p" |
380 have fd_in: "fds \<subseteq> proc_file_fds s' p" using os alive ffd_remain |
423 using os alive' p_in |
381 by (auto simp:Execve proc_file_fds_def) |
424 apply (erule_tac x = "E_fd p fd" in allE) |
382 have "os_grant s' e" using p_in f_in fd_in by (simp add:Execve) |
425 by (simp add:Open None) |
383 moreover have "grant s' e" |
426 have "os_grant s' e" using p_in f_in fd_not_in os |
|
427 by (simp add:Open None) |
|
428 moreover have "grant s' e" |
|
429 proof- |
384 proof- |
430 from grant obtain up rp tp uf rf tf |
385 from grant obtain up rp tp uf rf tf |
431 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
386 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
432 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
387 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
433 apply (simp add:Open None split:option.splits) |
388 by (simp add:Execve split:option.splits, blast) |
434 by (case_tac a, case_tac aa, blast) |
389 with grant obtain pu nr nt where p3: "npctxt_execve (up, rp, tp) (uf, rf, tf) = Some (pu, nr, nt)" |
435 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
390 by (simp add:Execve split:option.splits del:npctxt_execve.simps, blast) |
436 using os cp2sp |
391 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
437 apply (erule_tac x = p in allE) |
392 using os cp2sp |
438 by (auto simp:Open None co2sobj.simps cp2sproc_def split:option.splits) |
393 apply (erule_tac x = p in allE) |
439 from os have f_in': "is_file s f" by (simp add:Open None) |
394 by (auto simp:Execve co2sobj.simps cp2sproc_def split:option.splits) |
|
395 from os have f_in': "is_file s f" by (simp add:Execve) |
440 from vd os have "\<exists> sf. cf2sfile s f = Some sf" |
396 from vd os have "\<exists> sf. cf2sfile s f = Some sf" |
441 by (auto dest!:is_file_in_current current_file_has_sfile simp:Open None) |
397 by (auto dest!:is_file_in_current current_file_has_sfile simp:Execve) |
442 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf |
398 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf |
443 apply (erule_tac x = f in allE) |
399 apply (erule_tac x = f in allE) |
444 apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) |
400 apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) |
445 apply (case_tac f, simp) |
401 apply (case_tac f, simp) |
446 apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
402 apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
447 done |
403 done |
448 have "search_check s' (up, rp, tp) f" |
404 have "inherit_fds_check s' (pu, nr, nt) p fds" |
449 using p1 p2 p2' vd cf2sf f_in' grant Open None f_in |
405 proof- |
|
406 have "fds \<subseteq> proc_file_fds s' p" using os ffd_remain Execve |
|
407 by (auto simp:proc_file_fds_def) |
|
408 thus ?thesis using Execve grant vd cfd2sfd p1 p2 p3 os |
|
409 apply (rule_tac s = s in enrich_inherit_fds_check) |
|
410 by (simp_all split:option.splits) |
|
411 qed |
|
412 moreover have "search_check s' (pu, rp, tp) f" |
|
413 using p1 p2 p2' vd cf2sf f_in' grant Execve p3 f_in |
450 apply (rule_tac s = s in enrich_search_check) |
414 apply (rule_tac s = s in enrich_search_check) |
451 by (simp_all split:option.splits) |
415 by (simp_all split:option.splits) |
452 thus ?thesis using p1' p2' |
416 ultimately show ?thesis using p1' p2' p3 |
453 apply (simp add:Open None split:option.splits) |
417 apply (simp add:Execve split:option.splits) |
454 using grant Open None p1 p2 |
418 using grant Execve p1 p2 |
|
419 by (simp add:Execve grant p1 p2) |
|
420 qed |
|
421 ultimately show ?thesis using vs' |
|
422 by (erule_tac valid.intros(2), simp+) |
|
423 next |
|
424 case (Clone p p' fds) |
|
425 have p_in: "p \<in> current_procs s'" using os alive |
|
426 apply (erule_tac x = "O_proc p" in allE) |
|
427 by (auto simp:Clone) |
|
428 have p'_not_in: "p' \<notin> current_procs s'" using alive' notin_all os Clone |
|
429 apply (erule_tac x = "E_proc p'" in allE) |
|
430 apply (auto dest:not_all_procs_prop3) |
|
431 done |
|
432 have fd_in: "fds \<subseteq> proc_file_fds s' p" using os alive ffd_remain |
|
433 by (auto simp:Clone proc_file_fds_def) |
|
434 have "os_grant s' e" using p_in p'_not_in fd_in by (simp add:Clone) |
|
435 moreover have "grant s' e" |
|
436 proof- |
|
437 from grant obtain up rp tp |
|
438 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
439 apply (simp add:Clone split:option.splits) |
|
440 by (case_tac a, auto) |
|
441 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
442 using os cp2sp |
|
443 apply (erule_tac x = p in allE) |
|
444 by (auto simp:Clone co2sobj.simps cp2sproc_def split:option.splits) |
|
445 have p2: "inherit_fds_check s' (up, rp, tp) p fds" |
|
446 proof- |
|
447 have "fds \<subseteq> proc_file_fds s' p" using os ffd_remain Clone |
|
448 by (auto simp:proc_file_fds_def) |
|
449 thus ?thesis using Clone grant vd cfd2sfd p1 os |
|
450 apply (rule_tac s = s in enrich_inherit_fds_check) |
|
451 by (simp_all split:option.splits) |
|
452 qed |
|
453 show ?thesis using p1 p2 p1' grant |
|
454 by (simp add:Clone) |
|
455 qed |
|
456 ultimately show ?thesis using vs' |
|
457 by (erule_tac valid.intros(2), simp+) |
|
458 next |
|
459 case (Kill p p') |
|
460 have p_in: "p \<in> current_procs s'" using os alive |
|
461 apply (erule_tac x = "O_proc p" in allE) |
|
462 by (auto simp:Kill) |
|
463 have p'_in: "p' \<in> current_procs s'" using os alive |
|
464 apply (erule_tac x = "O_proc p'" in allE) |
|
465 by (auto simp:Kill) |
|
466 have "os_grant s' e" using p_in p'_in by (simp add:Kill) |
|
467 moreover have "grant s' e" |
|
468 proof- |
|
469 from grant obtain up rp tp up' rp' tp' |
|
470 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
471 and p'1: "sectxt_of_obj s (O_proc p') = Some (up', rp', tp')" |
|
472 apply (simp add:Kill split:option.splits) |
|
473 by (case_tac a, case_tac aa, blast) |
|
474 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
475 using os cp2sp |
|
476 apply (erule_tac x = p in allE) |
|
477 by (auto simp:Kill co2sobj.simps cp2sproc_def split:option.splits) |
|
478 from p'1 have p'1': "sectxt_of_obj s' (O_proc p') = Some (up', rp', tp')" |
|
479 using os cp2sp |
|
480 apply (erule_tac x = p' in allE) |
|
481 by (auto simp:Kill co2sobj.simps cp2sproc_def split:option.splits) |
|
482 show ?thesis using p1 p'1 p1' p'1' grant |
|
483 by (simp add:Kill) |
|
484 qed |
|
485 ultimately show ?thesis using vs' |
|
486 by (erule_tac valid.intros(2), simp+) |
|
487 next |
|
488 case (Ptrace p p') |
|
489 have p_in: "p \<in> current_procs s'" using os alive |
|
490 apply (erule_tac x = "O_proc p" in allE) |
|
491 by (auto simp:Ptrace) |
|
492 have p'_in: "p' \<in> current_procs s'" using os alive |
|
493 apply (erule_tac x = "O_proc p'" in allE) |
|
494 by (auto simp:Ptrace) |
|
495 have "os_grant s' e" using p_in p'_in by (simp add:Ptrace) |
|
496 moreover have "grant s' e" |
|
497 proof- |
|
498 from grant obtain up rp tp up' rp' tp' |
|
499 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
500 and p'1: "sectxt_of_obj s (O_proc p') = Some (up', rp', tp')" |
|
501 apply (simp add:Ptrace split:option.splits) |
|
502 by (case_tac a, case_tac aa, blast) |
|
503 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
504 using os cp2sp |
|
505 apply (erule_tac x = p in allE) |
|
506 by (auto simp:Ptrace co2sobj.simps cp2sproc_def split:option.splits) |
|
507 from p'1 have p'1': "sectxt_of_obj s' (O_proc p') = Some (up', rp', tp')" |
|
508 using os cp2sp |
|
509 apply (erule_tac x = p' in allE) |
|
510 by (auto simp:Ptrace co2sobj.simps cp2sproc_def split:option.splits) |
|
511 show ?thesis using p1 p'1 p1' p'1' grant |
|
512 by (simp add:Ptrace) |
|
513 qed |
|
514 ultimately show ?thesis using vs' |
|
515 by (erule_tac valid.intros(2), simp+) |
|
516 next |
|
517 case (Exit p) |
|
518 have p_in: "p \<in> current_procs s'" using os alive |
|
519 apply (erule_tac x = "O_proc p" in allE) |
|
520 by (auto simp:Exit) |
|
521 have "os_grant s' e" using p_in by (simp add:Exit) |
|
522 moreover have "grant s' e" |
|
523 by (simp add:Exit) |
|
524 ultimately show ?thesis using vs' |
|
525 by (erule_tac valid.intros(2), simp+) |
|
526 next |
|
527 case (Open p f flags fd opt) |
|
528 show ?thesis |
|
529 proof (cases opt) |
|
530 case None |
|
531 have p_in: "p \<in> current_procs s'" using os alive |
|
532 apply (erule_tac x = "O_proc p" in allE) |
|
533 by (auto simp:Open None) |
|
534 have f_in: "is_file s' f" using os alive |
|
535 apply (erule_tac x = "O_file f" in allE) |
|
536 by (auto simp:Open None) |
|
537 have fd_not_in: "fd \<notin> current_proc_fds s' p" |
|
538 using os alive' p_in notin_all Open None |
|
539 apply (erule_tac x = "E_fd p fd" in allE) |
|
540 apply (case_tac obj') |
|
541 by auto |
|
542 have "os_grant s' e" using p_in f_in fd_not_in os |
|
543 by (simp add:Open None) |
|
544 moreover have "grant s' e" |
|
545 proof- |
|
546 from grant obtain up rp tp uf rf tf |
|
547 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
548 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
549 apply (simp add:Open None split:option.splits) |
|
550 by (case_tac a, case_tac aa, blast) |
|
551 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
552 using os cp2sp |
|
553 apply (erule_tac x = p in allE) |
|
554 by (auto simp:Open None co2sobj.simps cp2sproc_def split:option.splits) |
|
555 from os have f_in': "is_file s f" by (simp add:Open None) |
|
556 from vd os have "\<exists> sf. cf2sfile s f = Some sf" |
|
557 by (auto dest!:is_file_in_current current_file_has_sfile simp:Open None) |
|
558 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf |
|
559 apply (erule_tac x = f in allE) |
|
560 apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) |
|
561 apply (case_tac f, simp) |
|
562 apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
|
563 done |
|
564 have "search_check s' (up, rp, tp) f" |
|
565 using p1 p2 p2' vd cf2sf f_in' grant Open None f_in |
|
566 apply (rule_tac s = s in enrich_search_check) |
|
567 by (simp_all split:option.splits) |
|
568 thus ?thesis using p1' p2' |
|
569 apply (simp add:Open None split:option.splits) |
|
570 using grant Open None p1 p2 |
|
571 by simp |
|
572 qed |
|
573 ultimately show ?thesis using vs' |
|
574 by (erule_tac valid.intros(2), simp+) |
|
575 next |
|
576 case (Some inum) |
|
577 from os obtain pf where pf_in_s: "is_dir s pf" and parent: "parent f = Some pf" |
|
578 by (auto simp:Open Some) |
|
579 have pf_in: "is_dir s' pf" using pf_in_s alive |
|
580 apply (erule_tac x = "O_dir pf" in allE) |
455 by simp |
581 by simp |
|
582 have p_in: "p \<in> current_procs s'" using os alive |
|
583 apply (erule_tac x = "O_proc p" in allE) |
|
584 by (auto simp:Open Some) |
|
585 have f_not_in: "f \<notin> current_files s'" using os alive' Open Some notin_all |
|
586 apply (erule_tac x = "E_file f" in allE) |
|
587 apply (case_tac obj') |
|
588 by auto |
|
589 have fd_not_in: "fd \<notin> current_proc_fds s' p" |
|
590 using os alive' p_in Open Some notin_all |
|
591 apply (erule_tac x = "E_fd p fd" in allE) |
|
592 apply (case_tac obj', auto) |
|
593 done |
|
594 have inum_not_in: "inum \<notin> current_inode_nums s'" |
|
595 using os alive' Open Some notin_all |
|
596 apply (erule_tac x = "E_inum inum" in allE) |
|
597 by (case_tac obj', auto) |
|
598 have "os_grant s' e" using p_in pf_in parent f_not_in fd_not_in inum_not_in os |
|
599 by (simp add:Open Some hungs) |
|
600 moreover have "grant s' e" |
|
601 proof- |
|
602 from grant parent obtain up rp tp uf rf tf |
|
603 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
604 and p2: "sectxt_of_obj s (O_dir pf) = Some (uf, rf, tf)" |
|
605 apply (simp add:Open Some split:option.splits) |
|
606 by (case_tac a, case_tac aa, blast) |
|
607 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
608 using os cp2sp |
|
609 apply (erule_tac x = p in allE) |
|
610 by (auto simp:Open Some co2sobj.simps cp2sproc_def split:option.splits) |
|
611 from vd os pf_in_s have "\<exists> sf. cf2sfile s pf = Some sf" |
|
612 by (auto dest!:is_dir_in_current current_file_has_sfile simp:Open Some) |
|
613 hence p2': "sectxt_of_obj s' (O_dir pf) = Some (uf, rf, tf)" using p2 cf2sf pf_in pf_in_s |
|
614 apply (erule_tac x = pf in allE) |
|
615 apply (erule exE, frule_tac s = s in is_dir_in_current, simp) |
|
616 apply (drule is_dir_not_file, drule is_dir_not_file) |
|
617 apply (auto simp:cf2sfile_def split:option.splits) |
|
618 apply (case_tac pf, simp_all) |
|
619 by (simp add:sroot_def root_sec_remains vd vs') |
|
620 have "search_check s' (up, rp, tp) pf" |
|
621 using p1 p2 p2' vd cf2sf pf_in grant Open Some pf_in_s parent vs' |
|
622 apply (rule_tac s = s in enrich_search_check') |
|
623 by (simp_all split:option.splits) |
|
624 thus ?thesis using p1' p2' parent |
|
625 apply (simp add:Open Some split:option.splits) |
|
626 using grant Open Some p1 p2 |
|
627 by simp |
|
628 qed |
|
629 ultimately show ?thesis using vs' |
|
630 by (erule_tac valid.intros(2), simp+) |
|
631 qed |
|
632 next |
|
633 case (ReadFile p fd) |
|
634 have p_in: "p \<in> current_procs s'" using os alive |
|
635 apply (erule_tac x = "O_proc p" in allE) |
|
636 by (auto simp:ReadFile) |
|
637 have fd_in: "fd \<in> current_proc_fds s' p" using os alive |
|
638 apply (erule_tac x = "O_fd p fd" in allE) |
|
639 by (auto simp:ReadFile) |
|
640 obtain f where ffd: "file_of_proc_fd s p fd = Some f" |
|
641 using os ReadFile by auto |
|
642 hence f_in_s: "is_file s f" using vd |
|
643 by (auto intro:file_of_pfd_is_file) |
|
644 obtain flags where fflag: "flags_of_proc_fd s p fd = Some flags" |
|
645 using os ReadFile by auto |
|
646 have ffd_in: "file_of_proc_fd s' p fd = Some f" |
|
647 using ffd_remain ffd by auto |
|
648 hence f_in: "is_file s' f" using vs' |
|
649 by (auto intro:file_of_pfd_is_file) |
|
650 have flags_in: "flags_of_proc_fd s' p fd = Some flags" |
|
651 using fflags_remain fflag by auto |
|
652 have "os_grant s' e" using p_in fd_in ffd_in flags_in fflag os f_in |
|
653 by (auto simp add:ReadFile is_file_in_current) |
|
654 moreover have "grant s' e" |
|
655 proof- |
|
656 from grant ffd obtain up rp tp uf rf tf ufd rfd tfd |
|
657 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
658 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
659 and p3: "sectxt_of_obj s (O_fd p fd) = Some (ufd, rfd, tfd)" |
|
660 apply (simp add:ReadFile split:option.splits) |
|
661 by (case_tac a, case_tac aa, case_tac ab, blast) |
|
662 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
663 using os cp2sp |
|
664 apply (erule_tac x = p in allE) |
|
665 by (auto simp:ReadFile co2sobj.simps cp2sproc_def split:option.splits) |
|
666 from vd f_in_s have "\<exists> sf. cf2sfile s f = Some sf" |
|
667 by (auto dest!:is_file_in_current current_file_has_sfile) |
|
668 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in_s p2 cf2sf |
|
669 apply (erule_tac x = f in allE) |
|
670 apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) |
|
671 apply (case_tac f, simp) |
|
672 apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
|
673 done |
|
674 have p3': "sectxt_of_obj s' (O_fd p fd) = Some (ufd, rfd, tfd)" |
|
675 using cfd2sfd ffd_in ffd p3 f_in f_in_s vd |
|
676 apply (erule_tac x = p in allE) |
|
677 apply (erule_tac x = fd in allE) |
|
678 apply (simp add:proc_file_fds_def) |
|
679 apply (auto simp:cfd2sfd_def fflag flags_in p3 split:option.splits |
|
680 dest!:current_file_has_sfile' simp:is_file_in_current) |
|
681 done |
|
682 show ?thesis using p1' p2' p3' ffd_in ffd |
|
683 apply (simp add:ReadFile split:option.splits) |
|
684 using grant p1 p2 p3 ReadFile |
|
685 by simp |
456 qed |
686 qed |
457 ultimately show ?thesis using vs' |
687 ultimately show ?thesis using vs' |
458 by (erule_tac valid.intros(2), simp+) |
688 by (erule_tac valid.intros(2), simp+) |
459 next |
689 next |
460 case (Some inum) |
690 case (WriteFile p fd) |
|
691 have p_in: "p \<in> current_procs s'" using os alive |
|
692 apply (erule_tac x = "O_proc p" in allE) |
|
693 by (auto simp:WriteFile) |
|
694 have fd_in: "fd \<in> current_proc_fds s' p" using os alive |
|
695 apply (erule_tac x = "O_fd p fd" in allE) |
|
696 by (auto simp:WriteFile) |
|
697 obtain f where ffd: "file_of_proc_fd s p fd = Some f" |
|
698 using os WriteFile by auto |
|
699 hence f_in_s: "is_file s f" using vd |
|
700 by (auto intro:file_of_pfd_is_file) |
|
701 obtain flags where fflag: "flags_of_proc_fd s p fd = Some flags" |
|
702 using os WriteFile by auto |
|
703 have ffd_in: "file_of_proc_fd s' p fd = Some f" |
|
704 using ffd_remain ffd by auto |
|
705 hence f_in: "is_file s' f" using vs' |
|
706 by (auto intro:file_of_pfd_is_file) |
|
707 have flags_in: "flags_of_proc_fd s' p fd = Some flags" |
|
708 using fflags_remain fflag by auto |
|
709 have "os_grant s' e" using p_in fd_in ffd_in flags_in fflag os f_in |
|
710 by (auto simp add:WriteFile is_file_in_current) |
|
711 moreover have "grant s' e" |
|
712 proof- |
|
713 from grant ffd obtain up rp tp uf rf tf ufd rfd tfd |
|
714 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
715 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
716 and p3: "sectxt_of_obj s (O_fd p fd) = Some (ufd, rfd, tfd)" |
|
717 apply (simp add:WriteFile split:option.splits) |
|
718 by (case_tac a, case_tac aa, case_tac ab, blast) |
|
719 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
720 using os cp2sp |
|
721 apply (erule_tac x = p in allE) |
|
722 by (auto simp:WriteFile co2sobj.simps cp2sproc_def split:option.splits) |
|
723 from vd f_in_s have "\<exists> sf. cf2sfile s f = Some sf" |
|
724 by (auto dest!:is_file_in_current current_file_has_sfile) |
|
725 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in_s p2 cf2sf |
|
726 apply (erule_tac x = f in allE) |
|
727 apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) |
|
728 apply (case_tac f, simp) |
|
729 apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
|
730 done |
|
731 have p3': "sectxt_of_obj s' (O_fd p fd) = Some (ufd, rfd, tfd)" |
|
732 using cfd2sfd ffd_in ffd p3 f_in f_in_s vd |
|
733 apply (erule_tac x = p in allE) |
|
734 apply (erule_tac x = fd in allE) |
|
735 apply (simp add:proc_file_fds_def) |
|
736 apply (auto simp:cfd2sfd_def fflag flags_in p3 split:option.splits |
|
737 dest!:current_file_has_sfile' simp:is_file_in_current) |
|
738 done |
|
739 show ?thesis using p1' p2' p3' ffd_in ffd |
|
740 apply (simp add:WriteFile split:option.splits) |
|
741 using grant p1 p2 p3 WriteFile |
|
742 by simp |
|
743 qed |
|
744 ultimately show ?thesis using vs' |
|
745 by (erule_tac valid.intros(2), simp+) |
|
746 next |
|
747 case (CloseFd p fd) |
|
748 have p_in: "p \<in> current_procs s'" using os alive |
|
749 apply (erule_tac x = "O_proc p" in allE) |
|
750 by (auto simp:CloseFd) |
|
751 have fd_in: "fd \<in> current_proc_fds s' p" using os alive |
|
752 apply (erule_tac x = "O_fd p fd" in allE) |
|
753 by (auto simp:CloseFd) |
|
754 have "os_grant s' e" using p_in fd_in |
|
755 by (auto simp add:CloseFd) |
|
756 moreover have "grant s' e" |
|
757 by(simp add:CloseFd) |
|
758 ultimately show ?thesis using vs' |
|
759 by (erule_tac valid.intros(2), simp+) |
|
760 next |
|
761 case (UnLink p f) |
|
762 have p_in: "p \<in> current_procs s'" using os alive |
|
763 apply (erule_tac x = "O_proc p" in allE) |
|
764 by (auto simp:UnLink) |
|
765 have f_in: "is_file s' f" using os alive |
|
766 apply (erule_tac x = "O_file f" in allE) |
|
767 by (auto simp:UnLink) |
|
768 from os vd obtain pf where pf_in_s: "is_dir s pf" |
|
769 and parent: "parent f = Some pf" |
|
770 by (auto simp:UnLink dest!:file_has_parent) |
|
771 from pf_in_s alive have pf_in: "is_dir s' pf" |
|
772 apply (erule_tac x = "O_dir pf" in allE) |
|
773 by (auto simp:UnLink) |
|
774 have "os_grant s' e" using p_in f_in os |
|
775 by (simp add:UnLink hungs) |
|
776 moreover have "grant s' e" |
|
777 proof- |
|
778 from grant parent obtain up rp tp uf rf tf upf rpf tpf |
|
779 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
780 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
781 and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" |
|
782 apply (simp add:UnLink split:option.splits) |
|
783 by (case_tac a, case_tac aa, case_tac ab, blast) |
|
784 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
785 using os cp2sp |
|
786 apply (erule_tac x = p in allE) |
|
787 by (auto simp:UnLink co2sobj.simps cp2sproc_def split:option.splits) |
|
788 from vd os pf_in_s have "\<exists> sf. cf2sfile s f = Some sf" |
|
789 by (auto dest!:is_file_in_current current_file_has_sfile simp:UnLink) |
|
790 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" |
|
791 using p2 cf2sf f_in os parent |
|
792 apply (erule_tac x = f in allE) |
|
793 apply (erule exE, clarsimp simp:UnLink) |
|
794 apply (frule_tac s = s in is_file_in_current, simp) |
|
795 by (auto simp:cf2sfile_def split:option.splits) |
|
796 from vd os pf_in_s have "\<exists> sf. cf2sfile s pf = Some sf" |
|
797 by (auto dest!:is_dir_in_current current_file_has_sfile simp:UnLink) |
|
798 hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s |
|
799 apply (erule_tac x = pf in allE) |
|
800 apply (erule exE, frule_tac s = s in is_dir_in_current, simp) |
|
801 apply (drule is_dir_not_file, drule is_dir_not_file) |
|
802 apply (auto simp:cf2sfile_def split:option.splits) |
|
803 apply (case_tac pf, simp_all) |
|
804 by (simp add:sroot_def root_sec_remains vd vs') |
|
805 have "search_check s' (up, rp, tp) f" |
|
806 using p1 p2 p2' vd cf2sf f_in grant UnLink os parent vs' |
|
807 apply (rule_tac s = s in enrich_search_check) |
|
808 by (simp_all split:option.splits) |
|
809 thus ?thesis using p1' p2' p3' parent |
|
810 apply (simp add:UnLink split:option.splits) |
|
811 using grant UnLink p1 p2 p3 |
|
812 by simp |
|
813 qed |
|
814 ultimately show ?thesis using vs' |
|
815 by (erule_tac valid.intros(2), simp+) |
|
816 next |
|
817 case (Rmdir p f) |
|
818 have p_in: "p \<in> current_procs s'" using os alive |
|
819 apply (erule_tac x = "O_proc p" in allE) |
|
820 by (auto simp:Rmdir) |
|
821 have f_in: "is_dir s' f" using os alive |
|
822 apply (erule_tac x = "O_dir f" in allE) |
|
823 by (auto simp:Rmdir dir_is_empty_def) |
|
824 have not_root: "f \<noteq> []" using os |
|
825 by (auto simp:Rmdir) |
|
826 from os vd obtain pf where pf_in_s: "is_dir s pf" |
|
827 and parent: "parent f = Some pf" |
|
828 apply (auto simp:Rmdir dir_is_empty_def) |
|
829 apply (case_tac f, simp+) |
|
830 apply (drule parentf_is_dir_prop1, auto) |
|
831 done |
|
832 from pf_in_s alive have pf_in: "is_dir s' pf" |
|
833 apply (erule_tac x = "O_dir pf" in allE) |
|
834 by (auto simp:Rmdir) |
|
835 have empty_in: "dir_is_empty s' f" using os Rmdir notin_all |
|
836 apply (clarsimp simp add:dir_is_empty_def f_in) |
|
837 using alive' |
|
838 apply (erule_tac x = "E_file f'" in allE) |
|
839 apply simp |
|
840 apply (erule disjE) |
|
841 apply (erule_tac x = f' in allE, simp) |
|
842 apply (case_tac obj', simp_all) |
|
843 apply (clarsimp) |
|
844 apply (drule_tac f' = f in parent_ancen) |
|
845 apply (simp, rule notI, simp add:noJ_Anc) |
|
846 apply (case_tac "f = pf") |
|
847 using vd' Rmdir |
|
848 apply (simp_all add:is_dir_rmdir) |
|
849 apply (erule_tac x = pf in allE) |
|
850 apply (drule_tac f = pf in is_dir_in_current) |
|
851 apply (simp add:noJ_Anc) |
|
852 done |
|
853 have "os_grant s' e" using p_in f_in os empty_in |
|
854 by (simp add:Rmdir hungs) |
|
855 moreover have "grant s' e" |
|
856 proof- |
|
857 from grant parent obtain up rp tp uf rf tf upf rpf tpf |
|
858 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
859 and p2: "sectxt_of_obj s (O_dir f) = Some (uf, rf, tf)" |
|
860 and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" |
|
861 apply (simp add:Rmdir split:option.splits) |
|
862 by (case_tac a, case_tac aa, case_tac ab, blast) |
|
863 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
864 using os cp2sp |
|
865 apply (erule_tac x = p in allE) |
|
866 by (auto simp:Rmdir co2sobj.simps cp2sproc_def split:option.splits) |
|
867 from vd os pf_in_s have "\<exists> sf. cf2sfile s f = Some sf" |
|
868 by (auto dest!:is_dir_in_current current_file_has_sfile simp:dir_is_empty_def Rmdir) |
|
869 hence p2': "sectxt_of_obj s' (O_dir f) = Some (uf, rf, tf)" |
|
870 using p2 cf2sf f_in os parent |
|
871 apply (erule_tac x = f in allE) |
|
872 apply (erule exE, clarsimp simp:Rmdir dir_is_empty_def) |
|
873 apply (frule_tac s = s in is_dir_in_current, simp) |
|
874 apply (drule is_dir_not_file, drule is_dir_not_file) |
|
875 by (auto simp:cf2sfile_def split:option.splits) |
|
876 from vd os pf_in_s have "\<exists> sf. cf2sfile s pf = Some sf" |
|
877 by (auto dest!:is_dir_in_current current_file_has_sfile simp:Rmdir) |
|
878 hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s |
|
879 apply (erule_tac x = pf in allE) |
|
880 apply (erule exE, frule_tac s = s in is_dir_in_current, simp) |
|
881 apply (drule is_dir_not_file, drule is_dir_not_file) |
|
882 apply (auto simp:cf2sfile_def split:option.splits) |
|
883 apply (case_tac pf, simp_all) |
|
884 by (simp add:sroot_def root_sec_remains vd vs') |
|
885 have "search_check s' (up, rp, tp) f" |
|
886 using p1 p2 p2' vd cf2sf f_in grant Rmdir os parent vs' |
|
887 apply (rule_tac s = s in enrich_search_check') |
|
888 by (simp_all add:dir_is_empty_def split:option.splits) |
|
889 thus ?thesis using p1' p2' p3' parent |
|
890 apply (simp add:Rmdir split:option.splits) |
|
891 using grant Rmdir p1 p2 p3 |
|
892 by simp |
|
893 qed |
|
894 ultimately show ?thesis using vs' |
|
895 by (erule_tac valid.intros(2), simp+) |
|
896 next |
|
897 case (Mkdir p f inum) |
461 from os obtain pf where pf_in_s: "is_dir s pf" and parent: "parent f = Some pf" |
898 from os obtain pf where pf_in_s: "is_dir s pf" and parent: "parent f = Some pf" |
462 by (auto simp:Open Some) |
899 by (auto simp:Mkdir) |
463 have pf_in: "is_dir s' pf" using pf_in_s alive |
900 have pf_in: "is_dir s' pf" using pf_in_s alive |
464 apply (erule_tac x = "O_dir pf" in allE) |
901 apply (erule_tac x = "O_dir pf" in allE) |
465 by simp |
902 by simp |
466 have p_in: "p \<in> current_procs s'" using os alive |
903 have p_in: "p \<in> current_procs s'" using os alive |
467 apply (erule_tac x = "O_proc p" in allE) |
904 apply (erule_tac x = "O_proc p" in allE) |
468 by (auto simp:Open Some) |
905 by (auto simp:Mkdir) |
469 have f_not_in: "f \<notin> current_files s'" using os alive' |
906 have f_not_in: "f \<notin> current_files s'" using os alive' Mkdir notin_all |
470 apply (erule_tac x = "E_file f" in allE) |
907 apply (erule_tac x = "E_file f" in allE) |
471 by (auto simp:Open Some) |
908 by (auto) |
472 have fd_not_in: "fd \<notin> current_proc_fds s' p" |
|
473 using os alive' p_in |
|
474 apply (erule_tac x = "E_fd p fd" in allE) |
|
475 by (simp add:Open Some) |
|
476 have inum_not_in: "inum \<notin> current_inode_nums s'" |
909 have inum_not_in: "inum \<notin> current_inode_nums s'" |
477 using os alive' |
910 using os alive' Mkdir notin_all |
478 apply (erule_tac x = "E_inum inum" in allE) |
911 apply (erule_tac x = "E_inum inum" in allE) |
479 by (simp add:Open Some) |
912 by (auto) |
480 have "os_grant s' e" using p_in pf_in parent f_not_in fd_not_in inum_not_in os |
913 have "os_grant s' e" using p_in pf_in parent f_not_in os inum_not_in |
481 by (simp add:Open Some hungs) |
914 by (simp add:Mkdir hungs) |
482 moreover have "grant s' e" |
915 moreover have "grant s' e" |
483 proof- |
916 proof- |
484 from grant parent obtain up rp tp uf rf tf |
917 from grant parent obtain up rp tp uf rf tf |
485 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
918 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
486 and p2: "sectxt_of_obj s (O_dir pf) = Some (uf, rf, tf)" |
919 and p2: "sectxt_of_obj s (O_dir pf) = Some (uf, rf, tf)" |
487 apply (simp add:Open Some split:option.splits) |
920 apply (simp add:Mkdir split:option.splits) |
488 by (case_tac a, case_tac aa, blast) |
921 by (case_tac a, case_tac aa, blast) |
489 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
922 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
490 using os cp2sp |
923 using os cp2sp |
491 apply (erule_tac x = p in allE) |
924 apply (erule_tac x = p in allE) |
492 by (auto simp:Open Some co2sobj.simps cp2sproc_def split:option.splits) |
925 by (auto simp:Mkdir co2sobj.simps cp2sproc_def split:option.splits) |
493 from vd os pf_in_s have "\<exists> sf. cf2sfile s pf = Some sf" |
926 from vd os pf_in_s have "\<exists> sf. cf2sfile s pf = Some sf" |
494 by (auto dest!:is_dir_in_current current_file_has_sfile simp:Open Some) |
927 by (auto dest!:is_dir_in_current current_file_has_sfile simp:Mkdir) |
495 hence p2': "sectxt_of_obj s' (O_dir pf) = Some (uf, rf, tf)" using p2 cf2sf pf_in pf_in_s |
928 hence p2': "sectxt_of_obj s' (O_dir pf) = Some (uf, rf, tf)" using p2 cf2sf pf_in pf_in_s |
496 apply (erule_tac x = pf in allE) |
929 apply (erule_tac x = pf in allE) |
497 apply (erule exE, frule_tac s = s in is_dir_in_current, simp) |
930 apply (erule exE, frule_tac s = s in is_dir_in_current, simp) |
498 apply (drule is_dir_not_file, drule is_dir_not_file) |
931 apply (drule is_dir_not_file, drule is_dir_not_file) |
499 apply (auto simp:cf2sfile_def split:option.splits) |
932 apply (auto simp:cf2sfile_def split:option.splits) |
500 apply (case_tac pf, simp_all) |
933 apply (case_tac pf, simp_all) |
501 by (simp add:sroot_def root_sec_remains vd vs') |
934 by (simp add:sroot_def root_sec_remains vd vs') |
502 have "search_check s' (up, rp, tp) pf" |
935 have "search_check s' (up, rp, tp) pf" |
503 using p1 p2 p2' vd cf2sf pf_in grant Open Some pf_in_s parent vs' |
936 using p1 p2 p2' vd cf2sf pf_in grant Mkdir pf_in_s parent vs' |
504 apply (rule_tac s = s in enrich_search_check') |
937 apply (rule_tac s = s in enrich_search_check') |
|
938 apply (simp_all split:option.splits) |
|
939 done |
|
940 thus ?thesis using p1' p2' parent |
|
941 apply (simp add:Mkdir split:option.splits) |
|
942 using grant Mkdir p1 p2 |
|
943 apply simp |
|
944 done |
|
945 qed |
|
946 ultimately show ?thesis using vs' |
|
947 by (erule_tac valid.intros(2), simp+) |
|
948 next |
|
949 case (LinkHard p f f') |
|
950 from os obtain pf where pf_in_s: "is_dir s pf" and parent: "parent f' = Some pf" |
|
951 by (auto simp:LinkHard) |
|
952 have pf_in: "is_dir s' pf" using pf_in_s alive |
|
953 apply (erule_tac x = "O_dir pf" in allE) |
|
954 by simp |
|
955 have p_in: "p \<in> current_procs s'" using os alive |
|
956 apply (erule_tac x = "O_proc p" in allE) |
|
957 by (auto simp:LinkHard) |
|
958 have f'_not_in: "f' \<notin> current_files s'" using os alive' LinkHard notin_all |
|
959 apply (erule_tac x = "E_file f'" in allE) |
|
960 by (auto simp:LinkHard) |
|
961 have f_in: "is_file s' f" using os alive |
|
962 apply (erule_tac x = "O_file f" in allE) |
|
963 by (auto simp:LinkHard) |
|
964 have "os_grant s' e" using p_in pf_in parent os f_in f'_not_in |
|
965 by (simp add:LinkHard hungs) |
|
966 moreover have "grant s' e" |
|
967 proof- |
|
968 from grant parent obtain up rp tp uf rf tf upf rpf tpf |
|
969 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
970 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
971 and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" |
|
972 apply (simp add:LinkHard split:option.splits) |
|
973 by (case_tac a, case_tac aa, case_tac ab, blast) |
|
974 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
975 using os cp2sp |
|
976 apply (erule_tac x = p in allE) |
|
977 by (auto simp:LinkHard co2sobj.simps cp2sproc_def split:option.splits) |
|
978 from vd os pf_in_s have "\<exists> sf. cf2sfile s f = Some sf" |
|
979 by (auto dest!:is_file_in_current current_file_has_sfile simp:LinkHard) |
|
980 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" |
|
981 using p2 cf2sf f_in os parent |
|
982 apply (erule_tac x = f in allE) |
|
983 apply (erule exE, clarsimp simp:LinkHard) |
|
984 apply (frule_tac s = s in is_file_in_current, simp) |
|
985 apply (auto simp:cf2sfile_def split:option.splits) |
|
986 apply (case_tac f, simp) |
|
987 by (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
|
988 from vd os pf_in_s have "\<exists> sf. cf2sfile s pf = Some sf" |
|
989 by (auto dest!:is_dir_in_current current_file_has_sfile simp:LinkHard) |
|
990 hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s |
|
991 apply (erule_tac x = pf in allE) |
|
992 apply (erule exE, frule_tac s = s in is_dir_in_current, simp) |
|
993 apply (drule is_dir_not_file, drule is_dir_not_file) |
|
994 apply (auto simp:cf2sfile_def split:option.splits) |
|
995 apply (case_tac pf, simp_all) |
|
996 by (simp add:sroot_def root_sec_remains vd vs') |
|
997 have "search_check s' (up, rp, tp) f" |
|
998 using p1 p2 p2' vd cf2sf f_in grant LinkHard os parent vs' |
|
999 apply (rule_tac s = s in enrich_search_check) |
505 by (simp_all split:option.splits) |
1000 by (simp_all split:option.splits) |
506 thus ?thesis using p1' p2' parent |
1001 moreover have "search_check s' (up, rp, tp) pf" |
507 apply (simp add:Open Some split:option.splits) |
1002 using p1 p3 p3' vd cf2sf pf_in grant LinkHard os parent vs' |
508 using grant Open Some p1 p2 |
1003 apply (rule_tac s = s in enrich_search_check') |
|
1004 apply (simp_all split:option.splits) |
|
1005 done |
|
1006 ultimately show ?thesis using p1' p2' p3' parent |
|
1007 apply (simp add:LinkHard split:option.splits) |
|
1008 using grant LinkHard p1 p2 p3 |
|
1009 apply simp |
|
1010 done |
|
1011 qed |
|
1012 ultimately show ?thesis using vs' |
|
1013 by (erule_tac valid.intros(2), simp+) |
|
1014 next |
|
1015 case (Truncate p f len) |
|
1016 have p_in: "p \<in> current_procs s'" using os alive |
|
1017 apply (erule_tac x = "O_proc p" in allE) |
|
1018 by (auto simp:Truncate) |
|
1019 have f_in: "is_file s' f" using os alive |
|
1020 apply (erule_tac x = "O_file f" in allE) |
|
1021 by (auto simp:Truncate) |
|
1022 have "os_grant s' e" using p_in f_in by (simp add:Truncate) |
|
1023 moreover have "grant s' e" |
|
1024 proof- |
|
1025 from grant obtain up rp tp uf rf tf |
|
1026 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
1027 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
1028 apply (simp add:Truncate split:option.splits) |
|
1029 by (case_tac a, case_tac aa, blast) |
|
1030 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
1031 using os cp2sp |
|
1032 apply (erule_tac x = p in allE) |
|
1033 by (auto simp:Truncate co2sobj.simps cp2sproc_def split:option.splits) |
|
1034 from os have f_in': "is_file s f" by (simp add:Truncate) |
|
1035 from vd os have "\<exists> sf. cf2sfile s f = Some sf" |
|
1036 by (auto dest!:is_file_in_current current_file_has_sfile simp:Truncate) |
|
1037 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf |
|
1038 apply (erule_tac x = f in allE) |
|
1039 apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) |
|
1040 apply (case_tac f, simp) |
|
1041 apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
|
1042 done |
|
1043 have "search_check s' (up, rp, tp) f" |
|
1044 using p1 p2 p2' vd cf2sf f_in' grant Truncate f_in |
|
1045 apply (rule_tac s = s in enrich_search_check) |
|
1046 by (simp_all split:option.splits) |
|
1047 thus ?thesis using p1' p2' |
|
1048 apply (simp add:Truncate split:option.splits) |
|
1049 using grant Truncate p1 p2 |
|
1050 by (simp add:Truncate grant p1 p2) |
|
1051 qed |
|
1052 ultimately show ?thesis using vs' |
|
1053 by (erule_tac valid.intros(2), simp+) |
|
1054 next |
|
1055 case (CreateMsgq p q) |
|
1056 have p_in: "p \<in> current_procs s'" using os alive |
|
1057 apply (erule_tac x = "O_proc p" in allE) |
|
1058 by (auto simp:CreateMsgq) |
|
1059 have q_not_in: "q \<notin> current_msgqs s'" using os alive' CreateMsgq notin_all |
|
1060 apply (erule_tac x = "E_msgq q" in allE) |
|
1061 by auto |
|
1062 have "os_grant s' e" using p_in q_not_in by (simp add:CreateMsgq) |
|
1063 moreover have "grant s' e" |
|
1064 proof- |
|
1065 from grant obtain up rp tp |
|
1066 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
1067 apply (simp add:CreateMsgq split:option.splits) |
|
1068 by (case_tac a, blast) |
|
1069 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
1070 using os cp2sp |
|
1071 apply (erule_tac x = p in allE) |
|
1072 by (auto simp:CreateMsgq co2sobj.simps cp2sproc_def split:option.splits) |
|
1073 show ?thesis using p1' |
|
1074 apply (simp add:CreateMsgq split:option.splits) |
|
1075 using grant CreateMsgq p1 |
509 by simp |
1076 by simp |
510 qed |
1077 qed |
511 ultimately show ?thesis using vs' |
1078 ultimately show ?thesis using vs' |
512 by (erule_tac valid.intros(2), simp+) |
1079 by (erule_tac valid.intros(2), simp+) |
513 qed |
1080 next |
514 next |
1081 case (RemoveMsgq p q) |
515 case (ReadFile p fd) |
1082 have p_in: "p \<in> current_procs s'" using os alive |
516 have p_in: "p \<in> current_procs s'" using os alive |
1083 apply (erule_tac x = "O_proc p" in allE) |
517 apply (erule_tac x = "O_proc p" in allE) |
1084 by (auto simp:RemoveMsgq) |
518 by (auto simp:ReadFile) |
1085 have q_in: "q \<in> current_msgqs s'" using os alive |
519 have fd_in: "fd \<in> current_proc_fds s' p" using os alive |
1086 apply (erule_tac x = "O_msgq q" in allE) |
520 apply (erule_tac x = "O_fd p fd" in allE) |
|
521 by (auto simp:ReadFile) |
|
522 obtain f where ffd: "file_of_proc_fd s p fd = Some f" |
|
523 using os ReadFile by auto |
|
524 hence f_in_s: "is_file s f" using vd |
|
525 by (auto intro:file_of_pfd_is_file) |
|
526 obtain flags where fflag: "flags_of_proc_fd s p fd = Some flags" |
|
527 using os ReadFile by auto |
|
528 have ffd_in: "file_of_proc_fd s' p fd = Some f" |
|
529 using ffd_remain ffd by auto |
|
530 hence f_in: "is_file s' f" using vs' |
|
531 by (auto intro:file_of_pfd_is_file) |
|
532 have flags_in: "flags_of_proc_fd s' p fd = Some flags" |
|
533 using fflags_remain fflag by auto |
|
534 have "os_grant s' e" using p_in fd_in ffd_in flags_in fflag os f_in |
|
535 by (auto simp add:ReadFile is_file_in_current) |
|
536 moreover have "grant s' e" |
|
537 proof- |
|
538 from grant ffd obtain up rp tp uf rf tf ufd rfd tfd |
|
539 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
540 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
541 and p3: "sectxt_of_obj s (O_fd p fd) = Some (ufd, rfd, tfd)" |
|
542 apply (simp add:ReadFile split:option.splits) |
|
543 by (case_tac a, case_tac aa, case_tac ab, blast) |
|
544 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
545 using os cp2sp |
|
546 apply (erule_tac x = p in allE) |
|
547 by (auto simp:ReadFile co2sobj.simps cp2sproc_def split:option.splits) |
|
548 from vd f_in_s have "\<exists> sf. cf2sfile s f = Some sf" |
|
549 by (auto dest!:is_file_in_current current_file_has_sfile) |
|
550 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in_s p2 cf2sf |
|
551 apply (erule_tac x = f in allE) |
|
552 apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) |
|
553 apply (case_tac f, simp) |
|
554 apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
|
555 done |
|
556 have p3': "sectxt_of_obj s' (O_fd p fd) = Some (ufd, rfd, tfd)" |
|
557 using cfd2sfd ffd_in ffd p3 f_in f_in_s vd |
|
558 apply (erule_tac x = p in allE) |
|
559 apply (erule_tac x = fd in allE) |
|
560 apply (simp add:proc_file_fds_def) |
|
561 apply (auto simp:cfd2sfd_def fflag flags_in p3 split:option.splits |
|
562 dest!:current_file_has_sfile' simp:is_file_in_current) |
|
563 done |
|
564 show ?thesis using p1' p2' p3' ffd_in ffd |
|
565 apply (simp add:ReadFile split:option.splits) |
|
566 using grant p1 p2 p3 ReadFile |
|
567 by simp |
|
568 qed |
|
569 ultimately show ?thesis using vs' |
|
570 by (erule_tac valid.intros(2), simp+) |
|
571 next |
|
572 case (WriteFile p fd) |
|
573 have p_in: "p \<in> current_procs s'" using os alive |
|
574 apply (erule_tac x = "O_proc p" in allE) |
|
575 by (auto simp:WriteFile) |
|
576 have fd_in: "fd \<in> current_proc_fds s' p" using os alive |
|
577 apply (erule_tac x = "O_fd p fd" in allE) |
|
578 by (auto simp:WriteFile) |
|
579 obtain f where ffd: "file_of_proc_fd s p fd = Some f" |
|
580 using os WriteFile by auto |
|
581 hence f_in_s: "is_file s f" using vd |
|
582 by (auto intro:file_of_pfd_is_file) |
|
583 obtain flags where fflag: "flags_of_proc_fd s p fd = Some flags" |
|
584 using os WriteFile by auto |
|
585 have ffd_in: "file_of_proc_fd s' p fd = Some f" |
|
586 using ffd_remain ffd by auto |
|
587 hence f_in: "is_file s' f" using vs' |
|
588 by (auto intro:file_of_pfd_is_file) |
|
589 have flags_in: "flags_of_proc_fd s' p fd = Some flags" |
|
590 using fflags_remain fflag by auto |
|
591 have "os_grant s' e" using p_in fd_in ffd_in flags_in fflag os f_in |
|
592 by (auto simp add:WriteFile is_file_in_current) |
|
593 moreover have "grant s' e" |
|
594 proof- |
|
595 from grant ffd obtain up rp tp uf rf tf ufd rfd tfd |
|
596 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
597 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
598 and p3: "sectxt_of_obj s (O_fd p fd) = Some (ufd, rfd, tfd)" |
|
599 apply (simp add:WriteFile split:option.splits) |
|
600 by (case_tac a, case_tac aa, case_tac ab, blast) |
|
601 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
602 using os cp2sp |
|
603 apply (erule_tac x = p in allE) |
|
604 by (auto simp:WriteFile co2sobj.simps cp2sproc_def split:option.splits) |
|
605 from vd f_in_s have "\<exists> sf. cf2sfile s f = Some sf" |
|
606 by (auto dest!:is_file_in_current current_file_has_sfile) |
|
607 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in_s p2 cf2sf |
|
608 apply (erule_tac x = f in allE) |
|
609 apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) |
|
610 apply (case_tac f, simp) |
|
611 apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
|
612 done |
|
613 have p3': "sectxt_of_obj s' (O_fd p fd) = Some (ufd, rfd, tfd)" |
|
614 using cfd2sfd ffd_in ffd p3 f_in f_in_s vd |
|
615 apply (erule_tac x = p in allE) |
|
616 apply (erule_tac x = fd in allE) |
|
617 apply (simp add:proc_file_fds_def) |
|
618 apply (auto simp:cfd2sfd_def fflag flags_in p3 split:option.splits |
|
619 dest!:current_file_has_sfile' simp:is_file_in_current) |
|
620 done |
|
621 show ?thesis using p1' p2' p3' ffd_in ffd |
|
622 apply (simp add:WriteFile split:option.splits) |
|
623 using grant p1 p2 p3 WriteFile |
|
624 by simp |
|
625 qed |
|
626 ultimately show ?thesis using vs' |
|
627 by (erule_tac valid.intros(2), simp+) |
|
628 next |
|
629 case (CloseFd p fd) |
|
630 have p_in: "p \<in> current_procs s'" using os alive |
|
631 apply (erule_tac x = "O_proc p" in allE) |
|
632 by (auto simp:CloseFd) |
|
633 have fd_in: "fd \<in> current_proc_fds s' p" using os alive |
|
634 apply (erule_tac x = "O_fd p fd" in allE) |
|
635 by (auto simp:CloseFd) |
|
636 have "os_grant s' e" using p_in fd_in |
|
637 by (auto simp add:CloseFd) |
|
638 moreover have "grant s' e" |
|
639 by(simp add:CloseFd) |
|
640 ultimately show ?thesis using vs' |
|
641 by (erule_tac valid.intros(2), simp+) |
|
642 next |
|
643 case (UnLink p f) |
|
644 have p_in: "p \<in> current_procs s'" using os alive |
|
645 apply (erule_tac x = "O_proc p" in allE) |
|
646 by (auto simp:UnLink) |
|
647 have f_in: "is_file s' f" using os alive |
|
648 apply (erule_tac x = "O_file f" in allE) |
|
649 by (auto simp:UnLink) |
|
650 from os vd obtain pf where pf_in_s: "is_dir s pf" |
|
651 and parent: "parent f = Some pf" |
|
652 by (auto simp:UnLink dest!:file_has_parent) |
|
653 from pf_in_s alive have pf_in: "is_dir s' pf" |
|
654 apply (erule_tac x = "O_dir pf" in allE) |
|
655 by (auto simp:UnLink) |
|
656 have "os_grant s' e" using p_in f_in os |
|
657 by (simp add:UnLink hungs) |
|
658 moreover have "grant s' e" |
|
659 proof- |
|
660 from grant parent obtain up rp tp uf rf tf upf rpf tpf |
|
661 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
662 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
663 and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" |
|
664 apply (simp add:UnLink split:option.splits) |
|
665 by (case_tac a, case_tac aa, case_tac ab, blast) |
|
666 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
667 using os cp2sp |
|
668 apply (erule_tac x = p in allE) |
|
669 by (auto simp:UnLink co2sobj.simps cp2sproc_def split:option.splits) |
|
670 from vd os pf_in_s have "\<exists> sf. cf2sfile s f = Some sf" |
|
671 by (auto dest!:is_file_in_current current_file_has_sfile simp:UnLink) |
|
672 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" |
|
673 using p2 cf2sf f_in os parent |
|
674 apply (erule_tac x = f in allE) |
|
675 apply (erule exE, clarsimp simp:UnLink) |
|
676 apply (frule_tac s = s in is_file_in_current, simp) |
|
677 by (auto simp:cf2sfile_def split:option.splits) |
|
678 from vd os pf_in_s have "\<exists> sf. cf2sfile s pf = Some sf" |
|
679 by (auto dest!:is_dir_in_current current_file_has_sfile simp:UnLink) |
|
680 hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s |
|
681 apply (erule_tac x = pf in allE) |
|
682 apply (erule exE, frule_tac s = s in is_dir_in_current, simp) |
|
683 apply (drule is_dir_not_file, drule is_dir_not_file) |
|
684 apply (auto simp:cf2sfile_def split:option.splits) |
|
685 apply (case_tac pf, simp_all) |
|
686 by (simp add:sroot_def root_sec_remains vd vs') |
|
687 have "search_check s' (up, rp, tp) f" |
|
688 using p1 p2 p2' vd cf2sf f_in grant UnLink os parent vs' |
|
689 apply (rule_tac s = s in enrich_search_check) |
|
690 by (simp_all split:option.splits) |
|
691 thus ?thesis using p1' p2' p3' parent |
|
692 apply (simp add:UnLink split:option.splits) |
|
693 using grant UnLink p1 p2 p3 |
|
694 by simp |
|
695 qed |
|
696 ultimately show ?thesis using vs' |
|
697 by (erule_tac valid.intros(2), simp+) |
|
698 next |
|
699 case (Rmdir p f) |
|
700 have p_in: "p \<in> current_procs s'" using os alive |
|
701 apply (erule_tac x = "O_proc p" in allE) |
|
702 by (auto simp:Rmdir) |
|
703 have f_in: "is_dir s' f" using os alive |
|
704 apply (erule_tac x = "O_dir f" in allE) |
|
705 by (auto simp:Rmdir dir_is_empty_def) |
|
706 have not_root: "f \<noteq> []" using os |
|
707 by (auto simp:Rmdir) |
|
708 from os vd obtain pf where pf_in_s: "is_dir s pf" |
|
709 and parent: "parent f = Some pf" |
|
710 apply (auto simp:Rmdir dir_is_empty_def) |
|
711 apply (case_tac f, simp+) |
|
712 apply (drule parentf_is_dir_prop1, auto) |
|
713 done |
|
714 from pf_in_s alive have pf_in: "is_dir s' pf" |
|
715 apply (erule_tac x = "O_dir pf" in allE) |
|
716 by (auto simp:Rmdir) |
|
717 have empty_in: "dir_is_empty s' f" using os |
|
718 apply (simp add:dir_is_empty_def f_in) |
|
719 apply auto using alive' |
|
720 apply (erule_tac x = "E_file f'" in allE) |
|
721 by (simp add:Rmdir dir_is_empty_def) |
|
722 have "os_grant s' e" using p_in f_in os empty_in |
|
723 by (simp add:Rmdir hungs) |
|
724 moreover have "grant s' e" |
|
725 proof- |
|
726 from grant parent obtain up rp tp uf rf tf upf rpf tpf |
|
727 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
728 and p2: "sectxt_of_obj s (O_dir f) = Some (uf, rf, tf)" |
|
729 and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" |
|
730 apply (simp add:Rmdir split:option.splits) |
|
731 by (case_tac a, case_tac aa, case_tac ab, blast) |
|
732 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
733 using os cp2sp |
|
734 apply (erule_tac x = p in allE) |
|
735 by (auto simp:Rmdir co2sobj.simps cp2sproc_def split:option.splits) |
|
736 from vd os pf_in_s have "\<exists> sf. cf2sfile s f = Some sf" |
|
737 by (auto dest!:is_dir_in_current current_file_has_sfile simp:dir_is_empty_def Rmdir) |
|
738 hence p2': "sectxt_of_obj s' (O_dir f) = Some (uf, rf, tf)" |
|
739 using p2 cf2sf f_in os parent |
|
740 apply (erule_tac x = f in allE) |
|
741 apply (erule exE, clarsimp simp:Rmdir dir_is_empty_def) |
|
742 apply (frule_tac s = s in is_dir_in_current, simp) |
|
743 apply (drule is_dir_not_file, drule is_dir_not_file) |
|
744 by (auto simp:cf2sfile_def split:option.splits) |
|
745 from vd os pf_in_s have "\<exists> sf. cf2sfile s pf = Some sf" |
|
746 by (auto dest!:is_dir_in_current current_file_has_sfile simp:Rmdir) |
|
747 hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s |
|
748 apply (erule_tac x = pf in allE) |
|
749 apply (erule exE, frule_tac s = s in is_dir_in_current, simp) |
|
750 apply (drule is_dir_not_file, drule is_dir_not_file) |
|
751 apply (auto simp:cf2sfile_def split:option.splits) |
|
752 apply (case_tac pf, simp_all) |
|
753 by (simp add:sroot_def root_sec_remains vd vs') |
|
754 have "search_check s' (up, rp, tp) f" |
|
755 using p1 p2 p2' vd cf2sf f_in grant Rmdir os parent vs' |
|
756 apply (rule_tac s = s in enrich_search_check') |
|
757 by (simp_all add:dir_is_empty_def split:option.splits) |
|
758 thus ?thesis using p1' p2' p3' parent |
|
759 apply (simp add:Rmdir split:option.splits) |
|
760 using grant Rmdir p1 p2 p3 |
|
761 by simp |
|
762 qed |
|
763 ultimately show ?thesis using vs' |
|
764 by (erule_tac valid.intros(2), simp+) |
|
765 next |
|
766 case (Mkdir p f inum) |
|
767 from os obtain pf where pf_in_s: "is_dir s pf" and parent: "parent f = Some pf" |
|
768 by (auto simp:Mkdir) |
|
769 have pf_in: "is_dir s' pf" using pf_in_s alive |
|
770 apply (erule_tac x = "O_dir pf" in allE) |
|
771 by simp |
|
772 have p_in: "p \<in> current_procs s'" using os alive |
|
773 apply (erule_tac x = "O_proc p" in allE) |
|
774 by (auto simp:Mkdir) |
|
775 have f_not_in: "f \<notin> current_files s'" using os alive' |
|
776 apply (erule_tac x = "E_file f" in allE) |
|
777 by (auto simp:Mkdir) |
|
778 have inum_not_in: "inum \<notin> current_inode_nums s'" |
|
779 using os alive' |
|
780 apply (erule_tac x = "E_inum inum" in allE) |
|
781 by (simp add:Mkdir) |
|
782 have "os_grant s' e" using p_in pf_in parent f_not_in os inum_not_in |
|
783 by (simp add:Mkdir hungs) |
|
784 moreover have "grant s' e" |
|
785 proof- |
|
786 from grant parent obtain up rp tp uf rf tf |
|
787 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
788 and p2: "sectxt_of_obj s (O_dir pf) = Some (uf, rf, tf)" |
|
789 apply (simp add:Mkdir split:option.splits) |
|
790 by (case_tac a, case_tac aa, blast) |
|
791 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
792 using os cp2sp |
|
793 apply (erule_tac x = p in allE) |
|
794 by (auto simp:Mkdir co2sobj.simps cp2sproc_def split:option.splits) |
|
795 from vd os pf_in_s have "\<exists> sf. cf2sfile s pf = Some sf" |
|
796 by (auto dest!:is_dir_in_current current_file_has_sfile simp:Mkdir) |
|
797 hence p2': "sectxt_of_obj s' (O_dir pf) = Some (uf, rf, tf)" using p2 cf2sf pf_in pf_in_s |
|
798 apply (erule_tac x = pf in allE) |
|
799 apply (erule exE, frule_tac s = s in is_dir_in_current, simp) |
|
800 apply (drule is_dir_not_file, drule is_dir_not_file) |
|
801 apply (auto simp:cf2sfile_def split:option.splits) |
|
802 apply (case_tac pf, simp_all) |
|
803 by (simp add:sroot_def root_sec_remains vd vs') |
|
804 have "search_check s' (up, rp, tp) pf" |
|
805 using p1 p2 p2' vd cf2sf pf_in grant Mkdir pf_in_s parent vs' |
|
806 apply (rule_tac s = s in enrich_search_check') |
|
807 apply (simp_all split:option.splits) |
|
808 done |
|
809 thus ?thesis using p1' p2' parent |
|
810 apply (simp add:Mkdir split:option.splits) |
|
811 using grant Mkdir p1 p2 |
|
812 apply simp |
|
813 done |
|
814 qed |
|
815 ultimately show ?thesis using vs' |
|
816 by (erule_tac valid.intros(2), simp+) |
|
817 next |
|
818 case (LinkHard p f f') |
|
819 from os obtain pf where pf_in_s: "is_dir s pf" and parent: "parent f' = Some pf" |
|
820 by (auto simp:LinkHard) |
|
821 have pf_in: "is_dir s' pf" using pf_in_s alive |
|
822 apply (erule_tac x = "O_dir pf" in allE) |
|
823 by simp |
|
824 have p_in: "p \<in> current_procs s'" using os alive |
|
825 apply (erule_tac x = "O_proc p" in allE) |
|
826 by (auto simp:LinkHard) |
|
827 have f'_not_in: "f' \<notin> current_files s'" using os alive' |
|
828 apply (erule_tac x = "E_file f'" in allE) |
|
829 by (auto simp:LinkHard) |
|
830 have f_in: "is_file s' f" using os alive |
|
831 apply (erule_tac x = "O_file f" in allE) |
|
832 by (auto simp:LinkHard) |
|
833 have "os_grant s' e" using p_in pf_in parent os f_in f'_not_in |
|
834 by (simp add:LinkHard hungs) |
|
835 moreover have "grant s' e" |
|
836 proof- |
|
837 from grant parent obtain up rp tp uf rf tf upf rpf tpf |
|
838 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
839 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
840 and p3: "sectxt_of_obj s (O_dir pf) = Some (upf, rpf, tpf)" |
|
841 apply (simp add:LinkHard split:option.splits) |
|
842 by (case_tac a, case_tac aa, case_tac ab, blast) |
|
843 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
844 using os cp2sp |
|
845 apply (erule_tac x = p in allE) |
|
846 by (auto simp:LinkHard co2sobj.simps cp2sproc_def split:option.splits) |
|
847 from vd os pf_in_s have "\<exists> sf. cf2sfile s f = Some sf" |
|
848 by (auto dest!:is_file_in_current current_file_has_sfile simp:LinkHard) |
|
849 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" |
|
850 using p2 cf2sf f_in os parent |
|
851 apply (erule_tac x = f in allE) |
|
852 apply (erule exE, clarsimp simp:LinkHard) |
|
853 apply (frule_tac s = s in is_file_in_current, simp) |
|
854 apply (auto simp:cf2sfile_def split:option.splits) |
|
855 apply (case_tac f, simp) |
|
856 by (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
|
857 from vd os pf_in_s have "\<exists> sf. cf2sfile s pf = Some sf" |
|
858 by (auto dest!:is_dir_in_current current_file_has_sfile simp:LinkHard) |
|
859 hence p3': "sectxt_of_obj s' (O_dir pf) = Some (upf, rpf, tpf)" using p3 cf2sf pf_in pf_in_s |
|
860 apply (erule_tac x = pf in allE) |
|
861 apply (erule exE, frule_tac s = s in is_dir_in_current, simp) |
|
862 apply (drule is_dir_not_file, drule is_dir_not_file) |
|
863 apply (auto simp:cf2sfile_def split:option.splits) |
|
864 apply (case_tac pf, simp_all) |
|
865 by (simp add:sroot_def root_sec_remains vd vs') |
|
866 have "search_check s' (up, rp, tp) f" |
|
867 using p1 p2 p2' vd cf2sf f_in grant LinkHard os parent vs' |
|
868 apply (rule_tac s = s in enrich_search_check) |
|
869 by (simp_all split:option.splits) |
|
870 moreover have "search_check s' (up, rp, tp) pf" |
|
871 using p1 p3 p3' vd cf2sf pf_in grant LinkHard os parent vs' |
|
872 apply (rule_tac s = s in enrich_search_check') |
|
873 apply (simp_all split:option.splits) |
|
874 done |
|
875 ultimately show ?thesis using p1' p2' p3' parent |
|
876 apply (simp add:LinkHard split:option.splits) |
|
877 using grant LinkHard p1 p2 p3 |
|
878 apply simp |
|
879 done |
|
880 qed |
|
881 ultimately show ?thesis using vs' |
|
882 by (erule_tac valid.intros(2), simp+) |
|
883 next |
|
884 case (Truncate p f len) |
|
885 have p_in: "p \<in> current_procs s'" using os alive |
|
886 apply (erule_tac x = "O_proc p" in allE) |
|
887 by (auto simp:Truncate) |
|
888 have f_in: "is_file s' f" using os alive |
|
889 apply (erule_tac x = "O_file f" in allE) |
|
890 by (auto simp:Truncate) |
|
891 have "os_grant s' e" using p_in f_in by (simp add:Truncate) |
|
892 moreover have "grant s' e" |
|
893 proof- |
|
894 from grant obtain up rp tp uf rf tf |
|
895 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
896 and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" |
|
897 apply (simp add:Truncate split:option.splits) |
|
898 by (case_tac a, case_tac aa, blast) |
|
899 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
900 using os cp2sp |
|
901 apply (erule_tac x = p in allE) |
|
902 by (auto simp:Truncate co2sobj.simps cp2sproc_def split:option.splits) |
|
903 from os have f_in': "is_file s f" by (simp add:Truncate) |
|
904 from vd os have "\<exists> sf. cf2sfile s f = Some sf" |
|
905 by (auto dest!:is_file_in_current current_file_has_sfile simp:Truncate) |
|
906 hence p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" using f_in f_in' p2 cf2sf |
|
907 apply (erule_tac x = f in allE) |
|
908 apply (auto dest:is_file_in_current simp:cf2sfile_def split:option.splits) |
|
909 apply (case_tac f, simp) |
|
910 apply (drule_tac s = s in root_is_dir', simp add:vd, simp+) |
|
911 done |
|
912 have "search_check s' (up, rp, tp) f" |
|
913 using p1 p2 p2' vd cf2sf f_in' grant Truncate f_in |
|
914 apply (rule_tac s = s in enrich_search_check) |
|
915 by (simp_all split:option.splits) |
|
916 thus ?thesis using p1' p2' |
|
917 apply (simp add:Truncate split:option.splits) |
|
918 using grant Truncate p1 p2 |
|
919 by (simp add:Truncate grant p1 p2) |
|
920 qed |
|
921 ultimately show ?thesis using vs' |
|
922 by (erule_tac valid.intros(2), simp+) |
|
923 next |
|
924 case (CreateMsgq p q) |
|
925 have p_in: "p \<in> current_procs s'" using os alive |
|
926 apply (erule_tac x = "O_proc p" in allE) |
|
927 by (auto simp:CreateMsgq) |
|
928 have q_not_in: "q \<notin> current_msgqs s'" using os alive' |
|
929 apply (erule_tac x = "E_msgq q" in allE) |
|
930 by (simp add:CreateMsgq) |
|
931 have "os_grant s' e" using p_in q_not_in by (simp add:CreateMsgq) |
|
932 moreover have "grant s' e" |
|
933 proof- |
|
934 from grant obtain up rp tp |
|
935 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
936 apply (simp add:CreateMsgq split:option.splits) |
|
937 by (case_tac a, blast) |
|
938 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
939 using os cp2sp |
|
940 apply (erule_tac x = p in allE) |
|
941 by (auto simp:CreateMsgq co2sobj.simps cp2sproc_def split:option.splits) |
|
942 show ?thesis using p1' |
|
943 apply (simp add:CreateMsgq split:option.splits) |
|
944 using grant CreateMsgq p1 |
|
945 by simp |
|
946 qed |
|
947 ultimately show ?thesis using vs' |
|
948 by (erule_tac valid.intros(2), simp+) |
|
949 next |
|
950 case (RemoveMsgq p q) |
|
951 have p_in: "p \<in> current_procs s'" using os alive |
|
952 apply (erule_tac x = "O_proc p" in allE) |
|
953 by (auto simp:RemoveMsgq) |
|
954 have q_in: "q \<in> current_msgqs s'" using os alive |
|
955 apply (erule_tac x = "O_msgq q" in allE) |
|
956 by (simp add:RemoveMsgq) |
|
957 have "os_grant s' e" using p_in q_in by (simp add:RemoveMsgq) |
|
958 moreover have "grant s' e" |
|
959 proof- |
|
960 from grant obtain up rp tp uq rq tq |
|
961 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
|
962 and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" |
|
963 apply (simp add:RemoveMsgq split:option.splits) |
|
964 by (case_tac a, case_tac aa, blast) |
|
965 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
|
966 using os cp2sp |
|
967 apply (erule_tac x = p in allE) |
|
968 by (auto simp:RemoveMsgq co2sobj.simps cp2sproc_def split:option.splits) |
|
969 from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" |
|
970 using os cq2sq vd |
|
971 apply (erule_tac x = q in allE) |
|
972 by (auto simp:RemoveMsgq co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) |
|
973 show ?thesis using p1' p2' grant p1 p2 |
|
974 by (simp add:RemoveMsgq) |
1087 by (simp add:RemoveMsgq) |
975 qed |
1088 have "os_grant s' e" using p_in q_in by (simp add:RemoveMsgq) |
976 ultimately show ?thesis using vs' |
1089 moreover have "grant s' e" |
977 by (erule_tac valid.intros(2), simp+) |
1090 proof- |
978 next |
1091 from grant obtain up rp tp uq rq tq |
979 case (SendMsg p q m) |
1092 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
980 have p_in: "p \<in> current_procs s'" using os alive |
1093 and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" |
981 apply (erule_tac x = "O_proc p" in allE) |
1094 apply (simp add:RemoveMsgq split:option.splits) |
982 by (auto simp:SendMsg) |
1095 by (case_tac a, case_tac aa, blast) |
983 have q_in: "q \<in> current_msgqs s'" using os alive |
1096 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
984 apply (erule_tac x = "O_msgq q" in allE) |
1097 using os cp2sp |
985 by (simp add:SendMsg) |
1098 apply (erule_tac x = p in allE) |
986 have m_not_in: "m \<notin> set (msgs_of_queue s' q)" using os alive' |
1099 by (auto simp:RemoveMsgq co2sobj.simps cp2sproc_def split:option.splits) |
987 apply (erule_tac x = "E_msg q m" in allE) |
1100 from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" |
988 by (simp add:SendMsg q_in) |
1101 using os cq2sq vd |
989 have "os_grant s' e" using p_in q_in m_not_in |
1102 apply (erule_tac x = q in allE) |
990 by (simp add:SendMsg) |
1103 by (auto simp:RemoveMsgq co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) |
991 moreover have "grant s' e" |
1104 show ?thesis using p1' p2' grant p1 p2 |
992 proof- |
1105 by (simp add:RemoveMsgq) |
993 from grant obtain up rp tp uq rq tq |
1106 qed |
994 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
1107 ultimately show ?thesis using vs' |
995 and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" |
1108 by (erule_tac valid.intros(2), simp+) |
996 apply (simp add:SendMsg split:option.splits) |
1109 next |
997 by (case_tac a, case_tac aa, blast) |
1110 case (SendMsg p q m) |
998 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
1111 have p_in: "p \<in> current_procs s'" using os alive |
999 using os cp2sp |
1112 apply (erule_tac x = "O_proc p" in allE) |
1000 apply (erule_tac x = p in allE) |
1113 by (auto simp:SendMsg) |
1001 by (auto simp:SendMsg co2sobj.simps cp2sproc_def split:option.splits) |
1114 have q_in: "q \<in> current_msgqs s'" using os alive |
1002 from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" |
1115 apply (erule_tac x = "O_msgq q" in allE) |
1003 using os cq2sq vd |
|
1004 apply (erule_tac x = q in allE) |
|
1005 by (auto simp:SendMsg co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) |
|
1006 show ?thesis using p1' p2' grant p1 p2 |
|
1007 by (simp add:SendMsg) |
1116 by (simp add:SendMsg) |
1008 qed |
1117 have m_not_in: "m \<notin> set (msgs_of_queue s' q)" using os alive' notin_all SendMsg q_in |
1009 ultimately show ?thesis using vs' |
1118 apply (erule_tac x = "E_msg q m" in allE) |
1010 by (erule_tac valid.intros(2), simp+) |
1119 apply (case_tac obj') |
1011 next |
1120 by auto |
1012 case (RecvMsg p q m) |
1121 have "os_grant s' e" using p_in q_in m_not_in |
1013 have p_in: "p \<in> current_procs s'" using os alive |
1122 by (simp add:SendMsg) |
1014 apply (erule_tac x = "O_proc p" in allE) |
1123 moreover have "grant s' e" |
1015 by (auto simp:RecvMsg) |
1124 proof- |
1016 have q_in: "q \<in> current_msgqs s'" using os alive |
1125 from grant obtain up rp tp uq rq tq |
1017 apply (erule_tac x = "O_msgq q" in allE) |
1126 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
1018 by (simp add:RecvMsg) |
1127 and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" |
1019 have m_in: "m = hd (msgs_of_queue s' q)" |
1128 apply (simp add:SendMsg split:option.splits) |
1020 and sms_not_empty: "msgs_of_queue s' q \<noteq> []" |
1129 by (case_tac a, case_tac aa, blast) |
1021 using os sms_remain |
1130 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
1022 by (auto simp:RecvMsg) |
1131 using os cp2sp |
1023 have "os_grant s' e" using p_in q_in m_in sms_not_empty os |
1132 apply (erule_tac x = p in allE) |
1024 by (simp add:RecvMsg) |
1133 by (auto simp:SendMsg co2sobj.simps cp2sproc_def split:option.splits) |
1025 moreover have "grant s' e" |
1134 from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" |
1026 proof- |
1135 using os cq2sq vd |
1027 from grant obtain up rp tp uq rq tq um rm tm |
1136 apply (erule_tac x = q in allE) |
1028 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
1137 by (auto simp:SendMsg co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) |
1029 and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" |
1138 show ?thesis using p1' p2' grant p1 p2 |
1030 and p3: "sectxt_of_obj s (O_msg q m) = Some (um, rm, tm)" |
1139 by (simp add:SendMsg) |
1031 apply (simp add:RecvMsg split:option.splits) |
1140 qed |
1032 by (case_tac a, case_tac aa, case_tac ab, blast) |
1141 ultimately show ?thesis using vs' |
1033 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
1142 by (erule_tac valid.intros(2), simp+) |
1034 using os cp2sp |
1143 next |
1035 apply (erule_tac x = p in allE) |
1144 case (RecvMsg p q m) |
1036 by (auto simp:RecvMsg co2sobj.simps cp2sproc_def split:option.splits) |
1145 have p_in: "p \<in> current_procs s'" using os alive |
1037 from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" |
1146 apply (erule_tac x = "O_proc p" in allE) |
1038 using os cq2sq vd |
1147 by (auto simp:RecvMsg) |
1039 apply (erule_tac x = q in allE) |
1148 have q_in: "q \<in> current_msgqs s'" using os alive |
1040 by (auto simp:RecvMsg co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) |
1149 apply (erule_tac x = "O_msgq q" in allE) |
1041 from p3 have p3': "sectxt_of_obj s' (O_msg q m) = Some (um, rm, tm)" |
1150 by (simp add:RecvMsg) |
1042 using sms_remain cq2sq vd os p2 p2' p3 |
1151 have m_in: "m = hd (msgs_of_queue s' q)" |
1043 apply (erule_tac x = q in allE) |
1152 and sms_not_empty: "msgs_of_queue s' q \<noteq> []" |
1044 apply (erule_tac x = q in allE) |
1153 using os sms_remain |
1045 apply (clarsimp simp:RecvMsg) |
1154 by (auto simp:RecvMsg) |
1046 apply (simp add:cq2smsgq_def split:option.splits if_splits) |
1155 have "os_grant s' e" using p_in q_in m_in sms_not_empty os |
1047 apply (drule current_has_sms', simp, simp) |
|
1048 apply (case_tac "msgs_of_queue s q", simp) |
|
1049 apply (simp add:cqm2sms.simps split:option.splits) |
|
1050 apply (auto simp add:cm2smsg_def split:option.splits if_splits)[1] |
|
1051 apply (case_tac "msgs_of_queue s q", simp) |
|
1052 apply (simp add:cqm2sms.simps split:option.splits) |
|
1053 apply (auto simp add:cm2smsg_def split:option.splits if_splits)[1] |
|
1054 done |
|
1055 show ?thesis using p1' p2' p3' grant p1 p2 p3 |
|
1056 by (simp add:RecvMsg) |
1156 by (simp add:RecvMsg) |
1057 qed |
1157 moreover have "grant s' e" |
1058 ultimately show ?thesis using vs' |
1158 proof- |
1059 by (erule_tac valid.intros(2), simp+) |
1159 from grant obtain up rp tp uq rq tq um rm tm |
1060 next |
1160 where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" |
1061 case (CreateSock p af st fd inum) |
1161 and p2: "sectxt_of_obj s (O_msgq q) = Some (uq, rq, tq)" |
1062 show ?thesis using grant |
1162 and p3: "sectxt_of_obj s (O_msg q m) = Some (um, rm, tm)" |
1063 by (simp add:CreateSock) |
1163 apply (simp add:RecvMsg split:option.splits) |
1064 next |
1164 by (case_tac a, case_tac aa, case_tac ab, blast) |
1065 case (Bind p fd addr) |
1165 from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" |
1066 show ?thesis using grant |
1166 using os cp2sp |
1067 by (simp add:Bind) |
1167 apply (erule_tac x = p in allE) |
1068 next |
1168 by (auto simp:RecvMsg co2sobj.simps cp2sproc_def split:option.splits) |
1069 case (Connect p fd addr) |
1169 from p2 have p2': "sectxt_of_obj s' (O_msgq q) = Some (uq, rq, tq)" |
1070 show ?thesis using grant |
1170 using os cq2sq vd |
1071 by (simp add:Connect) |
1171 apply (erule_tac x = q in allE) |
1072 next |
1172 by (auto simp:RecvMsg co2sobj.simps cq2smsgq_def dest!:current_has_sms' split:option.splits) |
1073 case (Listen p fd) |
1173 from p3 have p3': "sectxt_of_obj s' (O_msg q m) = Some (um, rm, tm)" |
1074 show ?thesis using grant |
1174 using sms_remain cq2sq vd os p2 p2' p3 |
1075 by (simp add:Listen) |
1175 apply (erule_tac x = q in allE) |
1076 next |
1176 apply (erule_tac x = q in allE) |
1077 case (Accept p fd addr port fd' inum) |
1177 apply (clarsimp simp:RecvMsg) |
1078 show ?thesis using grant |
1178 apply (simp add:cq2smsgq_def split:option.splits if_splits) |
1079 by (simp add:Accept) |
1179 apply (drule current_has_sms', simp, simp) |
1080 next |
1180 apply (case_tac "msgs_of_queue s q", simp) |
1081 case (SendSock p fd) |
1181 apply (simp add:cqm2sms.simps split:option.splits) |
1082 show ?thesis using grant |
1182 apply (auto simp add:cm2smsg_def split:option.splits if_splits)[1] |
1083 by (simp add:SendSock) |
1183 apply (case_tac "msgs_of_queue s q", simp) |
1084 next |
1184 apply (simp add:cqm2sms.simps split:option.splits) |
1085 case (RecvSock p fd) |
1185 apply (auto simp add:cm2smsg_def split:option.splits if_splits)[1] |
1086 show ?thesis using grant |
1186 done |
1087 by (simp add:RecvSock) |
1187 show ?thesis using p1' p2' p3' grant p1 p2 p3 |
1088 next |
1188 by (simp add:RecvMsg) |
1089 case (Shutdown p fd how) |
1189 qed |
1090 show ?thesis using grant |
1190 ultimately show ?thesis using vs' |
1091 by (simp add:Shutdown) |
1191 by (erule_tac valid.intros(2), simp+) |
1092 qed |
1192 next |
1093 |
1193 case (CreateSock p af st fd inum) |
1094 lemma not_all_procs_prop2: |
1194 show ?thesis using grant |
1095 "p' \<notin> all_procs s \<Longrightarrow> p' \<notin> init_procs" |
1195 by (simp add:CreateSock) |
1096 apply (induct s, simp) |
1196 next |
1097 by (case_tac a, auto) |
1197 case (Bind p fd addr) |
1098 |
1198 show ?thesis using grant |
1099 lemma not_all_procs_prop3: |
1199 by (simp add:Bind) |
1100 "p' \<notin> all_procs s \<Longrightarrow> p' \<notin> current_procs s" |
1200 next |
1101 apply (induct s, simp) |
1201 case (Connect p fd addr) |
1102 by (case_tac a, auto) |
1202 show ?thesis using grant |
1103 |
1203 by (simp add:Connect) |
1104 definition is_created_proc:: "t_state \<Rightarrow> t_process \<Rightarrow> bool" |
1204 next |
1105 where |
1205 case (Listen p fd) |
1106 "is_created_proc s p \<equiv> p \<in> current_procs s \<and> (p \<in> init_procs \<longrightarrow> died (O_proc p) s)" |
1206 show ?thesis using grant |
|
1207 by (simp add:Listen) |
|
1208 next |
|
1209 case (Accept p fd addr port fd' inum) |
|
1210 show ?thesis using grant |
|
1211 by (simp add:Accept) |
|
1212 next |
|
1213 case (SendSock p fd) |
|
1214 show ?thesis using grant |
|
1215 by (simp add:SendSock) |
|
1216 next |
|
1217 case (RecvSock p fd) |
|
1218 show ?thesis using grant |
|
1219 by (simp add:RecvSock) |
|
1220 next |
|
1221 case (Shutdown p fd how) |
|
1222 show ?thesis using grant |
|
1223 by (simp add:Shutdown) |
|
1224 qed |
|
1225 qed |
1107 |
1226 |
1108 lemma created_proc_clone: |
1227 lemma created_proc_clone: |
1109 "valid (Clone p p' fds # s) \<Longrightarrow> |
1228 "valid (Clone p p' fds # s) \<Longrightarrow> |
1110 is_created_proc (Clone p p' fds # s) tp = (if (tp = p') then True else is_created_proc s tp)" |
1229 is_created_proc (Clone p p' fds # s) tp = (if (tp = p') then True else is_created_proc s tp)" |
1111 apply (drule vt_grant_os) |
1230 apply (drule vt_grant_os) |