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