77
+ − 1
(*<*)
+ − 2
theory Co2sobj_prop
+ − 3
imports Main Flask Flask_type Static Static_type Sectxt_prop Init_prop Current_files_prop Current_sockets_prop Delete_prop Proc_fd_of_file_prop Tainted_prop Valid_prop Init_prop Alive_prop
+ − 4
begin
+ − 5
(*<*)
+ − 6
+ − 7
ML {*
+ − 8
fun print_ss ss =
+ − 9
let
+ − 10
val {simps, congs, procs, ...} = Raw_Simplifier.dest_ss ss
+ − 11
in
+ − 12
simps
+ − 13
end
+ − 14
*}
+ − 15
+ − 16
+ − 17
context tainting_s begin
+ − 18
+ − 19
(********************* cm2smsg simpset ***********************)
+ − 20
+ − 21
lemma cm2smsg_other:
+ − 22
"\<lbrakk>valid (e # s); \<forall> p q m. e \<noteq> SendMsg p q m; \<forall> p q m. e \<noteq> RecvMsg p q m; \<forall> p q. e \<noteq> RemoveMsgq p q\<rbrakk>
+ − 23
\<Longrightarrow> cm2smsg (e # s) = cm2smsg s"
+ − 24
apply (frule vt_grant_os, frule vd_cons, rule ext, rule ext)
+ − 25
unfolding cm2smsg_def
+ − 26
apply (case_tac e)
+ − 27
apply (auto simp:sectxt_of_obj_simps
+ − 28
split:t_object.splits option.splits if_splits
+ − 29
dest!:current_has_sec')
+ − 30
done
+ − 31
+ − 32
lemma cm2smsg_sendmsg:
+ − 33
"valid (SendMsg p q m # s) \<Longrightarrow> cm2smsg (SendMsg p q m # s) = (\<lambda> q' m'.
+ − 34
if (q' = q \<and> m' = m)
+ − 35
then (case (sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of
+ − 36
Some sec \<Rightarrow> Some (Created, sec, O_msg q m \<in> tainted (SendMsg p q m # s))
+ − 37
| _ \<Rightarrow> None)
+ − 38
else cm2smsg s q' m') "
+ − 39
apply (frule vd_cons, frule vt_grant_os, rule ext, rule ext)
+ − 40
apply (auto simp:cm2smsg_def sectxt_of_obj_simps
88
+ − 41
split:if_splits option.splits dest!:current_has_sec')
77
+ − 42
done
+ − 43
+ − 44
lemma cm2smsg_recvmsg1:
+ − 45
"\<lbrakk>q' \<noteq> q; valid (RecvMsg p q m # s)\<rbrakk> \<Longrightarrow> cm2smsg (RecvMsg p q m # s) q' = cm2smsg s q'"
+ − 46
apply (frule vd_cons, frule vt_grant_os, rule ext)
+ − 47
apply (auto simp:cm2smsg_def sectxt_of_obj_simps
+ − 48
split:if_splits option.splits)
+ − 49
done
+ − 50
+ − 51
lemma cm2smsg_recvmsg2:
+ − 52
"\<lbrakk>m' \<noteq> m; valid (RecvMsg p q m # s)\<rbrakk> \<Longrightarrow> cm2smsg (RecvMsg p q m # s) q m' = cm2smsg s q m'"
+ − 53
apply (frule vd_cons, frule vt_grant_os)
+ − 54
apply (auto simp:cm2smsg_def sectxt_of_obj_simps
+ − 55
split:if_splits option.splits)
+ − 56
done
+ − 57
+ − 58
lemma cm2smsg_recvmsg1':
+ − 59
"\<lbrakk>valid (RecvMsg p q m # s); q' \<noteq> q\<rbrakk> \<Longrightarrow> cm2smsg (RecvMsg p q m # s) q' = cm2smsg s q'"
+ − 60
apply (frule vd_cons, frule vt_grant_os, rule ext)
+ − 61
apply (auto simp:cm2smsg_def sectxt_of_obj_simps
+ − 62
split:if_splits option.splits)
+ − 63
done
+ − 64
+ − 65
lemma cm2smsg_recvmsg2':
+ − 66
"\<lbrakk>valid (RecvMsg p q m # s); m' \<noteq> m\<rbrakk> \<Longrightarrow> cm2smsg (RecvMsg p q m # s) q m' = cm2smsg s q m'"
+ − 67
apply (frule vd_cons, frule vt_grant_os)
+ − 68
apply (auto simp:cm2smsg_def sectxt_of_obj_simps
+ − 69
split:if_splits option.splits)
+ − 70
done
+ − 71
+ − 72
lemma cm2smsg_removemsgq:
+ − 73
"\<lbrakk>q' \<noteq> q; valid (RemoveMsgq p q # s)\<rbrakk> \<Longrightarrow> cm2smsg (RemoveMsgq p q # s) q' = cm2smsg s q'"
+ − 74
apply (frule vd_cons, frule vt_grant_os, rule ext)
+ − 75
apply (auto simp:cm2smsg_def sectxt_of_obj_simps
+ − 76
split:if_splits option.splits)
+ − 77
done
+ − 78
88
+ − 79
lemmas cm2smsg_simps = cm2smsg_sendmsg cm2smsg_recvmsg1' cm2smsg_recvmsg2'
77
+ − 80
cm2smsg_removemsgq cm2smsg_other
+ − 81
87
+ − 82
(*
77
+ − 83
lemma init_cm2smsg_has_smsg:
+ − 84
"\<lbrakk>m \<in> set (init_msgs_of_queue q); q \<in> init_msgqs\<rbrakk> \<Longrightarrow> \<exists> sm. init_cm2smsg q m = Some sm"
+ − 85
by (auto simp:init_cm2smsg_def split:option.splits dest:init_msg_has_ctxt)
87
+ − 86
*)
77
+ − 87
+ − 88
lemma hd_set_prop1:
+ − 89
"hd l \<notin> set l \<Longrightarrow> l = []"
+ − 90
by (case_tac l, auto)
+ − 91
+ − 92
lemma tl_set_prop1:
+ − 93
"a \<in> set (tl l) \<Longrightarrow> a \<in> set l"
+ − 94
by (case_tac l, auto)
+ − 95
+ − 96
lemma current_has_smsg:
+ − 97
"\<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sm. cm2smsg s q m = Some sm"
+ − 98
apply (induct s)
88
+ − 99
apply (simp add:msgs_of_queue.simps current_msgqs.simps)
77
+ − 100
+ − 101
apply (frule vt_grant_os, frule vd_cons, case_tac a)
+ − 102
apply (auto simp:cm2smsg_def sectxt_of_obj_simps split:if_splits option.splits
88
+ − 103
dest!:current_has_sec' hd_set_prop1 dest:tl_set_prop1)
77
+ − 104
done
+ − 105
+ − 106
lemma current_has_smsg':
+ − 107
"\<lbrakk>cm2smsg s q m = None; q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> m \<notin> set (msgs_of_queue s q)"
+ − 108
by (auto dest:current_has_smsg)
+ − 109
+ − 110
lemma cqm2sms_has_sms_aux:
+ − 111
"\<forall> m \<in> set ms. sectxt_of_obj s (O_msg q m) \<noteq> None \<Longrightarrow> (\<exists> sms. cqm2sms s q ms = Some sms)"
+ − 112
by (induct ms, auto split:option.splits simp:cm2smsg_def)
+ − 113
+ − 114
lemma current_has_sms:
+ − 115
"\<lbrakk>q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sms. cqm2sms s q (msgs_of_queue s q) = Some sms"
+ − 116
apply (rule cqm2sms_has_sms_aux)
+ − 117
apply (auto dest:current_msg_has_sec)
+ − 118
done
+ − 119
+ − 120
lemma current_has_sms':
+ − 121
"\<lbrakk>cqm2sms s q (msgs_of_queue s q) = None; valid s\<rbrakk> \<Longrightarrow> q \<notin> current_msgqs s"
+ − 122
by (auto dest:current_has_sms)
+ − 123
+ − 124
(********************* cqm2sms simpset ***********************)
+ − 125
+ − 126
lemma cqm2sms_other:
+ − 127
"\<lbrakk>valid (e # s); \<forall> p m. e \<noteq> CreateMsgq p q; \<forall> p q m. e \<noteq> SendMsg p q m;
+ − 128
\<forall> p q m. e \<noteq> RecvMsg p q m; \<forall> p q. e \<noteq> RemoveMsgq p q\<rbrakk>
+ − 129
\<Longrightarrow> cqm2sms (e # s) = cqm2sms s"
+ − 130
apply (rule ext, rule ext)
+ − 131
apply (induct_tac xa, simp)
+ − 132
apply (frule vt_grant_os, frule vd_cons, case_tac e)
+ − 133
apply (auto split:option.splits dest:cm2smsg_other)
+ − 134
done
+ − 135
+ − 136
lemma cqm2sms_createmsgq:
+ − 137
"valid (CreateMsgq p q # s) \<Longrightarrow> cqm2sms (CreateMsgq p q # s) = (\<lambda> q' ms'.
+ − 138
if (q' = q \<and> ms' = []) then Some []
+ − 139
else cqm2sms s q' ms')"
+ − 140
apply (rule ext, rule ext)
+ − 141
apply (frule vt_grant_os, frule vd_cons, induct_tac ms')
+ − 142
apply (auto split:if_splits option.splits dest:cm2smsg_other)
+ − 143
done
+ − 144
+ − 145
lemma cqm2sms_2:
+ − 146
"cqm2sms s q (ms @ [m]) =
+ − 147
(case (cqm2sms s q ms, cm2smsg s q m) of
+ − 148
(Some sms, Some sm) \<Rightarrow> Some (sms @ [sm])
+ − 149
| _ \<Rightarrow> None)"
+ − 150
apply (induct ms, simp split:option.splits)
+ − 151
by (auto split:option.splits)
+ − 152
+ − 153
lemmas cqm2sms_simps2 = cqm2sms.simps(1) cqm2sms_2
+ − 154
+ − 155
declare cqm2sms.simps [simp del]
+ − 156
+ − 157
lemma cqm2sms_prop1:
+ − 158
"cqm2sms s q ms = None \<Longrightarrow> \<exists> m \<in> set ms. cm2smsg s q m = None"
+ − 159
by (induct ms, auto simp:cqm2sms.simps split:option.splits)
+ − 160
+ − 161
lemma cqm2sms_sendmsg1:
+ − 162
"\<lbrakk>valid (SendMsg p q m # s); m \<notin> set ms\<rbrakk>
+ − 163
\<Longrightarrow> cqm2sms (SendMsg p q m # s) q' ms = cqm2sms s q' ms"
+ − 164
apply (frule vt_grant_os, frule vd_cons)
+ − 165
apply (frule cm2smsg_sendmsg)
+ − 166
apply (induct ms rule:rev_induct)
+ − 167
apply (auto split:if_splits option.splits simp:cqm2sms_simps2)
+ − 168
done
+ − 169
+ − 170
lemma cqm2sms_sendmsg2:
+ − 171
"\<lbrakk>valid (SendMsg p q m # s); q' \<noteq> q\<rbrakk>
+ − 172
\<Longrightarrow> cqm2sms (SendMsg p q m # s) q' ms = cqm2sms s q' ms"
+ − 173
apply (frule vt_grant_os, frule vd_cons)
+ − 174
apply (frule cm2smsg_sendmsg)
+ − 175
apply (induct ms rule:rev_induct)
+ − 176
apply (auto split:if_splits option.splits simp:cqm2sms_simps2)
+ − 177
done
+ − 178
+ − 179
lemma cqm2sms_sendmsg3:
+ − 180
"\<lbrakk>valid (SendMsg p q m # s); ms' = msgs_of_queue (SendMsg p q m # s) q\<rbrakk>
+ − 181
\<Longrightarrow> cqm2sms (SendMsg p q m # s) q ms' =
+ − 182
(case (cqm2sms s q (msgs_of_queue s q), sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of
+ − 183
(Some sms, Some sec) \<Rightarrow> Some (sms @ [(Created, sec, O_msg q m \<in> tainted (SendMsg p q m # s))])
+ − 184
| _ \<Rightarrow> None)"
+ − 185
apply (frule vt_grant_os, frule vd_cons)
+ − 186
apply (frule cm2smsg_sendmsg)
+ − 187
apply (auto split:if_splits option.splits simp:cqm2sms_simps2 cqm2sms_sendmsg1)
+ − 188
done
+ − 189
+ − 190
lemma cqm2sms_sendmsg:
+ − 191
"\<lbrakk>valid (SendMsg p q m # s); ms' = msgs_of_queue (SendMsg p q m # s) q'\<rbrakk>
+ − 192
\<Longrightarrow> cqm2sms (SendMsg p q m # s) q' ms' = (
+ − 193
if (q' = q)
+ − 194
then (case (cqm2sms s q (msgs_of_queue s q), sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of
+ − 195
(Some sms, Some sec) \<Rightarrow> Some (sms @ [(Created, sec, O_msg q m \<in> tainted (SendMsg p q m # s))])
+ − 196
| _ \<Rightarrow> None)
+ − 197
else cqm2sms s q' ms' )"
+ − 198
apply (simp split:if_splits add:cqm2sms_sendmsg2 cqm2sms_sendmsg3)
+ − 199
done
+ − 200
+ − 201
lemma cqm2sms_recvmsg1:
+ − 202
"\<lbrakk>valid (RecvMsg p q m # s); m \<notin> set ms\<rbrakk>
+ − 203
\<Longrightarrow> cqm2sms (RecvMsg p q m # s) q ms = cqm2sms s q ms"
+ − 204
apply (frule vt_grant_os, frule vd_cons)
+ − 205
apply (induct ms rule:rev_induct)
+ − 206
apply (auto split:if_splits option.splits simp:cqm2sms_simps2 cm2smsg_recvmsg2')
+ − 207
done
+ − 208
+ − 209
lemma cqm2sms_recvmsg2:
+ − 210
"\<lbrakk>valid (RecvMsg p q m # s); q' \<noteq> q\<rbrakk>
+ − 211
\<Longrightarrow> cqm2sms (RecvMsg p q m # s) q' ms = cqm2sms s q' ms"
+ − 212
apply (frule vt_grant_os, frule vd_cons)
+ − 213
apply (induct ms rule:rev_induct)
+ − 214
apply (auto split:if_splits option.splits simp:cqm2sms_simps2 cm2smsg_recvmsg1')
+ − 215
done
+ − 216
+ − 217
lemma cqm2sms_recvmsg:
+ − 218
"\<lbrakk>valid (RecvMsg p q m # s); ms = msgs_of_queue (RecvMsg p q m # s) q'\<rbrakk>
+ − 219
\<Longrightarrow> cqm2sms (RecvMsg p q m # s) q' ms = (
+ − 220
if (q' = q)
+ − 221
then (case (cqm2sms s q (msgs_of_queue s q)) of
+ − 222
Some sms \<Rightarrow> Some (tl sms)
+ − 223
| _ \<Rightarrow> None)
+ − 224
else cqm2sms s q' ms)"
+ − 225
apply (frule vt_grant_os, frule vd_cons)
+ − 226
apply (auto split:if_splits option.splits simp:cqm2sms_recvmsg1 cqm2sms_recvmsg2
+ − 227
dest!:current_has_sms')
+ − 228
apply (case_tac "msgs_of_queue s q", simp)
+ − 229
apply (frule_tac ms = "tl (msgs_of_queue s q)" in cqm2sms_recvmsg1)
+ − 230
apply (drule_tac q = q in distinct_queue_msgs, simp+)
+ − 231
apply (case_tac a, auto simp:cqm2sms.simps split:option.splits if_splits)
+ − 232
done
+ − 233
+ − 234
lemma cqm2sms_removemsgq:
+ − 235
"\<lbrakk>valid (RemoveMsgq p q # s); q' \<noteq> q; q' \<in> current_msgqs s\<rbrakk>
+ − 236
\<Longrightarrow> cqm2sms (RemoveMsgq p q # s) q' ms = cqm2sms s q' ms"
+ − 237
apply (frule vt_grant_os, frule vd_cons)
+ − 238
apply (induct ms rule:rev_induct)
+ − 239
apply (auto simp:cqm2sms_simps2 cm2smsg_removemsgq split:option.splits if_splits)
+ − 240
done
+ − 241
+ − 242
lemmas cqm2sms_simps = cqm2sms_other cqm2sms_createmsgq cqm2sms_sendmsg cqm2sms_recvmsg cqm2sms_removemsgq
+ − 243
+ − 244
(********************* cq2smsgq simpset ***********************)
+ − 245
+ − 246
lemma cq2smsgq_other:
+ − 247
"\<lbrakk>valid (e # s); \<forall> p q. e \<noteq> CreateMsgq p q; \<forall> p q m. e \<noteq> SendMsg p q m;
+ − 248
\<forall> p q m. e \<noteq> RecvMsg p q m; \<forall> p q. e \<noteq> RemoveMsgq p q\<rbrakk>
+ − 249
\<Longrightarrow> cq2smsgq (e # s) = cq2smsgq s"
+ − 250
apply (frule cqm2sms_other, simp+)
+ − 251
apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac e)
+ − 252
apply (auto simp:cq2smsgq_def sectxt_of_obj_simps
88
+ − 253
split:t_object.splits option.splits if_splits)
77
+ − 254
done
+ − 255
+ − 256
lemma cq2smsg_createmsgq:
+ − 257
"valid (CreateMsgq p q # s)
+ − 258
\<Longrightarrow> cq2smsgq (CreateMsgq p q # s) = (cq2smsgq s) (q :=
+ − 259
case (sectxt_of_obj s (O_proc p)) of
+ − 260
Some psec \<Rightarrow> Some (Created, (fst psec, R_object, (snd o snd) psec), [])
+ − 261
| None \<Rightarrow> None)"
+ − 262
apply (frule vd_cons, frule vt_grant_os)
+ − 263
apply (frule cqm2sms_createmsgq)
+ − 264
apply (rule ext, auto simp:cq2smsgq_def sec_createmsgq
87
+ − 265
split:option.splits if_splits)
77
+ − 266
done
+ − 267
+ − 268
lemma cq2smsg_sendmsg:
+ − 269
"valid (SendMsg p q m # s)
+ − 270
\<Longrightarrow> cq2smsgq (SendMsg p q m # s) = (cq2smsgq s) (q :=
+ − 271
case (cq2smsgq s q, sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of
+ − 272
(Some (qi, sec, sms), Some msec) \<Rightarrow>
+ − 273
Some (qi, sec, sms @ [(Created, msec, O_msg q m \<in> tainted (SendMsg p q m # s))])
+ − 274
| _ \<Rightarrow> None)"
+ − 275
apply (frule vd_cons, frule vt_grant_os)
+ − 276
apply (rule ext)
+ − 277
apply (frule_tac q' = x in cqm2sms_sendmsg, simp)
+ − 278
apply (auto simp:cq2smsgq_def sectxt_of_obj_simps split:option.splits if_splits
+ − 279
dest!:current_has_sms' current_has_sec')
+ − 280
done
+ − 281
+ − 282
lemma current_has_smsgq:
+ − 283
"\<lbrakk>q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sq. cq2smsgq s q = Some sq"
88
+ − 284
by (auto simp:cq2smsgq_def split:option.splits
+ − 285
dest!:current_has_sec' current_has_sms' dest:current_has_sec)
77
+ − 286
+ − 287
lemma current_has_smsgq':
+ − 288
"\<lbrakk>cq2smsgq s q = None; valid s\<rbrakk> \<Longrightarrow> q \<notin> current_msgqs s"
+ − 289
by (auto dest:current_has_smsgq)
+ − 290
+ − 291
lemma cq2smsg_recvmsg:
+ − 292
"valid (RecvMsg p q m # s)
+ − 293
\<Longrightarrow> cq2smsgq (RecvMsg p q m # s) = (cq2smsgq s) (q :=
+ − 294
case (cq2smsgq s q) of
+ − 295
Some (qi, sec, sms) \<Rightarrow> Some (qi, sec, tl sms)
+ − 296
| _ \<Rightarrow> None)"
+ − 297
apply (frule vd_cons, frule vt_grant_os)
+ − 298
apply (rule ext, frule_tac q' = x in cqm2sms_recvmsg, simp)
+ − 299
apply (auto simp add:cq2smsgq_def sectxt_of_obj_simps split:option.splits if_splits
+ − 300
dest!:current_has_sec' current_has_sms' current_has_smsgq')
+ − 301
done
+ − 302
+ − 303
lemma cq2smsg_removemsgq:
+ − 304
"\<lbrakk>valid (RemoveMsgq p q # s); q' \<noteq> q; q' \<in> current_msgqs s\<rbrakk>
+ − 305
\<Longrightarrow> cq2smsgq (RemoveMsgq p q # s) q' = cq2smsgq s q'"
+ − 306
apply (frule vd_cons, frule vt_grant_os)
+ − 307
apply (auto simp:cq2smsgq_def sectxt_of_obj_simps cqm2sms_removemsgq split:if_splits option.splits
+ − 308
dest!:current_has_sec' current_has_sms' current_has_smsgq')
+ − 309
done
+ − 310
+ − 311
lemmas cq2smsgq_simps = cq2smsgq_other cq2smsg_sendmsg cq2smsg_recvmsg cq2smsg_removemsgq
+ − 312
88
+ − 313
(*
77
+ − 314
lemma sm_in_sqsms_init_aux:
+ − 315
"\<lbrakk>m \<in> set ms; init_cm2smsg q m = Some sm; q \<in> init_msgqs;
+ − 316
init_cqm2sms q ms = Some sms\<rbrakk> \<Longrightarrow> sm \<in> set sms"
+ − 317
apply (induct ms arbitrary:m sm sms)
+ − 318
by (auto split:if_splits option.splits)
+ − 319
+ − 320
lemma sm_in_sqsms_init:
+ − 321
"\<lbrakk>m \<in> set (init_msgs_of_queue q); init_cm2smsg q m = Some sm; q \<in> init_msgqs;
+ − 322
init_cqm2sms q (init_msgs_of_queue q) = Some sms\<rbrakk> \<Longrightarrow> sm \<in> set sms"
+ − 323
by (simp add:sm_in_sqsms_init_aux)
88
+ − 324
*)
77
+ − 325
+ − 326
lemma cqm2sms_prop0:
+ − 327
"\<lbrakk>m \<in> set ms; cm2smsg s q m = Some sm; cqm2sms s q ms = Some sms\<rbrakk> \<Longrightarrow> sm \<in> set sms"
+ − 328
apply (induct ms arbitrary:m sm sms)
+ − 329
apply (auto simp:cqm2sms.simps split:option.splits)
+ − 330
done
+ − 331
+ − 332
lemma sm_in_sqsms:
+ − 333
"\<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s; valid s; cq2smsgq s q = Some (qi, qsec, sms);
+ − 334
cm2smsg s q m = Some sm\<rbrakk> \<Longrightarrow> sm \<in> set sms"
+ − 335
proof (induct s arbitrary:m q qi qsec sms sm)
+ − 336
case Nil
+ − 337
thus ?case
88
+ − 338
by (simp split:option.splits)
77
+ − 339
next
+ − 340
case (Cons e s)
+ − 341
hence p1:"\<And> m q qi qsec sms sm. \<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s; valid s;
+ − 342
cq2smsgq s q = Some (qi, qsec, sms); cm2smsg s q m = Some sm\<rbrakk>
+ − 343
\<Longrightarrow> sm \<in> set sms" and p2: "m \<in> set (msgs_of_queue (e # s) q)"
+ − 344
and p3: "q \<in> current_msgqs (e # s)" and p4: "valid (e # s)"
+ − 345
and p5: "cq2smsgq (e # s) q = Some (qi, qsec, sms)"
+ − 346
and p6: "cm2smsg (e # s) q m = Some sm" by auto
+ − 347
from p4 have os: "os_grant s e" and vd: "valid s" by (auto dest:vd_cons vt_grant_os)
+ − 348
(*
+ − 349
from p1 have p1':
+ − 350
"\<And> m q qi qsec sms sm ms. \<lbrakk>m \<in> set ms; set ms \<subseteq> set (msgs_of_queue s q); q \<in> current_msgqs s;
+ − 351
valid s; cq2smsgq s q = Some (qi, qsec, sms); cm2smsg s q m = Some sm\<rbrakk>
+ − 352
\<Longrightarrow> sm \<in> set "
+ − 353
*)
+ − 354
show ?case using p2 p3 p4 p5 p6 vd os
+ − 355
apply (case_tac e)
+ − 356
apply (auto simp:cq2smsgq_simps cm2smsg_simps split:if_splits option.splits intro:p1)
+ − 357
+ − 358
apply (rule_tac m = m and q = q and qi = qi and qsec = qsec in p1, simp+)
+ − 359
apply (case_tac "q = nat2", simp)
+ − 360
apply (drule cq2smsg_createmsgq, simp, simp)
+ − 361
+ − 362
apply (drule_tac m = m and q = q and qi = qi and qsec = "(aa,ab,b)" in p1, simp+)
+ − 363
apply (drule_tac m = m and q = q and qi = qi and qsec = "(aa,ab,b)" in p1, simp+)
+ − 364
+ − 365
apply (simp add:cq2smsgq_def split:option.splits)
+ − 366
apply (rule_tac m = m and sm = sm in cqm2sms_prop0, simp+)
+ − 367
apply (simp add:cqm2sms_recvmsg)
+ − 368
done
+ − 369
qed
+ − 370
+ − 371
(****************** cf2sfile path simpset ***************)
+ − 372
+ − 373
lemma sroot_only:
+ − 374
"cf2sfile s [] = Some sroot"
+ − 375
by (simp add:cf2sfile_def)
+ − 376
+ − 377
lemma not_file_is_dir:
+ − 378
"\<lbrakk>\<not> is_file s f; f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> is_dir s f"
+ − 379
by (auto simp:is_file_def current_files_def is_dir_def
+ − 380
dest:finum_has_itag finum_has_ftag' split:t_inode_tag.splits option.splits)
+ − 381
+ − 382
lemma not_dir_is_file:
+ − 383
"\<lbrakk>\<not> is_dir s f; f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> is_file s f"
+ − 384
by (auto simp:is_file_def current_files_def is_dir_def
+ − 385
dest:finum_has_itag finum_has_ftag' split:t_inode_tag.splits option.splits)
+ − 386
+ − 387
lemma is_file_or_dir:
+ − 388
"\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> is_file s f \<or> is_dir s f"
+ − 389
by (auto dest:not_dir_is_file)
+ − 390
+ − 391
lemma current_file_has_sfile:
+ − 392
"\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> sf. cf2sfile s f = Some sf"
+ − 393
apply (induct f)
+ − 394
apply (rule_tac x = "sroot" in exI, simp add:sroot_only)
+ − 395
apply (frule parentf_in_current', simp, clarsimp)
+ − 396
apply (frule parentf_is_dir'', simp)
+ − 397
apply (frule is_file_or_dir, simp)
+ − 398
apply (auto dest!:current_has_sec'
+ − 399
simp:cf2sfile_def split:option.splits if_splits dest!:get_pfs_secs_prop')
+ − 400
done
+ − 401
+ − 402
definition sectxt_of_pf :: "t_state \<Rightarrow> t_file \<Rightarrow> security_context_t option"
+ − 403
where
+ − 404
"sectxt_of_pf s f = (case f of [] \<Rightarrow> None | (a # pf) \<Rightarrow> sectxt_of_obj s (O_dir pf))"
+ − 405
+ − 406
definition get_parentfs_ctxts' :: "t_state \<Rightarrow> t_file \<Rightarrow> (security_context_t list) option"
+ − 407
where
+ − 408
"get_parentfs_ctxts' s f = (case f of [] \<Rightarrow> None | (a # pf) \<Rightarrow> get_parentfs_ctxts s pf)"
+ − 409
+ − 410
lemma is_file_has_sfile:
+ − 411
"\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> sec psec asecs. cf2sfile s f = Some
+ − 412
(if (\<not> died (O_file f) s \<and> is_init_file f) then Init f else Created,
+ − 413
sec, Some psec, set asecs) \<and> (sectxt_of_obj s (O_file f) = Some sec) \<and>
+ − 414
(sectxt_of_pf s f = Some psec) \<and> (get_parentfs_ctxts' s f = Some asecs)"
+ − 415
apply (case_tac f, simp, drule root_is_dir', simp, simp)
+ − 416
apply (frule is_file_in_current)
+ − 417
apply (drule current_file_has_sfile, simp)
+ − 418
apply (auto simp:cf2sfile_def sectxt_of_pf_def get_parentfs_ctxts'_def split:if_splits option.splits)
+ − 419
done
+ − 420
+ − 421
lemma is_file_has_sfile':
+ − 422
"\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> sf. cf2sfile s f = Some sf"
+ − 423
by (drule is_file_has_sfile, auto)
+ − 424
+ − 425
lemma is_dir_has_sfile:
+ − 426
"\<lbrakk>is_dir s f; valid s\<rbrakk> \<Longrightarrow> (case f of
+ − 427
[] \<Rightarrow> cf2sfile s f = Some sroot
+ − 428
| a # pf \<Rightarrow> (\<exists> sec psec asecs. cf2sfile s f = Some
+ − 429
(if (\<not> died (O_dir f) s \<and> is_init_dir f) then Init f else Created,
+ − 430
sec, Some psec, set asecs) \<and> (sectxt_of_obj s (O_dir f) = Some sec) \<and>
+ − 431
(sectxt_of_obj s (O_dir pf) = Some psec) \<and> (get_parentfs_ctxts s pf = Some asecs)))"
+ − 432
apply (case_tac f, simp add:sroot_only)
+ − 433
apply (frule is_dir_in_current, frule is_dir_not_file)
+ − 434
apply (drule current_file_has_sfile, simp)
+ − 435
apply (auto simp:cf2sfile_def split:if_splits option.splits)
+ − 436
done
+ − 437
+ − 438
lemma is_dir_has_sdir':
+ − 439
"\<lbrakk>is_dir s f; valid s\<rbrakk> \<Longrightarrow> \<exists> sf. cf2sfile s f = Some sf"
+ − 440
apply (case_tac f)
+ − 441
apply (rule_tac x = sroot in exI)
+ − 442
apply (simp add:sroot_only)
+ − 443
apply (drule is_dir_has_sfile, auto)
+ − 444
done
+ − 445
+ − 446
lemma sroot_set:
+ − 447
"valid s \<Longrightarrow> \<exists> sec. sroot = (Init [], sec, None, {}) \<and> sectxt_of_obj s (O_dir []) = Some sec"
+ − 448
apply (frule root_is_dir)
+ − 449
apply (drule is_dir_has_sec, simp)
+ − 450
apply (auto simp:sroot_def sec_of_root_def sectxt_of_obj_def type_of_obj.simps
+ − 451
root_type_remains root_user_remains
+ − 452
dest!:root_has_type' root_has_user' root_has_init_type' root_has_init_user'
+ − 453
split:option.splits)
+ − 454
done
+ − 455
+ − 456
lemma cf2sfile_path_file:
+ − 457
"\<lbrakk>is_file s (f # pf); valid s\<rbrakk>
+ − 458
\<Longrightarrow> cf2sfile s (f # pf) = (
+ − 459
case (cf2sfile s pf) of
+ − 460
Some (pfi, pfsec, psec, asecs) \<Rightarrow>
+ − 461
(case (sectxt_of_obj s (O_file (f # pf))) of
+ − 462
Some fsec \<Rightarrow> Some (if (\<not> died (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
+ − 463
else Created, fsec, Some pfsec, asecs \<union> {pfsec})
+ − 464
| None \<Rightarrow> None)
+ − 465
| _ \<Rightarrow> None)"
+ − 466
apply (frule is_file_in_current, drule parentf_is_dir'', simp)
+ − 467
apply (frule is_dir_has_sfile, simp, frule is_file_has_sfile, simp)
+ − 468
apply (frule sroot_set)
+ − 469
apply (case_tac pf, (clarsimp simp:get_parentfs_ctxts'_def sectxt_of_pf_def)+)
+ − 470
done
+ − 471
+ − 472
lemma cf2sfile_path_dir:
+ − 473
"\<lbrakk>is_dir s (f # pf); valid s\<rbrakk>
+ − 474
\<Longrightarrow> cf2sfile s (f # pf) = (
+ − 475
case (cf2sfile s pf) of
+ − 476
Some (pfi, pfsec, psec, asecs) \<Rightarrow>
+ − 477
(case (sectxt_of_obj s (O_dir (f # pf))) of
+ − 478
Some fsec \<Rightarrow> Some (if (\<not> died (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
+ − 479
else Created, fsec, Some pfsec, asecs \<union> {pfsec})
+ − 480
| None \<Rightarrow> None)
+ − 481
| _ \<Rightarrow> None)"
+ − 482
apply (frule is_dir_in_current, drule parentf_is_dir'', simp)
+ − 483
apply (frule_tac f = "f # pf" in is_dir_has_sfile, simp)
+ − 484
apply (frule_tac f = "pf" in is_dir_has_sfile, simp)
+ − 485
apply (frule sroot_set)
+ − 486
apply (case_tac pf, (clarsimp simp:get_parentfs_ctxts'_def sectxt_of_pf_def)+)
+ − 487
done
+ − 488
+ − 489
lemma cf2sfile_path:
+ − 490
"\<lbrakk>f # pf \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> cf2sfile s (f # pf) = (
+ − 491
case (cf2sfile s pf) of
+ − 492
Some (pfi, pfsec, psec, asecs) \<Rightarrow> (if (is_file s (f # pf))
+ − 493
then (case (sectxt_of_obj s (O_file (f # pf))) of
+ − 494
Some fsec \<Rightarrow> Some (if (\<not> died (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
+ − 495
else Created, fsec, Some pfsec, asecs \<union> {pfsec})
+ − 496
| None \<Rightarrow> None)
+ − 497
else (case (sectxt_of_obj s (O_dir (f # pf))) of
+ − 498
Some fsec \<Rightarrow> Some (if (\<not> died (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
+ − 499
else Created, fsec, Some pfsec, asecs \<union> {pfsec})
+ − 500
| None \<Rightarrow> None) )
+ − 501
| None \<Rightarrow> None)"
+ − 502
apply (drule is_file_or_dir, simp)
+ − 503
apply (erule disjE)
+ − 504
apply (frule cf2sfile_path_file, simp) defer
+ − 505
apply (frule cf2sfile_path_dir, simp, drule is_dir_not_file)
+ − 506
apply (auto split:option.splits)
+ − 507
done
+ − 508
+ − 509
lemma cf2sfile_path_file_prop1:
+ − 510
"\<lbrakk>is_file s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs); valid s\<rbrakk>
+ − 511
\<Longrightarrow> \<exists> fsec. cf2sfile s (f # pf) =
+ − 512
Some (if (\<not> died (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
+ − 513
else Created, fsec, Some pfsec, asecs \<union> {pfsec}) \<and>
+ − 514
sectxt_of_obj s (O_file (f # pf)) = Some fsec"
+ − 515
apply (frule is_file_has_sfile, simp)
+ − 516
by (auto simp:cf2sfile_path_file)
+ − 517
+ − 518
lemma cf2sfile_path_file_prop2:
+ − 519
"\<lbrakk>is_file s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs);
+ − 520
sectxt_of_obj s (O_file (f # pf)) = Some fsec; valid s\<rbrakk> \<Longrightarrow> cf2sfile s (f # pf) =
+ − 521
Some (if (\<not> died (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
+ − 522
else Created, fsec, Some pfsec, asecs \<union> {pfsec})"
+ − 523
by (drule cf2sfile_path_file_prop1, auto)
+ − 524
+ − 525
lemma cf2sfile_path_dir_prop1:
+ − 526
"\<lbrakk>is_dir s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs); valid s\<rbrakk>
+ − 527
\<Longrightarrow> \<exists> fsec. cf2sfile s (f # pf) =
+ − 528
Some (if (\<not> died (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
+ − 529
else Created, fsec, Some pfsec, asecs \<union> {pfsec}) \<and>
+ − 530
sectxt_of_obj s (O_dir (f # pf)) = Some fsec"
+ − 531
apply (frule is_dir_has_sfile, simp)
+ − 532
by (auto simp:cf2sfile_path_dir)
+ − 533
+ − 534
lemma cf2sfile_path_dir_prop2:
+ − 535
"\<lbrakk>is_dir s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs);
+ − 536
sectxt_of_obj s (O_dir (f # pf)) = Some fsec; valid s\<rbrakk> \<Longrightarrow> cf2sfile s (f # pf) =
+ − 537
Some (if (\<not> died (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
+ − 538
else Created, fsec, Some pfsec, asecs \<union> {pfsec})"
+ − 539
by (drule cf2sfile_path_dir_prop1, auto)
+ − 540
+ − 541
(**************** cf2sfile event list simpset ****************)
+ − 542
+ − 543
lemma cf2sfile_open_none':
+ − 544
"valid (Open p f flag fd None # s) \<Longrightarrow> cf2sfile (Open p f flag fd None # s) f'= cf2sfile s f'"
+ − 545
apply (frule vd_cons, frule vt_grant_os)
+ − 546
apply (induct f', simp add:cf2sfile_def)
+ − 547
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ − 548
get_parentfs_ctxts_simps)
+ − 549
done
+ − 550
+ − 551
lemma cf2sfile_open_none:
+ − 552
"valid (Open p f flag fd None # s) \<Longrightarrow> cf2sfile (Open p f flag fd None # s) = cf2sfile s"
+ − 553
apply (rule ext)
+ − 554
by (simp add:cf2sfile_open_none')
+ − 555
+ − 556
lemma cf2sfile_open_some1:
+ − 557
"\<lbrakk>valid (Open p f flag fd (Some inum) # s); f' \<in> current_files s\<rbrakk>
+ − 558
\<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'"
+ − 559
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+ − 560
apply (case_tac "f = f'", simp)
+ − 561
apply (induct f', simp add:sroot_only, simp)
+ − 562
apply (frule parentf_in_current', simp+)
+ − 563
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ − 564
get_parentfs_ctxts_simps)
+ − 565
done
+ − 566
+ − 567
lemma cf2sfile_open_some2:
+ − 568
"\<lbrakk>valid (Open p f flag fd (Some inum) # s); is_file s f'\<rbrakk>
+ − 569
\<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'"
+ − 570
apply (frule vd_cons, drule is_file_in_current)
+ − 571
by (simp add:cf2sfile_open_some1)
+ − 572
+ − 573
lemma cf2sfile_open_some3:
+ − 574
"\<lbrakk>valid (Open p f flag fd (Some inum) # s); is_dir s f'\<rbrakk>
+ − 575
\<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'"
+ − 576
apply (frule vd_cons, drule is_dir_in_current)
+ − 577
by (simp add:cf2sfile_open_some1)
+ − 578
+ − 579
lemma cf2sfile_open_some4:
+ − 580
"valid (Open p f flag fd (Some inum) # s) \<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f = (
+ − 581
case (parent f) of
+ − 582
Some pf \<Rightarrow> (case (sectxt_of_obj (Open p f flag fd (Some inum) # s) (O_file f), sectxt_of_obj s (O_dir pf),
+ − 583
get_parentfs_ctxts s pf) of
+ − 584
(Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ − 585
| _ \<Rightarrow> None)
+ − 586
| None \<Rightarrow> None)"
+ − 587
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+ − 588
apply (case_tac f, simp)
+ − 589
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ − 590
get_parentfs_ctxts_simps)
+ − 591
apply (rule impI, (erule conjE)+)
+ − 592
apply (drule not_died_init_file, simp+)
+ − 593
apply (simp add:is_file_in_current)
+ − 594
done
+ − 595
+ − 596
lemma cf2sfile_open:
+ − 597
"\<lbrakk>valid (Open p f flag fd opt # s); f' \<in> current_files (Open p f flag fd opt # s)\<rbrakk>
+ − 598
\<Longrightarrow> cf2sfile (Open p f flag fd opt # s) f' = (
+ − 599
if (opt = None) then cf2sfile s f'
+ − 600
else if (f' = f)
+ − 601
then (case (parent f) of
+ − 602
Some pf \<Rightarrow> (case (sectxt_of_obj (Open p f flag fd opt # s) (O_file f), sectxt_of_obj s (O_dir pf),
+ − 603
get_parentfs_ctxts s pf) of
+ − 604
(Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ − 605
| _ \<Rightarrow> None)
+ − 606
| None \<Rightarrow> None)
+ − 607
else cf2sfile s f')"
+ − 608
apply (case_tac opt)
+ − 609
apply (simp add:cf2sfile_open_none)
+ − 610
apply (case_tac "f = f'")
+ − 611
apply (simp add:cf2sfile_open_some4 split:option.splits)
+ − 612
apply (simp add:cf2sfile_open_some1 current_files_simps)
+ − 613
done
+ − 614
+ − 615
lemma cf2sfile_mkdir1:
+ − 616
"\<lbrakk>valid (Mkdir p f i # s); f' \<in> current_files s\<rbrakk>
+ − 617
\<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = cf2sfile s f'"
+ − 618
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+ − 619
apply (case_tac "f = f'", simp)
+ − 620
apply (induct f', simp add:sroot_only, simp)
+ − 621
apply (frule parentf_in_current', simp+)
+ − 622
apply (case_tac "f = f'", simp)
+ − 623
apply (simp add:cf2sfile_path is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ − 624
get_parentfs_ctxts_simps split:if_splits option.splits)
+ − 625
done
+ − 626
+ − 627
lemma cf2sfile_mkdir2:
+ − 628
"\<lbrakk>valid (Mkdir p f i # s); is_file s f'\<rbrakk>
+ − 629
\<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = cf2sfile s f'"
+ − 630
apply (frule vd_cons, drule is_file_in_current)
+ − 631
by (simp add:cf2sfile_mkdir1)
+ − 632
+ − 633
lemma cf2sfile_mkdir3:
+ − 634
"\<lbrakk>valid (Mkdir p f i # s); is_dir s f'\<rbrakk>
+ − 635
\<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = cf2sfile s f'"
+ − 636
apply (frule vd_cons, drule is_dir_in_current)
+ − 637
by (simp add:cf2sfile_mkdir1)
+ − 638
+ − 639
lemma cf2sfile_mkdir4:
+ − 640
"valid (Mkdir p f i # s)
+ − 641
\<Longrightarrow> cf2sfile (Mkdir p f i # s) f = (case (parent f) of
+ − 642
Some pf \<Rightarrow> (case (sectxt_of_obj (Mkdir p f i # s) (O_dir f), sectxt_of_obj s (O_dir pf),
+ − 643
get_parentfs_ctxts s pf) of
+ − 644
(Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ − 645
| _ \<Rightarrow> None)
+ − 646
| None \<Rightarrow> None)"
+ − 647
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+ − 648
apply (case_tac f, simp)
+ − 649
apply (clarsimp simp:os_grant.simps)
+ − 650
apply (simp add:sectxt_of_obj_simps)
+ − 651
apply (frule current_proc_has_sec, simp)
+ − 652
apply (frule is_dir_has_sec, simp)
+ − 653
apply (frule get_pfs_secs_prop, simp)
+ − 654
apply (frule is_dir_not_file)
+ − 655
apply (auto simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ − 656
get_parentfs_ctxts_simps split:option.splits if_splits
+ − 657
dest:not_died_init_dir is_dir_in_current not_died_init_file is_file_in_current)
+ − 658
done
+ − 659
+ − 660
lemma cf2sfile_mkdir:
+ − 661
"\<lbrakk>valid (Mkdir p f i # s); f' \<in> current_files (Mkdir p f i # s)\<rbrakk>
+ − 662
\<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = (
+ − 663
if (f' = f)
+ − 664
then (case (parent f) of
+ − 665
Some pf \<Rightarrow> (case (sectxt_of_obj (Mkdir p f i # s) (O_dir f), sectxt_of_obj s (O_dir pf),
+ − 666
get_parentfs_ctxts s pf) of
+ − 667
(Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ − 668
| _ \<Rightarrow> None)
+ − 669
| None \<Rightarrow> None)
+ − 670
else cf2sfile s f')"
+ − 671
apply (case_tac "f = f'")
+ − 672
apply (simp add:cf2sfile_mkdir4 split:option.splits)
+ − 673
apply (simp add:cf2sfile_mkdir1 current_files_simps)
+ − 674
done
+ − 675
+ − 676
lemma cf2sfile_linkhard1:
+ − 677
"\<lbrakk>valid (LinkHard p oldf f # s); f' \<in> current_files s\<rbrakk>
+ − 678
\<Longrightarrow> cf2sfile (LinkHard p oldf f# s) f' = cf2sfile s f'"
+ − 679
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+ − 680
apply (case_tac "f = f'", simp)
+ − 681
apply (induct f', simp add:sroot_only, simp)
+ − 682
apply (frule parentf_in_current', simp+)
+ − 683
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ − 684
get_parentfs_ctxts_simps split:if_splits option.splits)
+ − 685
done
+ − 686
+ − 687
lemma cf2sfile_linkhard2:
+ − 688
"\<lbrakk>valid (LinkHard p oldf f # s); is_file s f'\<rbrakk>
+ − 689
\<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f' = cf2sfile s f'"
+ − 690
apply (frule vd_cons, drule is_file_in_current)
+ − 691
by (simp add:cf2sfile_linkhard1)
+ − 692
+ − 693
lemma cf2sfile_linkhard3:
+ − 694
"\<lbrakk>valid (LinkHard p oldf f # s); is_dir s f'\<rbrakk>
+ − 695
\<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f' = cf2sfile s f'"
+ − 696
apply (frule vd_cons, drule is_dir_in_current)
+ − 697
by (simp add:cf2sfile_linkhard1)
+ − 698
+ − 699
lemma cf2sfile_linkhard4:
+ − 700
"valid (LinkHard p oldf f # s)
+ − 701
\<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f = (case (parent f) of
+ − 702
Some pf \<Rightarrow> (case (sectxt_of_obj (LinkHard p oldf f # s) (O_file f), sectxt_of_obj s (O_dir pf),
+ − 703
get_parentfs_ctxts s pf) of
+ − 704
(Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ − 705
| _ \<Rightarrow> None)
+ − 706
| None \<Rightarrow> None)"
+ − 707
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
+ − 708
apply (case_tac f, simp)
+ − 709
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ − 710
get_parentfs_ctxts_simps)
+ − 711
apply (rule impI, (erule conjE)+)
+ − 712
apply (drule not_died_init_file, simp+)
+ − 713
apply (simp add:is_file_in_current)
+ − 714
done
+ − 715
+ − 716
lemma cf2sfile_linkhard:
+ − 717
"\<lbrakk>valid (LinkHard p oldf f # s); f' \<in> current_files (LinkHard p oldf f # s)\<rbrakk>
+ − 718
\<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f' = (
+ − 719
if (f' = f)
+ − 720
then (case (parent f) of
+ − 721
Some pf \<Rightarrow> (case (sectxt_of_obj (LinkHard p oldf f # s) (O_file f), sectxt_of_obj s (O_dir pf),
+ − 722
get_parentfs_ctxts s pf) of
+ − 723
(Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ − 724
| _ \<Rightarrow> None)
+ − 725
| None \<Rightarrow> None)
+ − 726
else cf2sfile s f')"
+ − 727
apply (case_tac "f = f'")
+ − 728
apply (simp add:cf2sfile_linkhard4 split:option.splits)
+ − 729
apply (simp add:cf2sfile_linkhard1 current_files_simps)
+ − 730
done
+ − 731
+ − 732
lemma cf2sfile_other:
+ − 733
"\<lbrakk>ff \<in> current_files s;
+ − 734
\<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
+ − 735
\<forall> p fd. e \<noteq> CloseFd p fd;
+ − 736
\<forall> p f. e \<noteq> UnLink p f;
+ − 737
\<forall> p f. e \<noteq> Rmdir p f;
+ − 738
\<forall> p f i. e \<noteq> Mkdir p f i;
+ − 739
\<forall> p f f'. e \<noteq> LinkHard p f f';
+ − 740
valid (e # s)\<rbrakk> \<Longrightarrow> cf2sfile (e # s) ff = cf2sfile s ff"
+ − 741
apply (frule vd_cons, frule vt_grant_os)
+ − 742
apply (induct ff, simp add:sroot_only)
+ − 743
apply (frule parentf_in_current', simp+, case_tac e)
+ − 744
apply (auto simp:current_files_simps is_file_simps is_dir_simps sectxt_of_obj_simps cf2sfile_path
+ − 745
split:if_splits option.splits)
+ − 746
done
+ − 747
+ − 748
lemma cf2sfile_other':
+ − 749
"\<lbrakk>valid (e # s);
+ − 750
\<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
+ − 751
\<forall> p fd. e \<noteq> CloseFd p fd;
+ − 752
\<forall> p f. e \<noteq> UnLink p f;
+ − 753
\<forall> p f. e \<noteq> Rmdir p f;
+ − 754
\<forall> p f i. e \<noteq> Mkdir p f i;
+ − 755
\<forall> p f f'. e \<noteq> LinkHard p f f';
+ − 756
ff \<in> current_files s\<rbrakk> \<Longrightarrow> cf2sfile (e # s) ff = cf2sfile s ff"
+ − 757
by (auto intro!:cf2sfile_other)
+ − 758
+ − 759
lemma cf2sfile_unlink:
+ − 760
"\<lbrakk>valid (UnLink p f # s); f' \<in> current_files (UnLink p f # s)\<rbrakk>
+ − 761
\<Longrightarrow> cf2sfile (UnLink p f # s) f' = cf2sfile s f'"
+ − 762
apply (frule vd_cons, frule vt_grant_os)
+ − 763
apply (simp add:current_files_simps split:if_splits)
+ − 764
apply (auto simp:cf2sfile_def sectxt_of_obj_simps get_parentfs_ctxts_simps is_file_simps is_dir_simps
+ − 765
split:if_splits option.splits)
+ − 766
done
+ − 767
+ − 768
lemma cf2sfile_rmdir:
+ − 769
"\<lbrakk>valid (Rmdir p f # s); f' \<in> current_files (Rmdir p f # s)\<rbrakk>
+ − 770
\<Longrightarrow> cf2sfile (Rmdir p f # s) f' = cf2sfile s f'"
+ − 771
apply (frule vd_cons, frule vt_grant_os)
+ − 772
apply (simp add:current_files_simps split:if_splits)
+ − 773
apply (auto simp:cf2sfile_def sectxt_of_obj_simps get_parentfs_ctxts_simps is_file_simps is_dir_simps
+ − 774
split:if_splits option.splits)
+ − 775
done
+ − 776
+ − 777
lemma pfdof_simp5: "\<lbrakk>proc_fd_of_file s f = {(p, fd)}; file_of_proc_fd s p fd = None\<rbrakk> \<Longrightarrow> False"
+ − 778
apply (subgoal_tac "(p, fd) \<in> proc_fd_of_file s f")
+ − 779
by (simp add:pfdof_simp2, simp)
+ − 780
+ − 781
lemma pfdof_simp6: "proc_fd_of_file s f = {(p, fd)} \<Longrightarrow> file_of_proc_fd s p fd = Some f"
+ − 782
apply (subgoal_tac "(p, fd) \<in> proc_fd_of_file s f")
+ − 783
by (simp add:pfdof_simp2, simp)
+ − 784
+ − 785
lemma cf2sfile_closefd:
+ − 786
"\<lbrakk>valid (CloseFd p fd # s); f \<in> current_files (CloseFd p fd # s)\<rbrakk>
+ − 787
\<Longrightarrow> cf2sfile (CloseFd p fd # s) f = cf2sfile s f"
+ − 788
apply (frule vd_cons, frule vt_grant_os)
+ − 789
apply (simp add:current_files_simps split:if_splits option.splits)
+ − 790
(* costs too much time, but solved *)
+ − 791
(*
+ − 792
apply (auto simp:cf2sfile_def sectxt_of_obj_simps get_parentfs_ctxts_simps is_file_simps is_dir_simps
+ − 793
split:if_splits option.splits
+ − 794
dest:init_file_dir_conflict pfdof_simp5 pfdof_simp6 file_of_pfd_is_file
+ − 795
not_died_init_file not_died_init_dir is_file_not_dir is_dir_not_file
+ − 796
dest!:current_has_sec')
+ − 797
done
+ − 798
*)
+ − 799
sorry
+ − 800
+ − 801
lemmas cf2sfile_simps = cf2sfile_open cf2sfile_mkdir cf2sfile_linkhard cf2sfile_other
+ − 802
cf2sfile_unlink cf2sfile_rmdir cf2sfile_closefd
+ − 803
+ − 804
(*********** cfd2sfd simpset *********)
+ − 805
+ − 806
lemma cfd2sfd_open1:
+ − 807
"valid (Open p f flags fd opt # s)
+ − 808
\<Longrightarrow> cfd2sfd (Open p f flags fd opt # s) p fd =
+ − 809
(case (sectxt_of_obj (Open p f flags fd opt # s) (O_fd p fd), cf2sfile (Open p f flags fd opt # s) f) of
85
+ − 810
(Some sec, Some sf) \<Rightarrow> Some (sec, remove_create_flag flags, sf)
77
+ − 811
| _ \<Rightarrow> None)"
+ − 812
by (simp add:cfd2sfd_def sectxt_of_obj_simps split:if_splits)
+ − 813
+ − 814
lemma cfd2sfd_open_some2:
+ − 815
"\<lbrakk>valid (Open p f flags fd (Some inum) # s); file_of_proc_fd s p' fd' = Some f'\<rbrakk>
+ − 816
\<Longrightarrow> cfd2sfd (Open p f flags fd (Some inum) # s) p' fd' = cfd2sfd s p' fd'"
+ − 817
apply (frule vd_cons, frule vt_grant_os)
+ − 818
apply (frule proc_fd_in_fds, simp)
+ − 819
apply (frule file_of_proc_fd_in_curf, simp)
+ − 820
apply (case_tac "f = f'", simp)
+ − 821
apply (simp add:cfd2sfd_def sectxt_of_obj_simps cf2sfile_open_some1)
+ − 822
apply (case_tac "p = p'", simp)
+ − 823
apply (rule conjI, rule impI, simp)
+ − 824
apply (drule cf2sfile_open_some1, simp)
+ − 825
apply (auto split:option.splits)[1]
+ − 826
apply simp
+ − 827
apply (drule cf2sfile_open_some1, simp)
+ − 828
apply (auto split:option.splits)[1]
+ − 829
done
+ − 830
+ − 831
lemma cfd2sfd_open_none2:
+ − 832
"\<lbrakk>valid (Open p f flags fd None # s); file_of_proc_fd s p' fd' = Some f'\<rbrakk>
+ − 833
\<Longrightarrow> cfd2sfd (Open p f flags fd None # s) p' fd' = cfd2sfd s p' fd'"
+ − 834
apply (frule vd_cons, frule vt_grant_os)
+ − 835
apply (frule proc_fd_in_fds, simp)
+ − 836
apply (frule file_of_proc_fd_in_curf, simp)
+ − 837
apply (simp add:cfd2sfd_def sectxt_of_obj_simps cf2sfile_open_none)
+ − 838
apply (case_tac "p = p'", simp)
+ − 839
apply (rule conjI, rule impI, simp)
+ − 840
apply (drule cf2sfile_open_none)
+ − 841
apply (auto split:option.splits)[1]
+ − 842
apply simp
+ − 843
apply (drule cf2sfile_open_none)
+ − 844
apply (auto split:option.splits)[1]
+ − 845
done
+ − 846
+ − 847
lemma cfd2sfd_open2:
+ − 848
"\<lbrakk>valid (Open p f flags fd opt # s); file_of_proc_fd s p' fd' = Some f'\<rbrakk>
+ − 849
\<Longrightarrow> cfd2sfd (Open p f flags fd opt # s) p' fd' = cfd2sfd s p' fd'"
+ − 850
apply (case_tac opt)
+ − 851
apply (simp add:cfd2sfd_open_none2)
+ − 852
apply (simp add:cfd2sfd_open_some2)
+ − 853
done
+ − 854
+ − 855
lemma cfd2sfd_open:
+ − 856
"\<lbrakk>valid (Open p f flags fd opt # s); file_of_proc_fd (Open p f flags fd opt # s) p' fd' = Some f'\<rbrakk>
+ − 857
\<Longrightarrow> cfd2sfd (Open p f flags fd opt # s) p' fd' = (if (p' = p \<and> fd' = fd) then
+ − 858
(case (sectxt_of_obj (Open p f flags fd opt # s) (O_fd p fd), cf2sfile (Open p f flags fd opt # s) f) of
85
+ − 859
(Some sec, Some sf) \<Rightarrow> Some (sec, remove_create_flag flags, sf)
77
+ − 860
| _ \<Rightarrow> None) else cfd2sfd s p' fd')"
+ − 861
apply (simp split:if_splits)
+ − 862
apply (simp add:cfd2sfd_open1 split:option.splits)
+ − 863
apply (simp add:cfd2sfd_open2)
+ − 864
apply (rule impI, simp)
+ − 865
done
+ − 866
+ − 867
lemma cfd2sfd_closefd:
+ − 868
"\<lbrakk>valid (CloseFd p fd # s); file_of_proc_fd (CloseFd p fd # s) p' fd' = Some f\<rbrakk>
+ − 869
\<Longrightarrow> cfd2sfd (CloseFd p fd # s) p' fd' = cfd2sfd s p' fd'"
+ − 870
apply (frule vd_cons, frule vt_grant_os)
+ − 871
apply (frule proc_fd_in_fds, simp)
+ − 872
apply (frule file_of_proc_fd_in_curf, simp)
+ − 873
apply (frule cf2sfile_closefd, simp)
+ − 874
apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
+ − 875
apply (auto split:option.splits if_splits)
+ − 876
done
+ − 877
+ − 878
lemma cfd2sfd_clone:
+ − 879
"\<lbrakk>valid (Clone p p' fds # s); file_of_proc_fd (Clone p p' fds # s) p'' fd' = Some f\<rbrakk>
+ − 880
\<Longrightarrow> cfd2sfd (Clone p p' fds # s) p'' fd' = (
+ − 881
if (p'' = p') then cfd2sfd s p fd'
+ − 882
else cfd2sfd s p'' fd')"
+ − 883
apply (frule vd_cons, frule vt_grant_os)
+ − 884
apply (frule proc_fd_in_fds, simp)
+ − 885
apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
+ − 886
apply (frule_tac cf2sfile_other', simp+)
+ − 887
apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
+ − 888
apply (case_tac "p'' = p'", simp)
+ − 889
apply (auto split:option.splits if_splits)[1]
+ − 890
apply (simp)
+ − 891
apply (auto split:option.splits if_splits)[1]
+ − 892
done
+ − 893
+ − 894
lemma cfd2sfd_execve:
+ − 895
"\<lbrakk>valid (Execve p f fds # s); file_of_proc_fd (Execve p f fds # s) p' fd' = Some f'\<rbrakk>
+ − 896
\<Longrightarrow> cfd2sfd (Execve p f fds # s) p' fd' = cfd2sfd s p' fd'"
+ − 897
apply (frule vd_cons, frule vt_grant_os)
+ − 898
apply (frule proc_fd_in_fds, simp)
+ − 899
apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
+ − 900
apply (frule_tac cf2sfile_other', simp+)
+ − 901
apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
+ − 902
apply (case_tac "p' = p", simp)
+ − 903
apply (auto split:option.splits if_splits)[1]
+ − 904
apply (simp)
+ − 905
apply (auto split:option.splits if_splits)[1]
+ − 906
done
+ − 907
+ − 908
lemma cfd2sfd_kill:
+ − 909
"\<lbrakk>valid (Kill p p'' # s); file_of_proc_fd (Kill p p'' # s) p' fd' = Some f'\<rbrakk>
+ − 910
\<Longrightarrow> cfd2sfd (Kill p p'' # s) p' fd' = cfd2sfd s p' fd'"
+ − 911
apply (frule vd_cons, frule vt_grant_os)
+ − 912
apply (frule proc_fd_in_fds, simp)
+ − 913
apply (frule proc_fd_in_procs, simp)
+ − 914
apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
+ − 915
apply (frule_tac cf2sfile_other', simp+)
+ − 916
apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
+ − 917
apply (auto split:option.splits if_splits)
+ − 918
done
+ − 919
+ − 920
lemma cfd2sfd_exit:
+ − 921
"\<lbrakk>valid (Exit p # s); file_of_proc_fd (Exit p # s) p' fd' = Some f'\<rbrakk>
+ − 922
\<Longrightarrow> cfd2sfd (Exit p # s) p' fd' = cfd2sfd s p' fd'"
+ − 923
apply (frule vd_cons, frule vt_grant_os)
+ − 924
apply (frule proc_fd_in_fds, simp)
+ − 925
apply (frule proc_fd_in_procs, simp)
+ − 926
apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
+ − 927
apply (frule_tac cf2sfile_other', simp+)
+ − 928
apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
+ − 929
apply (auto split:option.splits if_splits)
+ − 930
done
+ − 931
+ − 932
lemma cfd2sfd_other:
+ − 933
"\<lbrakk>valid (e # s); file_of_proc_fd (e # s) p' fd' = Some f';
+ − 934
\<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ − 935
\<forall> p p'' fds. e \<noteq> Clone p p'' fds\<rbrakk>
+ − 936
\<Longrightarrow> cfd2sfd (e # s) p' fd' = cfd2sfd s p' fd'"
+ − 937
apply (frule vd_cons, frule vt_grant_os)
+ − 938
apply (frule proc_fd_in_fds, simp)
+ − 939
apply (frule proc_fd_in_procs, simp)
+ − 940
apply (frule file_of_proc_fd_in_curf, simp)
+ − 941
apply (case_tac e)
+ − 942
apply (auto intro!:cfd2sfd_execve cfd2sfd_closefd cfd2sfd_kill cfd2sfd_exit)
+ − 943
apply (auto simp:cfd2sfd_def sectxt_of_obj_simps current_files_simps cf2sfile_simps split:option.splits)
+ − 944
apply (auto dest!:current_has_sec' dest:file_of_proc_fd_in_curf proc_fd_in_fds)
+ − 945
done
+ − 946
+ − 947
lemmas cfd2sfd_simps = cfd2sfd_open cfd2sfd_clone cfd2sfd_other
+ − 948
+ − 949
(********** cpfd2sfds simpset **********)
+ − 950
+ − 951
lemma current_filefd_has_flags:
+ − 952
"\<lbrakk>file_of_proc_fd s p fd = Some f; valid s\<rbrakk> \<Longrightarrow> \<exists> flags. flags_of_proc_fd s p fd = Some flags"
+ − 953
apply (induct s arbitrary:p)
+ − 954
apply (simp only:flags_of_proc_fd.simps file_of_proc_fd.simps init_filefd_prop4)
+ − 955
apply (frule vd_cons, frule vt_grant_os, case_tac a)
+ − 956
apply (auto split:if_splits option.splits dest:proc_fd_in_fds)
+ − 957
done
+ − 958
+ − 959
lemma current_filefd_has_flags':
+ − 960
"\<lbrakk>flags_of_proc_fd s p fd = None; valid s\<rbrakk> \<Longrightarrow> file_of_proc_fd s p fd = None"
+ − 961
apply (case_tac "file_of_proc_fd s p fd")
+ − 962
apply (simp, drule current_filefd_has_flags, simp+)
+ − 963
done
+ − 964
+ − 965
lemma current_file_has_sfile':
+ − 966
"\<lbrakk>cf2sfile s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
+ − 967
by (rule notI, drule current_file_has_sfile, simp+)
+ − 968
+ − 969
lemma current_filefd_has_sfd:
+ − 970
"\<lbrakk>file_of_proc_fd s p fd = Some f; valid s\<rbrakk> \<Longrightarrow> \<exists>sfd. cfd2sfd s p fd = Some sfd"
+ − 971
by (auto simp:cfd2sfd_def split:option.splits dest!:current_has_sec' current_file_has_sfile'
+ − 972
dest:file_of_proc_fd_in_curf proc_fd_in_fds current_filefd_has_flags)
+ − 973
+ − 974
lemma current_filefd_has_sfd':
+ − 975
"\<lbrakk>cfd2sfd s p fd = None; valid s\<rbrakk> \<Longrightarrow> file_of_proc_fd s p fd = None"
+ − 976
by (case_tac "file_of_proc_fd s p fd", auto dest:current_filefd_has_sfd)
+ − 977
+ − 978
lemma cpfd2sfds_open1:
+ − 979
"valid (Open p f flags fd opt # s) \<Longrightarrow>
+ − 980
cpfd2sfds (Open p f flags fd opt # s) p = (
+ − 981
case (cfd2sfd (Open p f flags fd opt # s) p fd) of
+ − 982
Some sfd \<Rightarrow> (cpfd2sfds s p) \<union> {sfd}
+ − 983
| _ \<Rightarrow> cpfd2sfds s p)"
+ − 984
apply (frule vd_cons, frule vt_grant_os)
+ − 985
apply (split option.splits)
+ − 986
apply (rule conjI, rule impI, drule current_filefd_has_sfd', simp, simp)
+ − 987
apply (rule allI, rule impI)
+ − 988
apply (rule set_eqI, rule iffI)
+ − 989
apply (case_tac "x = a", simp)
+ − 990
unfolding cpfd2sfds_def
+ − 991
apply (erule CollectE, (erule conjE|erule bexE)+)
+ − 992
apply (simp add:proc_file_fds_def split:if_splits)
+ − 993
apply (erule exE, rule_tac x = fda in exI)
+ − 994
apply (simp add:cfd2sfd_open2)
+ − 995
apply (case_tac "x = a", simp add:proc_file_fds_def)
+ − 996
apply (rule_tac x = fd in exI, simp+)
+ − 997
apply (erule conjE|erule bexE)+
+ − 998
apply (rule_tac x = fda in bexI)
+ − 999
apply (simp add:proc_file_fds_def, erule exE)
+ − 1000
apply (simp add:cfd2sfd_open2)
+ − 1001
apply (simp add:proc_file_fds_def)
+ − 1002
done
+ − 1003
+ − 1004
lemma cpfd2sfds_open1':
+ − 1005
"valid (Open p f flags fd opt # s) \<Longrightarrow>
+ − 1006
cpfd2sfds (Open p f flags fd opt # s) p = (
+ − 1007
case (sectxt_of_obj (Open p f flags fd opt # s) (O_fd p fd), cf2sfile (Open p f flags fd opt # s) f) of
85
+ − 1008
(Some sec, Some sf) \<Rightarrow> (cpfd2sfds s p) \<union> {(sec, remove_create_flag flags, sf)}
77
+ − 1009
| _ \<Rightarrow> cpfd2sfds s p)"
+ − 1010
apply (frule cfd2sfd_open1)
+ − 1011
apply (auto dest:cpfd2sfds_open1 split:option.splits)
+ − 1012
done
+ − 1013
+ − 1014
lemma cpfd2sfds_open2:
+ − 1015
"\<lbrakk>valid (Open p f flags fd opt # s); p' \<noteq> p\<rbrakk> \<Longrightarrow> cpfd2sfds (Open p f flags fd opt # s) p' = cpfd2sfds s p'"
+ − 1016
apply (frule vt_grant_os, frule vd_cons)
+ − 1017
unfolding cpfd2sfds_def
+ − 1018
apply (rule set_eqI, rule iffI)
+ − 1019
apply (simp add:proc_file_fds_def)
+ − 1020
apply (erule exE|erule conjE)+
+ − 1021
apply (simp only:file_of_proc_fd.simps cfd2sfd_open2 split:if_splits)
+ − 1022
apply (rule_tac x = fda in exI, simp)
+ − 1023
apply (simp add:proc_file_fds_def)
+ − 1024
apply (erule exE|erule conjE)+
+ − 1025
apply (rule_tac x = fda in exI, simp add:cfd2sfd_open2)
+ − 1026
done
+ − 1027
+ − 1028
lemma cpfd2sfds_open:
+ − 1029
"valid (Open p f flags fd opt # s)
+ − 1030
\<Longrightarrow> cpfd2sfds (Open p f flags fd opt # s) = (cpfd2sfds s) (p := (
+ − 1031
case (sectxt_of_obj (Open p f flags fd opt # s) (O_fd p fd), cf2sfile (Open p f flags fd opt # s) f) of
85
+ − 1032
(Some sec, Some sf) \<Rightarrow> (cpfd2sfds s p) \<union> {(sec, remove_create_flag flags, sf)}
77
+ − 1033
| _ \<Rightarrow> cpfd2sfds s p))"
+ − 1034
apply (rule ext)
+ − 1035
apply (case_tac "x \<noteq> p")
+ − 1036
apply (simp add:cpfd2sfds_open2)
+ − 1037
apply (simp add:cpfd2sfds_open1')
+ − 1038
done
+ − 1039
+ − 1040
lemma cpfd2sfds_execve:
+ − 1041
"valid (Execve p f fds # s)
+ − 1042
\<Longrightarrow> cpfd2sfds (Execve p f fds # s) = (cpfd2sfds s) (p := {sfd. \<exists> fd \<in> fds. \<exists> f. file_of_proc_fd s p fd = Some f \<and> cfd2sfd s p fd = Some sfd})"
+ − 1043
apply (frule vd_cons, frule vt_grant_os)
+ − 1044
apply (rule ext)
+ − 1045
apply (rule set_eqI, rule iffI)
+ − 1046
unfolding cpfd2sfds_def proc_file_fds_def
+ − 1047
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+ − 1048
apply (simp split:if_splits)
+ − 1049
apply (frule_tac p' = p and fd' = fd in cfd2sfd_other, simp+)
+ − 1050
apply (rule_tac x = fd in bexI, simp+)
+ − 1051
apply (simp add:cpfd2sfds_def proc_file_fds_def)
+ − 1052
apply (frule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+ − 1053
apply (rule_tac x = fd in exI, simp)
+ − 1054
apply (simp split:if_splits)
+ − 1055
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+ − 1056
apply (rule_tac x = fd in exI, simp)
+ − 1057
apply (frule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+ − 1058
apply (simp add:cpfd2sfds_def proc_file_fds_def)
+ − 1059
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+ − 1060
apply (rule_tac x = fd in exI, simp)
+ − 1061
apply (frule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+ − 1062
done
+ − 1063
+ − 1064
lemma cpfd2sfds_clone:
+ − 1065
"valid (Clone p p' fds # s)
+ − 1066
\<Longrightarrow> cpfd2sfds (Clone p p' fds # s) = (cpfd2sfds s) (p' := {sfd. \<exists> fd \<in> fds. \<exists> f. file_of_proc_fd s p fd = Some f \<and> cfd2sfd s p fd = Some sfd})"
+ − 1067
apply (frule vd_cons, frule vt_grant_os)
+ − 1068
apply (rule ext)
+ − 1069
apply (rule set_eqI, rule iffI)
+ − 1070
unfolding cpfd2sfds_def proc_file_fds_def
+ − 1071
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+ − 1072
apply (simp split:if_splits)
+ − 1073
apply (frule_tac p'' = p' and fd' = fd in cfd2sfd_clone, simp+)
+ − 1074
apply (rule_tac x = fd in bexI, simp+)
+ − 1075
apply (simp add:cpfd2sfds_def proc_file_fds_def)
+ − 1076
apply (frule_tac p'' = x and fd' = fd in cfd2sfd_clone, simp+)
+ − 1077
apply (rule_tac x = fd in exI, simp)
+ − 1078
apply (simp split:if_splits)
+ − 1079
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+ − 1080
apply (rule_tac x = fd in exI, simp)
+ − 1081
apply (frule_tac p'' = p' and fd' = fd in cfd2sfd_clone, simp+)
+ − 1082
apply (simp add:cpfd2sfds_def proc_file_fds_def)
+ − 1083
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+ − 1084
apply (rule_tac x = fd in exI, simp)
+ − 1085
apply (frule_tac p'' = x and fd' = fd in cfd2sfd_clone, simp+)
+ − 1086
done
+ − 1087
+ − 1088
lemma cpfd2sfds_other:
+ − 1089
"\<lbrakk>valid (e # s);
+ − 1090
\<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ − 1091
\<forall> p f fds. e \<noteq> Execve p f fds;
+ − 1092
\<forall> p p'. e \<noteq> Kill p p';
+ − 1093
\<forall> p. e \<noteq> Exit p;
+ − 1094
\<forall> p fd. e \<noteq> CloseFd p fd;
+ − 1095
\<forall> p p' fds. e \<noteq> Clone p p' fds\<rbrakk> \<Longrightarrow> cpfd2sfds (e # s) = cpfd2sfds s"
+ − 1096
apply (frule vd_cons, frule vt_grant_os)
+ − 1097
apply (rule ext)
+ − 1098
unfolding cpfd2sfds_def proc_file_fds_def
+ − 1099
apply (case_tac e)
+ − 1100
using cfd2sfd_other
+ − 1101
by auto
+ − 1102
+ − 1103
lemma cpfd2sfds_kill:
+ − 1104
"valid (Kill p p' # s) \<Longrightarrow> cpfd2sfds (Kill p p' # s) = (cpfd2sfds s) (p' := {})"
+ − 1105
apply (frule vd_cons, frule vt_grant_os)
+ − 1106
apply (rule ext, rule set_eqI)
+ − 1107
unfolding cpfd2sfds_def proc_file_fds_def
+ − 1108
apply (rule iffI)
+ − 1109
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+ − 1110
apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
+ − 1111
apply (rule_tac x = fd in exI, simp)
+ − 1112
apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+ − 1113
apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
+ − 1114
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+ − 1115
apply (rule_tac x = fd in exI, simp)
+ − 1116
apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+ − 1117
done
+ − 1118
+ − 1119
lemma cpfd2sfds_exit:
+ − 1120
"valid (Exit p # s) \<Longrightarrow> cpfd2sfds (Exit p # s) = (cpfd2sfds s) (p := {})"
+ − 1121
apply (frule vd_cons, frule vt_grant_os)
+ − 1122
apply (rule ext, rule set_eqI)
+ − 1123
unfolding cpfd2sfds_def proc_file_fds_def
+ − 1124
apply (rule iffI)
+ − 1125
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+ − 1126
apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
+ − 1127
apply (rule_tac x = fd in exI, simp)
+ − 1128
apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+ − 1129
apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
+ − 1130
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+ − 1131
apply (rule_tac x = fd in exI, simp)
+ − 1132
apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
+ − 1133
done
+ − 1134
+ − 1135
lemma cpfd2sfds_closefd:
+ − 1136
"valid (CloseFd p fd # s) \<Longrightarrow> cpfd2sfds (CloseFd p fd # s) = (cpfd2sfds s) (p :=
+ − 1137
if (fd \<in> proc_file_fds s p)
+ − 1138
then (case (cfd2sfd s p fd) of
+ − 1139
Some sfd \<Rightarrow> (if (\<exists> fd' f'. fd' \<noteq> fd \<and> file_of_proc_fd s p fd' = Some f' \<and> cfd2sfd s p fd' = Some sfd)
+ − 1140
then cpfd2sfds s p else cpfd2sfds s p - {sfd})
+ − 1141
| _ \<Rightarrow> cpfd2sfds s p)
+ − 1142
else cpfd2sfds s p)"
+ − 1143
apply (frule vd_cons)
+ − 1144
apply (rule ext, rule set_eqI, rule iffI)
+ − 1145
unfolding cpfd2sfds_def proc_file_fds_def
+ − 1146
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
+ − 1147
apply (simp split:if_splits)
+ − 1148
apply (rule conjI, rule impI, rule conjI, rule impI, erule exE)
+ − 1149
apply (frule_tac p = p and fd = fd in current_filefd_has_sfd, simp)
+ − 1150
apply (erule exE, simp)
+ − 1151
apply (rule conjI, rule impI, (erule exE|erule conjE)+)
+ − 1152
apply (rule_tac x = fda in exI, simp, simp add:cfd2sfd_closefd)
+ − 1153
+ − 1154
apply (rule impI, rule conjI)
+ − 1155
apply (rule_tac x = fda in exI, simp, simp add:cfd2sfd_closefd)
+ − 1156
apply (rule notI, simp)
+ − 1157
apply (erule_tac x = fda in allE, simp add:cfd2sfd_closefd)
+ − 1158
+ − 1159
apply (rule impI, simp add:cpfd2sfds_def proc_file_fds_def)
+ − 1160
apply (erule exE, rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+ − 1161
+ − 1162
apply (rule impI| rule conjI)+
+ − 1163
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+ − 1164
+ − 1165
apply (rule impI, simp add:cpfd2sfds_def proc_file_fds_def)
+ − 1166
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+ − 1167
+ − 1168
apply (simp split:if_splits)
+ − 1169
apply (frule_tac p = p and fd = fd in current_filefd_has_sfd, simp)
+ − 1170
apply (erule exE, simp)
+ − 1171
apply (case_tac "\<exists>fd'. fd' \<noteq> fd \<and> (\<exists>f'. file_of_proc_fd s p fd' = Some f') \<and> cfd2sfd s p fd' = Some sfd")
+ − 1172
apply simp
+ − 1173
apply (case_tac "xa = sfd")
+ − 1174
apply (erule exE|erule conjE)+
+ − 1175
apply (rule_tac x = fd' in exI, simp add:cfd2sfd_closefd)
+ − 1176
apply (erule exE|erule conjE)+
+ − 1177
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+ − 1178
apply (rule notI, simp)
+ − 1179
apply (simp, (erule exE|erule conjE)+)
+ − 1180
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+ − 1181
apply (rule notI, simp)
+ − 1182
apply (erule exE|erule conjE)+
+ − 1183
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+ − 1184
apply (rule notI, simp)
+ − 1185
apply (simp add:cpfd2sfds_def proc_file_fds_def)
+ − 1186
apply (erule exE|erule conjE)+
+ − 1187
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
+ − 1188
done
+ − 1189
+ − 1190
lemmas cpfd2sfds_simps = cpfd2sfds_open cpfd2sfds_execve cpfd2sfds_clone cpfd2sfds_kill cpfd2sfds_exit
+ − 1191
cpfd2sfds_closefd cpfd2sfds_other
+ − 1192
+ − 1193
(********* ch2sshm simpset ********)
+ − 1194
(*
+ − 1195
lemma ch2sshm_createshm:
+ − 1196
"valid (CreateShM p h # s)
+ − 1197
\<Longrightarrow> ch2sshm (CreateShM p h # s) = (ch2sshm s) (h :=
+ − 1198
(case (sectxt_of_obj (CreateShM p h # s) (O_shm h)) of
+ − 1199
Some sec \<Rightarrow>
+ − 1200
Some (if (\<not> died (O_shm h) s \<and> h \<in> init_shms) then Init h else Created, sec)
+ − 1201
| _ \<Rightarrow> None))"
+ − 1202
apply (frule vd_cons, frule vt_grant_os)
+ − 1203
apply (auto simp:ch2sshm_def sectxt_of_obj_simps dest!:current_has_sec' split:option.splits if_splits)
+ − 1204
done
+ − 1205
+ − 1206
lemma ch2sshm_other:
+ − 1207
"\<lbrakk>valid (e # s);
+ − 1208
\<forall> p h. e \<noteq> CreateShM p h;
+ − 1209
h' \<in> current_shms (e # s)\<rbrakk> \<Longrightarrow> ch2sshm (e # s) h' = ch2sshm s h'"
+ − 1210
apply (frule vd_cons, frule vt_grant_os)
+ − 1211
apply (case_tac e)
+ − 1212
apply (auto simp:ch2sshm_def sectxt_of_obj_simps dest!:current_has_sec' split:option.splits if_splits)
+ − 1213
done
+ − 1214
+ − 1215
lemmas ch2sshm_simps = ch2sshm_createshm ch2sshm_other
+ − 1216
+ − 1217
lemma current_shm_has_sh:
+ − 1218
"\<lbrakk>h \<in> current_shms s; valid s\<rbrakk> \<Longrightarrow> \<exists> sh. ch2sshm s h = Some sh"
+ − 1219
by (auto simp:ch2sshm_def split:option.splits dest!:current_has_sec')
+ − 1220
+ − 1221
lemma current_shm_has_sh':
+ − 1222
"\<lbrakk>ch2sshm s h = None; valid s\<rbrakk> \<Longrightarrow> h \<notin> current_shms s"
+ − 1223
by (auto dest:current_shm_has_sh)
+ − 1224
+ − 1225
(********** cph2spshs simpset **********)
+ − 1226
+ − 1227
lemma cph2spshs_attach:
+ − 1228
"valid (Attach p h flag # s) \<Longrightarrow>
+ − 1229
cph2spshs (Attach p h flag # s) = (cph2spshs s) (p :=
+ − 1230
(case (ch2sshm s h) of
+ − 1231
Some sh \<Rightarrow> cph2spshs s p \<union> {(sh, flag)}
+ − 1232
| _ \<Rightarrow> cph2spshs s p) )"
+ − 1233
apply (frule vd_cons, frule vt_grant_os, rule ext)
+ − 1234
using ch2sshm_other[where e = "Attach p h flag" and s = s]
+ − 1235
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+ − 1236
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 simp:cph2spshs_def)
+ − 1237
apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
+ − 1238
apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
+ − 1239
apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
+ − 1240
apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
+ − 1241
apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
+ − 1242
apply (case_tac "ha = h", simp, frule procs_of_shm_prop1, simp)
+ − 1243
apply (rule_tac x = ha in exI, simp)
+ − 1244
apply (rule_tac x = ha in exI, simp, drule procs_of_shm_prop1, simp, simp)
+ − 1245
apply (rule_tac x = ha in exI, simp)
+ − 1246
apply (frule procs_of_shm_prop1, simp, simp)
+ − 1247
apply (rule impI, simp)
+ − 1248
done
+ − 1249
+ − 1250
lemma cph2spshs_detach: "valid (Detach p h # s) \<Longrightarrow>
+ − 1251
cph2spshs (Detach p h # s) = (cph2spshs s) (p :=
+ − 1252
(case (ch2sshm s h, flag_of_proc_shm s p h) of
+ − 1253
(Some sh, Some flag) \<Rightarrow> if (\<exists> h'. h' \<noteq> h \<and> (p,flag) \<in> procs_of_shm s h' \<and> ch2sshm s h' = Some sh)
+ − 1254
then cph2spshs s p else cph2spshs s p - {(sh, flag)}
+ − 1255
| _ \<Rightarrow> cph2spshs s p) )"
+ − 1256
apply (frule vd_cons, frule vt_grant_os, rule ext)
+ − 1257
apply (case_tac "x = p") defer
+ − 1258
using ch2sshm_other[where e = "Detach p h" and s = s]
+ − 1259
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+ − 1260
dest!:current_shm_has_sh' procs_of_shm_prop4' dest:procs_of_shm_prop1 procs_of_shm_prop3 simp:cph2spshs_def) [1]
+ − 1261
apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp, simp)
+ − 1262
apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp, simp)
+ − 1263
apply (rule impI, simp)
+ − 1264
+ − 1265
apply (clarsimp)
+ − 1266
apply (frule current_shm_has_sh, simp, erule exE)
+ − 1267
apply (frule procs_of_shm_prop4, simp, simp)
+ − 1268
apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
+ − 1269
+ − 1270
apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
+ − 1271
apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
+ − 1272
apply (case_tac "ha = h", simp)
+ − 1273
apply (rule_tac x = sha in exI, rule_tac x = flaga in exI, rule_tac x = ha in exI, simp)
+ − 1274
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1275
+ − 1276
apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
+ − 1277
apply (case_tac "ha = h", simp)
+ − 1278
apply (rule_tac x = h' in exI, simp)
+ − 1279
apply (frule_tac flag = flag and flag' = flaga in procs_of_shm_prop3, simp+)
+ − 1280
apply (frule_tac flag = flaga in procs_of_shm_prop4, simp+)
+ − 1281
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1282
apply (frule_tac h = h' in procs_of_shm_prop1, simp, simp)
+ − 1283
apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+ − 1284
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1285
+ − 1286
apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
+ − 1287
apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
+ − 1288
apply (erule CollectE | erule exE| erule conjE| rule DiffI |rule CollectI)+
+ − 1289
apply (simp split:if_splits)
+ − 1290
apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+ − 1291
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1292
apply (rule notI, simp split:if_splits)
+ − 1293
apply (erule_tac x = ha in allE, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+ − 1294
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1295
apply (erule CollectE | erule exE| erule conjE| erule DiffE |rule CollectI)+
+ − 1296
apply (simp split:if_splits)
+ − 1297
apply (erule_tac x = ha in allE, simp, rule_tac x = ha in exI, simp)
+ − 1298
apply (case_tac "ha = h", simp add:procs_of_shm_prop3, frule_tac h = ha in procs_of_shm_prop1, simp+)
+ − 1299
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1300
done
+ − 1301
+ − 1302
lemma cph2spshs_deleteshm:
+ − 1303
"valid (DeleteShM p h # s) \<Longrightarrow>
+ − 1304
cph2spshs (DeleteShM p h # s) = (\<lambda> p'.
+ − 1305
(case (ch2sshm s h, flag_of_proc_shm s p' h) of
+ − 1306
(Some sh, Some flag) \<Rightarrow> if (\<exists> h'. h' \<noteq> h \<and> (p', flag) \<in> procs_of_shm s h' \<and> ch2sshm s h' = Some sh)
+ − 1307
then cph2spshs s p' else cph2spshs s p' - {(sh, flag)}
+ − 1308
| _ \<Rightarrow> cph2spshs s p') )"
+ − 1309
apply (frule vd_cons, frule vt_grant_os, rule ext)
+ − 1310
+ − 1311
apply (clarsimp)
+ − 1312
apply (frule current_shm_has_sh, simp, erule exE, simp, simp split:option.splits)
+ − 1313
apply (rule conjI, rule impI)
+ − 1314
using ch2sshm_other[where e = "DeleteShM p h" and s = s]
+ − 1315
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+ − 1316
dest!:current_shm_has_sh' procs_of_shm_prop4' dest:procs_of_shm_prop1 procs_of_shm_prop3 simp:cph2spshs_def) [1]
+ − 1317
apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp+)
+ − 1318
apply (rule_tac x = ha in exI, simp)
+ − 1319
apply (case_tac "ha = h", simp+, frule_tac h = ha in procs_of_shm_prop1, simp+)
+ − 1320
+ − 1321
apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
+ − 1322
+ − 1323
apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
+ − 1324
apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
+ − 1325
apply (case_tac "ha = h", simp)
+ − 1326
apply (rule_tac x = sha in exI, rule_tac x = flag in exI, rule_tac x = ha in exI, simp)
+ − 1327
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1328
+ − 1329
apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
+ − 1330
apply (case_tac "ha = h", simp)
+ − 1331
apply (rule_tac x = h' in exI, simp)
+ − 1332
apply (frule_tac flag = flag in procs_of_shm_prop4, simp+)
+ − 1333
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1334
apply (frule_tac h = h' in procs_of_shm_prop1, simp, simp)
+ − 1335
apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+ − 1336
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1337
+ − 1338
apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
+ − 1339
apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
+ − 1340
apply (erule CollectE | erule exE| erule conjE| rule DiffI |rule CollectI)+
+ − 1341
apply (simp split:if_splits)
+ − 1342
apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+ − 1343
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1344
apply (rule notI, simp split:if_splits)
+ − 1345
apply (erule_tac x = ha in allE, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
+ − 1346
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1347
apply (erule CollectE | erule exE| erule conjE| erule DiffE |rule CollectI)+
+ − 1348
apply (simp split:if_splits)
+ − 1349
apply (erule_tac x = ha in allE, simp, rule_tac x = ha in exI, simp)
+ − 1350
apply (case_tac "ha = h", simp add:procs_of_shm_prop4)
+ − 1351
apply (frule_tac h = ha in procs_of_shm_prop1, simp+)
+ − 1352
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
+ − 1353
done
+ − 1354
+ − 1355
lemma flag_of_proc_shm_prop1:
+ − 1356
"\<lbrakk>flag_of_proc_shm s p h = Some flag; valid s\<rbrakk> \<Longrightarrow> (p, flag) \<in> procs_of_shm s h"
+ − 1357
apply (induct s arbitrary:p)
+ − 1358
apply (simp add:init_shmflag_has_proc)
+ − 1359
apply (frule vt_grant_os, frule vd_cons, case_tac a, auto split:if_splits option.splits)
+ − 1360
done
+ − 1361
+ − 1362
lemma cph2spshs_clone:
+ − 1363
"valid (Clone p p' fds shms # s) \<Longrightarrow>
+ − 1364
cph2spshs (Clone p p' fds shms # s) = (cph2spshs s) (p' :=
+ − 1365
{(sh, flag) | h sh flag. h \<in> shms \<and> ch2sshm s h = Some sh \<and> flag_of_proc_shm s p h = Some flag} )"
+ − 1366
apply (frule vd_cons, frule vt_grant_os, rule ext)
+ − 1367
using ch2sshm_other[where e = "Clone p p' fds shms" and s = s]
+ − 1368
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+ − 1369
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
+ − 1370
apply (erule_tac x = h in allE, frule procs_of_shm_prop1, simp, simp add:procs_of_shm_prop4)
+ − 1371
apply (rule_tac x = h in exI, simp, frule flag_of_proc_shm_prop1, simp+, simp add:procs_of_shm_prop1)
+ − 1372
apply (rule_tac x = h in exI, simp, frule procs_of_shm_prop1, simp+)+
+ − 1373
done
+ − 1374
+ − 1375
lemma cph2spshs_execve:
+ − 1376
"valid (Execve p f fds # s) \<Longrightarrow>
+ − 1377
cph2spshs (Execve p f fds # s) = (cph2spshs s) (p := {})"
+ − 1378
apply (frule vd_cons, frule vt_grant_os, rule ext)
+ − 1379
using ch2sshm_other[where e = "Execve p f fds" and s = s]
+ − 1380
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+ − 1381
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
+ − 1382
apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+
+ − 1383
done
+ − 1384
+ − 1385
lemma cph2spshs_kill:
+ − 1386
"valid (Kill p p' # s) \<Longrightarrow> cph2spshs (Kill p p' # s) = (cph2spshs s) (p' := {})"
+ − 1387
apply (frule vd_cons, frule vt_grant_os, rule ext)
+ − 1388
using ch2sshm_other[where e = "Kill p p'" and s = s]
+ − 1389
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+ − 1390
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
+ − 1391
apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+
+ − 1392
done
+ − 1393
+ − 1394
lemma cph2spshs_exit:
+ − 1395
"valid (Exit p # s) \<Longrightarrow> cph2spshs (Exit p # s) = (cph2spshs s) (p := {})"
+ − 1396
apply (frule vd_cons, frule vt_grant_os, rule ext)
+ − 1397
using ch2sshm_other[where e = "Exit p" and s = s]
+ − 1398
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+ − 1399
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
+ − 1400
apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+
+ − 1401
done
+ − 1402
+ − 1403
lemma cph2spshs_createshm:
+ − 1404
"valid (CreateShM p h # s) \<Longrightarrow> cph2spshs (CreateShM p h # s) = cph2spshs s"
+ − 1405
apply (frule vt_grant_os, frule vd_cons, rule ext)
+ − 1406
apply (auto simp:cph2spshs_def)
+ − 1407
using ch2sshm_createshm
+ − 1408
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+ − 1409
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
+ − 1410
apply (rule_tac x = ha in exI, auto simp:procs_of_shm_prop1)
+ − 1411
done
+ − 1412
+ − 1413
lemma cph2spshs_other:
+ − 1414
"\<lbrakk>valid (e # s);
+ − 1415
\<forall> p h flag. e \<noteq> Attach p h flag;
+ − 1416
\<forall> p h. e \<noteq> Detach p h;
+ − 1417
\<forall> p h. e \<noteq> DeleteShM p h;
+ − 1418
\<forall> p p' fds shms. e \<noteq> Clone p p' fds shms;
+ − 1419
\<forall> p f fds. e \<noteq> Execve p f fds;
+ − 1420
\<forall> p p'. e \<noteq> Kill p p';
+ − 1421
\<forall> p. e \<noteq> Exit p\<rbrakk> \<Longrightarrow> cph2spshs (e # s) = cph2spshs s"
+ − 1422
apply (frule vt_grant_os, frule vd_cons, rule ext)
+ − 1423
unfolding cph2spshs_def
+ − 1424
apply (rule set_eqI, rule iffI)
+ − 1425
apply (erule CollectE | erule conjE| erule exE| rule conjI| rule impI| rule CollectI)+
+ − 1426
apply (frule procs_of_shm_prop1, simp, rule_tac x= sh in exI, rule_tac x = flag in exI, rule_tac x = h in exI)
+ − 1427
apply (case_tac e)
+ − 1428
using ch2sshm_other[where e = e and s = s]
+ − 1429
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+ − 1430
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3
+ − 1431
simp:ch2sshm_createshm)
+ − 1432
apply (rule_tac x = h in exI, case_tac e)
+ − 1433
using ch2sshm_other[where e = e and s = s]
+ − 1434
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
+ − 1435
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3
+ − 1436
simp:ch2sshm_createshm)
+ − 1437
done
+ − 1438
+ − 1439
lemmas cph2spshs_simps = cph2spshs_attach cph2spshs_detach cph2spshs_deleteshm cph2spshs_clone cph2spshs_execve
+ − 1440
cph2spshs_exit cph2spshs_kill cph2spshs_other
+ − 1441
*)
+ − 1442
(******** cp2sproc simpset *********)
+ − 1443
+ − 1444
lemma cp2sproc_clone:
+ − 1445
"valid (Clone p p' fds # s) \<Longrightarrow> cp2sproc (Clone p p' fds # s) = (cp2sproc s) (p' :=
+ − 1446
case (sectxt_of_obj s (O_proc p)) of
+ − 1447
Some sec \<Rightarrow> Some (Created, sec,
+ − 1448
{sfd. \<exists> fd \<in> fds. \<exists> f. file_of_proc_fd s p fd = Some f \<and> cfd2sfd s p fd = Some sfd})
+ − 1449
| _ \<Rightarrow> None)" (* ,
+ − 1450
{(sh, flag) | h sh flag. h \<in> shms \<and> ch2sshm s h = Some sh \<and> flag_of_proc_shm s p h = Some flag} *)
+ − 1451
apply (frule cpfd2sfds_clone) (*
+ − 1452
apply (frule cph2spshs_clone, frule cpfd2sfds_clone) *)
+ − 1453
apply (frule vd_cons, frule vt_grant_os, simp only:os_grant.simps)
+ − 1454
apply ((erule exE| erule conjE)+, frule not_init_intro_proc, simp, rule ext, case_tac "x = p'", simp)
+ − 1455
apply (auto simp:cp2sproc_def sectxt_of_obj_simps dest!:current_has_sec'
+ − 1456
dest:current_proc_has_sec split:option.splits if_splits)
+ − 1457
done
+ − 1458
+ − 1459
lemma cp2sproc_other:
+ − 1460
"\<lbrakk>valid (e # s);
+ − 1461
\<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ − 1462
\<forall> p fd. e \<noteq> CloseFd p fd;
+ − 1463
\<forall> p p' fds. e \<noteq> Clone p p' fds;
+ − 1464
\<forall> p f fds. e \<noteq> Execve p f fds;
+ − 1465
\<forall> p p'. e \<noteq> Kill p p';
+ − 1466
\<forall> p. e \<noteq> Exit p\<rbrakk> \<Longrightarrow> cp2sproc (e # s) = cp2sproc s" (*
+ − 1467
\<forall> p h flag. e \<noteq> Attach p h flag;
+ − 1468
\<forall> p h. e \<noteq> Detach p h;
+ − 1469
\<forall> p h. e \<noteq> DeleteShM p h;*)
+ − 1470
apply (frule cpfd2sfds_other, simp+)(*
+ − 1471
apply (frule cph2spshs_other, simp+, frule cpfd2sfds_other, simp+)*)
+ − 1472
apply (frule vt_grant_os, frule vd_cons, rule ext, case_tac e, simp_all)
+ − 1473
apply (auto simp:cp2sproc_def sectxt_of_obj_simps dest!:current_has_sec'
+ − 1474
dest:current_proc_has_sec not_died_init_proc split:option.splits if_splits)
+ − 1475
done
+ − 1476
+ − 1477
lemma cp2sproc_open:
+ − 1478
"valid (Open p f flags fd opt # s) \<Longrightarrow>
+ − 1479
cp2sproc (Open p f flags fd opt # s) = (cp2sproc s) (p :=
+ − 1480
case (sectxt_of_obj s (O_proc p), sectxt_of_obj (Open p f flags fd opt # s) (O_fd p fd),
+ − 1481
cf2sfile (Open p f flags fd opt # s) f) of
+ − 1482
(Some sec, Some fdsec, Some sf) \<Rightarrow> Some (if (\<not> died (O_proc p) s \<and> p \<in> init_procs)
+ − 1483
then Init p else Created, sec,
85
+ − 1484
(cpfd2sfds s p) \<union> {(fdsec, remove_create_flag flags, sf)})
77
+ − 1485
| _ \<Rightarrow> None)" (*, cph2spshs s p
+ − 1486
apply (frule cph2spshs_other, simp, simp, simp, simp, simp, simp, simp) *)
+ − 1487
apply (frule cpfd2sfds_open, frule vt_grant_os, frule vd_cons, rule ext)
+ − 1488
apply (case_tac "x = p")
+ − 1489
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ − 1490
dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ − 1491
dest:current_proc_has_sec not_died_init_proc split:option.splits if_splits)
+ − 1492
done
+ − 1493
+ − 1494
lemma cp2sproc_closefd:
+ − 1495
"valid (CloseFd p fd # s) \<Longrightarrow>
+ − 1496
cp2sproc (CloseFd p fd # s) = (cp2sproc s) (p :=
+ − 1497
if (fd \<in> proc_file_fds s p)
+ − 1498
then (case (cfd2sfd s p fd) of
+ − 1499
Some sfd \<Rightarrow> (if (\<exists> fd' f'. fd' \<noteq> fd \<and> file_of_proc_fd s p fd' = Some f' \<and> cfd2sfd s p fd' = Some sfd)
+ − 1500
then cp2sproc s p
+ − 1501
else (case (sectxt_of_obj s (O_proc p)) of
+ − 1502
Some sec \<Rightarrow> Some (if (\<not> died (O_proc p) s \<and> p \<in> init_procs)
+ − 1503
then Init p else Created,
+ − 1504
sec, cpfd2sfds s p - {sfd})
+ − 1505
| _ \<Rightarrow> None))
+ − 1506
| _ \<Rightarrow> cp2sproc s p)
+ − 1507
else cp2sproc s p)"(*, cph2spshs s p *)
+ − 1508
apply (frule cpfd2sfds_closefd) (*
+ − 1509
apply (frule cpfd2sfds_closefd, frule cph2spshs_other, simp, simp, simp, simp, simp, simp, simp) *)
+ − 1510
apply (frule vt_grant_os, frule vd_cons, rule ext)
+ − 1511
apply (case_tac "x = p")
+ − 1512
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ − 1513
dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ − 1514
dest:current_proc_has_sec not_died_init_proc split:option.splits if_splits)
+ − 1515
done
+ − 1516
+ − 1517
(*
+ − 1518
lemma cp2sproc_attach:
+ − 1519
"valid (Attach p h flag # s) \<Longrightarrow>
+ − 1520
cp2sproc (Attach p h flag # s) = (cp2sproc s) (p :=
+ − 1521
(case (ch2sshm s h) of
+ − 1522
Some sh \<Rightarrow> (case (sectxt_of_obj s (O_proc p)) of
+ − 1523
Some sec \<Rightarrow> Some (if (\<not> died (O_proc p) s \<and> p \<in> init_procs)
+ − 1524
then Init p else Created,
+ − 1525
sec, cpfd2sfds s p, cph2spshs s p \<union> {(sh, flag)})
+ − 1526
| _ \<Rightarrow> None)
+ − 1527
| _ \<Rightarrow> cp2sproc s p) )"
+ − 1528
apply (frule cph2spshs_attach, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp)
+ − 1529
apply (frule vt_grant_os, frule vd_cons, rule ext)
+ − 1530
apply (case_tac "x = p")
+ − 1531
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ − 1532
dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ − 1533
dest:current_proc_has_sec not_died_init_proc split:option.splits if_splits)
+ − 1534
done
+ − 1535
+ − 1536
lemma cp2sproc_detach:
+ − 1537
"valid (Detach p h # s) \<Longrightarrow>
+ − 1538
cp2sproc (Detach p h # s) = (cp2sproc s) (p :=
+ − 1539
(case (ch2sshm s h, flag_of_proc_shm s p h) of
+ − 1540
(Some sh, Some flag) \<Rightarrow> if (\<exists> h'. h' \<noteq> h \<and> (p,flag) \<in> procs_of_shm s h' \<and> ch2sshm s h' = Some sh)
+ − 1541
then cp2sproc s p
+ − 1542
else (case (sectxt_of_obj s (O_proc p)) of
+ − 1543
Some sec \<Rightarrow> Some (if (\<not> died (O_proc p) s \<and> p \<in> init_procs)
+ − 1544
then Init p else Created, sec,
+ − 1545
cpfd2sfds s p, cph2spshs s p - {(sh, flag)})
+ − 1546
| None \<Rightarrow> None)
+ − 1547
| _ \<Rightarrow> cp2sproc s p) )"
+ − 1548
apply (frule cph2spshs_detach, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp)
+ − 1549
apply (frule vt_grant_os, frule vd_cons, rule ext)
+ − 1550
apply (case_tac "x = p")
+ − 1551
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ − 1552
dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ − 1553
dest:current_proc_has_sec not_died_init_proc split:option.splits if_splits)
+ − 1554
done
+ − 1555
+ − 1556
lemma cp2sproc_deleteshm:
+ − 1557
"valid (DeleteShM p h # s) \<Longrightarrow>
+ − 1558
cp2sproc (DeleteShM p h # s) = (\<lambda> p'.
+ − 1559
(case (ch2sshm s h, flag_of_proc_shm s p' h) of
+ − 1560
(Some sh, Some flag) \<Rightarrow> if (\<exists> h'. h' \<noteq> h \<and> (p', flag) \<in> procs_of_shm s h' \<and> ch2sshm s h' = Some sh)
+ − 1561
then cp2sproc s p'
+ − 1562
else (case (sectxt_of_obj s (O_proc p')) of
+ − 1563
Some sec \<Rightarrow> Some (if (\<not> died (O_proc p') s \<and> p' \<in> init_procs)
+ − 1564
then Init p' else Created, sec,
+ − 1565
cpfd2sfds s p', cph2spshs s p' - {(sh, flag)})
+ − 1566
| None \<Rightarrow> None)
+ − 1567
| _ \<Rightarrow> cp2sproc s p') )"
+ − 1568
apply (frule cph2spshs_deleteshm, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp)
+ − 1569
apply (frule vt_grant_os, frule vd_cons, rule ext)
+ − 1570
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ − 1571
dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ − 1572
dest:current_proc_has_sec not_died_init_proc split:option.splits if_splits)
+ − 1573
done
+ − 1574
*)
+ − 1575
+ − 1576
lemma cp2sproc_execve:
+ − 1577
"valid (Execve p f fds # s) \<Longrightarrow>
+ − 1578
cp2sproc (Execve p f fds # s) = (cp2sproc s) (p :=
+ − 1579
(case (sectxt_of_obj (Execve p f fds # s) (O_proc p)) of
+ − 1580
Some sec \<Rightarrow> Some (if (\<not> died (O_proc p) s \<and> p \<in> init_procs) then Init p else Created, sec,
+ − 1581
{sfd. \<exists> fd \<in> fds. \<exists> f. file_of_proc_fd s p fd = Some f \<and> cfd2sfd s p fd = Some sfd})
+ − 1582
| _ \<Rightarrow> None) )" (*, {}
+ − 1583
apply (frule cph2spshs_execve, frule cpfd2sfds_execve) *)
+ − 1584
apply (frule cpfd2sfds_execve)
+ − 1585
apply (frule vt_grant_os, frule vd_cons, rule ext)
+ − 1586
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ − 1587
dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ − 1588
dest:current_proc_has_sec not_died_init_proc split:option.splits if_splits)
+ − 1589
done
+ − 1590
+ − 1591
lemma cp2sproc_kill:
+ − 1592
"\<lbrakk>valid (Kill p p' # s); p'' \<noteq> p'\<rbrakk> \<Longrightarrow>
+ − 1593
cp2sproc (Kill p p' # s) p'' = (cp2sproc s) p''" (*
+ − 1594
apply (frule cph2spshs_kill, frule cpfd2sfds_kill) *)
+ − 1595
apply (frule cpfd2sfds_kill)
+ − 1596
apply (frule vt_grant_os, frule vd_cons)
+ − 1597
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ − 1598
dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ − 1599
dest:current_proc_has_sec not_died_init_proc split:option.splits if_splits)
+ − 1600
done
+ − 1601
+ − 1602
lemma cp2sproc_kill':
+ − 1603
"\<lbrakk>valid (Kill p p' # s); p'' \<in> current_procs (Kill p p' # s)\<rbrakk> \<Longrightarrow>
+ − 1604
cp2sproc (Kill p p' # s) p'' = (cp2sproc s) p''"
+ − 1605
by (drule_tac p'' = p'' in cp2sproc_kill, simp+)
+ − 1606
+ − 1607
lemma cp2sproc_exit:
+ − 1608
"\<lbrakk>valid (Exit p # s); p' \<noteq> p\<rbrakk> \<Longrightarrow>
+ − 1609
cp2sproc (Exit p # s) p' = (cp2sproc s) p'" (*
+ − 1610
apply (frule cph2spshs_exit, frule cpfd2sfds_exit) *)
+ − 1611
apply (frule cpfd2sfds_exit)
+ − 1612
apply (frule vt_grant_os, frule vd_cons)
+ − 1613
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps
+ − 1614
dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
+ − 1615
dest:current_proc_has_sec not_died_init_proc split:option.splits if_splits)
+ − 1616
done
+ − 1617
+ − 1618
lemma cp2sproc_exit':
+ − 1619
"\<lbrakk>valid (Exit p # s); p' \<in> current_procs (Exit p # s)\<rbrakk> \<Longrightarrow>
+ − 1620
cp2sproc (Exit p # s) p' = (cp2sproc s) p'"
+ − 1621
by (drule_tac p'= p' in cp2sproc_exit, simp+)
+ − 1622
+ − 1623
lemmas cp2sproc_simps = cp2sproc_open cp2sproc_closefd (* cp2sproc_attach cp2sproc_detach cp2sproc_deleteshm *)
+ − 1624
cp2sproc_clone cp2sproc_execve cp2sproc_kill cp2sproc_exit cp2sproc_other
+ − 1625
+ − 1626
lemma current_proc_has_sp:
+ − 1627
"\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sp. cp2sproc s p = Some sp"
+ − 1628
by (auto simp:cp2sproc_def split:option.splits dest!:current_has_sec')
+ − 1629
+ − 1630
lemma current_proc_has_sp':
+ − 1631
"\<lbrakk>cp2sproc s p = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
+ − 1632
by (auto dest:current_proc_has_sp)
+ − 1633
+ − 1634
(* simpset for cf2sfiles *)
+ − 1635
+ − 1636
lemma cf2sfiles_open:
+ − 1637
"\<lbrakk>valid (Open p f flag fd opt # s); f' \<in> current_files (Open p f flag fd opt # s)\<rbrakk>
+ − 1638
\<Longrightarrow> cf2sfiles (Open p f flag fd opt # s) f' = (
+ − 1639
if (f' = f \<and> opt \<noteq> None)
+ − 1640
then (case cf2sfile (Open p f flag fd opt # s) f of
+ − 1641
Some sf \<Rightarrow> {sf}
+ − 1642
| _ \<Rightarrow> {} )
+ − 1643
else cf2sfiles s f')"
+ − 1644
apply (frule vt_grant_os, frule vd_cons)
+ − 1645
apply (auto simp:cf2sfiles_def cf2sfile_open_none cf2sfile_simps same_inode_files_open
+ − 1646
split:if_splits option.splits dest!:current_file_has_sfile' dest:cf2sfile_open)
+ − 1647
apply (rule_tac x = "f'a" in bexI, drule same_inode_files_prop1, simp add:cf2sfile_open_some1, simp)+
+ − 1648
done
+ − 1649
+ − 1650
lemma cf2sfiles_other:
+ − 1651
"\<lbrakk>valid (e # s);
+ − 1652
\<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
+ − 1653
\<forall> p fd. e \<noteq> CloseFd p fd;
+ − 1654
\<forall> p f. e \<noteq> UnLink p f;
+ − 1655
\<forall> p f f'. e \<noteq> LinkHard p f f'\<rbrakk> \<Longrightarrow> cf2sfiles (e # s) = cf2sfiles s"
+ − 1656
apply (frule vt_grant_os, frule vd_cons, rule ext)
+ − 1657
apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI)
+ − 1658
apply (drule Set.CollectD, erule bexE, rule CollectI)
+ − 1659
apply (rule_tac x = f' in bexI, case_tac e)
+ − 1660
apply (auto simp:cf2sfiles_def cf2sfile_simps same_inode_files_simps current_files_simps
+ − 1661
split:if_splits option.splits dest!:current_file_has_sfile' dest:same_inode_files_prop1 cf2sfile_other')
+ − 1662
apply (drule_tac f' = f' in cf2sfile_rmdir)
+ − 1663
apply (simp add:current_files_simps same_inode_files_prop1 same_inode_files_prop3 dir_is_empty_def)+
+ − 1664
+ − 1665
apply (rule_tac x = f' in bexI, case_tac e)
+ − 1666
apply (auto simp:cf2sfiles_def cf2sfile_simps same_inode_files_simps current_files_simps
+ − 1667
split:if_splits option.splits dest!:current_file_has_sfile' dest:same_inode_files_prop1 cf2sfile_other')
+ − 1668
apply (drule_tac f' = f' in cf2sfile_rmdir)
+ − 1669
apply (simp add:current_files_simps same_inode_files_prop1 same_inode_files_prop3 dir_is_empty_def)+
+ − 1670
done
+ − 1671
+ − 1672
lemma cf2sfile_linkhard1':
+ − 1673
"\<lbrakk>valid (LinkHard p oldf f # s); f' \<in> same_inode_files s f''\<rbrakk>
+ − 1674
\<Longrightarrow> cf2sfile (LinkHard p oldf f# s) f' = cf2sfile s f'"
+ − 1675
apply (drule same_inode_files_prop1)
+ − 1676
by (simp add:cf2sfile_linkhard1)
+ − 1677
+ − 1678
lemma cf2sfiles_linkhard:
+ − 1679
"valid (LinkHard p oldf f # s) \<Longrightarrow> cf2sfiles (LinkHard p oldf f # s) = (\<lambda> f'.
+ − 1680
if (f' = f \<or> f' \<in> same_inode_files s oldf)
+ − 1681
then (case (cf2sfile (LinkHard p oldf f # s) f) of
+ − 1682
Some sf \<Rightarrow> cf2sfiles s oldf \<union> {sf}
+ − 1683
| _ \<Rightarrow> {})
+ − 1684
else cf2sfiles s f')"
+ − 1685
apply (frule vt_grant_os, frule vd_cons, rule ext)
+ − 1686
apply (auto simp:cf2sfiles_def cf2sfile_linkhard1' same_inode_files_linkhard current_files_linkhard
+ − 1687
split:if_splits option.splits dest!:current_file_has_sfile' current_has_sec' dest:same_inode_files_prop1)
+ − 1688
done
+ − 1689
+ − 1690
lemma cf2sfile_unlink':
+ − 1691
"\<lbrakk>valid (UnLink p f # s); f' \<in> same_inode_files (UnLink p f # s) f''\<rbrakk>
+ − 1692
\<Longrightarrow> cf2sfile (UnLink p f # s) f' = cf2sfile s f'"
+ − 1693
apply (drule same_inode_files_prop1)
+ − 1694
by (simp add:cf2sfile_unlink)
+ − 1695
+ − 1696
lemma cf2sfiles_unlink:
+ − 1697
"\<lbrakk>valid (UnLink p f # s); f' \<in> current_files (UnLink p f # s)\<rbrakk> \<Longrightarrow> cf2sfiles (UnLink p f # s) f' = (
+ − 1698
if (f' \<in> same_inode_files s f \<and> proc_fd_of_file s f = {} \<and>
+ − 1699
(\<forall> f'' \<in> same_inode_files s f. f'' \<noteq> f \<longrightarrow> cf2sfile s f'' \<noteq> cf2sfile s f)) then
+ − 1700
(case (cf2sfile s f) of
+ − 1701
Some sf \<Rightarrow> cf2sfiles s f' - {sf}
+ − 1702
| _ \<Rightarrow> {})
+ − 1703
else cf2sfiles s f')"
+ − 1704
apply (frule vt_grant_os, frule vd_cons, simp add:current_files_simps split:if_splits)
+ − 1705
apply (rule conjI, clarify, frule is_file_has_sfile', simp, erule exE, simp)
+ − 1706
apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+ − 1707
apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink)
+ − 1708
apply (simp add:current_files_unlink, simp, erule conjE)
+ − 1709
apply (erule_tac x = f'a in ballE, frule_tac f' = "f'a" in cf2sfile_unlink)
+ − 1710
apply (simp add:current_files_unlink same_inode_files_prop1, simp)
+ − 1711
apply (rule_tac x = f'a in bexI, simp, simp)
+ − 1712
apply (drule_tac f = f and f' = f' and f'' = f'a in same_inode_files_prop4, simp+)
+ − 1713
apply (erule conjE|erule exE|erule bexE)+
+ − 1714
apply (case_tac "f'a = f", simp)
+ − 1715
apply (frule_tac f' = f' in same_inode_files_unlink, simp add:current_files_unlink)
+ − 1716
apply (frule_tac f' = f'a in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1)
+ − 1717
apply (rule_tac x = f'a in bexI, simp, simp)
+ − 1718
+ − 1719
apply (rule impI)+
+ − 1720
apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+ − 1721
apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink)
+ − 1722
apply (simp add:current_files_unlink, simp, (erule conjE)+)
+ − 1723
apply (rule_tac x = f'a in bexI, frule_tac f' = "f'a" in cf2sfile_unlink)
+ − 1724
apply (simp add:current_files_unlink same_inode_files_prop1, simp, simp)
+ − 1725
apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_unlink)
+ − 1726
apply (simp add:current_files_unlink, simp)
+ − 1727
apply (case_tac "f'a = f", simp)
+ − 1728
apply (frule_tac f = f' and f' = f in same_inode_files_prop5, simp)
+ − 1729
apply (erule bexE, erule conjE)
+ − 1730
apply (rule_tac x = f'' in bexI)
+ − 1731
apply (drule_tac f' = f'' in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1)
+ − 1732
apply (simp, simp, erule same_inode_files_prop4, simp)
+ − 1733
apply (rule_tac x = f'a in bexI)
+ − 1734
apply (drule_tac f' = f'a in cf2sfile_unlink, simp add:current_files_unlink same_inode_files_prop1)
+ − 1735
apply (simp, simp)
+ − 1736
+ − 1737
+ − 1738
apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+ − 1739
apply (erule bexE, frule_tac f' = f' in same_inode_files_unlink)
+ − 1740
apply (simp add:current_files_unlink, simp)
+ − 1741
apply (rule_tac x = f'a in bexI)
+ − 1742
apply (frule_tac f' = f'a in cf2sfile_unlink)
+ − 1743
apply (simp add:same_inode_files_prop1 current_files_unlink, simp, simp)
+ − 1744
apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_unlink)
+ − 1745
apply (simp add:current_files_unlink, simp)
+ − 1746
apply (rule_tac x = f'a in bexI)
+ − 1747
apply (frule_tac f' = f'a in cf2sfile_unlink)
+ − 1748
apply (simp add:same_inode_files_prop1 current_files_unlink, simp, simp)
+ − 1749
done
+ − 1750
+ − 1751
lemma cf2sfiles_closefd:
+ − 1752
"\<lbrakk>valid (CloseFd p fd # s); f' \<in> current_files (CloseFd p fd # s)\<rbrakk> \<Longrightarrow> cf2sfiles (CloseFd p fd # s) f' = (
+ − 1753
case (file_of_proc_fd s p fd) of
+ − 1754
Some f \<Rightarrow> if (f' \<in> same_inode_files s f \<and> proc_fd_of_file s f = {(p, fd)} \<and> f \<in> files_hung_by_del s \<and>
+ − 1755
(\<forall> f'' \<in> same_inode_files s f. f'' \<noteq> f \<longrightarrow> cf2sfile s f'' \<noteq> cf2sfile s f))
+ − 1756
then (case (cf2sfile s f) of
+ − 1757
Some sf \<Rightarrow> cf2sfiles s f' - {sf}
+ − 1758
| _ \<Rightarrow> {})
+ − 1759
else cf2sfiles s f'
+ − 1760
| _ \<Rightarrow> cf2sfiles s f')"
+ − 1761
+ − 1762
apply (frule vt_grant_os, frule vd_cons, case_tac "file_of_proc_fd s p fd")
+ − 1763
apply (simp_all add:current_files_simps split:if_splits)
+ − 1764
+ − 1765
apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+ − 1766
apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+ − 1767
apply (simp add:current_files_closefd, simp)
+ − 1768
apply (rule_tac x = f'a in bexI)
+ − 1769
apply (frule_tac f = f'a in cf2sfile_closefd)
+ − 1770
apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+ − 1771
apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+ − 1772
apply (simp add:current_files_closefd, simp)
+ − 1773
apply (rule_tac x = f'a in bexI)
+ − 1774
apply (frule_tac f = f'a in cf2sfile_closefd)
+ − 1775
apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+ − 1776
+ − 1777
apply (rule conjI, clarify, frule file_of_pfd_is_file, simp)
+ − 1778
apply (frule is_file_has_sfile', simp, erule exE, simp)
+ − 1779
apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+ − 1780
apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+ − 1781
apply (simp add:current_files_closefd, simp, erule conjE)
+ − 1782
apply (erule_tac x = f'a in ballE, frule_tac f = "f'a" in cf2sfile_closefd)
+ − 1783
apply (simp add:current_files_closefd same_inode_files_prop1, simp)
+ − 1784
apply (rule_tac x = f'a in bexI, simp, simp)
+ − 1785
apply (drule_tac f = a and f' = f' and f'' = f'a in same_inode_files_prop4, simp+)
+ − 1786
apply (erule conjE|erule exE|erule bexE)+
+ − 1787
apply (case_tac "f'a = a", simp)
+ − 1788
apply (frule_tac f' = f' in same_inode_files_closefd, simp add:current_files_closefd, simp)
+ − 1789
apply (frule_tac f = f'a in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1)
+ − 1790
apply (rule_tac x = f'a in bexI, simp, simp)
+ − 1791
+ − 1792
apply (rule impI)+
+ − 1793
apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+ − 1794
apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+ − 1795
apply (simp add:current_files_closefd, simp, (erule conjE)+)
+ − 1796
apply (rule_tac x = f'a in bexI, frule_tac f = f'a in cf2sfile_closefd)
+ − 1797
apply (simp add:current_files_closefd same_inode_files_prop1, simp, simp)
+ − 1798
apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+ − 1799
apply (simp add:current_files_closefd, simp)
+ − 1800
apply (case_tac "f'a = a", simp)
+ − 1801
apply (frule_tac f = f' and f' = a in same_inode_files_prop5, simp)
+ − 1802
apply (erule bexE, erule conjE)
+ − 1803
apply (rule_tac x = f'' in bexI)
+ − 1804
apply (drule_tac f = f'' in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1)
+ − 1805
apply (simp, simp, erule same_inode_files_prop4, simp)
+ − 1806
apply (rule_tac x = f'a in bexI)
+ − 1807
apply (drule_tac f = f'a in cf2sfile_closefd, simp add:current_files_closefd same_inode_files_prop1)
+ − 1808
apply (simp, simp)
+ − 1809
+ − 1810
apply (rule conjI, clarify)
+ − 1811
+ − 1812
apply (rule impI)
+ − 1813
apply (case_tac "a \<in> files_hung_by_del s", simp_all)
+ − 1814
apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+ − 1815
apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+ − 1816
apply (simp add:current_files_closefd, simp)
+ − 1817
apply (rule_tac x = f'a in bexI)
+ − 1818
apply (frule_tac f = f'a in cf2sfile_closefd)
+ − 1819
apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+ − 1820
apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+ − 1821
apply (simp add:current_files_closefd, simp)
+ − 1822
apply (rule_tac x = f'a in bexI)
+ − 1823
apply (frule_tac f = f'a in cf2sfile_closefd)
+ − 1824
apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+ − 1825
apply (simp add:cf2sfiles_def, rule set_eqI, rule iffI, drule CollectD)
+ − 1826
apply (erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+ − 1827
apply (simp add:current_files_closefd, simp)
+ − 1828
apply (rule_tac x = f'a in bexI)
+ − 1829
apply (frule_tac f = f'a in cf2sfile_closefd)
+ − 1830
apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+ − 1831
apply (drule CollectD, erule bexE, frule_tac f' = f' in same_inode_files_closefd)
+ − 1832
apply (simp add:current_files_closefd, simp)
+ − 1833
apply (rule_tac x = f'a in bexI)
+ − 1834
apply (frule_tac f = f'a in cf2sfile_closefd)
+ − 1835
apply (simp add:same_inode_files_prop1 current_files_closefd, simp, simp)
+ − 1836
done
+ − 1837
+ − 1838
lemmas cf2sfiles_simps = cf2sfiles_open cf2sfiles_linkhard cf2sfiles_other
+ − 1839
cf2sfiles_unlink cf2sfiles_closefd
+ − 1840
+ − 1841
+ − 1842
(* simpset for co2sobj *)
+ − 1843
+ − 1844
lemma co2sobj_execve:
92
+ − 1845
"\<lbrakk>valid (Execve p f fds # s); dalive (Execve p f fds # s) obj\<rbrakk> \<Longrightarrow> co2sobj (Execve p f fds # s) obj = (
+ − 1846
if (obj = D_proc p)
77
+ − 1847
then (case (cp2sproc (Execve p f fds # s) p) of
+ − 1848
Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> tainted s \<or> O_file f \<in> tainted s))
+ − 1849
| _ \<Rightarrow> None)
+ − 1850
else co2sobj s obj )"
+ − 1851
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 1852
apply (simp_all add:current_files_simps cf2sfiles_other (* ch2sshm_other *) cq2smsgq_other)
+ − 1853
apply (case_tac "cp2sproc (Execve p f fds # s) p")
+ − 1854
apply (drule current_proc_has_sp', simp, simp)
+ − 1855
apply (simp (no_asm_simp) add:cp2sproc_execve split:option.splits)
+ − 1856
apply (simp add:is_dir_simps, frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp)
+ − 1857
apply (frule_tac ff = list in cf2sfile_other', simp_all)
+ − 1858
apply (simp add:is_dir_in_current)
+ − 1859
done
+ − 1860
+ − 1861
lemma co2sobj_execve':
92
+ − 1862
"\<lbrakk>valid (Execve p f fds # s); dalive s obj\<rbrakk> \<Longrightarrow> co2sobj (Execve p f fds # s) obj = (
+ − 1863
if (obj = D_proc p)
77
+ − 1864
then (case (cp2sproc (Execve p f fds # s) p) of
+ − 1865
Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> tainted s \<or> O_file f \<in> tainted s))
+ − 1866
| _ \<Rightarrow> None)
+ − 1867
else co2sobj s obj )"
+ − 1868
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 1869
apply (simp_all add:current_files_simps cf2sfiles_other (* ch2sshm_other *) cq2smsgq_other)
+ − 1870
apply (case_tac "cp2sproc (Execve p f fds # s) p")
+ − 1871
apply (drule current_proc_has_sp', simp, simp)
+ − 1872
apply (simp (no_asm_simp) add:cp2sproc_execve split:option.splits)
+ − 1873
apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp)
+ − 1874
apply (frule_tac ff = list in cf2sfile_other', simp_all)
+ − 1875
apply (simp add:is_dir_in_current)
+ − 1876
done
+ − 1877
+ − 1878
lemma co2sobj_clone':
92
+ − 1879
"\<lbrakk>valid (Clone p p' fds # s); dalive s obj\<rbrakk> \<Longrightarrow> co2sobj (Clone p p' fds # s) obj = (
+ − 1880
if (obj = D_proc p')
77
+ − 1881
then (case (cp2sproc (Clone p p' fds # s) p') of
+ − 1882
Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> tainted s))
+ − 1883
| _ \<Rightarrow> None)
+ − 1884
else co2sobj s obj )"
+ − 1885
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 1886
apply (simp_all add:current_files_simps cf2sfiles_other (* ch2sshm_other *) cq2smsgq_other)
+ − 1887
apply (case_tac "cp2sproc (Clone p p' fds # s) p'")
+ − 1888
apply (drule current_proc_has_sp', simp, simp)
+ − 1889
apply ((erule conjE)+, frule_tac p = p in current_proc_has_sec, simp, erule exE, simp)
+ − 1890
apply (simp (no_asm_simp) add:cp2sproc_clone split:option.splits)
+ − 1891
apply (case_tac "nat = p'", simp, simp)
+ − 1892
apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp)
+ − 1893
apply (frule_tac ff = list in cf2sfile_other', simp_all)
+ − 1894
apply (simp add:is_dir_in_current)
+ − 1895
done
+ − 1896
+ − 1897
lemma co2sobj_clone:
92
+ − 1898
"\<lbrakk>valid (Clone p p' fds # s); dalive (Clone p p' fds # s) obj\<rbrakk>
77
+ − 1899
\<Longrightarrow> co2sobj (Clone p p' fds # s) obj = (
92
+ − 1900
if (obj = D_proc p')
77
+ − 1901
then (case (cp2sproc (Clone p p' fds # s) p') of
+ − 1902
Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> tainted s))
+ − 1903
| _ \<Rightarrow> None)
+ − 1904
else co2sobj s obj )"
+ − 1905
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 1906
apply (simp_all add:current_files_simps cf2sfiles_other (* ch2sshm_other *) cq2smsgq_other)
+ − 1907
apply (rule conjI, rule impI, simp)
+ − 1908
apply (case_tac "cp2sproc (Clone p p' fds # s) p'")
+ − 1909
apply (drule current_proc_has_sp', simp, simp)
+ − 1910
apply ((erule conjE)+, frule_tac p = p in current_proc_has_sec, simp, erule exE, simp)
+ − 1911
apply (rule impI,rule notI, drule tainted_in_current, simp+)
+ − 1912
apply (rule impI, simp)
+ − 1913
apply (drule current_proc_has_sp, simp, (erule exE|erule conjE)+)
+ − 1914
apply (simp (no_asm_simp) add:cp2sproc_clone tainted_in_current split:option.splits)
+ − 1915
+ − 1916
apply (simp add:is_dir_simps, frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp)
+ − 1917
apply (frule_tac ff = list in cf2sfile_other', simp_all)
+ − 1918
apply (simp add:is_dir_in_current)
+ − 1919
done
+ − 1920
+ − 1921
(*
+ − 1922
lemma co2sobj_ptrace:
92
+ − 1923
"\<lbrakk>valid (Ptrace p p' # s); dalive s obj\<rbrakk>\<Longrightarrow> co2sobj (Ptrace p p' # s) obj = (
77
+ − 1924
case obj of
+ − 1925
O_proc p'' \<Rightarrow> if (info_flow_shm s p' p'')
+ − 1926
then (case (cp2sproc s p'') of
+ − 1927
Some sp \<Rightarrow> Some (S_proc sp (O_proc p'' \<in> tainted s \<or> O_proc p \<in> tainted s))
+ − 1928
| _ \<Rightarrow> None)
+ − 1929
else if (info_flow_shm s p p'')
+ − 1930
then (case (cp2sproc s p'') of
+ − 1931
Some sp \<Rightarrow> Some (S_proc sp (O_proc p'' \<in> tainted s \<or> O_proc p' \<in> tainted s))
+ − 1932
| _ \<Rightarrow> None)
+ − 1933
else co2sobj s (O_proc p'')
+ − 1934
| _ \<Rightarrow> co2sobj s obj)"
+ − 1935
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 1936
apply (simp_all add:current_files_simps ch2sshm_other cq2smsgq_other cf2sfiles_other)
+ − 1937
+ − 1938
apply (auto simp:cp2sproc_other split:option.splits
+ − 1939
dest!:current_proc_has_sec' current_proc_has_sp' intro:info_flow_shm_tainted)[1]
+ − 1940
+ − 1941
apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp)
+ − 1942
apply (frule_tac ff = list in cf2sfile_other', simp_all)
+ − 1943
apply (simp add:is_dir_in_current)
+ − 1944
done
+ − 1945
*)
+ − 1946
+ − 1947
lemma co2sobj_ptrace:
92
+ − 1948
"\<lbrakk>valid (Ptrace p p' # s); dalive s obj\<rbrakk>\<Longrightarrow> co2sobj (Ptrace p p' # s) obj = (
77
+ − 1949
case obj of
92
+ − 1950
D_proc p'' \<Rightarrow> if (p'' = p')
77
+ − 1951
then (case (cp2sproc s p'') of
+ − 1952
Some sp \<Rightarrow> Some (S_proc sp (O_proc p'' \<in> tainted s \<or> O_proc p \<in> tainted s))
+ − 1953
| _ \<Rightarrow> None)
+ − 1954
else if (p'' = p)
+ − 1955
then (case (cp2sproc s p'') of
+ − 1956
Some sp \<Rightarrow> Some (S_proc sp (O_proc p'' \<in> tainted s \<or> O_proc p' \<in> tainted s))
+ − 1957
| _ \<Rightarrow> None)
92
+ − 1958
else co2sobj s (D_proc p'')
77
+ − 1959
| _ \<Rightarrow> co2sobj s obj)"
+ − 1960
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 1961
apply (simp_all add:current_files_simps cq2smsgq_other cf2sfiles_other)
+ − 1962
+ − 1963
apply (auto simp:cp2sproc_other split:option.splits
+ − 1964
dest!:current_proc_has_sec' current_proc_has_sp')[1]
+ − 1965
+ − 1966
apply (frule_tac s = s in is_dir_has_sdir', simp, erule exE, simp)
+ − 1967
apply (frule_tac ff = list in cf2sfile_other', simp_all)
+ − 1968
apply (simp add:is_dir_in_current)
+ − 1969
done
+ − 1970
+ − 1971
+ − 1972
lemma co2sobj_open:
92
+ − 1973
"\<lbrakk>valid (Open p f flag fd opt # s); dalive (Open p f flag fd opt # s) obj\<rbrakk>
77
+ − 1974
\<Longrightarrow> co2sobj (Open p f flag fd opt # s) obj = (case obj of
92
+ − 1975
D_file f' \<Rightarrow> if (f' = f \<and> opt \<noteq> None)
77
+ − 1976
then (case (cf2sfile (Open p f flag fd opt # s) f) of
+ − 1977
Some sf \<Rightarrow> Some (S_file {sf} (O_proc p \<in> tainted s))
+ − 1978
| _ \<Rightarrow> None)
92
+ − 1979
else co2sobj s (D_file f')
+ − 1980
| D_proc p' \<Rightarrow> if (p' = p)
77
+ − 1981
then (case (cp2sproc (Open p f flag fd opt # s) p) of
+ − 1982
Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> tainted s))
+ − 1983
| _ \<Rightarrow> None)
92
+ − 1984
else co2sobj s (D_proc p')
77
+ − 1985
| _ \<Rightarrow> co2sobj s obj )"
+ − 1986
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 1987
apply (auto simp:cp2sproc_simps split:option.splits
+ − 1988
dest!:current_proc_has_sp' (* intro:info_flow_shm_tainted *))[1]
+ − 1989
+ − 1990
apply (simp split:if_splits t_object.splits)
+ − 1991
apply ((rule conjI | rule impI| erule conjE| erule exE)+)
+ − 1992
apply ( simp, (erule exE|erule conjE)+)
+ − 1993
apply (case_tac "cf2sfile (Open p f flag fd (Some y) # s) f")
+ − 1994
apply (drule current_file_has_sfile', simp, simp add:current_files_simps, simp)
+ − 1995
apply (frule_tac f' = f in cf2sfiles_open, simp add:current_files_simps, simp)
+ − 1996
apply (frule_tac f' = list in cf2sfiles_open, simp add:is_file_in_current)
+ − 1997
apply (rule impI)
+ − 1998
apply (case_tac "list = f", simp, simp split:option.splits)
+ − 1999
apply (case_tac "cf2sfile (Open p f flag fd opt # s) f")
+ − 2000
apply (drule current_file_has_sfile', simp)
+ − 2001
apply (simp add:current_files_simps is_file_in_current split:option.splits)
+ − 2002
apply (rule impI, simp add:cf2sfiles_open is_file_in_current split:option.splits)
+ − 2003
apply (rule impI, rule conjI)
+ − 2004
apply (drule_tac f' = f in cf2sfiles_open)
+ − 2005
apply (simp add:current_files_simps, simp)
+ − 2006
apply (rule notI, drule tainted_in_current, simp, simp add:is_file_in_current)
+ − 2007
+ − 2008
apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other) (* ch2sshm_other *)
+ − 2009
+ − 2010
apply (frule is_dir_in_current)
+ − 2011
apply (frule_tac f' = list in cf2sfile_open)
+ − 2012
apply (simp add:current_files_simps split:option.splits)
+ − 2013
apply (simp split:if_splits option.splits)
+ − 2014
done
+ − 2015
+ − 2016
(*
+ − 2017
lemma co2sobj_readfile:
92
+ − 2018
"\<lbrakk>valid (ReadFile p fd # s); dalive s obj\<rbrakk> \<Longrightarrow> co2sobj (ReadFile p fd # s) obj = (
77
+ − 2019
case obj of
+ − 2020
O_proc p' \<Rightarrow> (case (file_of_proc_fd s p fd) of
+ − 2021
Some f \<Rightarrow> (if (info_flow_shm s p p' \<and> O_file f \<in> tainted s)
+ − 2022
then (case (cp2sproc s p') of
+ − 2023
Some sp \<Rightarrow> Some (S_proc sp True)
+ − 2024
| _ \<Rightarrow> None)
+ − 2025
else co2sobj s obj)
+ − 2026
| _ \<Rightarrow> None)
+ − 2027
| _ \<Rightarrow> co2sobj s obj)"
+ − 2028
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2029
apply (auto simp:cp2sproc_simps split:option.splits
+ − 2030
dest!:current_proc_has_sp' intro:info_flow_shm_tainted)[1]
+ − 2031
apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other)
+ − 2032
+ − 2033
apply (auto split:if_splits option.splits dest!:current_file_has_sfile'
+ − 2034
simp:current_files_simps cf2sfiles_simps cf2sfile_simps
+ − 2035
dest:is_file_in_current is_dir_in_current)
+ − 2036
done
+ − 2037
*)
+ − 2038
lemma co2sobj_readfile:
92
+ − 2039
"\<lbrakk>valid (ReadFile p fd # s); dalive s obj\<rbrakk> \<Longrightarrow> co2sobj (ReadFile p fd # s) obj = (
77
+ − 2040
case obj of
92
+ − 2041
D_proc p' \<Rightarrow> (case (file_of_proc_fd s p fd) of
77
+ − 2042
Some f \<Rightarrow> (if (p' = p \<and> O_file f \<in> tainted s)
+ − 2043
then (case (cp2sproc s p') of
+ − 2044
Some sp \<Rightarrow> Some (S_proc sp True)
+ − 2045
| _ \<Rightarrow> None)
+ − 2046
else co2sobj s obj)
+ − 2047
| _ \<Rightarrow> None)
+ − 2048
| _ \<Rightarrow> co2sobj s obj)"
+ − 2049
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2050
apply (auto simp:cp2sproc_simps split:option.splits dest!:current_proc_has_sp')[1]
+ − 2051
apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other)
+ − 2052
+ − 2053
apply (auto split:if_splits option.splits dest!:current_file_has_sfile'
+ − 2054
simp:current_files_simps cf2sfile_simps cf2sfiles_simps
+ − 2055
dest:is_file_in_current is_dir_in_current)
+ − 2056
done
+ − 2057
+ − 2058
lemma co2sobj_writefile:
92
+ − 2059
"\<lbrakk>valid (WriteFile p fd # s); dalive s obj\<rbrakk> \<Longrightarrow> co2sobj (WriteFile p fd # s) obj = (
77
+ − 2060
case obj of
92
+ − 2061
D_file f' \<Rightarrow> (case (file_of_proc_fd s p fd) of
77
+ − 2062
Some f \<Rightarrow> (if (f \<in> same_inode_files s f')
+ − 2063
then Some (S_file (cf2sfiles s f')
+ − 2064
(O_file f' \<in> tainted s \<or> O_proc p \<in> tainted s))
+ − 2065
else co2sobj s obj)
+ − 2066
| _ \<Rightarrow> None)
+ − 2067
| _ \<Rightarrow> co2sobj s obj)"
+ − 2068
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2069
apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other) (* ch2sshm_other*)
+ − 2070
+ − 2071
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2072
simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps
+ − 2073
same_inode_files_prop6
+ − 2074
dest:is_file_in_current is_dir_in_current)
+ − 2075
+ − 2076
(* should delete has_same_inode !?!?*)
+ − 2077
by (auto simp:same_inode_files_def is_file_def has_same_inode_def split:if_splits option.splits)
+ − 2078
+ − 2079
lemma co2sobj_closefd:
92
+ − 2080
"\<lbrakk>valid (CloseFd p fd # s); dalive (CloseFd p fd # s) obj\<rbrakk> \<Longrightarrow> co2sobj (CloseFd p fd # s) obj = (
77
+ − 2081
case obj of
92
+ − 2082
D_file f' \<Rightarrow> (case (file_of_proc_fd s p fd) of
77
+ − 2083
Some f \<Rightarrow> (if (f' \<in> same_inode_files s f \<and> proc_fd_of_file s f = {(p, fd)} \<and>
+ − 2084
f \<in> files_hung_by_del s \<and> (\<forall> f'' \<in> same_inode_files s f.
+ − 2085
f'' \<noteq> f \<longrightarrow> cf2sfile s f'' \<noteq> cf2sfile s f))
92
+ − 2086
then (case (cf2sfile s f, co2sobj s (D_file f')) of
77
+ − 2087
(Some sf, Some (S_file sfs b)) \<Rightarrow> Some (S_file (sfs - {sf}) b)
+ − 2088
| _ \<Rightarrow> None)
+ − 2089
else co2sobj s obj)
+ − 2090
| _ \<Rightarrow> co2sobj s obj)
92
+ − 2091
| D_proc p' \<Rightarrow> (if (p = p')
77
+ − 2092
then (case (cp2sproc (CloseFd p fd # s) p) of
+ − 2093
Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> tainted s))
+ − 2094
| _ \<Rightarrow> None)
+ − 2095
else co2sobj s obj)
+ − 2096
| _ \<Rightarrow> co2sobj s obj) "
+ − 2097
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2098
apply (simp_all add:current_files_simps cq2smsgq_other) (* ch2sshm_other *)
+ − 2099
apply (auto simp:cp2sproc_simps cf2sfiles_simps split:option.splits if_splits
+ − 2100
dest!:current_proc_has_sp') [1](* intro:info_flow_shm_tainted)[1] *)
+ − 2101
+ − 2102
apply (frule is_file_in_current)
+ − 2103
apply (case_tac "file_of_proc_fd s p fd")
+ − 2104
apply (simp)
+ − 2105
apply (drule_tac f' = list in cf2sfiles_closefd, simp add:current_files_closefd, simp)
+ − 2106
apply (frule_tac f' = list in cf2sfiles_closefd, simp)
+ − 2107
apply (simp add:is_file_simps current_files_simps)
+ − 2108
apply (auto simp add: cf2sfile_closefd split:if_splits option.splits
+ − 2109
dest!:current_file_has_sfile' dest:hung_file_in_current)[1]
+ − 2110
+ − 2111
apply (simp add:is_dir_simps, frule is_dir_in_current)
+ − 2112
apply (frule_tac f = list in cf2sfile_closefd)
+ − 2113
apply (clarsimp simp:current_files_closefd split:option.splits)
+ − 2114
apply (drule file_of_pfd_is_file', simp+)
+ − 2115
done
+ − 2116
+ − 2117
lemma co2sobj_unlink:
92
+ − 2118
"\<lbrakk>valid (UnLink p f # s); dalive (UnLink p f # s) obj\<rbrakk> \<Longrightarrow> co2sobj (UnLink p f # s) obj = (
77
+ − 2119
case obj of
92
+ − 2120
D_file f' \<Rightarrow> if (f' \<in> same_inode_files s f \<and> proc_fd_of_file s f = {} \<and>
77
+ − 2121
(\<forall> f'' \<in> same_inode_files s f. f'' \<noteq> f \<longrightarrow> cf2sfile s f'' \<noteq> cf2sfile s f))
92
+ − 2122
then (case (cf2sfile s f, co2sobj s (D_file f')) of
77
+ − 2123
(Some sf, Some (S_file sfs b)) \<Rightarrow> Some (S_file (sfs - {sf}) b)
+ − 2124
| _ \<Rightarrow> None)
+ − 2125
else co2sobj s obj
+ − 2126
| _ \<Rightarrow> co2sobj s obj)"
+ − 2127
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2128
apply (simp_all add:current_files_simps cq2smsgq_other) (* ch2sshm_other*)
+ − 2129
apply (auto simp:cp2sproc_simps split:option.splits if_splits
+ − 2130
dest!:current_proc_has_sp')[1] (* intro:info_flow_shm_tainted) *)
+ − 2131
apply (frule is_file_in_current)
+ − 2132
apply (frule_tac f' = list in cf2sfile_unlink, simp)
+ − 2133
apply (frule_tac f' = list in cf2sfiles_unlink, simp)
+ − 2134
apply (simp add:is_file_simps current_files_simps)
+ − 2135
apply (auto simp add: is_file_in_current split:if_splits option.splits
+ − 2136
dest!:current_file_has_sfile')[1]
+ − 2137
+ − 2138
apply (simp add:is_dir_simps, frule is_dir_in_current)
+ − 2139
apply (frule_tac f' = list in cf2sfile_unlink)
+ − 2140
apply (clarsimp simp:current_files_unlink split:option.splits)
+ − 2141
apply (drule file_dir_conflict, simp+)
+ − 2142
done
+ − 2143
+ − 2144
lemma co2sobj_rmdir:
92
+ − 2145
"\<lbrakk>valid (Rmdir p f # s); dalive (Rmdir p f # s) obj\<rbrakk> \<Longrightarrow> co2sobj (Rmdir p f # s) obj = co2sobj s obj"
77
+ − 2146
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2147
apply (simp_all add:current_files_simps cq2smsgq_other) (* ch2sshm_other*)
+ − 2148
apply (auto simp:cp2sproc_simps split:option.splits if_splits
+ − 2149
dest!:current_proc_has_sp')[1] (* intro:info_flow_shm_tainted*)
+ − 2150
apply (simp add:is_file_simps dir_is_empty_def)
+ − 2151
apply (case_tac "f = list", drule file_dir_conflict, simp, simp)
+ − 2152
apply (simp add:cf2sfiles_other)
+ − 2153
apply (auto simp:cf2sfile_simps dest:is_dir_in_current)
+ − 2154
done
+ − 2155
+ − 2156
lemma co2sobj_mkdir:
92
+ − 2157
"\<lbrakk>valid (Mkdir p f i # s); dalive (Mkdir p f i # s) obj\<rbrakk> \<Longrightarrow> co2sobj (Mkdir p f i # s) obj = (
+ − 2158
if (obj = D_dir f)
77
+ − 2159
then (case (cf2sfile (Mkdir p f i # s) f) of
+ − 2160
Some sf \<Rightarrow> Some (S_dir sf)
+ − 2161
| _ \<Rightarrow> None)
+ − 2162
else co2sobj s obj)"
+ − 2163
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2164
apply (simp_all add:current_files_simps cq2smsgq_other) (* ch2sshm_other *)
+ − 2165
apply (auto simp:cp2sproc_simps split:option.splits if_splits
+ − 2166
dest!:current_proc_has_sp')[1] (* intro:info_flow_shm_tainted *)
+ − 2167
apply (frule_tac cf2sfiles_other, simp+)
+ − 2168
apply (frule is_dir_in_current, rule impI, simp add:current_files_mkdir)
+ − 2169
apply (frule current_file_has_sfile, simp, erule exE, simp)
+ − 2170
apply (drule cf2sfile_mkdir1, simp+)
+ − 2171
done
+ − 2172
+ − 2173
lemma co2sobj_linkhard:
92
+ − 2174
"\<lbrakk>valid (LinkHard p oldf f # s); dalive (LinkHard p oldf f # s) obj\<rbrakk>
77
+ − 2175
\<Longrightarrow> co2sobj (LinkHard p oldf f # s) obj = (
+ − 2176
case obj of
92
+ − 2177
D_file f' \<Rightarrow> if (f' = f \<or> f' \<in> same_inode_files s oldf)
77
+ − 2178
then (case (cf2sfile (LinkHard p oldf f # s) f) of
+ − 2179
Some sf \<Rightarrow> Some (S_file (cf2sfiles s oldf \<union> {sf}) (O_file oldf \<in> tainted s))
+ − 2180
| _ \<Rightarrow> None)
+ − 2181
else co2sobj s obj
+ − 2182
| _ \<Rightarrow> co2sobj s obj)"
+ − 2183
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2184
apply (simp_all add:current_files_simps cq2smsgq_other) (* ch2sshm_other *)
+ − 2185
apply (auto simp:cp2sproc_simps split:option.splits if_splits
+ − 2186
dest!:current_proc_has_sp')[1] (* intro:info_flow_shm_tainted *)
+ − 2187
apply (frule_tac cf2sfiles_linkhard, simp+)
+ − 2188
apply (frule_tac f' = f in cf2sfile_linkhard, simp add:current_files_linkhard)
+ − 2189
apply (auto simp:is_file_simps sectxt_of_obj_simps current_files_simps is_file_in_current same_inodes_tainted
+ − 2190
split:option.splits if_splits dest:tainted_in_current
+ − 2191
dest!:current_has_sec' current_file_has_sfile')[1]
+ − 2192
+ − 2193
apply (frule is_dir_in_current, simp add:current_files_linkhard is_dir_simps is_dir_in_current)
+ − 2194
apply (frule is_dir_in_current)
+ − 2195
apply (frule current_file_has_sfile, simp)
+ − 2196
apply (drule cf2sfile_linkhard1, simp+)
+ − 2197
done
+ − 2198
+ − 2199
lemma co2sobj_truncate:
92
+ − 2200
"\<lbrakk>valid (Truncate p f len # s); dalive s obj\<rbrakk> \<Longrightarrow> co2sobj (Truncate p f len # s) obj = (
77
+ − 2201
case obj of
92
+ − 2202
D_file f' \<Rightarrow> if (f' \<in> same_inode_files s f \<and> len > 0)
77
+ − 2203
then Some (S_file (cf2sfiles s f') (O_file f' \<in> tainted s \<or> O_proc p \<in> tainted s))
+ − 2204
else co2sobj s obj
+ − 2205
| _ \<Rightarrow> co2sobj s obj)"
+ − 2206
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2207
apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other) (* ch2sshm_other *)
+ − 2208
+ − 2209
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2210
simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps
+ − 2211
same_inode_files_prop6
+ − 2212
dest:is_file_in_current is_dir_in_current)
+ − 2213
done
+ − 2214
+ − 2215
lemma co2sobj_kill:
92
+ − 2216
"\<lbrakk>valid (Kill p p' # s); dalive (Kill p p' # s) obj\<rbrakk> \<Longrightarrow> co2sobj (Kill p p' # s) obj = co2sobj s obj"
77
+ − 2217
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2218
apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other) (* ch2sshm_other *)
+ − 2219
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2220
simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps
+ − 2221
same_inode_files_prop6
+ − 2222
dest:is_file_in_current is_dir_in_current)
+ − 2223
done
+ − 2224
+ − 2225
lemma co2sobj_exit:
92
+ − 2226
"\<lbrakk>valid (Exit p # s); dalive (Exit p # s) obj\<rbrakk> \<Longrightarrow> co2sobj (Exit p # s) obj = co2sobj s obj"
77
+ − 2227
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2228
apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other)
+ − 2229
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2230
simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps
+ − 2231
same_inode_files_prop6
+ − 2232
dest:is_file_in_current is_dir_in_current)
+ − 2233
done
+ − 2234
+ − 2235
lemma co2sobj_createmsgq:
92
+ − 2236
"\<lbrakk>valid (CreateMsgq p q # s); dalive (CreateMsgq p q # s) obj\<rbrakk> \<Longrightarrow> co2sobj (CreateMsgq p q # s) obj = (
77
+ − 2237
case obj of
92
+ − 2238
D_msgq q' \<Rightarrow> if (q' = q) then (case (cq2smsgq (CreateMsgq p q # s) q) of
77
+ − 2239
Some sq \<Rightarrow> Some (S_msgq sq)
+ − 2240
| _ \<Rightarrow> None)
+ − 2241
else co2sobj s obj
+ − 2242
| _ \<Rightarrow> co2sobj s obj)"
+ − 2243
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2244
apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other)
+ − 2245
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2246
simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps
+ − 2247
same_inode_files_prop6
+ − 2248
dest!:current_has_smsgq'
+ − 2249
dest:is_file_in_current is_dir_in_current cq2smsg_createmsgq)
+ − 2250
done
+ − 2251
+ − 2252
lemma co2sobj_sendmsg:
92
+ − 2253
"\<lbrakk>valid (SendMsg p q m # s); dalive (SendMsg p q m # s) obj\<rbrakk> \<Longrightarrow> co2sobj (SendMsg p q m # s) obj = (
77
+ − 2254
case obj of
92
+ − 2255
D_msgq q' \<Rightarrow> if (q' = q) then (case (cq2smsgq (SendMsg p q m # s) q) of
77
+ − 2256
Some sq \<Rightarrow> Some (S_msgq sq)
+ − 2257
| _ \<Rightarrow> None)
+ − 2258
else co2sobj s obj
+ − 2259
| _ \<Rightarrow> co2sobj s obj)"
+ − 2260
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2261
apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other)
+ − 2262
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2263
simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps
+ − 2264
same_inode_files_prop6
+ − 2265
dest!:current_has_smsgq'
+ − 2266
dest:is_file_in_current is_dir_in_current cq2smsg_sendmsg)
+ − 2267
done
+ − 2268
+ − 2269
lemma co2sobj_recvmsg:
92
+ − 2270
"\<lbrakk>valid (RecvMsg p q m # s); dalive (RecvMsg p q m # s) obj\<rbrakk> \<Longrightarrow> co2sobj (RecvMsg p q m # s) obj = (
77
+ − 2271
case obj of
92
+ − 2272
D_msgq q' \<Rightarrow> if (q' = q) then (case (cq2smsgq (RecvMsg p q m # s) q) of
77
+ − 2273
Some sq \<Rightarrow> Some (S_msgq sq)
+ − 2274
| _ \<Rightarrow> None)
+ − 2275
else co2sobj s obj
92
+ − 2276
| D_proc p' \<Rightarrow> if (p' = p \<and> O_msg q m \<in> tainted s)
77
+ − 2277
then (case (cp2sproc s p') of
+ − 2278
Some sp \<Rightarrow> Some (S_proc sp True)
+ − 2279
| _ \<Rightarrow> None)
+ − 2280
else co2sobj s obj
+ − 2281
| _ \<Rightarrow> co2sobj s obj)"
+ − 2282
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2283
apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other)
+ − 2284
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2285
simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps
+ − 2286
same_inode_files_prop6
+ − 2287
dest!:current_has_smsgq'
+ − 2288
dest:is_file_in_current is_dir_in_current cq2smsg_recvmsg)
+ − 2289
done
+ − 2290
(*
+ − 2291
lemma co2sobj_recvmsg:
92
+ − 2292
"\<lbrakk>valid (RecvMsg p q m # s); dalive (RecvMsg p q m # s) obj\<rbrakk> \<Longrightarrow> co2sobj (RecvMsg p q m # s) obj = (
77
+ − 2293
case obj of
+ − 2294
O_msgq q' \<Rightarrow> if (q' = q) then (case (cq2smsgq (RecvMsg p q m # s) q) of
+ − 2295
Some sq \<Rightarrow> Some (S_msgq sq)
+ − 2296
| _ \<Rightarrow> None)
+ − 2297
else co2sobj s obj
+ − 2298
| O_proc p' \<Rightarrow> if (info_flow_shm s p p' \<and> O_msg q m \<in> tainted s)
+ − 2299
then (case (cp2sproc s p') of
+ − 2300
Some sp \<Rightarrow> Some (S_proc sp True)
+ − 2301
| _ \<Rightarrow> None)
+ − 2302
else co2sobj s obj
+ − 2303
| _ \<Rightarrow> co2sobj s obj)"
+ − 2304
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2305
apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other)
+ − 2306
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2307
simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps
+ − 2308
same_inode_files_prop6
+ − 2309
dest!:current_has_smsgq'
+ − 2310
dest:is_file_in_current is_dir_in_current cq2smsg_recvmsg)
+ − 2311
done
+ − 2312
*)
+ − 2313
+ − 2314
lemma co2sobj_removemsgq:
92
+ − 2315
"\<lbrakk>valid (RemoveMsgq p q # s); dalive (RemoveMsgq p q # s) obj\<rbrakk>
77
+ − 2316
\<Longrightarrow> co2sobj (RemoveMsgq p q # s) obj = co2sobj s obj"
+ − 2317
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2318
apply (simp_all add:current_files_simps is_dir_simps cq2smsgq_other)
+ − 2319
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2320
simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps
+ − 2321
same_inode_files_prop6
+ − 2322
dest!:current_has_smsgq'
+ − 2323
dest:is_file_in_current is_dir_in_current cq2smsg_removemsgq)
+ − 2324
done
+ − 2325
+ − 2326
(*
+ − 2327
lemma co2sobj_createshm:
92
+ − 2328
"\<lbrakk>valid (CreateShM p h # s); dalive (CreateShM p h # s) obj\<rbrakk> \<Longrightarrow> co2sobj (CreateShM p h # s) obj = (
77
+ − 2329
case obj of
+ − 2330
O_shm h' \<Rightarrow> if (h' = h) then (case (ch2sshm (CreateShM p h # s) h) of
+ − 2331
Some sh \<Rightarrow> Some (S_shm sh)
+ − 2332
| _ \<Rightarrow> None)
+ − 2333
else co2sobj s obj
+ − 2334
| _ \<Rightarrow> co2sobj s obj)"
+ − 2335
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2336
apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other)
+ − 2337
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2338
simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_tainted
+ − 2339
same_inode_files_prop6 ch2sshm_simps
+ − 2340
dest!:current_shm_has_sh'
+ − 2341
dest:is_file_in_current is_dir_in_current)
+ − 2342
done
+ − 2343
+ − 2344
lemma co2sobj_detach:
92
+ − 2345
"\<lbrakk>valid (Detach p h # s); dalive s obj\<rbrakk> \<Longrightarrow> co2sobj (Detach p h # s) obj = (
77
+ − 2346
case obj of
+ − 2347
O_proc p' \<Rightarrow> if (p' = p) then (case (cp2sproc (Detach p h # s) p) of
+ − 2348
Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> tainted s))
+ − 2349
| _ \<Rightarrow> None)
+ − 2350
else co2sobj s obj
+ − 2351
| _ \<Rightarrow> co2sobj s obj)"
+ − 2352
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2353
apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other)
+ − 2354
+ − 2355
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2356
simp:current_files_simps cf2sfiles_simps cf2sfile_simps cp2sproc_simps tainted_eq_tainted
+ − 2357
same_inode_files_prop6 ch2sshm_simps
+ − 2358
dest!:current_shm_has_sh'
+ − 2359
dest:is_file_in_current is_dir_in_current)
+ − 2360
done
+ − 2361
+ − 2362
lemma co2sobj_deleteshm:
92
+ − 2363
"\<lbrakk>valid (DeleteShM p h # s); dalive (DeleteShM p h # s) obj\<rbrakk> \<Longrightarrow> co2sobj (DeleteShM p h # s) obj = (
77
+ − 2364
case obj of
+ − 2365
O_proc p' \<Rightarrow> (case (cp2sproc (DeleteShM p h # s) p') of
+ − 2366
Some sp \<Rightarrow> Some (S_proc sp (O_proc p' \<in> tainted s))
+ − 2367
| _ \<Rightarrow> None)
+ − 2368
| _ \<Rightarrow> co2sobj s obj)"
+ − 2369
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2370
apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other)
+ − 2371
+ − 2372
apply (auto split:if_splits option.splits dest!:current_file_has_sfile' current_proc_has_sp'
+ − 2373
simp:current_files_simps cf2sfiles_simps cf2sfile_simps tainted_eq_tainted
+ − 2374
same_inode_files_prop6 ch2sshm_simps
+ − 2375
dest!:current_shm_has_sh'
+ − 2376
dest:is_file_in_current is_dir_in_current)
+ − 2377
done
+ − 2378
*)
+ − 2379
+ − 2380
declare Product_Type.split_paired_Ex Product_Type.split_paired_All [simp del]
+ − 2381
+ − 2382
(*
+ − 2383
lemma info_flow_shm_prop1:
+ − 2384
"p \<in> current_procs s \<Longrightarrow> info_flow_shm s p p"
+ − 2385
by (simp add:info_flow_shm_def)
+ − 2386
+ − 2387
lemma co2sobj_attach:
92
+ − 2388
"\<lbrakk>valid (Attach p h flag # s); dalive s obj\<rbrakk> \<Longrightarrow> co2sobj (Attach p h flag # s) obj = (
77
+ − 2389
case obj of
+ − 2390
O_proc p' \<Rightarrow> if (info_flow_shm s p p')
+ − 2391
then (case (cp2sproc (Attach p h flag # s) p') of
+ − 2392
Some sp \<Rightarrow> Some (S_proc sp (O_proc p' \<in> tainted s \<or>
+ − 2393
(\<exists> p''. O_proc p'' \<in> tainted s \<and> (p'', SHM_RDWR) \<in> procs_of_shm s h)))
+ − 2394
| _ \<Rightarrow> None)
+ − 2395
else if (\<exists> p'' flag'. (p'', flag') \<in> procs_of_shm s h \<and> flag = SHM_RDWR \<and> O_proc p \<in> tainted s \<and>
+ − 2396
info_flow_shm s p'' p')
+ − 2397
then (case (cp2sproc s p') of
+ − 2398
Some sp \<Rightarrow> Some (S_proc sp True)
+ − 2399
| _ \<Rightarrow> None)
+ − 2400
else co2sobj s obj
+ − 2401
| _ \<Rightarrow> co2sobj s obj)"
+ − 2402
apply (frule vt_grant_os, frule vd_cons, case_tac obj)
+ − 2403
apply (simp_all add:current_files_simps is_dir_simps ch2sshm_other cq2smsgq_other)
+ − 2404
+ − 2405
apply (rule conjI|rule impI|erule exE)+
+ − 2406
apply (simp split:option.splits del:split_paired_Ex)
+ − 2407
apply (rule impI, frule current_proc_has_sp, simp)
+ − 2408
apply ((erule exE)+, auto simp:tainted_eq_tainted intro:info_flow_shm_tainted)[1]
+ − 2409
apply (rule impI, simp add:tainted_eq_tainted split:option.splits del:split_paired_Ex)
+ − 2410
apply (auto simp:info_flow_shm_prop1 cp2sproc_attach dest!:current_proc_has_sp')[1]
+ − 2411
+ − 2412
apply (case_tac "cp2sproc (Attach p h flag # s) nat")
+ − 2413
apply (drule current_proc_has_sp', simp+)
+ − 2414
+ − 2415
apply (rule conjI|erule exE|erule conjE|rule impI)+
+ − 2416
apply (simp add:tainted_eq_tainted)
+ − 2417
apply (auto simp:info_flow_shm_prop1 cp2sproc_attach intro:info_flow_shm_tainted dest!:current_proc_has_sp')[1]
+ − 2418
apply (auto simp:info_flow_shm_prop1 cp2sproc_attach intro:info_flow_shm_tainted dest!:current_proc_has_sp'
+ − 2419
split:option.splits if_splits)[1]
+ − 2420
+ − 2421
+ − 2422
apply (auto split:if_splits option.splits dest!:current_file_has_sfile'
+ − 2423
simp:current_files_simps cf2sfiles_simps cf2sfile_simps tainted_eq_tainted
+ − 2424
same_inode_files_prop6
+ − 2425
dest:is_file_in_current is_dir_in_current)
+ − 2426
done
+ − 2427
*)
+ − 2428
92
+ − 2429
lemma co2sobj_bind:
+ − 2430
"valid (Bind p fd addr # s) \<Longrightarrow> co2sobj (Bind p fd addr # s) = co2sobj s"
+ − 2431
apply (frule vd_cons, frule vt_grant, rule ext, case_tac x)
+ − 2432
by auto
+ − 2433
+ − 2434
lemma co2sobj_connect:
+ − 2435
"valid (Connect p fd addr # s) \<Longrightarrow> co2sobj (Connect p fd addr # s) = co2sobj s"
+ − 2436
apply (frule vd_cons, frule vt_grant, rule ext, case_tac x)
+ − 2437
by auto
+ − 2438
+ − 2439
lemma co2sobj_createsock:
+ − 2440
"valid (CreateSock p af st fd inum # s) \<Longrightarrow> co2sobj (CreateSock p af st fd inum # s) = co2sobj s"
+ − 2441
apply (frule vd_cons, frule vt_grant, rule ext, case_tac x)
+ − 2442
by auto
+ − 2443
+ − 2444
lemma co2sobj_accept:
+ − 2445
"valid (Accept p fd addr port fd' inum # s) \<Longrightarrow> co2sobj (Accept p fd addr port fd' inum # s) = co2sobj s"
+ − 2446
apply (frule vd_cons, frule vt_grant, rule ext, case_tac x)
+ − 2447
by auto
+ − 2448
+ − 2449
lemma co2sobj_listen:
+ − 2450
"valid (Listen p fd # s) \<Longrightarrow> co2sobj (Listen p fd # s) = co2sobj s"
+ − 2451
apply (frule vd_cons, frule vt_grant, rule ext, case_tac x)
+ − 2452
by auto
+ − 2453
lemma co2sobj_sendsock:
+ − 2454
"valid (SendSock p fd # s) \<Longrightarrow> co2sobj (SendSock p fd # s) = co2sobj s"
+ − 2455
apply (frule vd_cons, frule vt_grant, rule ext, case_tac x)
+ − 2456
by auto
+ − 2457
lemma co2sobj_recvsock:
+ − 2458
"valid (RecvSock p fd # s) \<Longrightarrow> co2sobj (RecvSock p fd # s) = co2sobj s"
+ − 2459
apply (frule vd_cons, frule vt_grant, rule ext, case_tac x)
+ − 2460
by auto
+ − 2461
lemma co2sobj_shutdown:
+ − 2462
"valid (Shutdown p fd addr # s) \<Longrightarrow> co2sobj (Shutdown p fd addr # s) = co2sobj s"
+ − 2463
apply (frule vd_cons, frule vt_grant, rule ext, case_tac x)
+ − 2464
by auto
+ − 2465
77
+ − 2466
lemma co2sobj_other:
92
+ − 2467
"\<lbrakk>valid (e # s); dalive (e # s) obj;
77
+ − 2468
\<forall> p f fds. e \<noteq> Execve p f fds;
+ − 2469
\<forall> p p' fds. e \<noteq> Clone p p' fds;
+ − 2470
\<forall> p p'. e \<noteq> Ptrace p p';
+ − 2471
\<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ − 2472
\<forall> p fd. e \<noteq> ReadFile p fd;
+ − 2473
\<forall> p fd. e \<noteq> WriteFile p fd;
+ − 2474
\<forall> p fd. e \<noteq> CloseFd p fd;
+ − 2475
\<forall> p f. e \<noteq> UnLink p f;
+ − 2476
\<forall> p f. e \<noteq> Rmdir p f;
+ − 2477
\<forall> p f i. e \<noteq> Mkdir p f i;
+ − 2478
\<forall> p f f'. e \<noteq> LinkHard p f f';
+ − 2479
\<forall> p f len. e \<noteq> Truncate p f len;
+ − 2480
\<forall> p q. e \<noteq> CreateMsgq p q;
+ − 2481
\<forall> p q m. e \<noteq> SendMsg p q m;
+ − 2482
\<forall> p q m. e \<noteq> RecvMsg p q m;
+ − 2483
\<forall> p q. e \<noteq> RemoveMsgq p q
+ − 2484
\<rbrakk> \<Longrightarrow> co2sobj (e # s) obj = co2sobj s obj" (*;
+ − 2485
\<forall> p h flag. e \<noteq> Attach p h flag;
+ − 2486
\<forall> p h. e \<noteq> Detach p h;
+ − 2487
\<forall> p h. e \<noteq> DeleteShM p h*)
+ − 2488
apply (frule vt_grant, case_tac e)
+ − 2489
apply (auto intro:co2sobj_kill co2sobj_exit)
+ − 2490
done
+ − 2491
+ − 2492
lemmas co2sobj_simps = co2sobj_execve co2sobj_clone co2sobj_ptrace co2sobj_open co2sobj_readfile
+ − 2493
co2sobj_writefile co2sobj_closefd co2sobj_unlink co2sobj_rmdir co2sobj_mkdir co2sobj_linkhard
+ − 2494
co2sobj_truncate co2sobj_kill co2sobj_exit co2sobj_createmsgq co2sobj_sendmsg co2sobj_recvmsg
+ − 2495
co2sobj_removemsgq (* co2sobj_attach co2sobj_detach co2sobj_createshm co2sobj_deleteshm *)
92
+ − 2496
co2sobj_createsock co2sobj_accept co2sobj_listen co2sobj_sendsock co2sobj_recvsock
+ − 2497
co2sobj_shutdown co2sobj_bind co2sobj_connect
77
+ − 2498
+ − 2499
end
+ − 2500
+ − 2501
(*<*)
+ − 2502
end
+ − 2503
(*>*)