diff -r b6ee5f35c41f -r e9c5594d5963 Delete_prop.thy --- a/Delete_prop.thy Wed May 08 10:00:38 2013 +0800 +++ b/Delete_prop.thy Thu May 09 11:19:44 2013 +0800 @@ -1,5 +1,5 @@ theory Deleted_prop -imports Main Flask Flask_type Init_prop Alive_prop +imports Main Flask Flask_type Init_prop Alive_prop Current_files_prop begin context flask begin @@ -10,21 +10,132 @@ lemma not_deleted_cons_D: "\ deleted obj (e # s) \ \ deleted obj s" by (auto dest:deleted_cons_I) -lemma not_deleted_imp_exists: +lemma not_deleted_init_file: + "\\ deleted (O_file f) s; valid s; is_init_file f\ \ is_file s f" +apply (induct s, simp add:is_init_file_prop1) +apply (frule vd_cons, frule vt_grant_os) +apply (case_tac a) prefer 6 apply (case_tac option) +apply (auto simp:is_file_simps split:option.splits) +done + +lemma not_deleted_init_dir: + "\\ deleted (O_dir f) s; valid s; is_init_dir f\ \ is_dir s f" +apply (induct s, simp add:is_init_dir_prop1) +apply (frule vd_cons, frule vt_grant_os) +apply (case_tac a) prefer 6 apply (case_tac option) +apply (auto simp:is_dir_simps split:option.splits) +done + +lemma not_deleted_init_proc: + "\\ deleted (O_proc p) s; p \ init_procs\ \ p \ current_procs s" +apply (induct s, simp) +by (case_tac a, auto) + +(**** ???*) +lemma current_fd_imp_current_proc: + "\fd \ current_proc_fds s p; valid s\ \ p \ current_procs s" +apply (induct s) +apply (simp add:init_fds_of_proc_prop1) +apply (frule vd_cons, drule vt_grant_os, case_tac a) +by (auto split:if_splits option.splits) + +lemma not_deleted_init_fd_aux: + "\\ deleted (O_fd p fd) s; valid s; fd \ init_fds_of_proc p\ + \ fd \ current_proc_fds s p \ \ deleted (O_proc p) s" +apply (induct s arbitrary: p, simp) +apply (frule vd_cons, drule vt_grant_os) +apply (case_tac a, auto dest:current_fd_imp_current_proc) +done + +lemma not_deleted_init_fd2: + "\\ deleted (O_fd p fd) s; fd \ init_fds_of_proc p; valid s\ \ \ deleted (O_proc p) s" +by (auto dest:not_deleted_init_fd_aux) + +lemma not_deleted_init_fd1: + "\\ deleted (O_fd p fd) s; valid s; fd \ init_fds_of_proc p\ \ fd \ current_proc_fds s p" +by (auto dest:not_deleted_init_fd_aux) + +lemma not_deleted_init_tcp_aux: + "\\ deleted (O_tcp_sock (p,fd)) s; valid s; is_init_tcp_sock (p,fd)\ + \ is_tcp_sock s (p,fd) \ \ deleted (O_proc p) s" +apply (induct s arbitrary:p, simp add:is_init_tcp_sock_prop1) +apply (frule vd_cons, frule vt_grant_os) +apply (case_tac a) prefer 6 apply (case_tac option) +by (auto simp:is_tcp_sock_simps split:option.splits t_socket_type.splits + dest:is_tcp_sock_imp_curernt_proc) + +lemma not_deleted_init_tcp1: + "\\ deleted (O_tcp_sock (p,fd)) s; valid s; is_init_tcp_sock (p,fd)\ + \ is_tcp_sock s (p,fd)" +by (auto dest:not_deleted_init_tcp_aux) + +lemma not_deleted_init_tcp2: + "\\ deleted (O_tcp_sock (p,fd)) s; valid s; is_init_tcp_sock (p,fd)\ + \ \ deleted (O_proc p) s" +by (auto dest:not_deleted_init_tcp_aux) + +lemma not_deleted_init_udp_aux: + "\\ deleted (O_udp_sock (p,fd)) s; valid s; is_init_udp_sock (p,fd)\ + \ is_udp_sock s (p,fd) \ \ deleted (O_proc p) s" +apply (induct s arbitrary:p, simp add:is_init_udp_sock_prop1) +apply (frule vd_cons, frule vt_grant_os) +apply (case_tac a) prefer 6 apply (case_tac option) +by (auto simp:is_udp_sock_simps split:option.splits t_socket_type.splits + dest:is_udp_sock_imp_curernt_proc) + +lemma not_deleted_init_udp1: + "\\ deleted (O_udp_sock (p,fd)) s; valid s; is_init_udp_sock (p,fd)\ + \ is_udp_sock s (p,fd)" +by (auto dest:not_deleted_init_udp_aux) + +lemma not_deleted_init_udp2: + "\\ deleted (O_udp_sock (p,fd)) s; valid s; is_init_udp_sock (p,fd)\ + \ \ deleted (O_proc p) s" +by (auto dest:not_deleted_init_udp_aux) + +lemma not_deleted_init_shm: + "\\ deleted (O_shm h) s; h \ init_shms\ \ h \ current_shms s" +apply (induct s, simp) +by (case_tac a, auto) + +lemma not_deleted_init_msgq: + "\\ deleted (O_msgq q) s; q \ init_msgqs\ \ q \ current_msgqs s" +apply (induct s, simp) +by (case_tac a, auto) + +lemma current_msg_imp_current_msgq: + "\m \ set (msgs_of_queue s q); valid s\ \ q \ current_msgqs s" +apply (induct s) +apply (simp add:init_msgs_valid) +apply (frule vd_cons, drule vt_grant_os) +apply (case_tac a, auto split:if_splits) +done + +lemma not_deleted_init_msg: + "\\ deleted (O_msg q m) s; valid s; m \ set (init_msgs_of_queue q)\ \ m \ set (msgs_of_queue s q)" +apply (induct s, simp) +apply (frule vd_cons, frule vt_grant_os) +apply (case_tac a, auto dest:current_msg_imp_current_msgq) +apply (case_tac "msgs_of_queue s q", simp+) +done + +lemma not_deleted_imp_alive: "\\ deleted obj s; valid s; init_alive obj\ \ alive s obj" -apply (induct s, simp add:init_alive_prop) -apply (frule vd_cons) -apply (case_tac a, auto simp:alive_simps split:t_object.splits) -pr 19 +apply (case_tac obj) +apply (auto dest!:not_deleted_init_msg not_deleted_init_file not_deleted_init_dir not_deleted_init_proc + not_deleted_init_fd1 not_deleted_init_tcp1 not_deleted_init_udp1 not_deleted_init_shm + not_deleted_init_msgq + intro:is_file_in_current is_dir_in_current is_tcp_in_current is_udp_in_current + current_msg_imp_current_msgq) done lemma cons_app_simp_aux: "(a # b) @ c = a # (b @ c)" by auto -lemma not_deleted_imp_exists': - "\\ deleted obj (s'@s); exists s obj\ \ exists (s'@s) obj" +lemma not_deleted_imp_alive_app: + "\\ deleted obj (s'@s); alive s obj\ \ alive (s'@s) obj" apply (induct s', simp, simp only:cons_app_simp_aux) -apply (frule not_deleted_cons_D) +apply (frule not_deleted_cons_D, simp) apply (case_tac a, case_tac [!] obj, auto) done