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