77
|
1 |
theory Temp
|
|
2 |
imports Static_type OS_type_def Flask_type Flask Static
|
|
3 |
begin
|
|
4 |
|
|
5 |
context tainting_s
|
|
6 |
|
|
7 |
begin
|
|
8 |
|
|
9 |
definition update_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
|
|
10 |
where
|
|
11 |
"update_ss ss so so' \<equiv> if (is_many so) then ss \<union> {so'} else (ss - {so}) \<union> {so'}"
|
|
12 |
|
|
13 |
definition add_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
|
|
14 |
where
|
|
15 |
"add_ss ss so \<equiv> ss \<union> {so}"
|
|
16 |
|
92
|
17 |
(*
|
77
|
18 |
definition del_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
|
|
19 |
where
|
|
20 |
"del_ss ss so \<equiv> if (is_many so) then ss else ss - {so}"
|
92
|
21 |
*)
|
|
22 |
|
77
|
23 |
|
|
24 |
(* all reachable static states(sobjects set) *)
|
|
25 |
inductive_set static :: "t_static_state set"
|
|
26 |
where
|
|
27 |
s_init: "init_static_state \<in> static"
|
|
28 |
| s_execve: "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds) tagp \<in> ss; S_file sfs tagf \<in> ss;
|
|
29 |
(fi,fsec,pfsec,asecs) \<in> sfs; npctxt_execve pctxt fsec = Some pctxt';
|
|
30 |
grant_execve pctxt fsec pctxt'; search_check_s pctxt (fi,fsec,pfsec,asecs) True;
|
|
31 |
inherit_fds_check_s pctxt' fds'; fds' \<subseteq> fds\<rbrakk>
|
|
32 |
\<Longrightarrow> (update_ss ss (S_proc (pi, pctxt, fds) tagp)
|
|
33 |
(S_proc (pi, pctxt', fds') (tagp \<or> tagf))) \<in> static"
|
|
34 |
| s_clone: "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds) tagp \<in> ss;
|
|
35 |
permission_check pctxt pctxt C_process P_fork;
|
|
36 |
inherit_fds_check_s pctxt fds'; fds' \<subseteq> fds\<rbrakk>
|
|
37 |
\<Longrightarrow> (add_ss ss (S_proc (Created, pctxt, fds') tagp)) \<in> static"
|
|
38 |
(*
|
|
39 |
| s_kill: "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds) tagp \<in> ss;
|
|
40 |
S_proc (pi', pctxt', fds') tagp' \<in> ss;
|
|
41 |
permission_check pctxt pctxt' C_process P_sigkill\<rbrakk>
|
|
42 |
\<Longrightarrow> (del_ss ss (S_proc (pi', pctxt', fds', shms') tagp')) \<in> static"
|
|
43 |
*)
|
|
44 |
| s_ptrace: "\<lbrakk>ss \<in> static; S_proc sp False \<in> ss; S_proc sp' True \<in> ss;
|
|
45 |
permission_check (sextxt_of_sproc sp) (sectxt_of_sproc sp') C_process P_ptrace\<rbrakk>
|
|
46 |
\<Longrightarrow> (update_ss ss (S_proc sp False) (S_proc sp True)) \<in> static"
|
|
47 |
| s_ptrace': "\<lbrakk>ss \<in> static; S_proc sp True \<in> ss; S_proc sp' False \<in> ss;
|
|
48 |
permission_check (sextxt_of_sproc sp) (sectxt_of_sproc sp') C_process P_ptrace\<rbrakk>
|
|
49 |
\<Longrightarrow> (update_ss ss (S_proc sp' False) (S_proc sp' True)) \<in> static"
|
|
50 |
(*
|
|
51 |
| s_exit: "\<lbrakk>ss \<in> static; S_proc sp tagp \<in> ss\<rbrakk> \<Longrightarrow> (del_ss ss (S_proc sp tagp)) \<in> static"
|
|
52 |
*)
|
|
53 |
| s_open: "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds) tagp \<in> ss; S_file sfs tagf \<in> ss; sf \<in> sfs;
|
|
54 |
search_check_s pctxt sf True; \<not> is_creat_excl_flag flags;
|
|
55 |
oflags_check flags pctxt (sectxt_of_sfile sf); permission_check pctxt pctxt C_fd P_setattr\<rbrakk>
|
|
56 |
\<Longrightarrow> (update_ss ss (S_proc (pi, pctxt, fds) tagp)
|
|
57 |
(S_proc (pi, pctxt, fds \<union> {(pctxt,flags,sf)}) tagp)) \<in> static"
|
|
58 |
| s_open': "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds) tagp \<in> ss; is_creat_excl_flag flags;
|
|
59 |
S_dir (pfi,fsec,pfsec,asecs) \<in> ss; search_check_s pctxt (pfi,fsec,pfsec,asecs) False;
|
|
60 |
nfsec = nfctxt_create pctxt fsec C_file; oflags_check flags pctxt nfsec;
|
|
61 |
permission_check pctxt fsec C_dir P_add_rename; permission_check pctxt pctxt C_fd P_setattr\<rbrakk>
|
|
62 |
\<Longrightarrow> (update_ss (add_ss ss (S_file {(Created, nfsec, Some fsec, asecs \<union> {fsec})} tagp))
|
|
63 |
(S_proc (pi, pctxt, fds) tagp)
|
|
64 |
(S_proc (pi, pctxt, fds \<union> {(pctxt, flags, (Created, nfsec, Some fsec, asecs \<union> {fsec}))}) tagp)
|
|
65 |
) \<in> static"
|
|
66 |
| S_readf: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) False \<in> ss; (fdctxt,flags,sf) \<in> fds;
|
|
67 |
permission_check pctxt fdctxt C_fd P_setattr; S_file sfs True \<in> ss; sf \<in> sfs;
|
|
68 |
permission_check pctxt (sectxt_of_sfile sf) C_file P_read; is_read_flag flags\<rbrakk>
|
|
69 |
\<Longrightarrow> (update_ss ss (S_proc (pi, pctxt,fds) False) (S_proc (pi, pctxt, fds) True)) \<in> static"
|
|
70 |
| S_writef: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) True \<in> ss; (fdctxt,flags,sf) \<in> fds;
|
|
71 |
permission_check pctxt fdctxt C_fd P_setattr; sf \<in> sfs; S_file sfs False \<in> ss;
|
|
72 |
permission_check pctxt (sectxt_of_sfile sf) C_file P_write; is_write_flag flags\<rbrakk>
|
|
73 |
\<Longrightarrow> (update_ss ss (S_file sfs False) (S_file sfs True)) \<in> static"
|
|
74 |
(*
|
|
75 |
| S_unlink: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_file sfs tagf \<in> ss;
|
|
76 |
(Init f,fsec,Some pfsec,asecs) \<in> sfs;
|
|
77 |
search_check_s pctxt (Init f,fsec,Some pfsec,asecs) True;
|
|
78 |
permission_check pctxt fsec C_file P_unlink;
|
|
79 |
permission_check pctxt pfsec C_dir P_remove_name\<rbrakk>
|
|
80 |
\<Longrightarrow> ((ss - {S_file sfs tagf}) \<union> {S_file (sfs - {(Init f,fsec,Some pfsec,asecs)}) tagf}) \<in> static"
|
|
81 |
| S_rmdir: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss;
|
|
82 |
S_dir (fi,fsec,Some pfsec,asecs) \<in> ss;
|
|
83 |
search_check_s pctxt (fi,fsec,Some pfsec,asecs) False;
|
|
84 |
permission_check pctxt fsec C_dir P_rmdir;
|
|
85 |
permission_check pctxt pfsec C_dir P_remove_name\<rbrakk>
|
|
86 |
\<Longrightarrow> (del_ss ss (S_dir (fi,fsec,Some pfsec,asecs))) \<in> static"
|
|
87 |
*)
|
|
88 |
| S_mkdir: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) tagp \<in> ss; S_dir (fi,fsec,pfsec,asecs) \<in> ss;
|
|
89 |
search_check_s pctxt (fi,fsec,pfsec,asecs) False;
|
|
90 |
permission_check pctxt (nfctxt_create pctxt fsec C_dir) C_dir P_create;
|
|
91 |
permission_check pctxt fsec C_dir P_add_name\<rbrakk>
|
|
92 |
\<Longrightarrow> (add_ss ss (S_dir (Created,nfctxt_create pctxt fsec C_dir,Some fsec,asecs \<union> {fsec}))) \<in> static"
|
|
93 |
| s_link: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) tagp \<in> ss; S_dir (pfi,pfsec,ppfsec,asecs) \<in> ss;
|
|
94 |
S_file sfs tagf \<in> ss; sf \<in> sfs; nfsec = nfctxt_create pctxt pfsec C_file;
|
|
95 |
search_check_s pctxt (pfi,pfsec,ppfsec,asecs) False; search_check_s pctxt sf True;
|
|
96 |
permission_check pctxt (sectxt_of_sfile sf) C_file P_link;
|
|
97 |
permission_check pctxt pfsec C_dir P_add_name\<rbrakk>
|
|
98 |
\<Longrightarrow> (update_ss ss (S_file sfs tagf)
|
|
99 |
(S_file (sfs \<union> {(Created,nfsec,Some pfsec, asecs \<union> {pfsec})}) tagf)) \<in> static"
|
|
100 |
| s_trunc: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) True \<in> ss; S_file sfs False \<in> ss; sf \<in> sfs;
|
|
101 |
search_check_s pctxt sf True; permission_check pctxt (sectxt_of_sfile sf) C_file P_setattr\<rbrakk>
|
|
102 |
\<Longrightarrow> (update_ss ss (S_file sfs False) (S_file sfs True)) \<in> static"
|
|
103 |
(*
|
|
104 |
| s_rename: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) tagp \<in> ss; S_file sfs tagf \<in> ss;
|
|
105 |
(sf#spf') \<in> sfs; S_dir spf tagpf \<in> ss; \<not>((sf#spf') \<preceq> (sf#spf));
|
|
106 |
search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True;
|
|
107 |
sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt;
|
|
108 |
permission_check pctxt fctxt C_file P_rename;
|
|
109 |
permission_check pctxt pfctxt C_dir P_add_name;
|
|
110 |
ss_rename ss (sf#spf') (sf#spf) ss';
|
|
111 |
ss_rename_no_same_fname ss (sf#spf') (sf#spf)\<rbrakk>
|
|
112 |
\<Longrightarrow> ss' \<in> static"
|
|
113 |
| s_rename': "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) tagp \<in> ss; S_dir (sf#spf') tagf \<in> ss;
|
|
114 |
S_dir spf tagpf \<in> ss; \<not>((sf#spf') \<preceq> (sf#spf));
|
|
115 |
search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True;
|
|
116 |
sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt;
|
|
117 |
permission_check pctxt fctxt C_dir P_reparent;
|
|
118 |
permission_check pctxt pfctxt C_dir P_add_name;
|
|
119 |
ss_rename ss (sf#spf') (sf#spf) ss';
|
|
120 |
ss_rename_no_same_fname ss (sf#spf') (sf#spf)\<rbrakk>
|
|
121 |
\<Longrightarrow> ss' \<in> static"
|
|
122 |
*)
|
|
123 |
| s_createq: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) tagp \<in> ss;
|
|
124 |
permission_check pctxt pctxt C_msgq P_associate;
|
|
125 |
permission_check pctxt pctxt C_msgq P_create\<rbrakk>
|
|
126 |
\<Longrightarrow> (add_ss ss (S_msgq (Created,pctxt,[]))) \<in> static"
|
|
127 |
| s_sendmsg: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) tagp \<in> ss; S_msgq (qi,qctxt,sms) \<in> ss;
|
|
128 |
permission_check pctxt qctxt C_msgq P_enqueue;
|
|
129 |
permission_check pctxt qctxt C_msgq P_write;
|
87
|
130 |
permission_check pctxt pctxt C_msg P_create; length sms < max_queue\<rbrakk>
|
77
|
131 |
\<Longrightarrow> (update_ss ss (S_msgq (qi,qctxt,sms))
|
|
132 |
(S_msgq (qi,qctxt,sms @ [(Created, pctxt, tagp)]))) \<in> static"
|
|
133 |
| s_recvmsg: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) tagp \<in> ss;
|
|
134 |
S_msgq (qi,qctxt,(mi,mctxt,tagm)#sms) \<in> ss;
|
|
135 |
permission_check pctxt qctxt C_msgq P_read;
|
|
136 |
permission_check pctxt mctxt C_msg P_receive\<rbrakk>
|
|
137 |
\<Longrightarrow> (update_ss (update_ss ss (S_proc (pi,pctxt,fds) tagp) (S_proc (pi, pctxt, fds) (tagp \<or> tagm)))
|
|
138 |
(S_msgq (qi, qctxt, (mi, mctxt, tagm)#sms))
|
|
139 |
(S_msgq (qi, qctxt, sms))) \<in> static"
|
|
140 |
(*
|
|
141 |
| s_removeq: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) tagp \<in> ss; S_msgq (qi,qctxt,sms) \<in> ss;
|
|
142 |
permission_check pctxt qctxt C_msgq P_destroy\<rbrakk>
|
|
143 |
\<Longrightarrow> (del_ss ss (S_msgq (qi,qctxt,sms))) \<in> static"
|
|
144 |
| s_createh: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) tagp \<in> ss;
|
|
145 |
permission_check pctxt pctxt C_shm P_associate;
|
|
146 |
permission_check pctxt pctxt C_shm P_create\<rbrakk>
|
|
147 |
\<Longrightarrow> (add_ss ss (S_shm (Created, pctxt))) \<in> static"
|
|
148 |
| s_attach: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) tagp \<in> ss; S_shm (hi,hctxt) \<in> ss;
|
|
149 |
if flag = SHM_RDONLY then permission_check pctxt hctxt C_shm P_read
|
|
150 |
else (permission_check pctxt hctxt C_shm P_read \<and>
|
|
151 |
permission_check pctxt hctxt C_shm P_write)\<rbrakk>
|
|
152 |
\<Longrightarrow> (update_ss ss (S_proc (pi,pctxt,fds) tagp)
|
|
153 |
(S_proc (pi,pctxt,fds,shms \<union> {((hi,hctxt),flag)}) tagp)) \<in> static"
|
|
154 |
| s_detach: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds) tagp \<in> ss; S_shm sh \<in> ss;
|
|
155 |
(sh,flag) \<in> shms; \<not> is_many_sshm sh\<rbrakk>
|
|
156 |
\<Longrightarrow> (update_ss ss (S_proc (pi,pctxt,fds) tagp)
|
|
157 |
(S_proc (pi,pctxt,fds,shms - {(sh,flag)}) tagp)) \<in> static"
|
|
158 |
| s_deleteh: "\<lbrakk>ss \<notin> static; S_proc (pi,pctxt,fds) tagp \<in> ss; S_shm (hi,hctxt) \<in> ss;
|
|
159 |
permission_check pctxt hctxt C_shm P_destroy; \<not> is_many_sshm sh\<rbrakk>
|
|
160 |
\<Longrightarrow> (remove_sproc_sshm (del_ss ss (S_shm (hi,hctxt))) (hi,hctxt)) \<in> static"
|
|
161 |
*)
|
|
162 |
|
|
163 |
fun tainted_s :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> bool"
|
|
164 |
where
|
|
165 |
"tainted_s ss (S_proc sp tag) = (S_proc sp tag \<in> ss \<and> tag = True)"
|
|
166 |
| "tainted_s ss (S_file sfs tag) = (S_file sfs tag \<in> ss \<and> tag = True)"
|
|
167 |
(*
|
|
168 |
| "tainted_s ss (S_msg (qi, sec, sms) (mi, secm, tag)) =
|
|
169 |
(S_msgq (qi, sec, sms) \<in> ss \<and> (mi,secm,tag) \<in> set sms \<and> tag = True)"
|
|
170 |
*)
|
|
171 |
| "tainted_s ss _ = False"
|
|
172 |
|
|
173 |
(*
|
|
174 |
fun tainted_s :: "t_object \<Rightarrow> t_static_state \<Rightarrow> bool"
|
|
175 |
where
|
|
176 |
"tainted_s (O_proc p) ss = (\<exists> sp. S_proc sp True \<in> ss \<and> sproc_related p sp)"
|
|
177 |
| "tainted_s (O_file f) ss = (\<exists> sfs sf. S_file sfs True \<in> ss \<and> sf \<in> sfs \<and> sfile_related f sf)"
|
|
178 |
| "tainted_s (O_msg q m) ss = (\<exists> sq. S_msgq sq \<in> ss \<and> smsgq_smsg_relatainted q m sq)"
|
|
179 |
| "tainted_s _ ss = False"
|
|
180 |
*)
|
|
181 |
|
|
182 |
definition taintable_s :: "t_object \<Rightarrow> bool"
|
|
183 |
where
|
|
184 |
"taintable_s obj \<equiv> \<exists> ss \<in> static. \<exists> sobj. tainted_s ss sobj \<and> init_obj_related sobj obj"
|
|
185 |
|
|
186 |
(* this definition is wrong, cause process can exited from static-world
|
|
187 |
definition deletable_s :: "t_object \<Rightarrow> bool"
|
|
188 |
where
|
|
189 |
"deletable_s obj \<equiv> init_alive obj \<and> (\<exists> ss \<in> static. \<forall> sobj \<in> ss. \<not> init_obj_related sobj obj)"
|
|
190 |
*)
|
|
191 |
|
|
192 |
lemma ss_init_sp_unique_initstate:
|
|
193 |
"\<lbrakk>S_proc (Init p, pctxt, fds) tagp \<in> init_static_state;
|
|
194 |
S_proc (Init p, pctxt', fds') tagp' \<in> init_static_state\<rbrakk>
|
|
195 |
\<Longrightarrow> pctxt' = pctxt \<and> fds' = fds \<and> tagp' = tagp"
|
|
196 |
apply (simp add:init_static_state_def)
|
|
197 |
apply (erule exE|erule conjE)+
|
|
198 |
apply (case_tac obj, case_tac [!] obja)
|
|
199 |
apply (auto simp:init_cp2sproc_def split:option.splits)
|
|
200 |
done
|
|
201 |
|
|
202 |
(*
|
|
203 |
lemma ss_init_sp_unique:
|
|
204 |
"ss \<in> static \<Longrightarrow> \<forall> p pctxt fds tagp pctxt' fds' tagp'. S_proc (Init p, pctxt, fds) tagp \<in> ss
|
|
205 |
\<and> S_proc (Init p, pctxt', fds') tagp' \<in> ss \<longrightarrow> pctxt' = pctxt \<and> fds' = fds \<and> tagp' = tagp"
|
|
206 |
apply (erule static.induct)
|
|
207 |
apply (rule allI|rule impI|erule conjE)+
|
|
208 |
apply (erule ss_init_sp_unique_initstate, simp)
|
|
209 |
sorry
|
|
210 |
|
|
211 |
lemma taintable_s_imp_init_alive_aux[rule_format]:
|
|
212 |
"ss \<in> static \<Longrightarrow> \<forall> sobj obj. sobj \<in> ss \<and> init_obj_related sobj obj \<longrightarrow> init_alive obj"
|
|
213 |
apply (erule static.induct)
|
|
214 |
apply (clarsimp simp:init_static_state_def)
|
|
215 |
apply (case_tac obj, simp)
|
|
216 |
sorry
|
|
217 |
|
|
218 |
lemma taintable_s_imp_init_alive:
|
|
219 |
"taintable_s obj \<Longrightarrow> init_alive obj"
|
|
220 |
apply (clarsimp simp add:taintable_s_def)
|
|
221 |
thm taintable_s_imp_init_alive_aux
|
|
222 |
apply (erule_tac sobj = sobj in taintable_s_imp_init_alive_aux)
|
|
223 |
apply (case_tac sobj)
|
|
224 |
apply simp_all
|
|
225 |
apply (case_tac [!] obj)
|
|
226 |
apply ( simp_all)
|
|
227 |
apply (case_tac bool, simp, simp)
|
|
228 |
apply (case_tac bool, simp+)
|
|
229 |
done
|
|
230 |
*)
|
|
231 |
|
|
232 |
fun deleted_s :: "t_static_state \<Rightarrow> t_object \<Rightarrow> bool"
|
|
233 |
where
|
|
234 |
"deleted_s ss (O_proc p) = (\<exists> pi pctxt fds tagp pctxt' fds' tagp'. S_proc (pi, pctxt, fds) tagp \<in> ss \<and>
|
|
235 |
S_proc (Init p, pctxt', fds') tagp' \<in> ss \<and> pi \<noteq> Init p \<and> permission_check pctxt pctxt' C_process P_sigkill)"
|
|
236 |
| "deleted_s ss (O_file f) = (\<exists> pi pctxt fds tagp sfs tagf fsec pfsec asecs. S_proc (pi, pctxt, fds) tagp \<in> ss \<and>
|
|
237 |
S_file sfs tagf \<in> ss \<and> (Init f, fsec, Some pfsec, asecs) \<in> sfs \<and>
|
|
238 |
search_check_s pctxt (Init f, fsec, Some pfsec, asecs) True \<and>
|
|
239 |
permission_check pctxt fsec C_file P_unlink \<and>
|
|
240 |
permission_check pctxt pfsec C_dir P_remove_name)"
|
|
241 |
| "deleted_s ss (O_dir f) = (\<exists> pi pctxt fds tagp fsec pfsec asecs. S_proc (pi, pctxt, fds) tagp \<in> ss \<and>
|
|
242 |
S_dir (Init f, fsec, Some pfsec, asecs) \<in> ss \<and> search_check_s pctxt (Init f, fsec, Some pfsec, asecs) False \<and>
|
|
243 |
permission_check pctxt fsec C_dir P_rmdir \<and> permission_check pctxt pfsec C_dir P_remove_name)"
|
|
244 |
| "deleted_s ss (O_msgq q) = (\<exists> pi pctxt fds tagp qctxt sms. S_proc (pi, pctxt, fds) tagp \<in> ss \<and>
|
|
245 |
S_msgq (Init q, qctxt, sms) \<in> ss \<and> permission_check pctxt qctxt C_msgq P_destroy)"
|
|
246 |
|
|
247 |
definition deletable_s :: "t_object \<Rightarrow> bool"
|
|
248 |
where
|
|
249 |
"deletable_s obj \<equiv> init_alive obj \<and> (\<exists> ss \<in> static. deleted_s ss obj)"
|
|
250 |
|
|
251 |
definition undeletable_s :: "t_object \<Rightarrow> bool"
|
|
252 |
where
|
|
253 |
"undeletable_s obj \<equiv> init_alive obj \<and> (\<forall> ss \<in> static. \<not> deleted_s ss obj)"
|
|
254 |
|
|
255 |
definition init_ss_eq:: "t_static_state \<Rightarrow> t_static_state \<Rightarrow> bool" (infix "\<doteq>" 100)
|
|
256 |
where
|
|
257 |
"ss \<doteq> ss' \<equiv> ss \<subseteq> ss' \<and> {sobj. is_init_sobj sobj \<and> sobj \<in> ss'} \<subseteq> ss"
|
|
258 |
|
|
259 |
lemma [simp]: "ss \<doteq> ss"
|
|
260 |
by (auto simp:init_ss_eq_def)
|
|
261 |
|
|
262 |
definition init_ss_in:: "t_static_state \<Rightarrow> t_static_state set \<Rightarrow> bool" (infix "\<propto>" 101)
|
|
263 |
where
|
|
264 |
"ss \<propto> sss \<equiv> \<exists> ss' \<in> sss. ss \<doteq> ss'"
|
|
265 |
|
|
266 |
lemma s2ss_included_sobj:
|
93
|
267 |
"\<lbrakk>dalive s obj; co2sobj s obj= Some sobj\<rbrakk> \<Longrightarrow> sobj \<in> (s2ss s)"
|
77
|
268 |
by (simp add:s2ss_def, rule_tac x = obj in exI, simp)
|
|
269 |
|
93
|
270 |
fun init_dobj_related :: "t_sobject \<Rightarrow> t_dobject \<Rightarrow> bool"
|
|
271 |
where
|
|
272 |
"init_dobj_related (S_proc (pi, sec, fds) tag) (D_proc p') = (pi = Init p')"
|
|
273 |
| "init_dobj_related (S_file sfs tag) (D_file f) = (\<exists> sf \<in> sfs. sfile_related sf f)"
|
|
274 |
| "init_dobj_related (S_dir sf) (D_dir f) = (sfile_related sf f)"
|
|
275 |
| "init_dobj_related _ _ = False"
|
|
276 |
|
77
|
277 |
lemma init_ss_in_prop:
|
93
|
278 |
"\<lbrakk>s2ss s \<propto> static; co2sobj s obj = Some sobj; dalive s obj; init_dobj_related sobj obj\<rbrakk>
|
77
|
279 |
\<Longrightarrow> \<exists> ss \<in> static. sobj \<in> ss"
|
|
280 |
apply (simp add:init_ss_in_def init_ss_eq_def)
|
|
281 |
apply (erule bexE, erule conjE)
|
|
282 |
apply (rule_tac x = ss' in bexI, auto dest!:s2ss_included_sobj)
|
|
283 |
done
|
|
284 |
|
|
285 |
|
|
286 |
end
|
|
287 |
|
|
288 |
end |