3 Temp Enrich Proc_fd_of_file_prop |
3 Temp Enrich Proc_fd_of_file_prop |
4 begin |
4 begin |
5 |
5 |
6 context tainting_s begin |
6 context tainting_s begin |
7 |
7 |
8 (* \<And> *) |
8 |
9 lemma current_proc_fds_in_curp: |
9 fun enrich_msgq :: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_msgq \<Rightarrow> t_state" |
10 "\<lbrakk>fd \<in> current_proc_fds s p; valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s" |
10 where |
|
11 "enrich_msgq [] tq dq = []" |
|
12 | "enrich_msgq (CreateMsgq p q # s) tq dq = |
|
13 (if (tq = q) |
|
14 then (CreateMsgq p dq # CreateMsgq p q # s) |
|
15 else CreateMsgq p q # (enrich_msgq s tq dq))" |
|
16 | "enrich_msgq (SendMsg p q m # s) tq dq = |
|
17 (if (tq = q) |
|
18 then (SendMsg p dq m # SendMsg p q m # (enrich_msgq s tq dq)) |
|
19 else SendMsg p q m # (enrich_msgq s tq dq))" |
|
20 | "enrich_msgq (RecvMsg p q m # s) tq dq = |
|
21 (if (tq = q) |
|
22 then (RecvMsg p dq m # RecvMsg p q m # (enrich_msgq s tq dq)) |
|
23 else RecvMsg p q m # (enrich_msgq s tq dq))" |
|
24 | "enrich_msgq (e # s) tq dq = e # (enrich_msgq s tq dq)" |
|
25 |
|
26 |
|
27 |
|
28 (* enrich s target_proc duplicated_pro *) |
|
29 fun enrich_proc :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> t_state" |
|
30 where |
|
31 "enrich_proc [] tp dp = []" |
|
32 | "enrich_proc (Execve p f fds # s) tp dp = ( |
|
33 if (tp = p) |
|
34 then Execve dp f (fds \<inter> proc_file_fds s p) # Execve p f fds # (enrich_proc s tp dp) |
|
35 else Execve p f fds # (enrich_proc s tp dp))" |
|
36 | "enrich_proc (Clone p p' fds # s) tp dp = ( |
|
37 if (tp = p') |
|
38 then Clone p dp (fds \<inter> proc_file_fds s p) # Clone p p' fds # s |
|
39 else Clone p p' fds # (enrich_proc s tp dp))" |
|
40 | "enrich_proc (Open p f flags fd opt # s) tp dp = ( |
|
41 if (tp = p) |
|
42 then Open dp f (remove_create_flag flags) fd None # Open p f flags fd opt # (enrich_proc s tp dp) |
|
43 else Open p f flags fd opt # (enrich_proc s tp dp))" |
|
44 | "enrich_proc (ReadFile p fd # s) tp dp = ( |
|
45 if (tp = p) |
|
46 then ReadFile dp fd # ReadFile p fd # (enrich_proc s tp dp) |
|
47 else ReadFile p fd # (enrich_proc s tp dp))" |
|
48 (* |
|
49 | "enrich_proc (CloseFd p fd # s) tp dp = ( |
|
50 if (tp = p \<and> fd \<in> proc_file_fds s p) |
|
51 then CloseFd dp fd # CloseFd p fd # (enrich_proc s tp dp) |
|
52 else CloseFd p fd # (enrich_proc s tp dp))" |
|
53 *) |
|
54 (* |
|
55 | "enrich_proc (Attach p h flag # s) tp dp = ( |
|
56 if (tp = p) |
|
57 then Attach dp h flag # Attach p h flag # (enrich_proc s tp dp) |
|
58 else Attach p h flag # (enrich_proc s tp dp))" |
|
59 | "enrich_proc (Detach p h # s) tp dp = ( |
|
60 if (tp = p) |
|
61 then Detach dp h # Detach p h # (enrich_proc s tp dp) |
|
62 else Detach p h # (enrich_proc s tp dp))" |
|
63 *) |
|
64 (* |
|
65 | "enrich_proc (Kill p p' # s) tp dp = ( |
|
66 if (tp = p') then Kill p p' # s |
|
67 else Kill p p' # (enrich_proc s tp dp))" |
|
68 | "enrich_proc (Exit p # s) tp dp = ( |
|
69 if (tp = p) then Exit p # s |
|
70 else Exit p # (enrich_proc s tp dp))" |
|
71 *) |
|
72 | "enrich_proc (e # s) tp dp = e # (enrich_proc s tp dp)" |
|
73 |
|
74 definition is_created_proc:: "t_state \<Rightarrow> t_process \<Rightarrow> bool" |
|
75 where |
|
76 "is_created_proc s p \<equiv> p \<in> current_procs s \<and> (p \<in> init_procs \<longrightarrow> died (O_proc p) s)" |
|
77 |
|
78 definition is_created_proc':: "t_state \<Rightarrow> t_process \<Rightarrow> bool" |
|
79 where |
|
80 "is_created_proc' s p \<equiv> p \<in> current_procs s \<and> p \<notin> init_procs" |
|
81 |
|
82 lemma created_proc_clone: |
|
83 "valid (Clone p p' fds # s) \<Longrightarrow> |
|
84 is_created_proc (Clone p p' fds # s) tp = (if (tp = p') then True else is_created_proc s tp)" |
|
85 apply (drule vt_grant_os) |
|
86 apply (auto simp:is_created_proc_def dest:not_all_procs_prop2) |
|
87 using not_died_init_proc |
|
88 by auto |
|
89 |
|
90 lemma created_proc_exit: |
|
91 "is_created_proc (Exit p # s) tp = (if (tp = p) then False else is_created_proc s tp)" |
|
92 by (simp add:is_created_proc_def) |
|
93 |
|
94 lemma created_proc_kill: |
|
95 "is_created_proc (Kill p p' # s) tp = (if (tp = p') then False else is_created_proc s tp)" |
|
96 by (simp add:is_created_proc_def) |
|
97 |
|
98 lemma created_proc_other: |
|
99 "\<lbrakk>\<And> p p' fds. e \<noteq> Clone p p' fds; |
|
100 \<And> p. e \<noteq> Exit p; |
|
101 \<And> p p'. e \<noteq> Kill p p'\<rbrakk> \<Longrightarrow> is_created_proc (e # s) tp = is_created_proc s tp" |
|
102 by (case_tac e, auto simp:is_created_proc_def) |
|
103 |
|
104 lemmas is_created_proc_simps = created_proc_clone created_proc_exit created_proc_kill created_proc_other |
|
105 |
|
106 lemma no_del_died: |
|
107 "\<lbrakk>no_del_event s; died obj s\<rbrakk> \<Longrightarrow> (\<exists> p fd. obj = O_fd p fd \<or> obj = O_tcp_sock (p, fd) \<or> obj = O_udp_sock (p, fd)) |
|
108 \<or> (\<exists> q m. obj = O_msg q m) " |
11 apply (induct s) |
109 apply (induct s) |
12 apply (simp add:init_fds_of_proc_prop1) |
110 apply simp |
|
111 apply (case_tac a) |
|
112 apply (auto split:option.splits) |
|
113 done |
|
114 |
|
115 lemma no_del_created_eq: |
|
116 "no_del_event s \<Longrightarrow> is_created_proc s p = is_created_proc' s p" |
|
117 apply (induct s) |
|
118 apply (simp add:is_created_proc_def is_created_proc'_def) |
|
119 apply (case_tac a) |
|
120 apply (auto simp add:is_created_proc_def is_created_proc'_def dest:no_del_died) |
|
121 done |
|
122 |
|
123 lemma enrich_proc_dup_in: |
|
124 "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk> |
|
125 \<Longrightarrow> p' \<in> current_procs (enrich_proc s p p')" |
|
126 apply (induct s, simp add:is_created_proc_def) |
13 apply (frule vt_grant_os, frule vd_cons) |
127 apply (frule vt_grant_os, frule vd_cons) |
14 apply (case_tac a, auto split:if_splits option.splits) |
128 apply (case_tac a, auto simp:is_created_proc_def dest:not_all_procs_prop3) |
15 done |
129 done |
16 |
130 |
17 lemma get_parentfs_ctxts_prop: |
131 lemma enrich_proc_dup_ffd: |
18 "\<lbrakk>get_parentfs_ctxts s (a # f) = Some ctxts; sectxt_of_obj s (O_dir f) = Some ctxt; valid s\<rbrakk> |
132 "\<lbrakk>file_of_proc_fd s p fd = Some f; is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk> |
19 \<Longrightarrow> ctxt \<in> set (ctxts)" |
133 \<Longrightarrow> file_of_proc_fd (enrich_proc s p p') p' fd = Some f" |
20 apply (induct f) |
134 apply (induct s, simp add:is_created_proc_def) |
21 apply (auto split:option.splits) |
135 apply (frule vt_grant_os, frule vd_cons) |
22 done |
136 apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def |
23 |
137 dest:not_all_procs_prop3 split:if_splits option.splits) |
24 lemma search_check_allp_intro: |
138 done |
25 "\<lbrakk>search_check s sp pf; get_parentfs_ctxts s pf = Some ctxts; valid s; is_dir s pf\<rbrakk> |
139 |
26 \<Longrightarrow> search_check_allp sp (set ctxts)" |
140 lemma enrich_proc_dup_ffd': |
27 apply (case_tac pf) |
141 "\<lbrakk>file_of_proc_fd (enrich_proc s p p') p' fd = Some f; is_created_proc s p; p' \<notin> all_procs s; |
28 apply (simp split:option.splits if_splits add:search_check_allp_def) |
142 no_del_event s; valid s\<rbrakk> |
29 apply (rule ballI) |
143 \<Longrightarrow> file_of_proc_fd s p fd = Some f" |
30 apply (auto simp:search_check_ctxt_def search_check_dir_def split:if_splits option.splits) |
144 apply (induct s, simp add:is_created_proc_def) |
31 apply (auto simp:search_check_allp_def search_check_file_def) |
145 apply (frule vt_grant_os, frule vd_cons) |
32 apply (frule is_dir_not_file, simp) |
146 apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def |
33 done |
147 dest:not_all_procs_prop3 split:if_splits option.splits) |
34 |
148 done |
35 lemma search_check_leveling_f: |
|
36 "\<lbrakk>search_check s sp pf; parent f = Some pf; is_file s f; valid s; |
|
37 sectxt_of_obj s (O_file f) = Some fctxt; search_check_file sp fctxt\<rbrakk> |
|
38 \<Longrightarrow> search_check s sp f" |
|
39 apply (case_tac f, simp+) |
|
40 apply (auto split:option.splits simp:search_check_ctxt_def) |
|
41 apply (frule parentf_is_dir_prop2, simp) |
|
42 apply (erule get_pfs_secs_prop, simp) |
|
43 apply (erule_tac search_check_allp_intro, simp_all) |
|
44 apply (simp add:parentf_is_dir_prop2) |
|
45 done |
|
46 |
149 |
47 lemma enrich_proc_dup_ffds': |
150 lemma enrich_proc_dup_ffds': |
48 "\<lbrakk>fd \<notin> current_proc_fds (enrich_proc s p p') p'; is_created_proc s p; p' \<notin> all_procs s; no_del_event s; valid s\<rbrakk> |
151 "\<lbrakk>fd \<notin> current_proc_fds (enrich_proc s p p') p'; is_created_proc s p; p' \<notin> all_procs s; no_del_event s; valid s\<rbrakk> |
49 \<Longrightarrow> fd \<notin> proc_file_fds s p \<and> file_of_proc_fd s p fd = None" |
152 \<Longrightarrow> fd \<notin> proc_file_fds s p \<and> file_of_proc_fd s p fd = None" |
50 apply (auto simp:enrich_proc_dup_ffds_eq_fds) |
153 apply (auto simp:enrich_proc_dup_ffds_eq_fds) |
51 apply (simp add:proc_file_fds_def) |
154 apply (simp add:proc_file_fds_def) |
|
155 done |
|
156 |
|
157 lemma enrich_proc_dup_ffd_eq: |
|
158 "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; no_del_event s; valid s\<rbrakk> |
|
159 \<Longrightarrow> file_of_proc_fd (enrich_proc s p p') p' fd = file_of_proc_fd s p fd" |
|
160 apply (case_tac "file_of_proc_fd s p fd") |
|
161 apply (case_tac[!] "file_of_proc_fd (enrich_proc s p p') p' fd") |
|
162 apply (auto dest:enrich_proc_dup_ffd enrich_proc_dup_ffd') |
|
163 done |
|
164 |
|
165 |
|
166 lemma enrich_proc_dup_fflags: |
|
167 "\<lbrakk>flags_of_proc_fd s p fd = Some flag; is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk> |
|
168 \<Longrightarrow> flags_of_proc_fd (enrich_proc s p p') p' fd = Some (remove_create_flag flag) \<or> |
|
169 flags_of_proc_fd (enrich_proc s p p') p' fd = Some flag" |
|
170 apply (induct s arbitrary:p, simp add:is_created_proc_def) |
|
171 apply (frule vt_grant_os, frule vd_cons) |
|
172 apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def is_creat_flag_def |
|
173 dest:not_all_procs_prop3 split:if_splits option.splits) |
|
174 done |
|
175 |
|
176 lemma enrich_proc_dup_ffds: |
|
177 "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; no_del_event s; valid s\<rbrakk> |
|
178 \<Longrightarrow> proc_file_fds (enrich_proc s p p') p' = proc_file_fds s p" |
|
179 apply (auto simp:proc_file_fds_def) |
|
180 apply (rule_tac x = f in exI) |
|
181 apply (erule enrich_proc_dup_ffd', simp+) |
|
182 apply (rule_tac x = f in exI) |
|
183 apply (erule enrich_proc_dup_ffd, simp+) |
|
184 done |
|
185 |
|
186 lemma enrich_proc_dup_ffds_eq_fds: |
|
187 "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; no_del_event s; valid s\<rbrakk> |
|
188 \<Longrightarrow> current_proc_fds (enrich_proc s p p') p' = proc_file_fds s p" |
|
189 apply (induct s arbitrary:p) |
|
190 apply (simp add: is_created_proc_def) |
|
191 apply (frule not_all_procs_prop3) |
|
192 apply (frule vd_cons, frule vt_grant_os, case_tac a) |
|
193 apply (auto split:if_splits option.splits dest:proc_fd_in_fds set_mp not_all_procs_prop3 |
|
194 simp:proc_file_fds_def is_created_proc_def) |
52 done |
195 done |
53 |
196 |
54 lemma enrich_proc_cur_inof: |
197 lemma enrich_proc_cur_inof: |
55 "\<lbrakk>valid s; no_del_event s\<rbrakk> \<Longrightarrow> inum_of_file (enrich_proc s p p') f = inum_of_file s f" |
198 "\<lbrakk>valid s; no_del_event s\<rbrakk> \<Longrightarrow> inum_of_file (enrich_proc s p p') f = inum_of_file s f" |
56 apply (induct s arbitrary:f) |
199 apply (induct s arbitrary:f) |