1303 "oflags_check flags sp sf \<Longrightarrow> oflags_check (remove_create_flag flags) sp sf" |
1299 "oflags_check flags sp sf \<Longrightarrow> oflags_check (remove_create_flag flags) sp sf" |
1304 apply (case_tac flags) |
1300 apply (case_tac flags) |
1305 apply (auto simp:oflags_check_def perms_of_flags_def perm_of_oflag_def split:bool.splits) |
1301 apply (auto simp:oflags_check_def perms_of_flags_def perm_of_oflag_def split:bool.splits) |
1306 done |
1302 done |
1307 |
1303 |
|
1304 fun enrich_msgq :: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_msgq \<Rightarrow> t_state" |
|
1305 where |
|
1306 "enrich_msgq [] tq dq = []" |
|
1307 | "enrich_msgq (CreateMsgq p q # s) tq dq = |
|
1308 (if (tq = q) |
|
1309 then (CreateMsgq p dq # CreateMsgq p q # s) |
|
1310 else CreateMsgq p q # (enrich_msgq s tq dq))" |
|
1311 | "enrich_msgq (SendMsg p q m # s) tq dq = |
|
1312 (if (tq = q) |
|
1313 then (SendMsg p dq m # SendMsg p q m # (enrich_msgq s tq dq)) |
|
1314 else SendMsg p q m # (enrich_msgq s tq dq))" |
|
1315 | "enrich_msgq (RecvMsg p q m # s) tq dq = |
|
1316 (if (tq = q) |
|
1317 then (RecvMsg p dq m # RecvMsg p q m # (enrich_msgq s tq dq)) |
|
1318 else RecvMsg p q m # (enrich_msgq s tq dq))" |
|
1319 | "enrich_msgq (e # s) tq dq = e # (enrich_msgq s tq dq)" |
|
1320 |
|
1321 lemma enrich_msgq_duq_in: |
|
1322 "\<lbrakk>q' \<notin> current_msgqs s; q \<in> current_msgqs s; no_del_event s; valid s\<rbrakk> |
|
1323 \<Longrightarrow> q' \<in> current_msgqs (enrich_msgq s q q')" |
|
1324 apply (induct s, simp) |
|
1325 apply (frule vt_grant_os, frule vd_cons, case_tac a) |
|
1326 apply auto |
|
1327 done |
|
1328 |
|
1329 lemma enrich_msgq_duq_sms: |
|
1330 "\<lbrakk>q' \<notin> current_msgqs s; q \<in> current_msgqs s; no_del_event s; valid s\<rbrakk> |
|
1331 \<Longrightarrow> msgs_of_queue (enrich_msgq s q q') q' = msgs_of_queue s q" |
|
1332 apply (induct s, simp) |
|
1333 apply (frule vt_grant_os, frule vd_cons, case_tac a) |
|
1334 apply auto |
|
1335 done |
|
1336 |
|
1337 lemma enrich_msgq_cur_inof: |
|
1338 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1339 \<Longrightarrow> inum_of_file (enrich_msgq s q q') f = inum_of_file s f" |
|
1340 apply (induct s arbitrary:f, simp) |
|
1341 apply (frule vt_grant_os, frule vd_cons, case_tac a) |
|
1342 apply (auto split:option.splits) |
|
1343 done |
|
1344 |
|
1345 lemma enrich_msgq_cur_inos: |
|
1346 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1347 \<Longrightarrow> inum_of_socket (enrich_msgq s q q') = inum_of_socket s" |
|
1348 apply (rule ext) |
|
1349 apply (induct s, simp) |
|
1350 apply (frule vt_grant_os, frule vd_cons, case_tac a) |
|
1351 apply (auto split:option.splits) |
|
1352 done |
|
1353 |
|
1354 lemma enrich_msgq_cur_inos': |
|
1355 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1356 \<Longrightarrow> inum_of_socket (enrich_msgq s q q') sock = inum_of_socket s sock" |
|
1357 apply (simp add:enrich_msgq_cur_inos) |
|
1358 done |
|
1359 |
|
1360 lemma enrich_msgq_cur_inums: |
|
1361 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1362 \<Longrightarrow> current_inode_nums (enrich_msgq s q q') = current_inode_nums s" |
|
1363 apply (auto simp:current_inode_nums_def current_file_inums_def |
|
1364 current_sock_inums_def enrich_msgq_cur_inof enrich_msgq_cur_inos) |
|
1365 done |
|
1366 |
|
1367 lemma enrich_msgq_cur_itag: |
|
1368 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1369 \<Longrightarrow> itag_of_inum (enrich_msgq s q q') = itag_of_inum s" |
|
1370 apply (rule ext) |
|
1371 apply (induct s, simp) |
|
1372 apply (frule vt_grant_os, frule vd_cons, case_tac a) |
|
1373 apply (auto split:option.splits t_socket_type.splits) |
|
1374 done |
|
1375 |
|
1376 lemma enrich_msgq_cur_tcps: |
|
1377 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1378 \<Longrightarrow> is_tcp_sock (enrich_msgq s q q') = is_tcp_sock s" |
|
1379 apply (rule ext) |
|
1380 apply (auto simp:is_tcp_sock_def enrich_msgq_cur_itag enrich_msgq_cur_inos |
|
1381 split:option.splits t_inode_tag.splits) |
|
1382 done |
|
1383 |
|
1384 lemma enrich_msgq_cur_udps: |
|
1385 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1386 \<Longrightarrow> is_udp_sock (enrich_msgq s q q') = is_udp_sock s" |
|
1387 apply (rule ext) |
|
1388 apply (auto simp:is_udp_sock_def enrich_msgq_cur_itag enrich_msgq_cur_inos |
|
1389 split:option.splits t_inode_tag.splits) |
|
1390 done |
|
1391 |
|
1392 lemma enrich_msgq_cur_msgqs: |
|
1393 "\<lbrakk>q \<in> current_msgqs s; no_del_event s; valid s\<rbrakk> |
|
1394 \<Longrightarrow> current_msgqs (enrich_msgq s q q') = current_msgqs s \<union> {q'}" |
|
1395 apply (induct s, simp) |
|
1396 apply (frule vt_grant_os, frule vd_cons) |
|
1397 apply (case_tac a, auto) |
|
1398 done |
|
1399 |
|
1400 lemma enrich_msgq_cur_msgs: |
|
1401 "\<lbrakk>q' \<notin> current_msgqs s; q \<in> current_msgqs s; no_del_event s; valid s\<rbrakk> |
|
1402 \<Longrightarrow> msgs_of_queue (enrich_msgq s q q') = (msgs_of_queue s) (q' := msgs_of_queue s q)" |
|
1403 apply (rule ext, simp, rule conjI, rule impI) |
|
1404 apply (simp add:enrich_msgq_duq_sms) |
|
1405 apply (rule impI) |
|
1406 apply (induct s, simp) |
|
1407 apply (frule vt_grant_os, frule vd_cons) |
|
1408 apply (case_tac a, auto) |
|
1409 done |
|
1410 |
|
1411 lemma enrich_msgq_cur_procs: |
|
1412 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1413 \<Longrightarrow> current_procs (enrich_msgq s q q') = current_procs s" |
|
1414 apply (induct s, simp) |
|
1415 apply (frule vt_grant_os, frule vd_cons) |
|
1416 apply (case_tac a, auto) |
|
1417 done |
|
1418 |
|
1419 lemma enrich_msgq_cur_files: |
|
1420 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1421 \<Longrightarrow> current_files (enrich_msgq s q q') = current_files s" |
|
1422 apply (auto simp:current_files_def) |
|
1423 apply (simp add:enrich_msgq_cur_inof)+ |
|
1424 done |
|
1425 |
|
1426 lemma enrich_msgq_cur_fds: |
|
1427 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1428 \<Longrightarrow> current_proc_fds (enrich_msgq s q q') = current_proc_fds s" |
|
1429 apply (induct s, simp) |
|
1430 apply (frule vt_grant_os, frule vd_cons, case_tac a) |
|
1431 apply auto |
|
1432 done |
|
1433 |
|
1434 lemma enrich_msgq_filefd: |
|
1435 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1436 \<Longrightarrow> file_of_proc_fd (enrich_msgq s q q') = file_of_proc_fd s" |
|
1437 apply (rule ext, rule ext) |
|
1438 apply (induct s, simp) |
|
1439 apply (frule vt_grant_os, frule vd_cons, case_tac a) |
|
1440 apply auto |
|
1441 done |
|
1442 |
|
1443 lemma enrich_msgq_flagfd: |
|
1444 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1445 \<Longrightarrow> flags_of_proc_fd (enrich_msgq s q q') = flags_of_proc_fd s" |
|
1446 apply (rule ext, rule ext) |
|
1447 apply (induct s, simp) |
|
1448 apply (frule vt_grant_os, frule vd_cons, case_tac a) |
|
1449 apply auto |
|
1450 done |
|
1451 |
|
1452 lemma enrich_msgq_proc_fds: |
|
1453 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1454 \<Longrightarrow> proc_file_fds (enrich_msgq s q q') = proc_file_fds s" |
|
1455 apply (auto simp:proc_file_fds_def enrich_msgq_filefd) |
|
1456 done |
|
1457 |
|
1458 lemma enrich_msgq_hungs: |
|
1459 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1460 \<Longrightarrow> files_hung_by_del (enrich_msgq s q q') = files_hung_by_del s" |
|
1461 apply (induct s, simp) |
|
1462 apply (frule vt_grant_os, frule vd_cons, case_tac a) |
|
1463 apply (auto simp:files_hung_by_del.simps) |
|
1464 done |
|
1465 |
|
1466 lemma enrich_msgq_is_file: |
|
1467 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1468 \<Longrightarrow> is_file (enrich_msgq s q q') = is_file s" |
|
1469 apply (rule ext) |
|
1470 apply (auto simp add:is_file_def enrich_msgq_cur_itag enrich_msgq_cur_inof |
|
1471 split:option.splits t_inode_tag.splits) |
|
1472 done |
|
1473 |
|
1474 lemma enrich_msgq_is_dir: |
|
1475 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1476 \<Longrightarrow> is_dir (enrich_msgq s q q') = is_dir s" |
|
1477 apply (rule ext) |
|
1478 apply (auto simp add:is_dir_def enrich_msgq_cur_itag enrich_msgq_cur_inof |
|
1479 split:option.splits t_inode_tag.splits) |
|
1480 done |
|
1481 |
|
1482 lemma enrich_msgq_sameinode: |
|
1483 "\<lbrakk>no_del_event s; valid s\<rbrakk> |
|
1484 \<Longrightarrow> (f \<in> same_inode_files (enrich_msgq s q q') f') = (f \<in> same_inode_files s f')" |
|
1485 apply (induct s, simp) |
|
1486 apply (simp add:same_inode_files_def) |
|
1487 apply (auto simp:enrich_msgq_is_file enrich_msgq_cur_inof) |
|
1488 done |
|
1489 |
|
1490 lemma enrich_msgq_tainted_aux1: |
|
1491 "\<lbrakk>no_del_event s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; valid s\<rbrakk> |
|
1492 \<Longrightarrow> (tainted s \<union> {O_msg q' m| m. O_msg q m \<in> tainted s}) \<subseteq> tainted (enrich_msgq s q q')" |
|
1493 apply (induct s, simp) |
|
1494 apply (frule vt_grant_os, frule vd_cons) |
|
1495 apply (case_tac a) |
|
1496 apply (auto split:option.splits if_splits dest:tainted_in_current |
|
1497 simp:enrich_msgq_filefd enrich_msgq_sameinode) |
|
1498 done |
|
1499 |
|
1500 lemma enrich_msgq_tainted_aux2: |
|
1501 "\<lbrakk>no_del_event s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; valid s; valid (enrich_msgq s q q')\<rbrakk> |
|
1502 \<Longrightarrow> tainted (enrich_msgq s q q') \<subseteq> (tainted s \<union> {O_msg q' m| m. O_msg q m \<in> tainted s})" |
|
1503 apply (induct s, simp) |
|
1504 apply (frule vt_grant_os, frule vd_cons) |
|
1505 apply (case_tac a) |
|
1506 apply (auto split:option.splits if_splits simp:enrich_msgq_filefd enrich_msgq_sameinode |
|
1507 dest:tainted_in_current vd_cons) |
|
1508 apply (drule_tac vd_cons)+ |
|
1509 apply (simp) |
|
1510 apply (drule set_mp) |
|
1511 apply simp |
|
1512 apply simp |
|
1513 apply (drule_tac s = s in tainted_in_current) |
|
1514 apply simp+ |
|
1515 done |
|
1516 |
|
1517 lemma enrich_msgq_alive: |
|
1518 "\<lbrakk>alive s obj; valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk> |
|
1519 \<Longrightarrow> alive (enrich_msgq s q q') obj" |
|
1520 apply (case_tac obj) |
|
1521 apply (simp_all add:enrich_msgq_is_file enrich_msgq_is_dir enrich_msgq_cur_msgqs |
|
1522 enrich_msgq_cur_msgs enrich_msgq_cur_procs enrich_msgq_cur_fds |
|
1523 enrich_msgq_cur_tcps enrich_msgq_cur_udps) |
|
1524 apply (rule impI, simp) |
|
1525 done |
|
1526 |
|
1527 lemma enrich_msgq_alive': |
|
1528 "\<lbrakk>alive (enrich_msgq s q q') obj; valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk> |
|
1529 \<Longrightarrow> alive s obj \<or> obj = O_msgq q' \<or> (\<exists> m. obj = O_msg q' m \<and> alive s (O_msg q m))" |
|
1530 apply (case_tac obj) |
|
1531 apply (simp_all add:enrich_msgq_is_file enrich_msgq_is_dir enrich_msgq_cur_msgqs |
|
1532 enrich_msgq_cur_msgs enrich_msgq_cur_procs enrich_msgq_cur_fds |
|
1533 enrich_msgq_cur_tcps enrich_msgq_cur_udps) |
|
1534 apply (auto split:if_splits) |
|
1535 done |
|
1536 |
|
1537 lemma enrich_msgq_not_alive: |
|
1538 "\<lbrakk>enrich_not_alive s (E_msgq q') obj; q' \<notin> current_msgqs s; q \<in> current_msgqs s; no_del_event s; valid s\<rbrakk> |
|
1539 \<Longrightarrow> enrich_not_alive (enrich_msgq s q q') (E_msgq q') obj" |
|
1540 apply (case_tac obj) |
|
1541 apply (auto simp:enrich_msgq_cur_fds enrich_msgq_cur_files |
|
1542 enrich_msgq_cur_procs enrich_msgq_cur_inums enrich_msgq_cur_msgqs enrich_msgq_cur_msgs) |
|
1543 done |
|
1544 |
|
1545 lemma enrich_msgq_nodel: |
|
1546 "no_del_event (enrich_msgq s q q') = no_del_event s" |
|
1547 apply (induct s, simp) |
|
1548 by (case_tac a, auto) |
|
1549 |
|
1550 lemma enrich_msgq_died_proc: |
|
1551 "died (O_proc p) (enrich_msgq s q q') = died (O_proc p) s" |
|
1552 apply (induct s, simp) |
|
1553 by (case_tac a, auto) |
|
1554 |
|
1555 lemma cf2sfile_execve: |
|
1556 "\<lbrakk>ff \<in> current_files (Execve p f fds # s); valid (Execve p f fds # s)\<rbrakk> |
|
1557 \<Longrightarrow> cf2sfile (Execve p f fds # s) ff= cf2sfile s ff" |
|
1558 by (auto dest:cf2sfile_other' simp:current_files_simps) |
|
1559 lemma cf2sfile_clone: |
|
1560 "\<lbrakk>ff \<in> current_files (Clone p p' fds # s); valid (Clone p p' fds # s)\<rbrakk> |
|
1561 \<Longrightarrow> cf2sfile (Clone p p' fds # s) ff= cf2sfile s ff" |
|
1562 by (auto dest:cf2sfile_other' simp:current_files_simps) |
|
1563 lemma cf2sfile_ptrace: |
|
1564 "\<lbrakk>ff \<in> current_files (Ptrace p p' # s); valid (Ptrace p p' # s)\<rbrakk> |
|
1565 \<Longrightarrow> cf2sfile (Ptrace p p' # s) ff= cf2sfile s ff" |
|
1566 by (auto dest:cf2sfile_other' simp:current_files_simps) |
|
1567 lemma cf2sfile_readfile: |
|
1568 "\<lbrakk>ff \<in> current_files (ReadFile p fd # s); valid (ReadFile p fd # s)\<rbrakk> |
|
1569 \<Longrightarrow> cf2sfile (ReadFile p fd # s) ff= cf2sfile s ff" |
|
1570 by (auto dest:cf2sfile_other' simp:current_files_simps) |
|
1571 lemma cf2sfile_writefile: |
|
1572 "\<lbrakk>ff \<in> current_files (WriteFile p fd # s); valid (WriteFile p fd # s)\<rbrakk> |
|
1573 \<Longrightarrow> cf2sfile (WriteFile p fd # s) ff= cf2sfile s ff" |
|
1574 by (auto dest:cf2sfile_other' simp:current_files_simps) |
|
1575 lemma cf2sfile_truncate: |
|
1576 "\<lbrakk>ff \<in> current_files (Truncate p f len # s); valid (Truncate p f len # s)\<rbrakk> |
|
1577 \<Longrightarrow> cf2sfile (Truncate p f len # s) ff= cf2sfile s ff" |
|
1578 by (auto dest:cf2sfile_other' simp:current_files_simps) |
|
1579 lemma cf2sfile_createmsgq: |
|
1580 "\<lbrakk>ff \<in> current_files (CreateMsgq p q # s); valid (CreateMsgq p q # s)\<rbrakk> |
|
1581 \<Longrightarrow> cf2sfile (CreateMsgq p q # s) ff= cf2sfile s ff" |
|
1582 by (auto dest:cf2sfile_other' simp:current_files_simps) |
|
1583 lemma cf2sfile_sendmsg: |
|
1584 "\<lbrakk>ff \<in> current_files (SendMsg p q m # s); valid (SendMsg p q m # s)\<rbrakk> |
|
1585 \<Longrightarrow> cf2sfile (SendMsg p q m # s) ff = cf2sfile s ff" |
|
1586 by (auto dest:cf2sfile_other' simp:current_files_simps) |
|
1587 lemma cf2sfile_recvmsg: |
|
1588 "\<lbrakk>ff \<in> current_files (RecvMsg p q m # s); valid (RecvMsg p q m # s)\<rbrakk> |
|
1589 \<Longrightarrow> cf2sfile (RecvMsg p q m # s) ff = cf2sfile s ff" |
|
1590 by (auto dest:cf2sfile_other' simp:current_files_simps) |
|
1591 lemmas cf2sfile_other'' = cf2sfile_recvmsg cf2sfile_sendmsg cf2sfile_createmsgq cf2sfile_truncate |
|
1592 cf2sfile_writefile cf2sfile_readfile cf2sfile_ptrace cf2sfile_clone cf2sfile_execve |
|
1593 |
|
1594 lemma enrich_msgq_prop: |
|
1595 "\<lbrakk>valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk> |
|
1596 \<Longrightarrow> valid (enrich_msgq s q q') \<and> |
|
1597 (\<forall> tp. tp \<in> current_procs s \<longrightarrow> cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp) \<and> |
|
1598 (\<forall> f. f \<in> current_files s \<longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f) \<and> |
|
1599 (\<forall> tq. tq \<in> current_msgqs s \<longrightarrow> cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq) \<and> |
|
1600 (\<forall> tp fd. fd \<in> proc_file_fds s tp \<longrightarrow> cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd) \<and> |
|
1601 (cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q) \<and> |
|
1602 (tainted (enrich_msgq s q q') = (tainted s \<union> {O_msg q' m| m. O_msg q m \<in> tainted s}))" |
|
1603 proof (induct s) |
|
1604 case Nil |
|
1605 thus ?case by (auto) |
|
1606 next |
|
1607 case (Cons e s) |
|
1608 hence vd_cons': "valid (e # s)" and curq_cons: "q \<in> current_msgqs (e # s)" |
|
1609 and curq'_cons: "q' \<notin> current_msgqs (e # s)" and nodel_cons: "no_del_event (e # s)" |
|
1610 and os: "os_grant s e" and grant: "grant s e" and vd: "valid s" |
|
1611 by (auto dest:vd_cons vt_grant_os vt_grant) |
|
1612 from curq'_cons nodel_cons have curq': "q' \<notin> current_msgqs s" by (case_tac e, auto) |
|
1613 from nodel_cons have nodel: "no_del_event s" by (case_tac e, auto) |
|
1614 from nodel curq' vd Cons |
|
1615 have pre: "q \<in> current_msgqs s \<Longrightarrow> valid (enrich_msgq s q q') \<and> |
|
1616 (\<forall>tp. tp \<in> current_procs s \<longrightarrow> cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp) \<and> |
|
1617 (\<forall>f. f \<in> current_files s \<longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f) \<and> |
|
1618 (\<forall>tq. tq \<in> current_msgqs s \<longrightarrow> cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq) \<and> |
|
1619 (\<forall>tp fd. fd \<in> proc_file_fds s tp \<longrightarrow> cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd) \<and> |
|
1620 (cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q) \<and> |
|
1621 (tainted (enrich_msgq s q q') = (tainted s \<union> {O_msg q' m| m. O_msg q m \<in> tainted s}))" |
|
1622 by auto |
|
1623 |
|
1624 from pre have pre_vd: "q \<in> current_msgqs s \<Longrightarrow> valid (enrich_msgq s q q')" by simp |
|
1625 from pre have pre_sp: "\<And> tp. \<lbrakk>tp \<in> current_procs s; q \<in> current_msgqs s\<rbrakk> |
|
1626 \<Longrightarrow> cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp" by auto |
|
1627 from pre have pre_sf: "\<And> f. \<lbrakk>f \<in> current_files s; q \<in> current_msgqs s\<rbrakk> |
|
1628 \<Longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f" by auto |
|
1629 from pre have pre_sq: "\<And> tq. \<lbrakk>tq \<in> current_msgqs s; q \<in> current_msgqs s\<rbrakk> |
|
1630 \<Longrightarrow> cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq" by auto |
|
1631 from pre have pre_sfd: "\<And> tp fd. \<lbrakk>fd \<in> proc_file_fds s tp; q \<in> current_msgqs s\<rbrakk> |
|
1632 \<Longrightarrow> cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" by auto |
|
1633 hence pre_sfd': "\<And> tp fd f. \<lbrakk>file_of_proc_fd s tp fd = Some f; q \<in> current_msgqs s\<rbrakk> |
|
1634 \<Longrightarrow> cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" by (auto simp:proc_file_fds_def) |
|
1635 from pre have pre_duq: "q \<in> current_msgqs s \<Longrightarrow> cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q" |
|
1636 by auto |
|
1637 have vd_enrich:"q \<in> current_msgqs s \<Longrightarrow> valid (e # enrich_msgq s q q')" |
|
1638 apply (frule pre_vd) |
|
1639 apply (erule_tac s = s and obj' = "E_msgq q'" in enrich_valid_intro_cons) |
|
1640 apply (simp_all add:pre nodel nodel_cons curq_cons vd_cons' vd enrich_msgq_hungs) |
|
1641 apply (clarsimp simp:nodel vd curq' enrich_msgq_alive) |
|
1642 apply (rule allI, rule impI, erule enrich_msgq_not_alive) |
|
1643 apply (simp_all add:curq' curq'_cons nodel vd enrich_msgq_cur_msgs enrich_msgq_filefd enrich_msgq_flagfd) |
|
1644 done |
|
1645 |
|
1646 have q_neq_q': "q' \<noteq> q" using curq'_cons curq_cons by auto |
|
1647 |
|
1648 have vd_enrich_cons: "valid (enrich_msgq (e # s) q q')" |
|
1649 proof- |
|
1650 have "\<And> p q''. e = CreateMsgq p q'' \<Longrightarrow> valid (enrich_msgq (e # s) q q')" |
|
1651 proof- |
|
1652 fix p q'' assume ev: "e = CreateMsgq p q''" |
|
1653 show "valid (enrich_msgq (e # s) q q')" |
|
1654 proof (cases "q'' = q") |
|
1655 case False with ev vd_enrich curq_cons |
|
1656 show ?thesis by simp |
|
1657 next |
|
1658 case True |
|
1659 have "os_grant (CreateMsgq p q # s) (CreateMsgq p q')" |
|
1660 using os ev |
|
1661 by (simp add:q_neq_q' curq') |
|
1662 moreover have "grant (CreateMsgq p q # s) (CreateMsgq p q')" |
|
1663 using grant ev |
|
1664 by (auto simp add:sectxt_of_obj_def split:option.splits) |
|
1665 ultimately |
|
1666 show ?thesis using ev vd_cons' True |
|
1667 by (auto intro: valid.intros(2)) |
|
1668 qed |
|
1669 qed |
|
1670 moreover have "\<And> p q'' m. \<lbrakk>e = SendMsg p q'' m; q \<in> current_msgqs s\<rbrakk> |
|
1671 \<Longrightarrow> valid (enrich_msgq (e # s) q q')" |
|
1672 proof- |
|
1673 fix p q'' m assume ev: "e = SendMsg p q'' m" and q_in: "q \<in> current_msgqs s" |
|
1674 show "valid (enrich_msgq (e # s) q q')" |
|
1675 proof (cases "q'' = q") |
|
1676 case False with ev vd_enrich q_in |
|
1677 show ?thesis by simp |
|
1678 next |
|
1679 case True |
|
1680 from grant os ev True obtain psec qsec |
|
1681 where psec: "sectxt_of_obj s (O_proc p) = Some psec" |
|
1682 and qsec: "sectxt_of_obj s (O_msgq q) = Some qsec" |
|
1683 by (auto split:option.splits) |
|
1684 from psec q_in os ev |
|
1685 have psec':"sectxt_of_obj (enrich_msgq s q q') (O_proc p) = Some psec" |
|
1686 by (auto dest!:pre_sp simp:cp2sproc_def split:option.splits) |
|
1687 from qsec q_in pre_duq vd |
|
1688 have qsec': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = Some qsec" |
|
1689 by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms') |
|
1690 show ?thesis using ev True vd_cons' q_in vd_enrich nodel vd |
|
1691 curq' psec psec' qsec qsec' grant os q_neq_q' |
|
1692 apply (simp, erule_tac valid.intros(2)) |
|
1693 apply (auto simp:q_neq_q' enrich_msgq_cur_msgqs enrich_msgq_cur_procs |
|
1694 enrich_msgq_cur_msgs sectxt_of_obj_simps) |
|
1695 done |
|
1696 qed |
|
1697 qed |
|
1698 moreover have "\<And> p q'' m. \<lbrakk>e = RecvMsg p q'' m; q \<in> current_msgqs s\<rbrakk> |
|
1699 \<Longrightarrow> valid (enrich_msgq (e # s) q q')" |
|
1700 proof- |
|
1701 fix p q'' m assume ev: "e = RecvMsg p q'' m" and q_in: "q \<in> current_msgqs s" |
|
1702 show "valid (enrich_msgq (e # s) q q')" |
|
1703 proof (cases "q'' = q") |
|
1704 case False with ev vd_enrich q_in |
|
1705 show ?thesis by simp |
|
1706 next |
|
1707 case True |
|
1708 from grant os ev True obtain psec qsec msec |
|
1709 where psec: "sectxt_of_obj s (O_proc p) = Some psec" |
|
1710 and qsec: "sectxt_of_obj s (O_msgq q) = Some qsec" |
|
1711 and msec: "sectxt_of_obj s (O_msg q (hd (msgs_of_queue s q))) = Some msec" |
|
1712 by (auto split:option.splits) |
|
1713 from psec q_in os ev |
|
1714 have psec':"sectxt_of_obj (enrich_msgq s q q') (O_proc p) = Some psec" |
|
1715 by (auto dest!:pre_sp simp:cp2sproc_def split:option.splits) |
|
1716 from qsec q_in pre_duq vd |
|
1717 have qsec': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = Some qsec" |
|
1718 by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms') |
|
1719 from qsec q_in vd |
|
1720 have qsec'': "sectxt_of_obj (enrich_msgq s q q') (O_msgq q) = Some qsec" |
|
1721 apply (frule_tac pre_sq, simp_all) |
|
1722 by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sms') |
|
1723 from msec q_in pre_duq vd qsec qsec' qsec'' curq' nodel |
|
1724 have msec': "sectxt_of_obj (enrich_msgq s q q') (O_msg q' (hd (msgs_of_queue s q))) = Some msec" |
|
1725 apply (auto simp:cq2smsgq_def enrich_msgq_cur_msgs |
|
1726 split:option.splits dest!:current_has_sms') |
|
1727 apply (case_tac "msgs_of_queue s q") |
|
1728 using os ev True apply simp |
|
1729 apply (simp add:cqm2sms.simps split:option.splits) |
|
1730 apply (auto simp:cm2smsg_def split:option.splits) |
|
1731 done |
|
1732 show ?thesis using ev True vd_cons' q_in vd_enrich nodel vd |
|
1733 curq' grant os q_neq_q' psec psec' msec msec' qsec qsec' |
|
1734 apply (simp, erule_tac valid.intros(2)) |
|
1735 apply (auto simp:enrich_msgq_cur_msgqs enrich_msgq_cur_procs |
|
1736 enrich_msgq_cur_msgs sectxt_of_obj_simps) |
|
1737 done |
|
1738 qed |
|
1739 qed |
|
1740 ultimately |
|
1741 show ?thesis using vd_enrich curq_cons vd_cons' |
|
1742 apply (case_tac e) |
|
1743 apply (auto simp del:enrich_msgq.simps) |
|
1744 apply (auto split:if_splits ) |
|
1745 done |
|
1746 qed |
|
1747 |
|
1748 have curpsec: "\<And> tp. \<lbrakk>tp \<in> current_procs s; q \<in> current_msgqs s\<rbrakk> \<Longrightarrow> |
|
1749 sectxt_of_obj (enrich_msgq s q q') (O_proc tp) = sectxt_of_obj s (O_proc tp)" |
|
1750 using pre_vd vd |
|
1751 apply (frule_tac pre_sp, simp) |
|
1752 by (auto simp:cp2sproc_def split:option.splits if_splits dest!:current_has_sec') |
|
1753 have curfsec: "\<And> f. \<lbrakk>is_file s f; q \<in> current_msgqs s\<rbrakk> \<Longrightarrow> |
|
1754 sectxt_of_obj (enrich_msgq s q q') (O_file f) = sectxt_of_obj s (O_file f)" |
|
1755 proof- |
|
1756 fix f |
|
1757 assume a1: "is_file s f" and a2: "q \<in> current_msgqs s" |
|
1758 from a2 pre_sf pre_vd |
|
1759 have pre': "\<And>f. f \<in> current_files s \<Longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f" |
|
1760 and vd_enrich: "valid (enrich_msgq s q q')" |
|
1761 by auto |
|
1762 hence csf: "cf2sfile (enrich_msgq s q q') f = cf2sfile s f" |
|
1763 using a1 by (auto simp:is_file_in_current) |
|
1764 from a1 obtain sf where csf_some: "cf2sfile s f = Some sf" |
|
1765 apply (case_tac "cf2sfile s f") |
|
1766 apply (drule current_file_has_sfile') |
|
1767 apply (simp add:vd, simp add:is_file_in_current, simp) |
|
1768 done |
|
1769 from a1 have a1': "is_file (enrich_msgq s q q') f" |
|
1770 using vd nodel by (simp add:enrich_msgq_is_file) |
|
1771 show "sectxt_of_obj (enrich_msgq s q q') (O_file f) = sectxt_of_obj s (O_file f)" |
|
1772 using csf csf_some vd_enrich vd a1 a1' |
|
1773 apply (auto simp:cf2sfile_def split:option.splits if_splits) |
|
1774 apply (case_tac f, simp_all) |
|
1775 apply (drule root_is_dir', simp+) |
|
1776 done |
|
1777 qed |
|
1778 have curdsec: "\<And> tf. \<lbrakk>is_dir s tf; q \<in> current_msgqs s\<rbrakk> |
|
1779 \<Longrightarrow> sectxt_of_obj (enrich_msgq s q q') (O_dir tf) = sectxt_of_obj s (O_dir tf)" |
|
1780 proof- |
|
1781 fix tf |
|
1782 assume a1: "is_dir s tf" and a2: "q \<in> current_msgqs s" |
|
1783 from a2 pre_sf pre_vd |
|
1784 have pre': "\<And>f. f \<in> current_files s \<Longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f" |
|
1785 and vd_enrich: "valid (enrich_msgq s q q')" |
|
1786 by auto |
|
1787 hence csf: "cf2sfile (enrich_msgq s q q') tf = cf2sfile s tf" |
|
1788 using a1 by (auto simp:is_dir_in_current) |
|
1789 from a1 obtain sf where csf_some: "cf2sfile s tf = Some sf" |
|
1790 apply (case_tac "cf2sfile s tf") |
|
1791 apply (drule current_file_has_sfile') |
|
1792 apply (simp add:vd, simp add:is_dir_in_current, simp) |
|
1793 done |
|
1794 from a1 have a1': "is_dir (enrich_msgq s q q') tf" |
|
1795 using enrich_msgq_is_dir vd nodel by simp |
|
1796 from a1 have a3: "\<not> is_file s tf" using vd by (simp add:is_dir_not_file) |
|
1797 from a1' vd have a3': "\<not> is_file (enrich_msgq s q q') tf" by (simp add:is_dir_not_file) |
|
1798 show "sectxt_of_obj (enrich_msgq s q q') (O_dir tf) = sectxt_of_obj s (O_dir tf)" |
|
1799 using csf csf_some a3 a3' vd_enrich vd |
|
1800 apply (auto simp:cf2sfile_def split:option.splits) |
|
1801 apply (case_tac tf) |
|
1802 apply (simp add:root_sec_remains, simp) |
|
1803 done |
|
1804 qed |
|
1805 have curpsecs: "\<And> tf ctxts'. \<lbrakk>is_dir s tf; q \<in> current_msgqs s; get_parentfs_ctxts s tf = Some ctxts'\<rbrakk> |
|
1806 \<Longrightarrow> \<exists> ctxts. get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts \<and> set ctxts = set ctxts'" |
|
1807 proof- |
|
1808 fix tf ctxts' |
|
1809 assume a1: "is_dir s tf" and a2: "q \<in> current_msgqs s" |
|
1810 and a4: "get_parentfs_ctxts s tf = Some ctxts'" |
|
1811 from a2 pre |
|
1812 have pre': "\<And>f. f \<in> current_files s \<Longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f" |
|
1813 and vd_enrich': "valid (enrich_msgq s q q')" |
|
1814 by auto |
|
1815 hence csf: "cf2sfile (enrich_msgq s q q') tf = cf2sfile s tf" |
|
1816 using a1 by (auto simp:is_dir_in_current) |
|
1817 from a1 obtain sf where csf_some: "cf2sfile s tf = Some sf" |
|
1818 apply (case_tac "cf2sfile s tf") |
|
1819 apply (drule current_file_has_sfile') |
|
1820 apply (simp add:vd, simp add:is_dir_in_current, simp) |
|
1821 done |
|
1822 from a1 have a1': "is_dir (enrich_msgq s q q') tf" |
|
1823 using enrich_msgq_is_dir vd nodel by simp |
|
1824 from a1 have a5: "\<not> is_file s tf" using vd by (simp add:is_dir_not_file) |
|
1825 from a1' vd have a5': "\<not> is_file (enrich_msgq s q q') tf" by (simp add:is_dir_not_file) |
|
1826 |
|
1827 from a1' pre_vd a2 obtain ctxts |
|
1828 where a3: "get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts" |
|
1829 apply (case_tac "get_parentfs_ctxts (enrich_msgq s q q') tf") |
|
1830 apply (drule get_pfs_secs_prop', simp+) |
|
1831 done |
|
1832 moreover have "set ctxts = set ctxts'" |
|
1833 proof (cases tf) |
|
1834 case Nil |
|
1835 with a3 a4 vd vd_enrich' |
|
1836 show ?thesis |
|
1837 by (simp add:get_parentfs_ctxts.simps root_sec_remains split:option.splits) |
|
1838 next |
|
1839 case (Cons a ff) |
|
1840 with csf csf_some a5 a5' vd_enrich' vd a3 a4 |
|
1841 show ?thesis |
|
1842 apply (auto simp:cf2sfile_def split:option.splits if_splits) |
|
1843 done |
|
1844 qed |
|
1845 ultimately |
|
1846 show "\<exists> ctxts. get_parentfs_ctxts (enrich_msgq s q q') tf = Some ctxts \<and> set ctxts = set ctxts'" |
|
1847 by auto |
|
1848 qed |
|
1849 |
|
1850 have psec_cons: "\<And> tp. tp \<in> current_procs (e # s) \<Longrightarrow> |
|
1851 sectxt_of_obj (enrich_msgq (e # s) q q') (O_proc tp) = sectxt_of_obj (e # s) (O_proc tp)" |
|
1852 using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd |
|
1853 apply (case_tac e) |
|
1854 apply (auto intro:curpsec simp:sectxt_of_obj_simps) |
|
1855 apply (frule curpsec, simp, frule curfsec, simp) |
|
1856 apply (auto split:option.splits)[1] |
|
1857 apply (frule vd_cons) defer apply (frule vd_cons) |
|
1858 apply (auto intro:curpsec simp:sectxt_of_obj_simps) |
|
1859 done |
|
1860 |
|
1861 |
|
1862 have sf_cons: "\<And> f. f \<in> current_files (e # s) \<Longrightarrow> cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" |
|
1863 proof- |
|
1864 fix f |
|
1865 assume a1: "f \<in> current_files (e # s)" |
|
1866 hence a1': "f \<in> current_files (enrich_msgq (e # s) q q')" |
|
1867 using nodel_cons os vd vd_cons' vd_enrich_cons |
|
1868 apply (case_tac e) |
|
1869 apply (auto simp:current_files_simps enrich_msgq_cur_files dest:is_file_in_current split:option.splits) |
|
1870 done |
|
1871 have b1: "\<And> p f' flags fd opt. e = Open p f' flags fd opt \<Longrightarrow> |
|
1872 cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" |
|
1873 proof- |
|
1874 fix p f' flags fd opt |
|
1875 assume ev: "e = Open p f' flags fd opt" |
|
1876 show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" |
|
1877 proof (cases opt) |
|
1878 case None |
|
1879 with a1 a1' ev vd_cons' vd_enrich_cons curq_cons |
|
1880 show ?thesis |
|
1881 apply (simp add:cf2sfile_open_none) |
|
1882 apply (simp add:pre_sf current_files_simps) |
|
1883 done |
|
1884 next |
|
1885 case (Some inum) |
|
1886 show ?thesis |
|
1887 proof (cases "f = f'") |
|
1888 case False |
|
1889 with a1 a1' ev vd_cons' vd_enrich_cons curq_cons Some |
|
1890 show ?thesis |
|
1891 apply (simp add:cf2sfile_open) |
|
1892 apply (simp add:pre_sf current_files_simps) |
|
1893 done |
|
1894 next |
|
1895 case True |
|
1896 with vd_cons' ev os Some |
|
1897 obtain pf where pf: "parent f = Some pf" by auto |
|
1898 then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts" |
|
1899 using os vd ev Some True |
|
1900 apply (case_tac "get_parentfs_ctxts s pf") |
|
1901 apply (drule get_pfs_secs_prop', simp, simp) |
|
1902 apply auto |
|
1903 done |
|
1904 |
|
1905 have "sectxt_of_obj (Open p f' flags fd (Some inum) # enrich_msgq s q q') (O_file f') = |
|
1906 sectxt_of_obj (Open p f' flags fd (Some inum) # s) (O_file f')" |
|
1907 using Some vd_enrich_cons vd_cons' ev pf True os curq_cons |
|
1908 by (simp add:sectxt_of_obj_simps curpsec curdsec) |
|
1909 moreover |
|
1910 have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)" |
|
1911 using curq_cons ev pf Some True os |
|
1912 by (simp add:curdsec) |
|
1913 moreover |
|
1914 have "\<exists> ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \<and> set ctxts' = set ctxts" |
|
1915 using curq_cons ev pf Some True os vd psecs |
|
1916 apply (case_tac "get_parentfs_ctxts s pf") |
|
1917 apply (drule get_pfs_secs_prop', simp+) |
|
1918 apply (rule curpsecs, simp+) |
|
1919 done |
|
1920 then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'" |
|
1921 and psecs_eq: "set ctxts' = set ctxts" by auto |
|
1922 ultimately show ?thesis |
|
1923 using a1 a1' ev vd_cons' vd_enrich_cons Some True pf psecs |
|
1924 by (simp add:cf2sfile_open split:option.splits) |
|
1925 qed |
|
1926 qed |
|
1927 qed |
|
1928 have b2: "\<And> p f' inum. e = Mkdir p f' inum \<Longrightarrow> |
|
1929 cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" |
|
1930 proof- |
|
1931 fix p f' inum |
|
1932 assume ev: "e = Mkdir p f' inum" |
|
1933 show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" |
|
1934 proof (cases "f' = f") |
|
1935 case False |
|
1936 with a1 a1' ev vd_cons' vd_enrich_cons curq_cons |
|
1937 show ?thesis |
|
1938 apply (simp add:cf2sfile_mkdir) |
|
1939 apply (simp add:pre_sf current_files_simps) |
|
1940 done |
|
1941 next |
|
1942 case True |
|
1943 with vd_cons' ev os |
|
1944 obtain pf where pf: "parent f = Some pf" by auto |
|
1945 then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts" |
|
1946 using os vd ev True |
|
1947 apply (case_tac "get_parentfs_ctxts s pf") |
|
1948 apply (drule get_pfs_secs_prop', simp, simp) |
|
1949 apply auto |
|
1950 done |
|
1951 |
|
1952 have "sectxt_of_obj (Mkdir p f' inum # enrich_msgq s q q') (O_dir f') = |
|
1953 sectxt_of_obj (Mkdir p f' inum # s) (O_dir f')" |
|
1954 using vd_enrich_cons vd_cons' ev pf True os curq_cons |
|
1955 by (simp add:sectxt_of_obj_simps curpsec curdsec) |
|
1956 moreover |
|
1957 have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)" |
|
1958 using curq_cons ev pf True os |
|
1959 by (simp add:curdsec) |
|
1960 moreover |
|
1961 have "\<exists> ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \<and> set ctxts' = set ctxts" |
|
1962 using curq_cons ev pf True os vd psecs |
|
1963 apply (case_tac "get_parentfs_ctxts s pf") |
|
1964 apply (drule get_pfs_secs_prop', simp+) |
|
1965 apply (rule curpsecs, simp+) |
|
1966 done |
|
1967 then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'" |
|
1968 and psecs_eq: "set ctxts' = set ctxts" by auto |
|
1969 ultimately show ?thesis |
|
1970 using a1 a1' ev vd_cons' vd_enrich_cons True pf psecs |
|
1971 apply (simp add:cf2sfile_mkdir split:option.splits) |
|
1972 done |
|
1973 qed |
|
1974 qed |
|
1975 have b3: "\<And> p f' f''. e = LinkHard p f' f'' \<Longrightarrow> |
|
1976 cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" |
|
1977 proof- |
|
1978 fix p f' f'' |
|
1979 assume ev: "e = LinkHard p f' f''" |
|
1980 show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" |
|
1981 proof (cases "f'' = f") |
|
1982 case False |
|
1983 with a1 a1' ev vd_cons' vd_enrich_cons curq_cons |
|
1984 show ?thesis |
|
1985 apply (simp add:cf2sfile_linkhard) |
|
1986 apply (simp add:pre_sf current_files_simps) |
|
1987 done |
|
1988 next |
|
1989 case True |
|
1990 with vd_cons' ev os |
|
1991 obtain pf where pf: "parent f = Some pf" by auto |
|
1992 then obtain ctxts where psecs: "get_parentfs_ctxts s pf = Some ctxts" |
|
1993 using os vd ev True |
|
1994 apply (case_tac "get_parentfs_ctxts s pf") |
|
1995 apply (drule get_pfs_secs_prop', simp, simp) |
|
1996 apply auto |
|
1997 done |
|
1998 |
|
1999 have "sectxt_of_obj (LinkHard p f' f'' # enrich_msgq s q q') (O_file f) = |
|
2000 sectxt_of_obj (LinkHard p f' f'' # s) (O_file f)" |
|
2001 using vd_enrich_cons vd_cons' ev pf True os curq_cons |
|
2002 by (simp add:sectxt_of_obj_simps curpsec curdsec) |
|
2003 moreover |
|
2004 have "sectxt_of_obj (enrich_msgq s q q') (O_dir pf) = sectxt_of_obj s (O_dir pf)" |
|
2005 using curq_cons ev pf True os |
|
2006 by (simp add:curdsec) |
|
2007 moreover |
|
2008 have "\<exists> ctxts'. get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts' \<and> set ctxts' = set ctxts" |
|
2009 using curq_cons ev pf True os vd psecs |
|
2010 apply (case_tac "get_parentfs_ctxts s pf") |
|
2011 apply (drule get_pfs_secs_prop', simp+) |
|
2012 apply (rule curpsecs, simp+) |
|
2013 done |
|
2014 then obtain ctxts' where psecs': "get_parentfs_ctxts (enrich_msgq s q q') pf = Some ctxts'" |
|
2015 and psecs_eq: "set ctxts' = set ctxts" by auto |
|
2016 ultimately show ?thesis |
|
2017 using a1 a1' ev vd_cons' vd_enrich_cons True pf psecs |
|
2018 apply (simp add:cf2sfile_linkhard split:option.splits) |
|
2019 done |
|
2020 qed |
|
2021 qed |
|
2022 show "cf2sfile (enrich_msgq (e # s) q q') f = cf2sfile (e # s) f" |
|
2023 apply (case_tac e) |
|
2024 prefer 6 apply (erule b1) |
|
2025 prefer 11 apply (erule b2) |
|
2026 prefer 11 apply (erule b3) |
|
2027 apply (simp_all only:b1 b2 b3) |
|
2028 using a1' a1 vd_enrich_cons vd_cons' curq_cons nodel_cons |
|
2029 apply (simp_all add:cf2sfile_other'' cf2sfile_simps enrich_msgq.simps no_del_event.simps split:if_splits) |
|
2030 apply (simp_all add:pre_sf cf2sfile_other' current_files_simps split:if_splits) |
|
2031 apply (drule vd_cons, simp add:cf2sfile_other', drule pre_sf, simp+)+ |
|
2032 done |
|
2033 qed |
|
2034 |
|
2035 have sfd_cons:"\<And> tp fd f. file_of_proc_fd (e # s) tp fd = Some f \<Longrightarrow> |
|
2036 cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" |
|
2037 proof- |
|
2038 fix tp fd f |
|
2039 assume a1: "file_of_proc_fd (e # s) tp fd = Some f" |
|
2040 hence a1': "file_of_proc_fd (enrich_msgq (e # s) q q') tp fd = Some f" |
|
2041 using nodel_cons vd_enrich os vd_cons' |
|
2042 apply (case_tac e, auto simp:enrich_msgq_filefd simp del:enrich_msgq.simps) |
|
2043 done |
|
2044 have b1: "\<And> p f' flags fd' opt. e = Open p f' flags fd' opt \<Longrightarrow> |
|
2045 cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" |
|
2046 proof- |
|
2047 fix p f' flags fd' opt |
|
2048 assume ev: "e = Open p f' flags fd' opt" |
|
2049 have c1': "file_of_proc_fd (Open p f' flags fd' opt # s) tp fd = Some f" |
|
2050 using a1' ev a1 by (simp split:if_splits) |
|
2051 show "cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" thm cfd2sfd_open |
|
2052 proof (cases "tp = p \<and> fd = fd'") |
|
2053 case False |
|
2054 show ?thesis using ev vd_enrich_cons vd_cons' a1' a1 False curq_cons |
|
2055 apply (simp add:cfd2sfd_open split:if_splits del:file_of_proc_fd.simps) |
|
2056 apply (rule conjI, rule impI, simp) |
|
2057 apply (rule conjI, rule impI, simp) |
|
2058 apply (auto simp: False intro!:pre_sfd' split:if_splits) |
|
2059 done |
|
2060 next |
|
2061 case True |
|
2062 have "f' \<in> current_files (Open p f' flags fd' opt # s)" using ev vd_cons' os |
|
2063 by (auto simp:current_files_simps is_file_in_current split:option.splits) |
|
2064 hence "cf2sfile (Open p f' flags fd' opt # enrich_msgq s q q') f' |
|
2065 = cf2sfile (Open p f' flags fd' opt # s) f'" |
|
2066 using sf_cons ev by auto |
|
2067 moreover have "sectxt_of_obj (enrich_msgq s q q') (O_proc p) = sectxt_of_obj s (O_proc p)" |
|
2068 apply (rule curpsec) |
|
2069 using os ev curq_cons |
|
2070 by (auto split:option.splits) |
|
2071 ultimately show ?thesis |
|
2072 using ev True a1 a1' vd_enrich_cons vd_cons' |
|
2073 apply (simp add:cfd2sfd_open sectxt_of_obj_simps del:file_of_proc_fd.simps) |
|
2074 done |
|
2075 qed |
|
2076 qed |
|
2077 show "cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" |
|
2078 apply (case_tac e) |
|
2079 prefer 6 apply (erule b1) |
|
2080 using a1' a1 vd_enrich_cons vd_cons' curq_cons |
|
2081 apply (simp_all only:cfd2sfd_simps enrich_msgq.simps) |
|
2082 apply (auto simp:cfd2sfd_simps pre_sfd' dest:vd_cons cfd2sfd_other split:if_splits) |
|
2083 done |
|
2084 qed |
|
2085 |
|
2086 have pfds_cons: "\<And> tp. tp \<in> current_procs (e # s) \<Longrightarrow> |
|
2087 cpfd2sfds (enrich_msgq (e # s) q q') tp = cpfd2sfds (e # s) tp" |
|
2088 apply (auto simp add:cpfd2sfds_def proc_file_fds_def) |
|
2089 apply (rule_tac x = fd in exI, rule conjI, rule_tac x = f in exI) |
|
2090 prefer 3 |
|
2091 apply (rule_tac x = fd in exI, rule conjI, rule_tac x = f in exI) |
|
2092 apply (auto simp:sfd_cons enrich_msgq_filefd nodel_cons vd_cons') |
|
2093 done |
|
2094 |
|
2095 have tainted_cons: "tainted (enrich_msgq (e # s) q q') = |
|
2096 (tainted (e # s) \<union> {O_msg q' m | m. O_msg q m \<in> tainted (e # s)})" |
|
2097 apply (rule equalityI) |
|
2098 using nodel_cons curq_cons curq'_cons vd_cons' vd_enrich_cons |
|
2099 apply (rule enrich_msgq_tainted_aux2) |
|
2100 using nodel_cons curq_cons curq'_cons vd_cons' |
|
2101 apply (rule enrich_msgq_tainted_aux1) |
|
2102 done |
|
2103 have pre_tainted: "q \<in> current_msgqs s \<Longrightarrow> tainted (enrich_msgq s q q') = |
|
2104 (tainted s \<union> {O_msg q' m| m. O_msg q m \<in> tainted s})" by (simp add:pre) |
|
2105 |
|
2106 have "\<forall>tp fd. fd \<in> proc_file_fds (e # s) tp \<longrightarrow> cfd2sfd (enrich_msgq (e # s) q q') tp fd = cfd2sfd (e # s) tp fd" |
|
2107 by (auto simp:proc_file_fds_def elim!:sfd_cons) |
|
2108 moreover |
|
2109 have "\<forall>tp. tp \<in> current_procs (e # s) \<longrightarrow> cp2sproc (enrich_msgq (e # s) q q') tp = cp2sproc (e # s) tp" |
|
2110 by (auto simp:cp2sproc_def pfds_cons psec_cons enrich_msgq_died_proc split:option.splits) |
|
2111 moreover |
|
2112 have "\<forall>tq. tq \<in> current_msgqs (e # s) \<longrightarrow> cq2smsgq (enrich_msgq (e # s) q q') tq = cq2smsgq (e # s) tq" |
|
2113 proof clarify |
|
2114 fix tq assume a1: "tq \<in> current_msgqs (e # s)" |
|
2115 |
|
2116 have curqsec: "\<And> tq. \<lbrakk>tq \<in> current_msgqs s; q \<in> current_msgqs s\<rbrakk> \<Longrightarrow> |
|
2117 sectxt_of_obj (enrich_msgq s q q') (O_msgq tq) = sectxt_of_obj s (O_msgq tq)" |
|
2118 using pre_vd vd |
|
2119 apply (frule_tac pre_sq, simp) |
|
2120 by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') |
|
2121 have cursms: "\<And> q''. \<lbrakk>q'' \<in> current_msgqs s; q \<in> current_msgqs s\<rbrakk> \<Longrightarrow> |
|
2122 cqm2sms (enrich_msgq s q q') q'' (msgs_of_queue (enrich_msgq s q q') q'') = |
|
2123 cqm2sms s q'' (msgs_of_queue s q'')" |
|
2124 using pre_vd vd |
|
2125 apply (frule_tac pre_sq, simp) |
|
2126 by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') |
|
2127 have qsec_cons: "sectxt_of_obj (enrich_msgq (e # s) q q') (O_msgq tq) = sectxt_of_obj (e # s) (O_msgq tq)" |
|
2128 using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons curq_cons a1 |
|
2129 apply (case_tac e) |
|
2130 apply (auto intro:curqsec simp:sectxt_of_obj_simps curpsec split:option.splits dest!:current_has_sec') |
|
2131 apply (frule vd_cons) defer apply (frule vd_cons) |
|
2132 apply (auto intro:curqsec simp:sectxt_of_obj_simps) |
|
2133 done |
|
2134 have sms_cons: "cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = |
|
2135 cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" |
|
2136 proof- |
|
2137 have b1: "\<And> p q'' m. e = SendMsg p q'' m \<Longrightarrow> |
|
2138 cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = |
|
2139 cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" |
|
2140 apply (case_tac e) |
|
2141 using a1 curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons |
|
2142 apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms pre_tainted |
|
2143 split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) |
|
2144 apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*}) |
|
2145 apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms pre_tainted |
|
2146 split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) |
|
2147 done |
|
2148 have b2: "\<And> p q''. e = CreateMsgq p q'' \<Longrightarrow> |
|
2149 cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = |
|
2150 cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" |
|
2151 using a1 curq_cons curq'_cons vd_enrich_cons vd_cons' |
|
2152 apply (auto simp:cqm2sms_simps intro:cursms) |
|
2153 apply (auto simp:cqm2sms.simps) |
|
2154 done |
|
2155 have b3: "\<And> p q'' m. e = RecvMsg p q'' m \<Longrightarrow> |
|
2156 cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = |
|
2157 cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" |
|
2158 using a1 curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons |
|
2159 apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms |
|
2160 split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) |
|
2161 apply (frule vd_cons) defer apply (frule vd_cons) |
|
2162 apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec cursms |
|
2163 split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) |
|
2164 done |
|
2165 show "cqm2sms (enrich_msgq (e # s) q q') tq (msgs_of_queue (enrich_msgq (e # s) q q') tq) = |
|
2166 cqm2sms (e # s) tq (msgs_of_queue (e # s) tq)" |
|
2167 apply (case_tac e) |
|
2168 prefer 15 apply (erule b2) |
|
2169 prefer 15 apply (erule b1) |
|
2170 prefer 15 apply (erule b3) |
|
2171 using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons a1 |
|
2172 apply (auto intro:cursms simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons curqsec |
|
2173 split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) |
|
2174 done |
|
2175 qed |
|
2176 |
|
2177 show "cq2smsgq (enrich_msgq (e # s) q q') tq = cq2smsgq (e # s) tq" |
|
2178 using a1 curq_cons |
|
2179 apply (simp add:cq2smsgq_def qsec_cons sms_cons) |
|
2180 done |
|
2181 qed |
|
2182 moreover |
|
2183 have "cq2smsgq (enrich_msgq (e # s) q q') q' = cq2smsgq (e # s) q" |
|
2184 proof- |
|
2185 have duqsec: "q \<in> current_msgqs s \<Longrightarrow> |
|
2186 sectxt_of_obj (enrich_msgq s q q') (O_msgq q') = sectxt_of_obj s (O_msgq q)" |
|
2187 apply (frule pre_duq) using vd |
|
2188 by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') |
|
2189 have duqsms: "q \<in> current_msgqs s \<Longrightarrow> |
|
2190 cqm2sms (enrich_msgq s q q') q' (msgs_of_queue (enrich_msgq s q q') q') = |
|
2191 cqm2sms s q (msgs_of_queue s q)" |
|
2192 apply (frule pre_duq) using vd |
|
2193 by (auto simp:cq2smsgq_def split:option.splits if_splits dest!:current_has_sec' current_has_sms') |
|
2194 have qsec_cons: "sectxt_of_obj (enrich_msgq (e # s) q q') (O_msgq q') = sectxt_of_obj (e # s) (O_msgq q)" |
|
2195 using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons curq_cons |
|
2196 apply (case_tac e) |
|
2197 apply (auto simp:duqsec sectxt_of_obj_simps curpsec split:option.splits dest!:current_has_sec') |
|
2198 apply (frule vd_cons) defer apply (frule vd_cons) |
|
2199 apply (auto intro:duqsec simp:sectxt_of_obj_simps) |
|
2200 done |
|
2201 have sms_cons: "cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = |
|
2202 cqm2sms (e # s) q (msgs_of_queue (e # s) q)" |
|
2203 proof- |
|
2204 have b1: "\<And> p q'' m. e = SendMsg p q'' m \<Longrightarrow> |
|
2205 cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = |
|
2206 cqm2sms (e # s) q (msgs_of_queue (e # s) q)" |
|
2207 apply (case_tac e) |
|
2208 using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons |
|
2209 apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted |
|
2210 enrich_msgq_cur_procs enrich_msgq_cur_msgqs |
|
2211 split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) |
|
2212 apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*}) |
|
2213 apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted |
|
2214 enrich_msgq_cur_procs enrich_msgq_cur_msgqs dest:tainted_in_current |
|
2215 split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) |
|
2216 done |
|
2217 have b2: "\<And> p q''. e = CreateMsgq p q'' \<Longrightarrow> |
|
2218 cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = |
|
2219 cqm2sms (e # s) q (msgs_of_queue (e # s) q)" |
|
2220 using curq_cons curq'_cons vd_enrich_cons vd_cons' |
|
2221 apply (auto simp:cqm2sms_simps intro:duqsms) |
|
2222 apply (auto simp:cqm2sms.simps) |
|
2223 done |
|
2224 have b3: "\<And> p q'' m. e = RecvMsg p q'' m \<Longrightarrow> |
|
2225 cqm2sms (enrich_msgq (e # s) q q') q' (msgs_of_queue (enrich_msgq (e # s) q q') q') = |
|
2226 cqm2sms (e # s) q (msgs_of_queue (e # s) q)" |
|
2227 using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons |
|
2228 apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted |
|
2229 enrich_msgq_cur_procs enrich_msgq_cur_msgqs |
|
2230 split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) |
|
2231 apply (tactic {*ALLGOALS (ftac @{thm vd_cons})*}) |
|
2232 apply (auto simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsec duqsms pre_tainted |
|
2233 enrich_msgq_cur_procs enrich_msgq_cur_msgqs dest:tainted_in_current |
|
2234 split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) |
|
2235 done |
|
2236 show ?thesis |
|
2237 apply (case_tac e) |
|
2238 prefer 15 apply (erule b2) |
|
2239 prefer 15 apply (erule b1) |
|
2240 prefer 15 apply (erule b3) |
|
2241 using curq_cons vd_enrich_cons vd_cons' os pre_vd nodel_cons vd curq'_cons |
|
2242 apply (auto intro:duqsms simp:sectxt_of_obj_simps cqm2sms_simps curpsec qsec_cons duqsms |
|
2243 split:option.splits dest!:current_has_sec' current_has_sms' simp del:cqm2sms.simps) |
|
2244 done |
|
2245 qed |
|
2246 show ?thesis |
|
2247 using curq_cons |
|
2248 apply (simp add:cq2smsgq_def qsec_cons sms_cons) |
|
2249 done |
|
2250 qed |
|
2251 ultimately show ?case using vd_enrich_cons sf_cons tainted_cons |
|
2252 by auto |
|
2253 qed |
|
2254 |
|
2255 lemma enrich_msgq_vd: |
|
2256 "\<lbrakk>q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s; valid s\<rbrakk> \<Longrightarrow> |
|
2257 valid (enrich_msgq s q q')" |
|
2258 by (auto dest:enrich_msgq_prop) |
|
2259 |
|
2260 lemma enrich_msgq_sp: |
|
2261 "\<lbrakk>tp \<in> current_procs s; valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk> |
|
2262 \<Longrightarrow> cp2sproc (enrich_msgq s q q') tp = cp2sproc s tp" |
|
2263 by (auto dest:enrich_msgq_prop) |
|
2264 |
|
2265 lemma enrich_msgq_sf: |
|
2266 "\<lbrakk>f \<in> current_files s; valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk> |
|
2267 \<Longrightarrow> cf2sfile (enrich_msgq s q q') f = cf2sfile s f" |
|
2268 by (auto dest:enrich_msgq_prop) |
|
2269 |
|
2270 lemma enrich_msgq_sfs: |
|
2271 "\<lbrakk>is_file s f; valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk> |
|
2272 \<Longrightarrow> cf2sfiles (enrich_msgq s q q') f = cf2sfiles s f" |
|
2273 apply (auto simp add:cf2sfiles_def) |
|
2274 apply (rule_tac x = f' in bexI) defer |
|
2275 apply (simp add:enrich_msgq_sameinode) |
|
2276 apply (rule_tac x = f' in bexI) defer |
|
2277 apply (simp add:enrich_msgq_sameinode)+ |
|
2278 apply (drule same_inode_files_prop11, drule_tac f = f' in is_file_in_current) |
|
2279 apply (simp add:enrich_msgq_sf) |
|
2280 apply (drule same_inode_files_prop11, drule_tac f = f' in is_file_in_current) |
|
2281 apply (simp add:enrich_msgq_sf) |
|
2282 done |
|
2283 |
|
2284 lemma enrich_msgq_sq: |
|
2285 "\<lbrakk>tq \<in> current_msgqs s; valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk> |
|
2286 \<Longrightarrow> cq2smsgq (enrich_msgq s q q') tq = cq2smsgq s tq" |
|
2287 by (auto dest:enrich_msgq_prop) |
|
2288 |
|
2289 lemma enrich_msgq_sfd': |
|
2290 "\<lbrakk>fd \<in> proc_file_fds s tp; valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk> |
|
2291 \<Longrightarrow> cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" |
|
2292 by (auto dest:enrich_msgq_prop) |
|
2293 |
|
2294 lemma enrich_msgq_sfd: |
|
2295 "\<lbrakk>file_of_proc_fd s tp fd = Some f; valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk> |
|
2296 \<Longrightarrow> cfd2sfd (enrich_msgq s q q') tp fd = cfd2sfd s tp fd" |
|
2297 by (auto intro:enrich_msgq_sfd' simp:proc_file_fds_def) |
|
2298 |
|
2299 lemma enrich_msgq_duq: |
|
2300 "\<lbrakk>valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk> |
|
2301 \<Longrightarrow> cq2smsgq (enrich_msgq s q q') q' = cq2smsgq s q" |
|
2302 by (auto dest:enrich_msgq_prop) |
|
2303 |
|
2304 lemma enrich_msgq_tainted: |
|
2305 "\<lbrakk>valid s; q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s\<rbrakk> |
|
2306 \<Longrightarrow> tainted (enrich_msgq s q q') = (tainted s \<union> {O_msg q' m| m. O_msg q m \<in> tainted s})" |
|
2307 by (auto dest:enrich_msgq_prop) |
|
2308 |
|
2309 lemma enrich_msgq_dalive: |
|
2310 "\<lbrakk>q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s; valid s\<rbrakk> |
|
2311 \<Longrightarrow> dalive (enrich_msgq s q q') obj = (dalive s obj \<or> obj = D_msgq q')" |
|
2312 apply (case_tac obj) |
|
2313 apply (auto simp:enrich_msgq_is_file enrich_msgq_is_dir enrich_msgq_cur_msgqs enrich_msgq_cur_procs) |
|
2314 done |
|
2315 |
|
2316 lemma enrich_msgq_s2ss: |
|
2317 "\<lbrakk>q \<in> current_msgqs s; q' \<notin> current_msgqs s; no_del_event s; valid s\<rbrakk> \<Longrightarrow> |
|
2318 s2ss (enrich_msgq s q q') = s2ss s" |
|
2319 apply (auto simp add:s2ss_def) |
|
2320 apply (simp add:enrich_msgq_dalive) |
|
2321 apply (erule disjE) |
|
2322 apply (rule_tac x = obj in exI) defer |
|
2323 apply (rule_tac x = "D_msgq q" in exI) defer |
|
2324 apply (rule_tac x = obj in exI) |
|
2325 apply (case_tac[!] obj) |
|
2326 apply (auto simp:enrich_msgq_duq enrich_msgq_tainted enrich_msgq_sq enrich_msgq_sf |
|
2327 enrich_msgq_sp co2sobj.simps enrich_msgq_is_file enrich_msgq_is_dir |
|
2328 enrich_msgq_cur_procs enrich_msgq_cur_msgqs enrich_msgq_sfs |
|
2329 split:option.splits dest:is_dir_in_current) |
|
2330 done |
|
2331 |
1308 end |
2332 end |
1309 |
2333 |
1310 end |
2334 end |