137 |
137 |
138 lemma file_has_parent: "\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> pf. is_dir s pf \<and> parent f = Some pf" |
138 lemma file_has_parent: "\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> pf. is_dir s pf \<and> parent f = Some pf" |
139 apply (case_tac f) |
139 apply (case_tac f) |
140 apply (simp, drule root_is_dir', simp+) |
140 apply (simp, drule root_is_dir', simp+) |
141 apply (simp add:parentf_is_dir_prop2) |
141 apply (simp add:parentf_is_dir_prop2) |
142 done |
|
143 |
|
144 (* enrich s target_proc duplicated_pro *) |
|
145 fun enrich_proc :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> t_state" |
|
146 where |
|
147 "enrich_proc [] tp dp = []" |
|
148 | "enrich_proc (Execve p f fds # s) tp dp = ( |
|
149 if (tp = p) |
|
150 then Execve dp f (fds \<inter> proc_file_fds s p) # Execve p f fds # (enrich_proc s tp dp) |
|
151 else Execve p f fds # (enrich_proc s tp dp))" |
|
152 | "enrich_proc (Clone p p' fds # s) tp dp = ( |
|
153 if (tp = p') |
|
154 then Clone p dp (fds \<inter> proc_file_fds s p) # Clone p p' fds # s |
|
155 else Clone p p' fds # (enrich_proc s tp dp))" |
|
156 | "enrich_proc (Open p f flags fd opt # s) tp dp = ( |
|
157 if (tp = p) |
|
158 then Open dp f (remove_create_flag flags) fd None # Open p f flags fd opt # (enrich_proc s tp dp) |
|
159 else Open p f flags fd opt # (enrich_proc s tp dp))" |
|
160 | "enrich_proc (ReadFile p fd # s) tp dp = ( |
|
161 if (tp = p) |
|
162 then ReadFile dp fd # ReadFile p fd # (enrich_proc s tp dp) |
|
163 else ReadFile p fd # (enrich_proc s tp dp))" |
|
164 (* |
|
165 | "enrich_proc (CloseFd p fd # s) tp dp = ( |
|
166 if (tp = p \<and> fd \<in> proc_file_fds s p) |
|
167 then CloseFd dp fd # CloseFd p fd # (enrich_proc s tp dp) |
|
168 else CloseFd p fd # (enrich_proc s tp dp))" |
|
169 *) |
|
170 (* |
|
171 | "enrich_proc (Attach p h flag # s) tp dp = ( |
|
172 if (tp = p) |
|
173 then Attach dp h flag # Attach p h flag # (enrich_proc s tp dp) |
|
174 else Attach p h flag # (enrich_proc s tp dp))" |
|
175 | "enrich_proc (Detach p h # s) tp dp = ( |
|
176 if (tp = p) |
|
177 then Detach dp h # Detach p h # (enrich_proc s tp dp) |
|
178 else Detach p h # (enrich_proc s tp dp))" |
|
179 *) |
|
180 (* |
|
181 | "enrich_proc (Kill p p' # s) tp dp = ( |
|
182 if (tp = p') then Kill p p' # s |
|
183 else Kill p p' # (enrich_proc s tp dp))" |
|
184 | "enrich_proc (Exit p # s) tp dp = ( |
|
185 if (tp = p) then Exit p # s |
|
186 else Exit p # (enrich_proc s tp dp))" |
|
187 *) |
|
188 | "enrich_proc (e # s) tp dp = e # (enrich_proc s tp dp)" |
|
189 |
|
190 definition is_created_proc:: "t_state \<Rightarrow> t_process \<Rightarrow> bool" |
|
191 where |
|
192 "is_created_proc s p \<equiv> p \<in> current_procs s \<and> (p \<in> init_procs \<longrightarrow> died (O_proc p) s)" |
|
193 |
|
194 definition is_created_proc':: "t_state \<Rightarrow> t_process \<Rightarrow> bool" |
|
195 where |
|
196 "is_created_proc' s p \<equiv> p \<in> current_procs s \<and> p \<notin> init_procs" |
|
197 |
|
198 lemma no_del_died: |
|
199 "\<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)) |
|
200 \<or> (\<exists> q m. obj = O_msg q m) " |
|
201 apply (induct s) |
|
202 apply simp |
|
203 apply (case_tac a) |
|
204 apply (auto split:option.splits) |
|
205 done |
|
206 |
|
207 lemma no_del_created_eq: |
|
208 "no_del_event s \<Longrightarrow> is_created_proc s p = is_created_proc' s p" |
|
209 apply (induct s) |
|
210 apply (simp add:is_created_proc_def is_created_proc'_def) |
|
211 apply (case_tac a) |
|
212 apply (auto simp add:is_created_proc_def is_created_proc'_def dest:no_del_died) |
|
213 done |
142 done |
214 |
143 |
215 lemma enrich_search_check: |
144 lemma enrich_search_check: |
216 assumes grant: "search_check s (up, rp, tp) f" |
145 assumes grant: "search_check s (up, rp, tp) f" |
217 and cf2sf: "\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile s' f = cf2sfile s f" |
146 and cf2sf: "\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile s' f = cf2sfile s f" |
1246 show ?thesis using grant |
1175 show ?thesis using grant |
1247 by (simp add:Shutdown) |
1176 by (simp add:Shutdown) |
1248 qed |
1177 qed |
1249 qed |
1178 qed |
1250 |
1179 |
1251 lemma created_proc_clone: |
1180 lemma current_proc_fds_in_curp: |
1252 "valid (Clone p p' fds # s) \<Longrightarrow> |
1181 "\<lbrakk>fd \<in> current_proc_fds s p; valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s" |
1253 is_created_proc (Clone p p' fds # s) tp = (if (tp = p') then True else is_created_proc s tp)" |
1182 apply (induct s) |
1254 apply (drule vt_grant_os) |
1183 apply (simp add:init_fds_of_proc_prop1) |
1255 apply (auto simp:is_created_proc_def dest:not_all_procs_prop2) |
|
1256 using not_died_init_proc |
|
1257 by auto |
|
1258 |
|
1259 lemma created_proc_exit: |
|
1260 "is_created_proc (Exit p # s) tp = (if (tp = p) then False else is_created_proc s tp)" |
|
1261 by (simp add:is_created_proc_def) |
|
1262 |
|
1263 lemma created_proc_kill: |
|
1264 "is_created_proc (Kill p p' # s) tp = (if (tp = p') then False else is_created_proc s tp)" |
|
1265 by (simp add:is_created_proc_def) |
|
1266 |
|
1267 lemma created_proc_other: |
|
1268 "\<lbrakk>\<And> p p' fds. e \<noteq> Clone p p' fds; |
|
1269 \<And> p. e \<noteq> Exit p; |
|
1270 \<And> p p'. e \<noteq> Kill p p'\<rbrakk> \<Longrightarrow> is_created_proc (e # s) tp = is_created_proc s tp" |
|
1271 by (case_tac e, auto simp:is_created_proc_def) |
|
1272 |
|
1273 lemmas is_created_proc_simps = created_proc_clone created_proc_exit created_proc_kill created_proc_other |
|
1274 (* |
|
1275 |
|
1276 |
|
1277 |
|
1278 |
|
1279 (p \<in> current_procs s \<longrightarrow> co2sobj (enrich_proc s p p') (O_proc p') = co2sobj (enrich_proc s p p') (O_proc p)) \<and> |
|
1280 (\<forall> obj. alive s obj \<longrightarrow> alive (enrich_proc s p p') obj) \<and> |
|
1281 (\<forall> p'. p' \<in> current_procs s \<longrightarrow> cp2sproc (enrich_proc s p p') p' = cp2sproc s p) \<and> |
|
1282 (\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile (enrich_proc s p p') f = cf2sfile s f) \<and> |
|
1283 (Tainted (enrich_proc s p p') = (Tainted s \<union> (if (O_proc p \<in> Tainted s) then {O_proc p'} else {})))"*) |
|
1284 |
|
1285 lemma enrich_proc_dup_in: |
|
1286 "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk> |
|
1287 \<Longrightarrow> p' \<in> current_procs (enrich_proc s p p')" |
|
1288 apply (induct s, simp add:is_created_proc_def) |
|
1289 apply (frule vt_grant_os, frule vd_cons) |
1184 apply (frule vt_grant_os, frule vd_cons) |
1290 apply (case_tac a, auto simp:is_created_proc_def dest:not_all_procs_prop3) |
1185 apply (case_tac a, auto split:if_splits option.splits) |
1291 done |
1186 done |
1292 |
1187 |
1293 lemma enrich_proc_dup_ffd: |
1188 lemma get_parentfs_ctxts_prop: |
1294 "\<lbrakk>file_of_proc_fd s p fd = Some f; is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk> |
1189 "\<lbrakk>get_parentfs_ctxts s (a # f) = Some ctxts; sectxt_of_obj s (O_dir f) = Some ctxt; valid s\<rbrakk> |
1295 \<Longrightarrow> file_of_proc_fd (enrich_proc s p p') p' fd = Some f" |
1190 \<Longrightarrow> ctxt \<in> set (ctxts)" |
1296 apply (induct s, simp add:is_created_proc_def) |
1191 apply (induct f) |
1297 apply (frule vt_grant_os, frule vd_cons) |
1192 apply (auto split:option.splits) |
1298 apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def |
|
1299 dest:not_all_procs_prop3 split:if_splits option.splits) |
|
1300 done |
|
1301 |
|
1302 lemma enrich_proc_dup_ffd': |
|
1303 "\<lbrakk>file_of_proc_fd (enrich_proc s p p') p' fd = Some f; is_created_proc s p; p' \<notin> all_procs s; |
|
1304 no_del_event s; valid s\<rbrakk> |
|
1305 \<Longrightarrow> file_of_proc_fd s p fd = Some f" |
|
1306 apply (induct s, simp add:is_created_proc_def) |
|
1307 apply (frule vt_grant_os, frule vd_cons) |
|
1308 apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def |
|
1309 dest:not_all_procs_prop3 split:if_splits option.splits) |
|
1310 done |
|
1311 |
|
1312 lemma enrich_proc_dup_ffd_eq: |
|
1313 "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; no_del_event s; valid s\<rbrakk> |
|
1314 \<Longrightarrow> file_of_proc_fd (enrich_proc s p p') p' fd = file_of_proc_fd s p fd" |
|
1315 apply (case_tac "file_of_proc_fd s p fd") |
|
1316 apply (case_tac[!] "file_of_proc_fd (enrich_proc s p p') p' fd") |
|
1317 apply (auto dest:enrich_proc_dup_ffd enrich_proc_dup_ffd') |
|
1318 done |
1193 done |
|
1194 |
|
1195 lemma search_check_allp_intro: |
|
1196 "\<lbrakk>search_check s sp pf; get_parentfs_ctxts s pf = Some ctxts; valid s; is_dir s pf\<rbrakk> |
|
1197 \<Longrightarrow> search_check_allp sp (set ctxts)" |
|
1198 apply (case_tac pf) |
|
1199 apply (simp split:option.splits if_splits add:search_check_allp_def) |
|
1200 apply (rule ballI) |
|
1201 apply (auto simp:search_check_ctxt_def search_check_dir_def split:if_splits option.splits) |
|
1202 apply (auto simp:search_check_allp_def search_check_file_def) |
|
1203 apply (frule is_dir_not_file, simp) |
|
1204 done |
|
1205 |
|
1206 lemma search_check_leveling_f: |
|
1207 "\<lbrakk>search_check s sp pf; parent f = Some pf; is_file s f; valid s; |
|
1208 sectxt_of_obj s (O_file f) = Some fctxt; search_check_file sp fctxt\<rbrakk> |
|
1209 \<Longrightarrow> search_check s sp f" |
|
1210 apply (case_tac f, simp+) |
|
1211 apply (auto split:option.splits simp:search_check_ctxt_def) |
|
1212 apply (frule parentf_is_dir_prop2, simp) |
|
1213 apply (erule get_pfs_secs_prop, simp) |
|
1214 apply (erule_tac search_check_allp_intro, simp_all) |
|
1215 apply (simp add:parentf_is_dir_prop2) |
|
1216 done |
|
1217 |
1319 |
1218 |
1320 lemma current_fflag_in_fds: |
1219 lemma current_fflag_in_fds: |
1321 "\<lbrakk>flags_of_proc_fd s p fd = Some flag; valid s\<rbrakk> \<Longrightarrow> fd \<in> current_proc_fds s p" |
1220 "\<lbrakk>flags_of_proc_fd s p fd = Some flag; valid s\<rbrakk> \<Longrightarrow> fd \<in> current_proc_fds s p" |
1322 apply (induct s arbitrary:p) |
1221 apply (induct s arbitrary:p) |
1323 apply (simp add:flags_of_proc_fd.simps file_of_proc_fd.simps init_oflags_prop2) |
1222 apply (simp add:flags_of_proc_fd.simps file_of_proc_fd.simps init_oflags_prop2) |
1331 apply (simp add: file_of_proc_fd.simps init_fileflag_valid) |
1230 apply (simp add: file_of_proc_fd.simps init_fileflag_valid) |
1332 apply (frule vd_cons, frule vt_grant_os, case_tac a) |
1231 apply (frule vd_cons, frule vt_grant_os, case_tac a) |
1333 apply (auto split:if_splits option.splits dest:proc_fd_in_fds) |
1232 apply (auto split:if_splits option.splits dest:proc_fd_in_fds) |
1334 done |
1233 done |
1335 |
1234 |
1336 lemma enrich_proc_dup_fflags: |
|
1337 "\<lbrakk>flags_of_proc_fd s p fd = Some flag; is_created_proc s p; p' \<notin> all_procs s; valid s\<rbrakk> |
|
1338 \<Longrightarrow> flags_of_proc_fd (enrich_proc s p p') p' fd = Some (remove_create_flag flag) \<or> |
|
1339 flags_of_proc_fd (enrich_proc s p p') p' fd = Some flag" |
|
1340 apply (induct s arbitrary:p, simp add:is_created_proc_def) |
|
1341 apply (frule vt_grant_os, frule vd_cons) |
|
1342 apply (case_tac a, auto simp:is_created_proc_def proc_file_fds_def is_creat_flag_def |
|
1343 dest:not_all_procs_prop3 split:if_splits option.splits) |
|
1344 done |
|
1345 |
|
1346 lemma enrich_proc_dup_ffds: |
|
1347 "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; no_del_event s; valid s\<rbrakk> |
|
1348 \<Longrightarrow> proc_file_fds (enrich_proc s p p') p' = proc_file_fds s p" |
|
1349 apply (auto simp:proc_file_fds_def) |
|
1350 apply (rule_tac x = f in exI) |
|
1351 apply (erule enrich_proc_dup_ffd', simp+) |
|
1352 apply (rule_tac x = f in exI) |
|
1353 apply (erule enrich_proc_dup_ffd, simp+) |
|
1354 done |
|
1355 |
|
1356 lemma enrich_proc_dup_ffds_eq_fds: |
|
1357 "\<lbrakk>is_created_proc s p; p' \<notin> all_procs s; no_del_event s; valid s\<rbrakk> |
|
1358 \<Longrightarrow> current_proc_fds (enrich_proc s p p') p' = proc_file_fds s p" |
|
1359 apply (induct s arbitrary:p) |
|
1360 apply (simp add: is_created_proc_def) |
|
1361 apply (frule not_all_procs_prop3) |
|
1362 apply (frule vd_cons, frule vt_grant_os, case_tac a) |
|
1363 apply (auto split:if_splits option.splits dest:proc_fd_in_fds set_mp not_all_procs_prop3 |
|
1364 simp:proc_file_fds_def is_created_proc_def) |
|
1365 done |
|
1366 |
|
1367 lemma oflags_check_remove_create: |
1235 lemma oflags_check_remove_create: |
1368 "oflags_check flags sp sf \<Longrightarrow> oflags_check (remove_create_flag flags) sp sf" |
1236 "oflags_check flags sp sf \<Longrightarrow> oflags_check (remove_create_flag flags) sp sf" |
1369 apply (case_tac flags) |
1237 apply (case_tac flags) |
1370 apply (auto simp:oflags_check_def perms_of_flags_def perm_of_oflag_def split:bool.splits) |
1238 apply (auto simp:oflags_check_def perms_of_flags_def perm_of_oflag_def split:bool.splits) |
1371 done |
1239 done |