| author | chunhan | 
| Tue, 03 Sep 2013 09:31:40 +0800 | |
| changeset 38 | 9dfc8c72565a | 
| parent 37 | 029cccce84b4 | 
| child 39 | 13bba99ca090 | 
| permissions | -rw-r--r-- | 
| 32 | 1 | (*<*) | 
| 2 | theory S2ss_prop | |
| 3 | imports Main Flask Flask_type Static Static_type Init_prop Tainted_prop Valid_prop Alive_prop Co2sobj_prop | |
| 4 | begin | |
| 5 | (*>*) | |
| 6 | ||
| 33 
6884b3c9284b
fix bug of static.thy for the static of recvmsg case
 chunhan parents: 
32diff
changeset | 7 | context tainting_s begin | 
| 
6884b3c9284b
fix bug of static.thy for the static of recvmsg case
 chunhan parents: 
32diff
changeset | 8 | |
| 35 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 9 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 10 | lemma current_proc_has_sp: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 11 | "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sp. cp2sproc s p = Some sp" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 12 | by (auto simp:cp2sproc_def split:option.splits dest!:current_has_sec') | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 13 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 14 | lemma current_proc_has_sp': | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 15 | "\<lbrakk>cp2sproc s p = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 16 | by (auto dest:current_proc_has_sp) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 17 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 18 | lemma is_dir_has_sdir': | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 19 | "\<lbrakk>is_dir s f; valid s\<rbrakk> \<Longrightarrow> \<exists> sf. cf2sfile s f = Some sf" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 20 | apply (case_tac f) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 21 | apply (rule_tac x = sroot in exI) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 22 | apply (simp add:sroot_only) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 23 | apply (drule is_dir_has_sfile, auto) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 24 | done | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 25 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 26 | lemma is_file_has_sfile': | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 27 | "\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> sf. cf2sfile s f = Some sf" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 28 | by (drule is_file_has_sfile, auto) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 29 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 30 | (* simpset for same_inode_files: Current_files_prop.thy *) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 31 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 32 | lemma same_inode_files_nil: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 33 | "same_inode_files [] = init_same_inode_files" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 34 | by (rule ext, simp add:same_inode_files_def init_same_inode_files_def is_file_nil) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 35 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 36 | lemma iof's_im_in_cim': "Some im = inum_of_file \<tau> f \<Longrightarrow> im \<in> current_inode_nums \<tau>" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 37 | by (auto simp add:current_inode_nums_def current_file_inums_def) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 38 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 39 | lemma same_inode_files_open: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 40 | "valid (Open p f flags fd opt # s) \<Longrightarrow> same_inode_files (Open p f flags fd opt # s) = (\<lambda> f'. | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 41 |      if (f' = f \<and> opt \<noteq> None) then {f} else same_inode_files s f')"
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 42 | apply (frule vt_grant_os, frule vd_cons, rule ext) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 43 | apply (auto simp:same_inode_files_def is_file_simps split:if_splits option.splits | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 44 | dest:iof's_im_in_cim iof's_im_in_cim') | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 45 | apply (drule is_file_in_current) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 46 | apply (simp add:current_files_def) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 47 | done | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 48 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 49 | lemma same_inode_files_linkhard: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 50 | "valid (LinkHard p oldf f # s) \<Longrightarrow> same_inode_files (LinkHard p oldf f # s) = (\<lambda> f'. | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 51 | if (f' = f \<or> f' \<in> same_inode_files s oldf) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 52 |      then same_inode_files s oldf \<union> {f}
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 53 | else same_inode_files s f')" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 54 | apply (frule vt_grant_os, frule vd_cons, rule ext) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 55 | apply (auto simp:same_inode_files_def is_file_simps split:if_splits option.splits | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 56 | dest:iof's_im_in_cim iof's_im_in_cim') | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 57 | apply (drule is_file_in_current) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 58 | apply (simp add:current_files_def is_file_def) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 59 | apply (simp add:is_file_def) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 60 | done | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 61 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 62 | lemma inum_of_file_none_prop: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 63 | "\<lbrakk>inum_of_file s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 64 | by (simp add:current_files_def) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 65 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 66 | lemma same_inode_files_closefd: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 67 | "\<lbrakk>valid (CloseFd p fd # s); f' \<in> current_files (CloseFd p fd # s)\<rbrakk> \<Longrightarrow> | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 68 | same_inode_files (CloseFd p fd # s) f' = ( | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 69 | case (file_of_proc_fd s p fd) of | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 70 |        Some f \<Rightarrow> (if ((proc_fd_of_file s f = {(p, fd)}) \<and> (f \<in> files_hung_by_del s))
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 71 |                  then same_inode_files s f' - {f}
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 72 | else same_inode_files s f' ) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 73 | | None \<Rightarrow> same_inode_files s f' )" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 74 | apply (frule vt_grant_os, frule vd_cons) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 75 | apply (auto simp:same_inode_files_def is_file_closefd current_files_closefd | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 76 | split:if_splits option.splits | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 77 | dest:iof's_im_in_cim iof's_im_in_cim' inum_of_file_none_prop) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 78 | done | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 79 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 80 | lemma same_inode_files_unlink: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 81 | "\<lbrakk>valid (UnLink p f # s); f' \<in> current_files (UnLink p f # s)\<rbrakk> | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 82 | \<Longrightarrow> same_inode_files (UnLink p f # s) f' = ( | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 83 |      if (proc_fd_of_file s f = {}) 
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 84 |      then same_inode_files s f' - {f}
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 85 | else same_inode_files s f')" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 86 | apply (frule vt_grant_os, frule vd_cons) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 87 | apply (auto simp:same_inode_files_def is_file_unlink current_files_unlink | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 88 | split:if_splits option.splits | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 89 | dest:iof's_im_in_cim iof's_im_in_cim' inum_of_file_none_prop) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 90 | done | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 91 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 92 | lemma same_inode_files_mkdir: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 93 | "valid (Mkdir p f inum # s) \<Longrightarrow> same_inode_files (Mkdir p f inum # s) = (same_inode_files s)" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 94 | apply (frule vt_grant_os, frule vd_cons, rule ext) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 95 | apply (auto simp:same_inode_files_def is_file_simps current_files_simps | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 96 | split:if_splits option.splits | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 97 | dest:iof's_im_in_cim iof's_im_in_cim' inum_of_file_none_prop is_file_in_current) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 98 | apply (simp add:current_files_def is_file_def) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 99 | done | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 100 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 101 | lemma same_inode_files_other: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 102 | "\<lbrakk>valid (e # s); | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 103 | \<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt; | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 104 | \<forall> p fd. e \<noteq> CloseFd p fd; | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 105 | \<forall> p f. e \<noteq> UnLink p f; | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 106 | \<forall> p f f'. e \<noteq> LinkHard p f f'\<rbrakk> \<Longrightarrow> same_inode_files (e # s) = same_inode_files s" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 107 | apply (frule vt_grant_os, frule vd_cons, rule ext, case_tac e) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 108 | apply (auto simp:same_inode_files_def is_file_simps current_files_simps dir_is_empty_def | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 109 | split:if_splits option.splits | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 110 | dest:iof's_im_in_cim iof's_im_in_cim' inum_of_file_none_prop is_file_not_dir) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 111 | apply (simp add:is_file_def is_dir_def current_files_def split:option.splits t_inode_tag.splits)+ | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 112 | done | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 113 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 114 | lemmas same_inode_files_simps = same_inode_files_nil same_inode_files_open same_inode_files_linkhard | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 115 | same_inode_files_closefd same_inode_files_unlink same_inode_files_mkdir same_inode_files_other | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 116 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 117 | lemma same_inode_files_prop1: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 118 | "f \<in> same_inode_files s f' \<Longrightarrow> f \<in> current_files s" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 119 | by (simp add:same_inode_files_def is_file_def current_files_def split:if_splits option.splits) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 120 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 121 | lemma same_inode_files_prop2: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 122 | "\<lbrakk>f \<in> same_inode_files s f'; f'' \<notin> current_files s\<rbrakk> \<Longrightarrow> f \<noteq> f''" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 123 | by (auto dest:same_inode_files_prop1) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 124 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 125 | lemma same_inode_files_prop3: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 126 | "\<lbrakk>f \<in> same_inode_files s f'; is_dir s f''\<rbrakk> \<Longrightarrow> f \<noteq> f''" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 127 | apply (rule notI) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 128 | apply (simp add:same_inode_files_def is_file_def is_dir_def | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 129 | split:if_splits option.splits t_inode_tag.splits) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 130 | done | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 131 | |
| 36 | 132 | lemma same_inode_files_prop4: | 
| 133 | "\<lbrakk>f' \<in> same_inode_files s f; f'' \<in> same_inode_files s f'\<rbrakk> \<Longrightarrow> f'' \<in> same_inode_files s f" | |
| 134 | by (auto simp:same_inode_files_def split:if_splits) | |
| 135 | ||
| 136 | lemma same_inode_files_prop5: | |
| 137 | "f' \<in> same_inode_files s f \<Longrightarrow> f \<in> same_inode_files s f'" | |
| 138 | by (auto simp:same_inode_files_def is_file_def split:if_splits) | |
| 139 | ||
| 38 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 140 | lemma same_inode_files_prop6: | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 141 | "f' \<in> same_inode_files s f \<Longrightarrow> same_inode_files s f' = same_inode_files s f" | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 142 | by (auto simp:same_inode_files_def is_file_def split:if_splits) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 143 | |
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 144 | lemma same_inode_files_prop7: | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 145 | "f' \<in> same_inode_files s f \<Longrightarrow> has_same_inode s f f'" | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 146 | thm has_same_inode_def | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 147 | apply (auto simp:same_inode_files_def is_file_def has_same_inode_def split:if_splits) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 148 | |
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 149 | |
| 35 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 150 | (* simpset for cf2sfiles *) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 151 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 152 | lemma cf2sfiles_open: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 153 | "\<lbrakk>valid (Open p f flag fd opt # s); f' \<in> current_files (Open p f flag fd opt # s)\<rbrakk> | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 154 | \<Longrightarrow> cf2sfiles (Open p f flag fd opt # s) f' = ( | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 155 | if (f' = f \<and> opt \<noteq> None) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 156 | then (case cf2sfile (Open p f flag fd opt # s) f of | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 157 |              Some sf \<Rightarrow> {sf}  
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 158 |            | _       \<Rightarrow> {} )
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 159 | else cf2sfiles s f')" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 160 | apply (frule vt_grant_os, frule vd_cons) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 161 | apply (auto simp:cf2sfiles_def cf2sfile_open_none cf2sfile_simps same_inode_files_open | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 162 | split:if_splits option.splits dest!:current_file_has_sfile' dest:cf2sfile_open) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 163 | apply (rule_tac x = "f'a" in bexI, drule same_inode_files_prop1, simp add:cf2sfile_open_some1, simp)+ | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 164 | done | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 165 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 166 | lemma cf2sfiles_other: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 167 | "\<lbrakk>valid (e # s); | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 168 | \<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt; | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 169 | \<forall> p fd. e \<noteq> CloseFd p fd; | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 170 | \<forall> p f. e \<noteq> UnLink p f; | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 171 | \<forall> p f f'. e \<noteq> LinkHard p f f'\<rbrakk> \<Longrightarrow> cf2sfiles (e # s) = cf2sfiles s" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 172 | apply (frule vt_grant_os, frule vd_cons, rule ext) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 173 | apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 174 | apply (drule Set.CollectD, erule bexE, rule CollectI) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 175 | apply (rule_tac x = f' in bexI, case_tac e) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 176 | apply (auto simp:cf2sfiles_def cf2sfile_simps same_inode_files_simps current_files_simps | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 177 | split:if_splits option.splits dest!:current_file_has_sfile' dest:same_inode_files_prop1 cf2sfile_other') | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 178 | apply (drule_tac f' = f' in cf2sfile_rmdir) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 179 | apply (simp add:current_files_simps same_inode_files_prop1 same_inode_files_prop3 dir_is_empty_def)+ | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 180 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 181 | apply (rule_tac x = f' in bexI, case_tac e) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 182 | apply (auto simp:cf2sfiles_def cf2sfile_simps same_inode_files_simps current_files_simps | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 183 | split:if_splits option.splits dest!:current_file_has_sfile' dest:same_inode_files_prop1 cf2sfile_other') | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 184 | apply (drule_tac f' = f' in cf2sfile_rmdir) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 185 | apply (simp add:current_files_simps same_inode_files_prop1 same_inode_files_prop3 dir_is_empty_def)+ | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 186 | done | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 187 | |
| 36 | 188 | lemma cf2sfile_linkhard1': | 
| 189 | "\<lbrakk>valid (LinkHard p oldf f # s); f' \<in> same_inode_files s f''\<rbrakk> | |
| 190 | \<Longrightarrow> cf2sfile (LinkHard p oldf f# s) f' = cf2sfile s f'" | |
| 191 | apply (drule same_inode_files_prop1) | |
| 192 | by (simp add:cf2sfile_linkhard1) | |
| 193 | ||
| 35 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 194 | lemma cf2sfiles_linkhard: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 195 | "valid (LinkHard p oldf f # s) \<Longrightarrow> cf2sfiles (LinkHard p oldf f # s) = (\<lambda> f'. | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 196 | if (f' = f \<or> f' \<in> same_inode_files s oldf) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 197 | then (case (cf2sfile (LinkHard p oldf f # s) f) of | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 198 |              Some sf \<Rightarrow> cf2sfiles s oldf \<union> {sf}
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 199 |            | _       \<Rightarrow> {})
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 200 | else cf2sfiles s f')" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 201 | apply (frule vt_grant_os, frule vd_cons, rule ext) | 
| 36 | 202 | apply (auto simp:cf2sfiles_def cf2sfile_linkhard1' same_inode_files_linkhard current_files_linkhard | 
| 203 | split:if_splits option.splits dest!:current_file_has_sfile' current_has_sec' dest:same_inode_files_prop1) | |
| 204 | done | |
| 205 | ||
| 206 | lemma cf2sfile_unlink': | |
| 207 | "\<lbrakk>valid (UnLink p f # s); f' \<in> same_inode_files (UnLink p f # s) f''\<rbrakk> | |
| 208 | \<Longrightarrow> cf2sfile (UnLink p f # s) f' = cf2sfile s f'" | |
| 209 | apply (drule same_inode_files_prop1) | |
| 210 | by (simp add:cf2sfile_unlink) | |
| 35 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 211 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 212 | lemma cf2sfiles_unlink: | 
| 36 | 213 | "\<lbrakk>valid (UnLink p f # s); f' \<in> current_files (UnLink p f # s)\<rbrakk> \<Longrightarrow> cf2sfiles (UnLink p f # s) f' = ( | 
| 214 |      if (f' \<in> same_inode_files s f \<and> proc_fd_of_file s f = {} \<and> 
 | |
| 215 | (\<forall> f'' \<in> same_inode_files s f. f'' \<noteq> f \<longrightarrow> cf2sfile s f'' \<noteq> cf2sfile s f)) then | |
| 216 | (case (cf2sfile s f) of | |
| 217 |            Some sf \<Rightarrow> cf2sfiles s f' - {sf}
 | |
| 218 |          | _       \<Rightarrow> {})
 | |
| 219 | else cf2sfiles s f')" | |
| 220 | apply (frule vt_grant_os, frule vd_cons, simp add:current_files_simps split:if_splits) | |
| 221 | apply (rule conjI, clarify, frule is_file_has_sfile', simp, erule exE, simp) | |
| 222 | apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) | |
| 223 | apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink) | |
| 224 | apply (simp add:current_files_unlink, simp, erule conjE) | |
| 225 | apply (erule_tac x = f'a in ballE, frule_tac f' = "f'a" in cf2sfile_unlink) | |
| 226 | apply (simp add:current_files_unlink same_inode_files_prop1, simp) | |
| 227 | apply (rule_tac x = f'a in bexI, simp, simp) | |
| 228 | apply (drule_tac f = f and f' = f' and f'' = f'a in same_inode_files_prop4, simp+) | |
| 229 | apply (erule conjE|erule exE|erule bexE)+ | |
| 230 | apply (case_tac "f'a = f", simp) | |
| 231 | apply (frule_tac f' = f' in same_inode_files_unlink, simp add:current_files_unlink) | |
| 232 | apply (frule_tac f' = f'a in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1) | |
| 233 | apply (rule_tac x = f'a in bexI, simp, simp) | |
| 234 | ||
| 235 | apply (rule impI)+ | |
| 236 | apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) | |
| 237 | apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink) | |
| 238 | apply (simp add:current_files_unlink, simp, (erule conjE)+) | |
| 239 | apply (rule_tac x = f'a in bexI, frule_tac f' = "f'a" in cf2sfile_unlink) | |
| 240 | apply (simp add:current_files_unlink same_inode_files_prop1, simp, simp) | |
| 241 | apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_unlink) | |
| 242 | apply (simp add:current_files_unlink, simp) | |
| 243 | apply (case_tac "f'a = f", simp) | |
| 244 | apply (frule_tac f = f' and f' = f in same_inode_files_prop5, simp) | |
| 245 | apply (erule bexE, erule conjE) | |
| 246 | apply (rule_tac x = f'' in bexI) | |
| 247 | apply (drule_tac f' = f'' in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1) | |
| 248 | apply (simp, simp, erule same_inode_files_prop4, simp) | |
| 249 | apply (rule_tac x = f'a in bexI) | |
| 250 | apply (drule_tac f' = f'a in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1) | |
| 251 | apply (simp, simp) | |
| 252 | ||
| 253 | ||
| 254 | apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) | |
| 255 | apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink) | |
| 256 | apply (simp add:current_files_unlink, simp) | |
| 257 | apply (rule_tac x = f'a in bexI) | |
| 258 | apply (frule_tac f' = f'a in cf2sfile_unlink) | |
| 259 | apply (simp add:same_inode_files_prop1 current_files_unlink, simp, simp) | |
| 260 | apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_unlink) | |
| 261 | apply (simp add:current_files_unlink, simp) | |
| 262 | apply (rule_tac x = f'a in bexI) | |
| 263 | apply (frule_tac f' = f'a in cf2sfile_unlink) | |
| 264 | apply (simp add:same_inode_files_prop1 current_files_unlink, simp, simp) | |
| 265 | done | |
| 35 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 266 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 267 | lemma cf2sfiles_closefd: | 
| 36 | 268 | "\<lbrakk>valid (CloseFd p fd # s); f' \<in> current_files (CloseFd p fd # s)\<rbrakk> \<Longrightarrow> cf2sfiles (CloseFd p fd # s) f' = ( | 
| 269 | case (file_of_proc_fd s p fd) of | |
| 270 |        Some f \<Rightarrow> if (f' \<in> same_inode_files s f \<and> proc_fd_of_file s f = {(p, fd)} \<and> f \<in> files_hung_by_del s \<and> 
 | |
| 271 | (\<forall> f'' \<in> same_inode_files s f. f'' \<noteq> f \<longrightarrow> cf2sfile s f'' \<noteq> cf2sfile s f)) | |
| 272 | then (case (cf2sfile s f) of | |
| 273 |                          Some sf \<Rightarrow> cf2sfiles s f' - {sf}
 | |
| 274 |                        | _       \<Rightarrow> {})
 | |
| 275 | else cf2sfiles s f' | |
| 276 | | _ \<Rightarrow> cf2sfiles s f')" | |
| 277 | ||
| 278 | apply (frule vt_grant_os, frule vd_cons, case_tac "file_of_proc_fd s p fd") | |
| 279 | apply (simp_all add:current_files_simps split:if_splits) | |
| 280 | ||
| 281 | apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) | |
| 282 | apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd) | |
| 283 | apply (simp add:current_files_closefd, simp) | |
| 284 | apply (rule_tac x = f'a in bexI) | |
| 285 | apply (frule_tac f = f'a in cf2sfile_closefd) | |
| 286 | apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) | |
| 287 | apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd) | |
| 288 | apply (simp add:current_files_closefd, simp) | |
| 289 | apply (rule_tac x = f'a in bexI) | |
| 290 | apply (frule_tac f = f'a in cf2sfile_closefd) | |
| 291 | apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) | |
| 292 | ||
| 293 | apply (rule conjI, clarify, frule file_of_pfd_is_file, simp) | |
| 294 | apply (frule is_file_has_sfile', simp, erule exE, simp) | |
| 295 | apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) | |
| 296 | apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd) | |
| 297 | apply (simp add:current_files_closefd, simp, erule conjE) | |
| 298 | apply (erule_tac x = f'a in ballE, frule_tac f = "f'a" in cf2sfile_closefd) | |
| 299 | apply (simp add:current_files_closefd same_inode_files_prop1, simp) | |
| 300 | apply (rule_tac x = f'a in bexI, simp, simp) | |
| 301 | apply (drule_tac f = a and f' = f' and f'' = f'a in same_inode_files_prop4, simp+) | |
| 302 | apply (erule conjE|erule exE|erule bexE)+ | |
| 303 | apply (case_tac "f'a = a", simp) | |
| 304 | apply (frule_tac f' = f' in same_inode_files_closefd, simp add:current_files_closefd, simp) | |
| 305 | apply (frule_tac f = f'a in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1) | |
| 306 | apply (rule_tac x = f'a in bexI, simp, simp) | |
| 307 | ||
| 308 | apply (rule impI)+ | |
| 309 | apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) | |
| 310 | apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd) | |
| 311 | apply (simp add:current_files_closefd, simp, (erule conjE)+) | |
| 312 | apply (rule_tac x = f'a in bexI, frule_tac f = f'a in cf2sfile_closefd) | |
| 313 | apply (simp add:current_files_closefd same_inode_files_prop1, simp, simp) | |
| 314 | apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd) | |
| 315 | apply (simp add:current_files_closefd, simp) | |
| 316 | apply (case_tac "f'a = a", simp) | |
| 317 | apply (frule_tac f = f' and f' = a in same_inode_files_prop5, simp) | |
| 318 | apply (erule bexE, erule conjE) | |
| 319 | apply (rule_tac x = f'' in bexI) | |
| 320 | apply (drule_tac f = f'' in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1) | |
| 321 | apply (simp, simp, erule same_inode_files_prop4, simp) | |
| 322 | apply (rule_tac x = f'a in bexI) | |
| 323 | apply (drule_tac f = f'a in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1) | |
| 324 | apply (simp, simp) | |
| 325 | ||
| 326 | apply (rule conjI, clarify) | |
| 327 | ||
| 328 | apply (rule impI) | |
| 329 | apply (case_tac "a \<in> files_hung_by_del s", simp_all) | |
| 330 | apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) | |
| 331 | apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd) | |
| 332 | apply (simp add:current_files_closefd, simp) | |
| 333 | apply (rule_tac x = f'a in bexI) | |
| 334 | apply (frule_tac f = f'a in cf2sfile_closefd) | |
| 335 | apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) | |
| 336 | apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd) | |
| 337 | apply (simp add:current_files_closefd, simp) | |
| 338 | apply (rule_tac x = f'a in bexI) | |
| 339 | apply (frule_tac f = f'a in cf2sfile_closefd) | |
| 340 | apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) | |
| 341 | apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD) | |
| 342 | apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd) | |
| 343 | apply (simp add:current_files_closefd, simp) | |
| 344 | apply (rule_tac x = f'a in bexI) | |
| 345 | apply (frule_tac f = f'a in cf2sfile_closefd) | |
| 346 | apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) | |
| 347 | apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd) | |
| 348 | apply (simp add:current_files_closefd, simp) | |
| 349 | apply (rule_tac x = f'a in bexI) | |
| 350 | apply (frule_tac f = f'a in cf2sfile_closefd) | |
| 351 | apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp) | |
| 352 | done | |
| 35 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 353 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 354 | lemmas cf2sfiles_simps = cf2sfiles_open cf2sfiles_linkhard cf2sfiles_other | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 355 | cf2sfiles_unlink cf2sfiles_closefd | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 356 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 357 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 358 | (* simpset for co2sobj *) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 359 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 360 | lemma co2sobj_execve: | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 361 | "\<lbrakk>valid (Execve p f fds # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (Execve p f fds # s) obj = ( | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 362 | if (obj = O_proc p) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 363 | then (case (cp2sproc (Execve p f fds # s) p) of | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 364 | Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s \<or> O_file f \<in> Tainted s)) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 365 | | _ \<Rightarrow> None) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 366 | else co2sobj s obj )" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 367 | apply (frule vt_grant_os, frule vd_cons, case_tac obj) | 
| 38 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 368 | apply (simp_all add:current_files_simps cf2sfiles_other ch2sshm_other cq2smsgq_other tainted_eq_Tainted) | 
| 35 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 369 | apply (case_tac "cp2sproc (Execve p f fds # s) p") | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 370 | apply (drule current_proc_has_sp', simp, simp) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 371 | apply (simp (no_asm_simp) add:cp2sproc_execve tainted_eq_Tainted split:option.splits) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 372 | apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 373 | apply (frule_tac ff = list in cf2sfile_other', simp_all) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 374 | apply (simp add:is_dir_in_current) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 375 | done | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 376 | |
| 37 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 377 | lemma co2sobj_clone: | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 378 | "\<lbrakk>valid (Clone p p' fds shms # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (Clone p p' fds shms # s) obj = ( | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 379 | if (obj = O_proc p') | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 380 | then (case (cp2sproc (Clone p p' fds shms # s) p') of | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 381 | Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s)) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 382 | | _ \<Rightarrow> None) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 383 | else co2sobj s obj )" | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 384 | apply (frule vt_grant_os, frule vd_cons, case_tac obj) | 
| 38 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 385 | apply (simp_all add:current_files_simps cf2sfiles_other ch2sshm_other cq2smsgq_other tainted_eq_Tainted) | 
| 37 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 386 | apply (case_tac "cp2sproc (Clone p p' fds shms # s) p'") | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 387 | apply (drule current_proc_has_sp', simp, simp) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 388 | apply ((erule conjE)+, frule_tac p = p in current_proc_has_sec, simp, erule exE, simp) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 389 | apply (rule conjI, rule impI, simp) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 390 | apply (simp (no_asm_simp) add:cp2sproc_clone tainted_eq_Tainted split:option.splits) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 391 | |
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 392 | apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 393 | apply (frule_tac ff = list in cf2sfile_other', simp_all) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 394 | apply (simp add:is_dir_in_current) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 395 | done | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 396 | |
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 397 | lemma co2sobj_ptrace: | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 398 | "\<lbrakk>valid (Ptrace p p' # s); alive s obj\<rbrakk>\<Longrightarrow> co2sobj (Ptrace p p' # s) obj = ( | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 399 | case obj of | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 400 | O_proc p'' \<Rightarrow> if (info_flow_shm s p' p'') | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 401 | then (case (cp2sproc s p'') of | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 402 | Some sp \<Rightarrow> Some (S_proc sp (O_proc p'' \<in> Tainted s \<or> O_proc p \<in> Tainted s)) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 403 | | _ \<Rightarrow> None) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 404 | else if (info_flow_shm s p p'') | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 405 | then (case (cp2sproc s p'') of | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 406 | Some sp \<Rightarrow> Some (S_proc sp (O_proc p'' \<in> Tainted s \<or> O_proc p' \<in> Tainted s)) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 407 | | _ \<Rightarrow> None) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 408 | else co2sobj s (O_proc p'') | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 409 | | _ \<Rightarrow> co2sobj s obj)" | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 410 | apply (frule vt_grant_os, frule vd_cons, case_tac obj) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 411 | apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other cf2sfiles_other tainted_eq_Tainted) | 
| 38 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 412 | |
| 37 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 413 | apply (auto simp:cp2sproc_other tainted_eq_Tainted split:option.splits | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 414 | dest!:current_proc_has_sec' current_proc_has_sp' intro:info_flow_shm_Tainted)[1] | 
| 36 | 415 | |
| 37 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 416 | apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 417 | apply (frule_tac ff = list in cf2sfile_other', simp_all) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 418 | apply (simp add:is_dir_in_current) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 419 | done | 
| 36 | 420 | |
| 37 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 421 | lemma co2sobj_open: | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 422 | "\<lbrakk>valid (Open p f flag fd opt # s); alive (Open p f flag fd opt # s) obj\<rbrakk> | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 423 | \<Longrightarrow> co2sobj (Open p f flag fd opt # s) obj = (case obj of | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 424 | O_file f' \<Rightarrow> if (f' = f \<and> opt \<noteq> None) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 425 | then (case (cf2sfile (Open p f flag fd opt # s) f) of | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 426 |                           Some sf \<Rightarrow> Some (S_file {sf} (O_proc p \<in> Tainted s))
 | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 427 | | _ \<Rightarrow> None) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 428 | else co2sobj s (O_file f') | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 429 | | O_proc p' \<Rightarrow> if (p' = p) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 430 | then (case (cp2sproc (Open p f flag fd opt # s) p) of | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 431 | Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> Tainted s)) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 432 | | _ \<Rightarrow> None) | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 433 | else co2sobj s (O_proc p') | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 434 | | _ \<Rightarrow> co2sobj s obj )" | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 435 | apply (frule vt_grant_os, frule vd_cons, case_tac obj) | 
| 38 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 436 | apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 437 | dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1] | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 438 | |
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 439 | apply (simp split:if_splits t_object.splits) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 440 | apply (rule conjI, rule impI, erule conjE, erule exE, simp, (erule exE|erule conjE)+) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 441 | apply (case_tac "cf2sfile (Open p f flag fd (Some y) # s) f") | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 442 | apply (drule current_file_has_sfile', simp, simp add:current_files_simps, simp) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 443 | apply (frule_tac f' = f in cf2sfiles_open, simp add:current_files_simps) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 444 | apply (simp add:tainted_eq_Tainted) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 445 | apply (rule impI, rule notI, drule Tainted_in_current, simp, simp add:is_file_in_current) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 446 | apply (rule impI, simp add:tainted_eq_Tainted cf2sfiles_open is_file_in_current split:option.splits) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 447 | |
| 37 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 448 | apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) | 
| 38 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 449 | |
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 450 | apply (frule is_dir_in_current) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 451 | apply (frule_tac f' = list in cf2sfile_open) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 452 | apply (simp add:current_files_simps split:option.splits) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 453 | apply (simp split:if_splits option.splits) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 454 | done | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 455 | |
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 456 | lemma co2sobj_readfile: | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 457 | "\<lbrakk>valid (ReadFile p fd # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (ReadFile p fd # s) obj = ( | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 458 | case obj of | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 459 | O_proc p' \<Rightarrow> (case (file_of_proc_fd s p fd) of | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 460 | Some f \<Rightarrow> (if (info_flow_shm s p p' \<and> O_file f \<in> Tainted s) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 461 | then (case (cp2sproc s p') of | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 462 | Some sp \<Rightarrow> Some (S_proc sp True) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 463 | | _ \<Rightarrow> None) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 464 | else co2sobj s obj) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 465 | | _ \<Rightarrow> None) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 466 | | _ \<Rightarrow> co2sobj s obj)" | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 467 | apply (frule vt_grant_os, frule vd_cons, case_tac obj) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 468 | apply (auto simp:cp2sproc_simps tainted_eq_Tainted split:option.splits | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 469 | dest!:current_proc_has_sp' intro:info_flow_shm_Tainted)[1] | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 470 | apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 471 | |
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 472 | apply (auto split:if_splits option.splits dest!:current_file_has_sfile' | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 473 | simp:current_files_simps cf2sfiles_simps cf2sfile_simps | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 474 | dest:is_file_in_current is_dir_in_current) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 475 | done | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 476 | |
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 477 | lemma co2sobj_writefile: | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 478 | "\<lbrakk>valid (WriteFile p fd # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (WriteFile p fd # s) obj = ( | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 479 | case obj of | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 480 | O_file f' \<Rightarrow> (case (file_of_proc_fd s p fd) of | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 481 | Some f \<Rightarrow> (if (f \<in> same_inode_files s f') | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 482 | then Some (S_file (cf2sfiles s f') | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 483 | (O_file f' \<in> Tainted s \<or> O_proc p \<in> Tainted s)) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 484 | else co2sobj s obj) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 485 | | _ \<Rightarrow> None) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 486 | | _ \<Rightarrow> co2sobj s obj)" | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 487 | apply (frule vt_grant_os, frule vd_cons, case_tac obj) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 488 | apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other tainted_eq_Tainted) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 489 | |
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 490 | apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp' | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 491 | simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_Tainted | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 492 | same_inode_files_prop6 | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 493 | dest:is_file_in_current is_dir_in_current) | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 494 | thm has_same_inode_def | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 495 | done | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 496 | |
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 497 | lemma co2sobj_closefd: | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 498 | "\<lbrakk>valid (CloseFd p fd # s); alive s obj\<rbrakk> \<Longrightarrow> co2sobj (CloseFd p fd # s) obj = ( | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 499 | case obj of | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 500 | O_file f' \<Rightarrow> (case (file_of_proc_fd s p fd) of | 
| 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 501 | Some f \<Rightarrow> (if (f " | 
| 37 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 502 | |
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 503 | lemma co2sobj_other: | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 504 | "\<lbrakk>valid (e # s); alive (e # s) obj; | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 505 | \<forall> p f fds. e \<noteq> Execve p f fds; | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 506 | \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms; | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 507 | \<forall> p p'. e \<noteq> Ptrace p p'; | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 508 | \<forall> | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 509 | \<rbrakk> \<Longrightarrow> co2sobj (e # s) obj = co2sobj s obj" | 
| 
029cccce84b4
remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
 chunhan parents: 
36diff
changeset | 510 | |
| 38 
9dfc8c72565a
found a bug: has_same_inode is a redundant concept of same_inode_files, which is the general form, has_same_inode should be deleted
 chunhan parents: 
37diff
changeset | 511 | lemmas co2sobj_simps = co2sobj_execve co2sobj_clone co2sobj_ptrace co2sobj_open | 
| 35 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 512 | |
| 33 
6884b3c9284b
fix bug of static.thy for the static of recvmsg case
 chunhan parents: 
32diff
changeset | 513 | (* simpset for s2ss*) | 
| 
6884b3c9284b
fix bug of static.thy for the static of recvmsg case
 chunhan parents: 
32diff
changeset | 514 | |
| 
6884b3c9284b
fix bug of static.thy for the static of recvmsg case
 chunhan parents: 
32diff
changeset | 515 | lemma s2ss_execve: | 
| 34 | 516 | "valid (Execve p f fds # s) \<Longrightarrow> s2ss (Execve p f fds # s) = ( | 
| 517 | if (\<exists> p'. p' \<noteq> p \<and> p' \<in> current_procs s \<and> co2sobj s (O_proc p') = co2sobj s (O_proc p)) | |
| 35 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 518 | then (case (cp2sproc (Execve p f fds # s) p) of | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 519 |              Some sp \<Rightarrow> s2ss s \<union> {S_proc sp (O_proc p \<in> Tainted s \<or> O_file f \<in> Tainted s)}
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 520 | | _ \<Rightarrow> s2ss s) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 521 | else (case (cp2sproc (Execve p f fds # s) p) of | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 522 |              Some sp \<Rightarrow> s2ss s - {S_proc sp (O_proc p \<in> Tainted s)}
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 523 |                               \<union> {S_proc sp (O_proc p \<in> Tainted s \<or> O_file f \<in> Tainted s)}
 | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 524 | | _ \<Rightarrow> s2ss s) )" | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 525 | apply (frule vd_cons, frule vt_grant_os, simp split:if_splits) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 526 | apply (rule conjI, clarify) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 527 | apply (frule_tac p = p in current_proc_has_sp, simp, erule exE) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 528 | apply (frule_tac p = p' in current_proc_has_sp, simp, erule exE) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 529 | apply (simp, (erule conjE)+) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 530 | apply (split option.splits, rule conjI, rule impI, drule current_proc_has_sp', simp, simp) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 531 | apply (rule allI, rule impI) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 532 | apply (rule set_eqI, rule iffI) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 533 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 534 | apply (simp split:option.splits) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 535 | apply (frule_tac p = p and s = "Execve p f fds # s" in current_proc_has_sp) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 536 | thm current_proc_has_sp | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 537 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 538 | |
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 539 | apply (simp split:option.splits) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 540 | apply (drule current_proc_has_sp', simp, simp) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 541 | apply (rule conjI, rule impI, drule current_proc_has_sp', simp, simp) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 542 | apply (simp add:s2ss_def) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 543 | apply (rule allI|rule impI)+ | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 544 | apply (rule set_eqI, rule iffI) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 545 | apply (auto simp:alive_simps) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 546 | apply (case_tac obj, auto split:option.splits simp:cp2sproc_execve) | 
| 
f2e620d799cf
in the middle of building simpset for co2sobj, therefore we can deal with s2ss
 chunhan parents: 
34diff
changeset | 547 | apply (auto split:if_splits) |