Delete_prop.thy
author chunhan
Thu, 16 Jan 2014 11:04:04 +0800
changeset 95 b7fd75d104bf
parent 77 6f7b9039715f
permissions -rw-r--r--
update
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 5
diff changeset
     1
theory Delete_prop
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
     2
imports Main Flask Flask_type Init_prop Alive_prop Current_files_prop
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     3
begin
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     4
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     5
context flask begin
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     6
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     7
lemma deleted_cons_I: "deleted obj s \<Longrightarrow> deleted obj (e # s)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     8
by (case_tac e, auto)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     9
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    10
lemma not_deleted_cons_D: "\<not> deleted obj (e # s) \<Longrightarrow> \<not> deleted obj s" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    11
by (auto dest:deleted_cons_I)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    12
5
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    13
lemma cons_app_simp_aux:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    14
  "(a # b) @ c = a # (b @ c)" by auto
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    15
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    16
lemma not_deleted_app_I:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    17
  "deleted obj s \<Longrightarrow> deleted obj (s' @ s)"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    18
apply (induct s', simp)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    19
by (simp add:cons_app_simp_aux deleted_cons_I)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    20
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    21
lemma not_deleted_app_D:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    22
  "\<not> deleted obj (s' @ s) \<Longrightarrow> \<not> deleted obj s"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    23
apply (rule notI)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    24
by (simp add:not_deleted_app_I)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
    25
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    26
lemma not_deleted_init_file:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    27
  "\<lbrakk>\<not> deleted (O_file f) s; valid s; is_init_file f\<rbrakk> \<Longrightarrow> is_file s f"
7
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
    28
apply (induct s, simp add:is_file_nil)
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    29
apply (frule vd_cons, frule vt_grant_os)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    30
apply (case_tac a) prefer 6 apply (case_tac option)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    31
apply (auto simp:is_file_simps split:option.splits)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    32
done
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    33
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    34
lemma not_deleted_init_dir:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    35
  "\<lbrakk>\<not> deleted (O_dir f) s; valid s; is_init_dir f\<rbrakk> \<Longrightarrow> is_dir s f"
7
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
    36
apply (induct s, simp add:is_dir_nil)
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    37
apply (frule vd_cons, frule vt_grant_os)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    38
apply (case_tac a) prefer 6 apply (case_tac option)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    39
apply (auto simp:is_dir_simps split:option.splits)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    40
done
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    41
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    42
lemma not_deleted_init_proc:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    43
  "\<lbrakk>\<not> deleted (O_proc p) s; p \<in> init_procs\<rbrakk> \<Longrightarrow> p \<in> current_procs s"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    44
apply (induct s, simp)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    45
by (case_tac a, auto)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    46
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    47
(**** ???*)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    48
lemma current_fd_imp_current_proc:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    49
  "\<lbrakk>fd \<in> current_proc_fds s p; valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    50
apply (induct s)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    51
apply (simp add:init_fds_of_proc_prop1)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    52
apply (frule vd_cons, drule vt_grant_os, case_tac a)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    53
by (auto split:if_splits option.splits)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    54
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    55
lemma not_deleted_init_fd_aux:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    56
  "\<lbrakk>\<not> deleted (O_fd p fd) s; valid s; fd \<in> init_fds_of_proc p\<rbrakk> 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    57
   \<Longrightarrow> fd \<in> current_proc_fds s p \<and> \<not> deleted (O_proc p) s"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    58
apply (induct s arbitrary: p, simp)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    59
apply (frule vd_cons, drule vt_grant_os)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    60
apply (case_tac a, auto dest:current_fd_imp_current_proc)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    61
done
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    62
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    63
lemma not_deleted_init_fd2:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    64
  "\<lbrakk>\<not> deleted (O_fd p fd) s; fd \<in> init_fds_of_proc p; valid s\<rbrakk> \<Longrightarrow> \<not> deleted (O_proc p) s"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    65
by (auto dest:not_deleted_init_fd_aux)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    66
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    67
lemma not_deleted_init_fd1:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    68
  "\<lbrakk>\<not> deleted (O_fd p fd) s; valid s; fd \<in> init_fds_of_proc p\<rbrakk> \<Longrightarrow> fd \<in> current_proc_fds s p"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    69
by (auto dest:not_deleted_init_fd_aux)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    70
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    71
lemma not_deleted_init_tcp_aux:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    72
  "\<lbrakk>\<not> deleted (O_tcp_sock (p,fd)) s; valid s; is_init_tcp_sock (p,fd)\<rbrakk> 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    73
   \<Longrightarrow> is_tcp_sock s (p,fd) \<and> \<not> deleted (O_proc p) s"
7
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
    74
apply (induct s arbitrary:p, simp add:is_tcp_sock_nil)
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    75
apply (frule vd_cons, frule vt_grant_os)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    76
apply (case_tac a) prefer 6 apply (case_tac option)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    77
by (auto simp:is_tcp_sock_simps split:option.splits t_socket_type.splits 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    78
         dest:is_tcp_sock_imp_curernt_proc)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    79
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    80
lemma not_deleted_init_tcp1:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    81
  "\<lbrakk>\<not> deleted (O_tcp_sock (p,fd)) s; valid s; is_init_tcp_sock (p,fd)\<rbrakk> 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    82
   \<Longrightarrow> is_tcp_sock s (p,fd)"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    83
by (auto dest:not_deleted_init_tcp_aux)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    84
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    85
lemma not_deleted_init_tcp2:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    86
  "\<lbrakk>\<not> deleted (O_tcp_sock (p,fd)) s; valid s; is_init_tcp_sock (p,fd)\<rbrakk> 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    87
   \<Longrightarrow> \<not> deleted (O_proc p) s"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    88
by (auto dest:not_deleted_init_tcp_aux)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    89
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    90
lemma not_deleted_init_udp_aux:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    91
  "\<lbrakk>\<not> deleted (O_udp_sock (p,fd)) s; valid s; is_init_udp_sock (p,fd)\<rbrakk> 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    92
   \<Longrightarrow> is_udp_sock s (p,fd) \<and> \<not> deleted (O_proc p) s"
7
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
    93
apply (induct s arbitrary:p, simp add:is_udp_sock_nil)
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    94
apply (frule vd_cons, frule vt_grant_os)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    95
apply (case_tac a) prefer 6 apply (case_tac option)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    96
by (auto simp:is_udp_sock_simps split:option.splits t_socket_type.splits 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    97
         dest:is_udp_sock_imp_curernt_proc)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    98
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
    99
lemma not_deleted_init_udp1:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   100
  "\<lbrakk>\<not> deleted (O_udp_sock (p,fd)) s; valid s; is_init_udp_sock (p,fd)\<rbrakk> 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   101
   \<Longrightarrow> is_udp_sock s (p,fd)"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   102
by (auto dest:not_deleted_init_udp_aux)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   103
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   104
lemma not_deleted_init_udp2:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   105
  "\<lbrakk>\<not> deleted (O_udp_sock (p,fd)) s; valid s; is_init_udp_sock (p,fd)\<rbrakk> 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   106
   \<Longrightarrow> \<not> deleted (O_proc p) s"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   107
by (auto dest:not_deleted_init_udp_aux)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   108
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   109
lemma not_deleted_init_shm:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   110
  "\<lbrakk>\<not> deleted (O_shm h) s; h \<in> init_shms\<rbrakk> \<Longrightarrow> h \<in> current_shms s"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   111
apply (induct s, simp)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   112
by (case_tac a, auto)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   113
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   114
lemma not_deleted_init_msgq:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   115
  "\<lbrakk>\<not> deleted (O_msgq q) s; q \<in> init_msgqs\<rbrakk> \<Longrightarrow> q \<in> current_msgqs s"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   116
apply (induct s, simp)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   117
by (case_tac a, auto)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   118
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   119
lemma current_msg_imp_current_msgq:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   120
  "\<lbrakk>m \<in> set (msgs_of_queue s q); valid s\<rbrakk> \<Longrightarrow> q \<in> current_msgqs s"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   121
apply (induct s)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   122
apply (simp add:init_msgs_valid)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   123
apply (frule vd_cons, drule vt_grant_os)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   124
apply (case_tac a, auto split:if_splits)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   125
done
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   126
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   127
lemma not_deleted_init_msg:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   128
  "\<lbrakk>\<not> deleted (O_msg q m) s; valid s; m \<in> set (init_msgs_of_queue q)\<rbrakk> \<Longrightarrow> m \<in> set (msgs_of_queue s q)"
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   129
apply (induct s, simp)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   130
apply (frule vd_cons, frule vt_grant_os)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   131
apply (case_tac a, auto dest:current_msg_imp_current_msgq)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   132
apply (case_tac "msgs_of_queue s q", simp+)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   133
done
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   134
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   135
lemma not_deleted_imp_alive:
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   136
  "\<lbrakk>\<not> deleted obj s; valid s; init_alive obj\<rbrakk> \<Longrightarrow> alive s obj"
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   137
apply (case_tac obj)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   138
apply (auto dest!:not_deleted_init_msg not_deleted_init_file not_deleted_init_dir not_deleted_init_proc
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   139
  not_deleted_init_fd1 not_deleted_init_tcp1 not_deleted_init_udp1 not_deleted_init_shm 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   140
  not_deleted_init_msgq 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   141
           intro:is_file_in_current is_dir_in_current is_tcp_in_current is_udp_in_current 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   142
  current_msg_imp_current_msgq)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   143
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   144
5
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   145
lemma not_deleted_cur_file_app:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   146
  "\<lbrakk>\<not> deleted (O_file f) (s' @ s); valid (s' @ s); is_file s f\<rbrakk> \<Longrightarrow> is_file (s' @ s) f"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   147
apply (induct s', simp, simp add:cons_app_simp_aux)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   148
apply (frule vd_cons, frule vt_grant_os, simp)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   149
apply (case_tac a) prefer 6 apply (case_tac option)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   150
apply (auto simp:is_file_simps split:option.splits)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   151
done
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   152
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   153
lemma not_deleted_cur_dir_app:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   154
  "\<lbrakk>\<not> deleted (O_dir f) (s' @ s); valid (s' @ s); is_dir s f\<rbrakk> \<Longrightarrow> is_dir (s' @ s) f"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   155
apply (induct s', simp, simp add:cons_app_simp_aux)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   156
apply (frule vd_cons, frule vt_grant_os, simp)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   157
apply (case_tac a) prefer 6 apply (case_tac option)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   158
apply (auto simp:is_dir_simps split:option.splits)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   159
done
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   160
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   161
lemma not_deleted_cur_proc_app:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   162
  "\<lbrakk>\<not> deleted (O_proc p) (s' @ s); p \<in> current_procs s\<rbrakk> \<Longrightarrow> p \<in> current_procs (s' @ s)"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   163
apply (induct s', simp, simp add:cons_app_simp_aux)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   164
by (case_tac a, auto)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   165
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   166
lemma not_deleted_cur_fd_app:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   167
  "\<lbrakk>\<not> deleted (O_fd p fd) (s' @ s); valid (s' @ s); fd \<in> current_proc_fds s p\<rbrakk> 
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   168
   \<Longrightarrow> fd \<in> current_proc_fds (s' @ s) p"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   169
apply (induct s' arbitrary: p, simp, simp add:cons_app_simp_aux)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   170
apply (frule vd_cons, drule vt_grant_os)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   171
apply (case_tac a, auto dest:current_fd_imp_current_proc)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   172
done
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   173
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   174
lemma not_deleted_cur_tcp_app:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   175
  "\<lbrakk>\<not> deleted (O_tcp_sock (p,fd)) (s' @ s); valid (s' @ s); is_tcp_sock s (p,fd)\<rbrakk> 
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   176
   \<Longrightarrow> is_tcp_sock (s' @ s) (p,fd)"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   177
apply (induct s' arbitrary:p, simp, simp add:cons_app_simp_aux)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   178
apply (frule vd_cons, frule vt_grant_os)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   179
apply (case_tac a) prefer 6 apply (case_tac option)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   180
by (auto simp:is_tcp_sock_simps split:option.splits t_socket_type.splits 
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   181
         dest:is_tcp_sock_imp_curernt_proc)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   182
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   183
lemma not_deleted_cur_udp_app:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   184
  "\<lbrakk>\<not> deleted (O_udp_sock (p,fd)) (s' @ s); valid (s' @ s); is_udp_sock s (p,fd)\<rbrakk> 
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   185
   \<Longrightarrow> is_udp_sock (s' @ s) (p,fd)"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   186
apply (induct s' arbitrary:p, simp, simp add:cons_app_simp_aux)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   187
apply (frule vd_cons, frule vt_grant_os)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   188
apply (case_tac a) prefer 6 apply (case_tac option)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   189
by (auto simp:is_udp_sock_simps split:option.splits t_socket_type.splits 
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   190
         dest:is_udp_sock_imp_curernt_proc)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   191
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   192
lemma not_deleted_cur_shm_app:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   193
  "\<lbrakk>\<not> deleted (O_shm h) (s' @ s); h \<in> current_shms s\<rbrakk> \<Longrightarrow> h \<in> current_shms (s' @ s)"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   194
apply (induct s', simp, simp add:cons_app_simp_aux)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   195
by (case_tac a, auto)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   196
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   197
lemma not_deleted_cur_msgq_app:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   198
  "\<lbrakk>\<not> deleted (O_msgq q) (s' @ s); q \<in> current_msgqs s\<rbrakk> \<Longrightarrow> q \<in> current_msgqs (s' @ s)"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   199
apply (induct s', simp, simp add:cons_app_simp_aux)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   200
by (case_tac a, auto)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   201
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   202
lemma not_deleted_cur_msg_app:
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   203
  "\<lbrakk>\<not> deleted (O_msg q m) (s' @ s); valid (s' @ s); m \<in> set (msgs_of_queue s q)\<rbrakk> 
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   204
   \<Longrightarrow> m \<in> set (msgs_of_queue (s' @ s) q)"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   205
apply (induct s', simp, simp add:cons_app_simp_aux)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   206
apply (frule vd_cons, frule vt_grant_os)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   207
apply (case_tac a, auto dest:current_msg_imp_current_msgq)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   208
apply (case_tac "msgs_of_queue (s' @ s) q", simp+)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   209
done
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   210
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 1
diff changeset
   211
lemma not_deleted_imp_alive_app:
5
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   212
  "\<lbrakk>\<not> deleted obj (s' @ s); valid (s' @ s); alive s obj\<rbrakk> \<Longrightarrow> alive (s' @ s) obj"
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   213
apply (case_tac obj)
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   214
apply (auto dest!:not_deleted_cur_msg_app not_deleted_cur_file_app not_deleted_cur_dir_app
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   215
  not_deleted_cur_proc_app not_deleted_cur_fd_app not_deleted_cur_tcp_app 
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   216
  not_deleted_cur_udp_app not_deleted_cur_shm_app not_deleted_cur_msgq_app
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   217
           intro:is_file_in_current is_dir_in_current is_tcp_in_current is_udp_in_current 
0c209a3e2647 finish delete_prop.thy
chunhan
parents: 4
diff changeset
   218
  current_msg_imp_current_msgq)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   219
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   220
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 7
diff changeset
   221
lemma not_deleted_imp_alive_cons:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 7
diff changeset
   222
  "\<lbrakk>\<not> deleted obj (e # s); valid (e # s); alive s obj\<rbrakk> \<Longrightarrow> alive (e # s) obj"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 7
diff changeset
   223
using not_deleted_imp_alive_app[where s = s and s' = "[e]" and obj = obj]
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 7
diff changeset
   224
by auto
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 7
diff changeset
   225
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   226
(*
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   227
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   228
lemma nodel_imp_un_deleted:
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   229
  "no_del_event s \<Longrightarrow> \<not> deleted obj s"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   230
by (induct s, simp, case_tac a,auto)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   231
77
chunhan
parents: 19
diff changeset
   232
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   233
lemma nodel_exists_remains:
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   234
  "\<lbrakk>no_del_event (s'@s); exists s obj\<rbrakk> \<Longrightarrow> exists (s'@s) obj"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   235
apply (drule_tac obj = obj in nodel_imp_un_deleted)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   236
by (simp add:not_deleted_imp_exists')
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   237
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   238
lemma nodel_imp_exists:
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   239
  "\<lbrakk>no_del_event s; exists [] obj\<rbrakk> \<Longrightarrow> exists s obj"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   240
apply (drule_tac obj = obj in nodel_imp_un_deleted)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   241
by (simp add:not_deleted_imp_exists)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   242
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   243
lemma no_del_event_cons_D:
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   244
  "no_del_event (e # s) \<Longrightarrow> no_del_event s"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   245
by (case_tac e, auto)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   246
*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   247
end
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   248
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   249
end