Co2sobj_prop.thy
author chunhan
Wed, 28 Aug 2013 08:59:12 +0800
changeset 32 705e1e41faf7
parent 31 aa1375b6c0eb
child 34 e7f850d1e08e
permissions -rw-r--r--
add event-trace simpset for s2ss
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     1
(*<*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     2
theory Co2sobj_prop
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
     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
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     4
begin
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     5
(*<*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     6
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
     7
ML {*
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
     8
fun print_ss ss =
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
     9
let
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    10
val {simps, congs, procs, ...} = Raw_Simplifier.dest_ss ss
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    11
in
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    12
simps
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    13
end
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    14
*}
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    15
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    16
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    17
context tainting_s begin
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    18
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    19
(********************* cm2smsg simpset ***********************)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    20
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    21
lemma cm2smsg_other: 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    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> 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    23
   \<Longrightarrow> cm2smsg (e # s) = cm2smsg s"
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    24
apply (frule vt_grant_os, frule vd_cons, rule ext, rule ext)
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    25
unfolding cm2smsg_def
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    26
apply (case_tac e)
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    27
apply (auto simp:sectxt_of_obj_simps tainted_eq_Tainted 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    28
           split:t_object.splits option.splits if_splits 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    29
            dest:not_deleted_init_msg  dest!:current_has_sec')
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    30
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    32
lemma cm2smsg_sendmsg:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    33
  "valid (SendMsg p q m # s) \<Longrightarrow> cm2smsg (SendMsg p q m # s) = (\<lambda> q' m'. 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    34
     if (q' = q \<and> m' = m)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    35
     then (case (sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    36
             Some sec \<Rightarrow> Some (Created, sec, O_msg q m \<in> tainted (SendMsg p q m # s))
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    37
           | _ \<Rightarrow> None)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    38
     else cm2smsg s q' m') "
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    39
apply (frule vd_cons, frule vt_grant_os, rule ext, rule ext)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    40
apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    41
  split:if_splits option.splits
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    42
  dest:not_deleted_init_msg  dest!:current_has_sec')
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    43
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    44
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    45
lemma cm2smsg_recvmsg1:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    46
  "\<lbrakk>q' \<noteq> q; valid (RecvMsg p q m # s)\<rbrakk> \<Longrightarrow> cm2smsg (RecvMsg p q m # s) q' = cm2smsg s q'"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    47
apply (frule vd_cons, frule vt_grant_os, rule ext)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    48
apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    49
  split:if_splits option.splits)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    50
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    51
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    52
lemma cm2smsg_recvmsg2:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    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'"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    54
apply (frule vd_cons, frule vt_grant_os)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    55
apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    56
  split:if_splits option.splits)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    57
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    58
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    59
lemma cm2smsg_recvmsg1':
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    60
  "\<lbrakk>valid (RecvMsg p q m # s); q' \<noteq> q\<rbrakk> \<Longrightarrow> cm2smsg (RecvMsg p q m # s) q' = cm2smsg s q'"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    61
apply (frule vd_cons, frule vt_grant_os, rule ext)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    62
apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    63
  split:if_splits option.splits)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    64
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    65
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    66
lemma cm2smsg_recvmsg2':
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    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'"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    68
apply (frule vd_cons, frule vt_grant_os)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    69
apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    70
  split:if_splits option.splits)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    71
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    72
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    73
lemma cm2smsg_removemsgq:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    74
  "\<lbrakk>q' \<noteq> q; valid (RemoveMsgq p q # s)\<rbrakk> \<Longrightarrow> cm2smsg (RemoveMsgq p q # s) q' = cm2smsg s q'"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    75
apply (frule vd_cons, frule vt_grant_os, rule ext)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    76
apply (auto simp:cm2smsg_def tainted_eq_Tainted sectxt_of_obj_simps 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    77
  split:if_splits option.splits)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    78
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    79
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    80
lemmas cm2smsg_simps = cm2smsg_nil_prop cm2smsg_sendmsg cm2smsg_recvmsg1' cm2smsg_recvmsg2'
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    81
  cm2smsg_removemsgq cm2smsg_other
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    82
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    83
lemma init_cm2smsg_has_smsg:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    84
  "\<lbrakk>m \<in> set (init_msgs_of_queue q); q \<in> init_msgqs\<rbrakk> \<Longrightarrow> \<exists> sm. init_cm2smsg q m = Some sm"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    85
by (auto simp:init_cm2smsg_def split:option.splits dest:init_msg_has_ctxt)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    86
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    87
lemma hd_set_prop1:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    88
  "hd l \<notin> set l \<Longrightarrow> l = []"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    89
by (case_tac l, auto)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    90
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    91
lemma tl_set_prop1:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    92
  "a \<in> set (tl l) \<Longrightarrow> a \<in> set l"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    93
by (case_tac l, auto)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    94
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    95
lemma current_has_smsg:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    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"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    97
apply (induct s)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    98
apply (simp only:cm2smsg_nil_prop msgs_of_queue.simps current_msgqs.simps init_cm2smsg_has_smsg)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
    99
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   100
apply (frule vt_grant_os, frule vd_cons, case_tac a)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   101
apply (auto simp:cm2smsg_def sectxt_of_obj_simps split:if_splits option.splits
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   102
  dest!:current_has_sec' hd_set_prop1 dest:not_deleted_init_msg tl_set_prop1)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   103
done 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   104
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   105
lemma current_has_smsg':
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   106
  "\<lbrakk>cm2smsg s q m = None; q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> m \<notin> set (msgs_of_queue s q)"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   107
by (auto dest:current_has_smsg)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   108
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   109
lemma cqm2sms_has_sms_aux:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   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)"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   111
by (induct ms, auto split:option.splits simp:cm2smsg_def)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   112
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   113
lemma current_has_sms:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   114
  "\<lbrakk>q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sms. cqm2sms s q (msgs_of_queue s q) = Some sms"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   115
apply (rule cqm2sms_has_sms_aux)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   116
apply (auto dest:current_msg_has_sec)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   117
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   118
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   119
lemma current_has_sms':
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   120
  "\<lbrakk>cqm2sms s q (msgs_of_queue s q) = None; valid s\<rbrakk> \<Longrightarrow> q \<notin> current_msgqs s"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   121
by (auto dest:current_has_sms)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   122
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   123
(********************* cqm2sms simpset ***********************) 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   124
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   125
lemma cqm2sms_other:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   126
  "\<lbrakk>valid (e # s); \<forall> p m. e \<noteq> CreateMsgq p q; \<forall> p q m. e \<noteq> SendMsg p q m; 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   127
    \<forall> p q m. e \<noteq> RecvMsg p q m; \<forall> p q. e \<noteq> RemoveMsgq p q\<rbrakk> 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   128
   \<Longrightarrow> cqm2sms (e # s) = cqm2sms s"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   129
apply (rule ext, rule ext)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   130
apply (induct_tac xa, simp)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   131
apply (frule vt_grant_os, frule vd_cons, case_tac e)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   132
apply (auto split:option.splits dest:cm2smsg_other) 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   133
done
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   134
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   135
lemma cqm2sms_createmsgq:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   136
  "valid (CreateMsgq p q # s) \<Longrightarrow> cqm2sms (CreateMsgq p q # s) = (\<lambda> q' ms'. 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   137
     if (q' = q \<and> ms' = []) then Some []
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   138
     else cqm2sms s q' ms')"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   139
apply (rule ext, rule ext)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   140
apply (frule vt_grant_os, frule vd_cons, induct_tac ms')
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   141
apply (auto split:if_splits option.splits dest:cm2smsg_other)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   142
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   143
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   144
lemma cqm2sms_2:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   145
  "cqm2sms s q (ms @ [m]) = 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   146
     (case (cqm2sms s q ms, cm2smsg s q m) of 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   147
       (Some sms, Some sm) \<Rightarrow> Some (sms @ [sm]) 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   148
     | _ \<Rightarrow> None)"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   149
apply (induct ms, simp split:option.splits)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   150
by (auto split:option.splits)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   151
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   152
lemmas cqm2sms_simps2 = cqm2sms.simps(1) cqm2sms_2
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   153
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   154
declare cqm2sms.simps [simp del]
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   155
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   156
lemma cqm2sms_prop1:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   157
  "cqm2sms s q ms = None \<Longrightarrow> \<exists> m \<in> set ms. cm2smsg s q m = None"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   158
by (induct ms, auto simp:cqm2sms.simps split:option.splits)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   159
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   160
lemma cqm2sms_sendmsg1:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   161
  "\<lbrakk>valid (SendMsg p q m # s); m \<notin> set ms\<rbrakk>
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   162
   \<Longrightarrow> cqm2sms (SendMsg p q m # s) q' ms = cqm2sms s q' ms"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   163
apply (frule vt_grant_os, frule vd_cons)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   164
apply (frule cm2smsg_sendmsg)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   165
apply (induct ms rule:rev_induct)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   166
apply (auto split:if_splits option.splits  simp:cqm2sms_simps2)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   167
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   168
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   169
lemma cqm2sms_sendmsg2:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   170
  "\<lbrakk>valid (SendMsg p q m # s); q' \<noteq> q\<rbrakk>
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   171
   \<Longrightarrow> cqm2sms (SendMsg p q m # s) q' ms = cqm2sms s q' ms"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   172
apply (frule vt_grant_os, frule vd_cons)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   173
apply (frule cm2smsg_sendmsg)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   174
apply (induct ms rule:rev_induct)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   175
apply (auto split:if_splits option.splits  simp:cqm2sms_simps2)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   176
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   177
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   178
lemma cqm2sms_sendmsg3:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   179
  "\<lbrakk>valid (SendMsg p q m # s); ms' = msgs_of_queue (SendMsg p q m # s) q\<rbrakk>
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   180
   \<Longrightarrow> cqm2sms (SendMsg p q m # s) q ms' = 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   181
     (case (cqm2sms s q (msgs_of_queue s q), sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   182
       (Some sms, Some sec) \<Rightarrow> Some (sms @ [(Created, sec, O_msg q m \<in> tainted (SendMsg p q m # s))])
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   183
     | _ \<Rightarrow> None)"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   184
apply (frule vt_grant_os, frule vd_cons)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   185
apply (frule cm2smsg_sendmsg)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   186
apply (auto split:if_splits option.splits  simp:cqm2sms_simps2 cqm2sms_sendmsg1)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   187
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   188
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   189
lemma cqm2sms_sendmsg:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   190
  "\<lbrakk>valid (SendMsg p q m # s); ms' = msgs_of_queue (SendMsg p q m # s) q'\<rbrakk> 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   191
   \<Longrightarrow> cqm2sms (SendMsg p q m # s) q' ms' = (
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   192
     if (q' = q) 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   193
     then (case (cqm2sms s q (msgs_of_queue s q), sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   194
             (Some sms, Some sec) \<Rightarrow> Some (sms @ [(Created, sec, O_msg q m \<in> tainted (SendMsg p q m # s))])
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   195
           | _ \<Rightarrow> None)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   196
     else cqm2sms s q' ms' )"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   197
apply (simp split:if_splits add:cqm2sms_sendmsg2 cqm2sms_sendmsg3)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   198
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   199
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   200
lemma cqm2sms_recvmsg1:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   201
  "\<lbrakk>valid (RecvMsg p q m # s); m \<notin> set ms\<rbrakk> 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   202
   \<Longrightarrow> cqm2sms (RecvMsg p q m # s) q ms = cqm2sms s q ms"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   203
apply (frule vt_grant_os, frule vd_cons)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   204
apply (induct ms rule:rev_induct)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   205
apply (auto split:if_splits option.splits simp:cqm2sms_simps2 cm2smsg_recvmsg2')
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   206
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   207
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   208
lemma cqm2sms_recvmsg2:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   209
  "\<lbrakk>valid (RecvMsg p q m # s); q' \<noteq> q\<rbrakk>
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   210
   \<Longrightarrow> cqm2sms (RecvMsg p q m # s) q' ms = cqm2sms s q' ms"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   211
apply (frule vt_grant_os, frule vd_cons)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   212
apply (induct ms rule:rev_induct)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   213
apply (auto split:if_splits option.splits simp:cqm2sms_simps2 cm2smsg_recvmsg1')
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   214
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   215
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   216
lemma cqm2sms_recvmsg:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   217
  "\<lbrakk>valid (RecvMsg p q m # s); ms = msgs_of_queue (RecvMsg p q m # s) q'\<rbrakk>
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   218
   \<Longrightarrow> cqm2sms (RecvMsg p q m # s) q' ms = (
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   219
     if (q' = q) 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   220
     then (case (cqm2sms s q (msgs_of_queue s q)) of
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   221
             Some sms \<Rightarrow> Some (tl sms)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   222
           | _ \<Rightarrow> None)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   223
     else cqm2sms s q' ms)"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   224
apply (frule vt_grant_os, frule vd_cons)
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   225
apply (auto split:if_splits option.splits simp:cqm2sms_recvmsg1 cqm2sms_recvmsg2 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   226
     dest!:current_has_sms')
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   227
apply (case_tac "msgs_of_queue s q", simp)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   228
apply (frule_tac ms = "tl (msgs_of_queue s q)" in cqm2sms_recvmsg1)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   229
apply (drule_tac q = q in distinct_queue_msgs, simp+)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   230
apply (case_tac a, auto simp:cqm2sms.simps split:option.splits if_splits)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   231
done
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   232
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   233
lemma cqm2sms_removemsgq:
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   234
  "\<lbrakk>valid (RemoveMsgq p q # s); q' \<noteq> q; q' \<in> current_msgqs s\<rbrakk> 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   235
   \<Longrightarrow> cqm2sms (RemoveMsgq p q # s) q' ms = cqm2sms s q' ms"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   236
apply (frule vt_grant_os, frule vd_cons)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   237
apply (induct ms rule:rev_induct)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   238
apply (auto simp:cqm2sms_simps2 cm2smsg_removemsgq split:option.splits if_splits)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   239
done
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   240
  
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   241
lemmas cqm2sms_simps = cqm2sms_other cqm2sms_createmsgq cqm2sms_sendmsg cqm2sms_recvmsg cqm2sms_removemsgq
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   242
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   243
(********************* cq2smsgq simpset ***********************) 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   244
 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   245
lemma cq2smsgq_other: 
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   246
  "\<lbrakk>valid (e # s); \<forall> p q. e \<noteq> CreateMsgq p q; \<forall> p q m. e \<noteq> SendMsg p q m; 
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   247
    \<forall> p q m. e \<noteq> RecvMsg p q m; \<forall> p q. e \<noteq> RemoveMsgq p q\<rbrakk> 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   248
   \<Longrightarrow> cq2smsgq (e # s) = cq2smsgq s" 
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   249
apply (frule cqm2sms_other, simp+)
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   250
apply (frule vd_cons, frule vt_grant_os, rule ext, case_tac e) 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   251
apply (auto simp:cq2smsgq_def sectxt_of_obj_simps 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 19
diff changeset
   252
           split:t_object.splits option.splits if_splits 
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   253
            dest:not_deleted_init_msg) 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   254
done
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   255
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   256
lemma cq2smsg_createmsgq:
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   257
  "valid (CreateMsgq p q # s)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   258
   \<Longrightarrow> cq2smsgq (CreateMsgq p q # s) = (cq2smsgq s) (q := 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   259
     case (sectxt_of_obj s (O_proc p)) of
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   260
       Some psec \<Rightarrow> Some (Created, (fst psec, R_object, (snd o snd) psec), [])
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   261
     | None     \<Rightarrow> None)"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   262
apply (frule vd_cons, frule vt_grant_os)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   263
apply (frule cqm2sms_createmsgq)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   264
apply (rule ext, auto simp:cq2smsgq_def sec_createmsgq 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   265
  split:option.splits if_splits dest:not_deleted_init_msgq)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   266
done
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   267
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   268
lemma cq2smsg_sendmsg:
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   269
  "valid (SendMsg p q m # s) 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   270
   \<Longrightarrow> cq2smsgq (SendMsg p q m # s) = (cq2smsgq s) (q := 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   271
     case (cq2smsgq s q, sectxt_of_obj (SendMsg p q m # s) (O_msg q m)) of 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   272
       (Some (qi, sec, sms), Some msec) \<Rightarrow> 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   273
          Some (qi, sec, sms @ [(Created, msec, O_msg q m \<in> tainted (SendMsg p q m # s))])
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   274
     | _ \<Rightarrow> None)"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   275
apply (frule vd_cons, frule vt_grant_os)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   276
apply (rule ext)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   277
apply (frule_tac q' = x in cqm2sms_sendmsg, simp)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   278
apply (auto simp:cq2smsgq_def sectxt_of_obj_simps  split:option.splits if_splits
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   279
  dest!:current_has_sms' current_has_sec')
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   280
done
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   281
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   282
lemma current_has_smsgq:
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   283
  "\<lbrakk>q \<in> current_msgqs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sq. cq2smsgq s q = Some sq"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   284
by (auto simp:cq2smsgq_def split:option.splits dest!:current_has_sec' current_has_sms')
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   285
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   286
lemma current_has_smsgq':
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   287
  "\<lbrakk>cq2smsgq s q = None; valid s\<rbrakk> \<Longrightarrow> q \<notin> current_msgqs s"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   288
by (auto dest:current_has_smsgq)
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   289
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   290
lemma cq2smsg_recvmsg:
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   291
  "valid (RecvMsg p q m # s) 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   292
   \<Longrightarrow> cq2smsgq (RecvMsg p q m # s) = (cq2smsgq s) (q := 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   293
    case (cq2smsgq s q) of
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   294
      Some (qi, sec, sms) \<Rightarrow> Some (qi, sec, tl sms)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   295
    | _ \<Rightarrow> None)"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   296
apply (frule vd_cons, frule vt_grant_os)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   297
apply (rule ext, frule_tac q' = x in cqm2sms_recvmsg, simp)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   298
apply (auto simp add:cq2smsgq_def sectxt_of_obj_simps split:option.splits if_splits 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   299
  dest!:current_has_sec' current_has_sms' current_has_smsgq')
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   300
done
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   301
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   302
lemma cq2smsg_removemsgq:
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   303
  "\<lbrakk>valid (RemoveMsgq p q # s); q' \<noteq> q; q' \<in> current_msgqs s\<rbrakk>
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   304
   \<Longrightarrow> cq2smsgq (RemoveMsgq p q # s) q' = cq2smsgq s q'"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   305
apply (frule vd_cons, frule vt_grant_os)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   306
apply (auto simp:cq2smsgq_def sectxt_of_obj_simps cqm2sms_removemsgq split:if_splits option.splits 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   307
  dest!:current_has_sec' current_has_sms' current_has_smsgq')
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   308
done
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   309
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   310
lemmas cq2smsgq_simps = cq2smsgq_other cq2smsg_sendmsg cq2smsg_recvmsg cq2smsg_removemsgq
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   311
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   312
lemma sm_in_sqsms_init_aux:
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   313
  "\<lbrakk>m \<in> set ms; init_cm2smsg q m = Some sm; q \<in> init_msgqs; 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   314
    init_cqm2sms q ms = Some sms\<rbrakk> \<Longrightarrow> sm \<in> set sms"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   315
apply (induct ms arbitrary:m sm sms)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   316
by (auto split:if_splits option.splits)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   317
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   318
lemma sm_in_sqsms_init:
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   319
  "\<lbrakk>m \<in> set (init_msgs_of_queue q); init_cm2smsg q m = Some sm; q \<in> init_msgqs; 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   320
    init_cqm2sms q (init_msgs_of_queue q) = Some sms\<rbrakk> \<Longrightarrow> sm \<in> set sms"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   321
by (simp add:sm_in_sqsms_init_aux)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   322
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   323
lemma cqm2sms_prop0:
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   324
  "\<lbrakk>m \<in> set ms; cm2smsg s q m = Some sm; cqm2sms s q ms = Some sms\<rbrakk> \<Longrightarrow> sm \<in> set sms"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   325
apply (induct ms arbitrary:m sm sms)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   326
apply (auto simp:cqm2sms.simps split:option.splits)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   327
done
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   328
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   329
lemma sm_in_sqsms:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   330
  "\<lbrakk>m \<in> set (msgs_of_queue s q); q \<in> current_msgqs s; valid s; cq2smsgq s q = Some (qi, qsec, sms);
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   331
    cm2smsg s q m = Some sm\<rbrakk> \<Longrightarrow> sm \<in> set sms"
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   332
proof (induct s arbitrary:m q qi qsec sms sm)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   333
  case Nil
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   334
  thus ?case 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   335
    by (simp add:cq2smsga_nil_prop cm2smsg_nil_prop init_cq2smsgq_def sm_in_sqsms_init 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   336
           split:option.splits)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   337
next
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   338
  case (Cons e s)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   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; 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   340
    cq2smsgq s q = Some (qi, qsec, sms); cm2smsg s q m = Some sm\<rbrakk>
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   341
    \<Longrightarrow> sm \<in> set sms" and p2: "m \<in> set (msgs_of_queue (e # s) q)" 
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   342
    and p3: "q \<in> current_msgqs (e # s)" and p4: "valid (e # s)"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   343
    and p5: "cq2smsgq (e # s) q = Some (qi, qsec, sms)"
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   344
    and p6: "cm2smsg (e # s) q m = Some sm" by auto
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   345
  from p4 have os: "os_grant s e" and vd: "valid s" by (auto dest:vd_cons vt_grant_os)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   346
(*
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   347
  from p1 have p1':
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   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;
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   349
    valid s; cq2smsgq s q = Some (qi, qsec, sms); cm2smsg s q m = Some sm\<rbrakk>
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   350
    \<Longrightarrow> sm \<in> set "
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   351
*)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   352
  show ?case using p2 p3 p4 p5 p6 vd os
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   353
    apply (case_tac e)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   354
    apply (auto simp:cq2smsgq_simps cm2smsg_simps split:if_splits option.splits intro:p1)
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   355
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   356
    apply (rule_tac m = m and q = q and qi = qi and qsec = qsec in p1, simp+)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   357
    apply (case_tac "q = nat2", simp)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   358
    apply (drule cq2smsg_createmsgq, simp, simp)
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   359
32
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   360
    apply (drule_tac m = m and q = q and qi = qi and qsec = "(aa,ab,b)" in p1, simp+)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   361
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   362
    apply (simp add:cq2smsgq_def split:option.splits)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   363
    apply (rule_tac m = m and sm = sm in cqm2sms_prop0, simp+)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   364
    apply (simp add:cqm2sms_recvmsg)
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   365
    done
705e1e41faf7 add event-trace simpset for s2ss
chunhan
parents: 31
diff changeset
   366
qed
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   367
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   368
(****************** cf2sfile path simpset ***************)
11
chunhan
parents: 10
diff changeset
   369
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   370
lemma sroot_only:
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   371
  "cf2sfile s [] = Some sroot"
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   372
by (simp add:cf2sfile_def)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   373
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   374
lemma not_file_is_dir:
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   375
  "\<lbrakk>\<not> is_file s f; f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> is_dir s f"
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   376
by (auto simp:is_file_def current_files_def is_dir_def 
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   377
         dest:finum_has_itag finum_has_ftag' split:t_inode_tag.splits option.splits)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   378
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   379
lemma not_dir_is_file:
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   380
  "\<lbrakk>\<not> is_dir s f; f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> is_file s f"
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   381
by (auto simp:is_file_def current_files_def is_dir_def 
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   382
         dest:finum_has_itag finum_has_ftag' split:t_inode_tag.splits option.splits)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   383
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   384
lemma is_file_or_dir:
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   385
  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> is_file s f \<or> is_dir s f"
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   386
by (auto dest:not_dir_is_file)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   387
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   388
lemma current_file_has_sfile:
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   389
  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> sf. cf2sfile s f = Some sf"
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   390
apply (induct f)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   391
apply (rule_tac x = "sroot" in exI, simp add:sroot_only)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   392
apply (frule parentf_in_current', simp, clarsimp)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   393
apply (frule parentf_is_dir'', simp)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   394
apply (frule is_file_or_dir, simp)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   395
apply (auto dest!:current_has_sec'
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   396
            simp:cf2sfile_def split:option.splits if_splits dest!:get_pfs_secs_prop')
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   397
done
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   398
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   399
definition sectxt_of_pf :: "t_state \<Rightarrow> t_file \<Rightarrow> security_context_t option"
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   400
where
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   401
  "sectxt_of_pf s f = (case f of [] \<Rightarrow> None | (a # pf) \<Rightarrow> sectxt_of_obj s (O_dir pf))"
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   402
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   403
definition get_parentfs_ctxts' :: "t_state \<Rightarrow> t_file \<Rightarrow> (security_context_t list) option"
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   404
where
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   405
  "get_parentfs_ctxts' s f = (case f of [] \<Rightarrow> None | (a # pf) \<Rightarrow> get_parentfs_ctxts s pf)"
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   406
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   407
lemma is_file_has_sfile:
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   408
  "\<lbrakk>is_file s f; valid s\<rbrakk> \<Longrightarrow> \<exists> sec psec asecs. cf2sfile s f = Some 
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   409
      (if (\<not> deleted (O_file f) s \<and> is_init_file f) then Init f else Created,
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   410
       sec, Some psec, set asecs) \<and> (sectxt_of_obj s (O_file f) = Some sec) \<and>
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   411
       (sectxt_of_pf s f = Some psec) \<and> (get_parentfs_ctxts' s f = Some asecs)"
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   412
apply (case_tac f, simp, drule root_is_dir', simp, simp)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   413
apply (frule is_file_in_current)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   414
apply (drule current_file_has_sfile, simp)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   415
apply (auto simp:cf2sfile_def sectxt_of_pf_def get_parentfs_ctxts'_def split:if_splits option.splits)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   416
done
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   417
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   418
lemma is_dir_has_sfile:
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   419
  "\<lbrakk>is_dir s f; valid s\<rbrakk> \<Longrightarrow> (case f of
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   420
      [] \<Rightarrow> cf2sfile s f = Some sroot
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   421
    | a # pf \<Rightarrow> (\<exists> sec psec asecs. cf2sfile s f = Some 
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   422
      (if (\<not> deleted (O_dir f) s \<and> is_init_dir f) then Init f else Created,
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   423
       sec, Some psec, set asecs) \<and> (sectxt_of_obj s (O_dir f) = Some sec) \<and>
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   424
       (sectxt_of_obj s (O_dir pf) = Some psec) \<and> (get_parentfs_ctxts s pf = Some asecs)))"
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   425
apply (case_tac f, simp add:sroot_only)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   426
apply (frule is_dir_in_current, frule is_dir_not_file)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   427
apply (drule current_file_has_sfile, simp)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   428
apply (auto simp:cf2sfile_def split:if_splits option.splits)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   429
done
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   430
7
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   431
lemma sroot_set:
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   432
  "valid s \<Longrightarrow> \<exists> sec. sroot = (Init [], sec, None, {}) \<and> sectxt_of_obj s (O_dir []) = Some sec"
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   433
apply (frule root_is_dir)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   434
apply (drule is_dir_has_sec, simp)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   435
apply (auto simp:sroot_def sec_of_root_def sectxt_of_obj_def type_of_obj.simps 
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   436
                 root_type_remains root_user_remains
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   437
           dest!:root_has_type' root_has_user' root_has_init_type' root_has_init_user'
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   438
           split:option.splits)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   439
done
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   440
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   441
lemma cf2sfile_path_file:
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   442
  "\<lbrakk>is_file s (f # pf); valid s\<rbrakk>
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   443
   \<Longrightarrow> cf2sfile s (f # pf) = (
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   444
     case (cf2sfile s pf) of
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   445
       Some (pfi, pfsec, psec, asecs) \<Rightarrow> 
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   446
          (case (sectxt_of_obj s (O_file (f # pf))) of
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   447
              Some fsec \<Rightarrow> Some (if (\<not> deleted (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   448
                                else Created, fsec, Some pfsec, asecs \<union> {pfsec})
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   449
           | None \<Rightarrow> None)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   450
     | _ \<Rightarrow> None)"
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   451
apply (frule is_file_in_current, drule parentf_is_dir'', simp)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   452
apply (frule is_dir_has_sfile, simp, frule is_file_has_sfile, simp)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   453
apply (frule sroot_set)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   454
apply (case_tac pf, (clarsimp simp:get_parentfs_ctxts'_def sectxt_of_pf_def)+) 
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   455
done
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   456
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   457
lemma cf2sfile_path_dir:
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   458
  "\<lbrakk>is_dir s (f # pf); valid s\<rbrakk>
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   459
   \<Longrightarrow> cf2sfile s (f # pf) = (
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   460
     case (cf2sfile s pf) of
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   461
       Some (pfi, pfsec, psec, asecs) \<Rightarrow> 
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   462
          (case (sectxt_of_obj s (O_dir (f # pf))) of
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   463
              Some fsec \<Rightarrow> Some (if (\<not> deleted (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   464
                                else Created, fsec, Some pfsec, asecs \<union> {pfsec})
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   465
           | None \<Rightarrow> None)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   466
     | _ \<Rightarrow> None)"
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   467
apply (frule is_dir_in_current, drule parentf_is_dir'', simp)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   468
apply (frule_tac f = "f # pf" in is_dir_has_sfile, simp)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   469
apply (frule_tac f = "pf" in is_dir_has_sfile, simp)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   470
apply (frule sroot_set)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   471
apply (case_tac pf, (clarsimp simp:get_parentfs_ctxts'_def sectxt_of_pf_def)+) 
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   472
done  
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   473
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   474
lemma cf2sfile_path:
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   475
  "\<lbrakk>f # pf \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> cf2sfile s (f # pf) = (
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   476
     case (cf2sfile s pf) of
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   477
       Some (pfi, pfsec, psec, asecs) \<Rightarrow> (if (is_file s (f # pf))
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   478
         then (case (sectxt_of_obj s (O_file (f # pf))) of
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   479
                 Some fsec \<Rightarrow> Some (if (\<not> deleted (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   480
                                   else Created, fsec, Some pfsec, asecs \<union> {pfsec})
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   481
               | None \<Rightarrow> None)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   482
         else (case (sectxt_of_obj s (O_dir (f # pf))) of
7
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   483
                 Some fsec \<Rightarrow> Some (if (\<not> deleted (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   484
                                   else Created, fsec, Some pfsec, asecs \<union> {pfsec})
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   485
               | None \<Rightarrow> None)           )
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   486
     | None \<Rightarrow> None)"
7
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   487
apply (drule is_file_or_dir, simp)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   488
apply (erule disjE)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   489
apply (frule cf2sfile_path_file, simp) defer
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   490
apply (frule cf2sfile_path_dir, simp, drule is_dir_not_file)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   491
apply (auto split:option.splits)
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   492
done
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   493
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   494
lemma cf2sfile_path_file_prop1:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   495
  "\<lbrakk>is_file s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs); valid s\<rbrakk>
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   496
   \<Longrightarrow> \<exists> fsec. cf2sfile s (f # pf) = 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   497
                 Some (if (\<not> deleted (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   498
                       else Created, fsec, Some pfsec, asecs \<union> {pfsec}) \<and> 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   499
               sectxt_of_obj s (O_file (f # pf)) = Some fsec"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   500
apply (frule is_file_has_sfile, simp)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   501
by (auto simp:cf2sfile_path_file)
7
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   502
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   503
lemma cf2sfile_path_file_prop2:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   504
  "\<lbrakk>is_file s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs); 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   505
    sectxt_of_obj s (O_file (f # pf)) = Some fsec; valid s\<rbrakk> \<Longrightarrow> cf2sfile s (f # pf) = 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   506
      Some (if (\<not> deleted (O_file (f # pf)) s \<and> is_init_file (f # pf)) then Init (f # pf)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   507
            else Created, fsec, Some pfsec, asecs \<union> {pfsec})"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   508
by (drule cf2sfile_path_file_prop1, auto)
7
f27882976251 finally get cf2sfile path property done
chunhan
parents: 6
diff changeset
   509
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   510
lemma cf2sfile_path_dir_prop1:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   511
  "\<lbrakk>is_dir s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs); valid s\<rbrakk>
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   512
   \<Longrightarrow> \<exists> fsec. cf2sfile s (f # pf) = 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   513
                 Some (if (\<not> deleted (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   514
                       else Created, fsec, Some pfsec, asecs \<union> {pfsec}) \<and> 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   515
               sectxt_of_obj s (O_dir (f # pf)) = Some fsec"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   516
apply (frule is_dir_has_sfile, simp)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   517
by (auto simp:cf2sfile_path_dir)
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   518
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   519
lemma cf2sfile_path_dir_prop2:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   520
  "\<lbrakk>is_dir s (f # pf); cf2sfile s pf = Some (pfi, pfsec, psec, asecs); 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   521
    sectxt_of_obj s (O_dir (f # pf)) = Some fsec; valid s\<rbrakk> \<Longrightarrow> cf2sfile s (f # pf) = 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   522
      Some (if (\<not> deleted (O_dir (f # pf)) s \<and> is_init_dir (f # pf)) then Init (f # pf)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   523
            else Created, fsec, Some pfsec, asecs \<union> {pfsec})"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   524
by (drule cf2sfile_path_dir_prop1, auto)
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   525
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   526
(**************** cf2sfile event list simpset ****************)
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   527
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   528
lemma cf2sfile_open_none':
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   529
  "valid (Open p f flag fd None # s) \<Longrightarrow> cf2sfile (Open p f flag fd None # s) f'= cf2sfile s f'"
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   530
apply (frule vd_cons, frule vt_grant_os)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   531
apply (induct f', simp add:cf2sfile_def)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   532
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   533
               get_parentfs_ctxts_simps)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   534
done
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   535
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   536
lemma cf2sfile_open_none:
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   537
  "valid (Open p f flag fd None # s) \<Longrightarrow> cf2sfile (Open p f flag fd None # s) = cf2sfile s"
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   538
apply (rule ext)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   539
by (simp add:cf2sfile_open_none')
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   540
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   541
lemma cf2sfile_open_some1:
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   542
  "\<lbrakk>valid (Open p f flag fd (Some inum) # s); f' \<in> current_files s\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   543
   \<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'"
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   544
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   545
apply (case_tac "f = f'", simp)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   546
apply (induct f', simp add:sroot_only, simp)
3
chunhan
parents: 1
diff changeset
   547
apply (frule parentf_in_current', simp+)
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   548
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps 
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   549
               get_parentfs_ctxts_simps)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   550
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   551
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   552
lemma cf2sfile_open_some2:
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   553
  "\<lbrakk>valid (Open p f flag fd (Some inum) # s); is_file s f'\<rbrakk>
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   554
   \<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   555
apply (frule vd_cons, drule is_file_in_current)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   556
by (simp add:cf2sfile_open_some1)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   557
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   558
lemma cf2sfile_open_some3:
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   559
  "\<lbrakk>valid (Open p f flag fd (Some inum) # s); is_dir s f'\<rbrakk>
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   560
   \<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   561
apply (frule vd_cons, drule is_dir_in_current)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   562
by (simp add:cf2sfile_open_some1)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   563
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   564
lemma cf2sfile_open_some4:
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   565
  "valid (Open p f flag fd (Some inum) # s) \<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f = (
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   566
     case (parent f) of
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   567
       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), 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   568
                         get_parentfs_ctxts s pf) of
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   569
                    (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   570
                  | _ \<Rightarrow> None)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   571
     | None \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   572
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
4
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   573
apply (case_tac f, simp)
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   574
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps 
e9c5594d5963 fixed bugs in deleted definition
chunhan
parents: 3
diff changeset
   575
               get_parentfs_ctxts_simps)
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   576
apply (rule impI, (erule conjE)+)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   577
apply (drule not_deleted_init_file, simp+)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   578
apply (simp add:is_file_in_current)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   579
done
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   580
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   581
lemma cf2sfile_open:
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   582
  "\<lbrakk>valid (Open p f flag fd opt # s); f' \<in> current_files (Open p f flag fd opt # s)\<rbrakk>
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   583
   \<Longrightarrow> cf2sfile (Open p f flag fd opt # s) f' = (
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   584
     if (opt = None) then cf2sfile s f'
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   585
     else if (f' = f) 
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   586
     then (case (parent f) of
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   587
             Some pf \<Rightarrow> (case (sectxt_of_obj (Open p f flag fd opt # s) (O_file f), sectxt_of_obj s (O_dir pf), 
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   588
                 get_parentfs_ctxts s pf) of
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   589
                          (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   590
                        | _ \<Rightarrow> None)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   591
           | None \<Rightarrow> None)
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   592
     else cf2sfile s f')"
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   593
apply (case_tac opt)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   594
apply (simp add:cf2sfile_open_none)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   595
apply (case_tac "f = f'")
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   596
apply (simp add:cf2sfile_open_some4 split:option.splits)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   597
apply (simp add:cf2sfile_open_some1 current_files_simps)
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   598
done
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   599
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   600
lemma cf2sfile_mkdir1:
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   601
  "\<lbrakk>valid (Mkdir p f i # s); f' \<in> current_files s\<rbrakk>
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   602
   \<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = cf2sfile s f'"
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   603
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   604
apply (case_tac "f = f'", simp)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   605
apply (induct f', simp add:sroot_only, simp)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   606
apply (frule parentf_in_current', simp+)
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   607
apply (case_tac "f = f'", simp)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   608
apply (simp add:cf2sfile_path is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps 
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   609
               get_parentfs_ctxts_simps split:if_splits option.splits)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   610
done
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   611
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   612
lemma cf2sfile_mkdir2:
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   613
  "\<lbrakk>valid (Mkdir p f i # s); is_file s f'\<rbrakk>
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   614
   \<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = cf2sfile s f'"
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   615
apply (frule vd_cons, drule is_file_in_current)
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   616
by (simp add:cf2sfile_mkdir1)
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   617
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   618
lemma cf2sfile_mkdir3:
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   619
  "\<lbrakk>valid (Mkdir p f i # s); is_dir s f'\<rbrakk>
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   620
   \<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = cf2sfile s f'"
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   621
apply (frule vd_cons, drule is_dir_in_current)
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   622
by (simp add:cf2sfile_mkdir1)
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   623
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   624
lemma cf2sfile_mkdir4:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   625
  "valid (Mkdir p f i # s)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   626
  \<Longrightarrow> cf2sfile (Mkdir p f i # s) f = (case (parent f) of
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   627
         Some pf \<Rightarrow> (case (sectxt_of_obj (Mkdir p f i # s) (O_dir f), sectxt_of_obj s (O_dir pf), 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   628
                           get_parentfs_ctxts s pf) of
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   629
                      (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   630
                    | _ \<Rightarrow> None)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   631
       | None \<Rightarrow> None)"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   632
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   633
apply (case_tac f, simp)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   634
apply (clarsimp simp:os_grant.simps)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   635
apply (simp add:sectxt_of_obj_simps)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   636
apply (frule current_proc_has_sec, simp)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   637
apply (frule is_dir_has_sec, simp)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   638
apply (frule get_pfs_secs_prop, simp)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   639
apply (frule is_dir_not_file)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   640
apply (auto simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   641
               get_parentfs_ctxts_simps split:option.splits if_splits 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   642
            dest:not_deleted_init_dir is_dir_in_current not_deleted_init_file is_file_in_current)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   643
done
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   644
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   645
lemma cf2sfile_mkdir:
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   646
  "\<lbrakk>valid (Mkdir p f i # s); f' \<in> current_files (Mkdir p f i # s)\<rbrakk>
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   647
  \<Longrightarrow> cf2sfile (Mkdir p f i # s) f' = (
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   648
    if (f' = f) 
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   649
    then (case (parent f) of
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   650
             Some pf \<Rightarrow> (case (sectxt_of_obj (Mkdir p f i # s) (O_dir f), sectxt_of_obj s (O_dir pf), 
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   651
                 get_parentfs_ctxts s pf) of
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   652
                          (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   653
                        | _ \<Rightarrow> None)
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   654
           | None \<Rightarrow> None)
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   655
     else cf2sfile s f')"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   656
apply (case_tac "f = f'")
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   657
apply (simp add:cf2sfile_mkdir4 split:option.splits)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   658
apply (simp add:cf2sfile_mkdir1 current_files_simps)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   659
done
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   660
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   661
lemma cf2sfile_linkhard1:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   662
  "\<lbrakk>valid (LinkHard p oldf f # s); f' \<in> current_files s\<rbrakk>
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   663
   \<Longrightarrow> cf2sfile (LinkHard p oldf f# s) f' = cf2sfile s f'"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   664
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   665
apply (case_tac "f = f'", simp)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   666
apply (induct f', simp add:sroot_only, simp)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   667
apply (frule parentf_in_current', simp+)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   668
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   669
               get_parentfs_ctxts_simps split:if_splits option.splits)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   670
done
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   671
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   672
lemma cf2sfile_linkhard2:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   673
  "\<lbrakk>valid (LinkHard p oldf f # s); is_file s f'\<rbrakk>
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   674
   \<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f' = cf2sfile s f'"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   675
apply (frule vd_cons, drule is_file_in_current)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   676
by (simp add:cf2sfile_linkhard1)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   677
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   678
lemma cf2sfile_linkhard3:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   679
  "\<lbrakk>valid (LinkHard p oldf f # s); is_dir s f'\<rbrakk>
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   680
   \<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f' = cf2sfile s f'"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   681
apply (frule vd_cons, drule is_dir_in_current)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   682
by (simp add:cf2sfile_linkhard1)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   683
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   684
lemma cf2sfile_linkhard4:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   685
  "valid (LinkHard p oldf f # s)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   686
  \<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f = (case (parent f) of
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   687
         Some pf \<Rightarrow> (case (sectxt_of_obj (LinkHard p oldf f # s) (O_file f), sectxt_of_obj s (O_dir pf), 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   688
                           get_parentfs_ctxts s pf) of
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   689
                      (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   690
                    | _ \<Rightarrow> None)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   691
       | None \<Rightarrow> None)"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   692
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   693
apply (case_tac f, simp)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   694
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   695
               get_parentfs_ctxts_simps)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   696
apply (rule impI, (erule conjE)+)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   697
apply (drule not_deleted_init_file, simp+)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   698
apply (simp add:is_file_in_current)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   699
done
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   700
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   701
lemma cf2sfile_linkhard:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   702
  "\<lbrakk>valid (LinkHard p oldf f # s); f' \<in> current_files (LinkHard p oldf f # s)\<rbrakk>
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   703
  \<Longrightarrow> cf2sfile (LinkHard p oldf f # s) f' = (
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   704
    if (f' = f) 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   705
    then (case (parent f) of
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   706
             Some pf \<Rightarrow> (case (sectxt_of_obj (LinkHard p oldf f # s) (O_file f), sectxt_of_obj s (O_dir pf), 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   707
                 get_parentfs_ctxts s pf) of
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   708
                          (Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   709
                        | _ \<Rightarrow> None)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   710
           | None \<Rightarrow> None)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   711
     else cf2sfile s f')"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   712
apply (case_tac "f = f'")
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   713
apply (simp add:cf2sfile_linkhard4 split:option.splits)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   714
apply (simp add:cf2sfile_linkhard1 current_files_simps)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   715
done
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   716
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   717
lemma cf2sfile_other:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   718
  "\<lbrakk>ff \<in> current_files s;
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   719
    \<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   720
    \<forall> p fd. e \<noteq> CloseFd p fd;
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   721
    \<forall> p f. e \<noteq> UnLink p f;
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   722
    \<forall> p f. e \<noteq> Rmdir p f;
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   723
    \<forall> p f i. e \<noteq> Mkdir p f i;
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   724
    \<forall> p f f'. e \<noteq> LinkHard p f f'; 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   725
    valid (e # s)\<rbrakk> \<Longrightarrow> cf2sfile (e # s) ff = cf2sfile s ff"
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   726
apply (frule vd_cons, frule vt_grant_os)
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   727
apply (induct ff, simp add:sroot_only)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   728
apply (frule parentf_in_current', simp+, case_tac e)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   729
apply (auto simp:current_files_simps is_file_simps is_dir_simps sectxt_of_obj_simps cf2sfile_path 
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   730
           split:if_splits option.splits)                     
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   731
done     
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 4
diff changeset
   732
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   733
lemma cf2sfile_other':
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   734
  "\<lbrakk>valid (e # s); 
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   735
    \<forall> p f flag fd opt. e \<noteq> Open p f flag fd opt;
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   736
    \<forall> p fd. e \<noteq> CloseFd p fd;
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   737
    \<forall> p f. e \<noteq> UnLink p f;
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   738
    \<forall> p f. e \<noteq> Rmdir p f;
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   739
    \<forall> p f i. e \<noteq> Mkdir p f i;
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   740
    \<forall> p f f'. e \<noteq> LinkHard p f f'; 
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   741
    ff \<in> current_files s\<rbrakk> \<Longrightarrow> cf2sfile (e # s) ff = cf2sfile s ff"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   742
by (auto intro!:cf2sfile_other)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   743
  
8
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   744
lemma cf2sfile_unlink:
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   745
  "\<lbrakk>valid (UnLink p f # s); f' \<in> current_files (UnLink p f # s)\<rbrakk>
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   746
   \<Longrightarrow> cf2sfile (UnLink p f # s) f' = cf2sfile s f'"
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   747
apply (frule vd_cons, frule vt_grant_os)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   748
apply (simp add:current_files_simps split:if_splits)
289a30c4cfb7 find bugs in deleted & inum_of_file
chunhan
parents: 7
diff changeset
   749
apply (auto simp:cf2sfile_def sectxt_of_obj_simps get_parentfs_ctxts_simps is_file_simps is_dir_simps
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   750
           split:if_splits option.splits)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   751
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   752
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   753
lemma cf2sfile_rmdir:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   754
  "\<lbrakk>valid (Rmdir p f # s); f' \<in> current_files (Rmdir p f # s)\<rbrakk>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   755
   \<Longrightarrow> cf2sfile (Rmdir p f # s) f' = cf2sfile s f'"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   756
apply (frule vd_cons, frule vt_grant_os)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   757
apply (simp add:current_files_simps split:if_splits)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   758
apply (auto simp:cf2sfile_def sectxt_of_obj_simps get_parentfs_ctxts_simps is_file_simps is_dir_simps
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   759
           split:if_splits option.splits)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   760
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   761
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   762
lemma pfdof_simp5: "\<lbrakk>proc_fd_of_file s f = {(p, fd)}; file_of_proc_fd s p fd = None\<rbrakk> \<Longrightarrow> False"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   763
apply (subgoal_tac "(p, fd) \<in> proc_fd_of_file s f")
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   764
by (simp add:pfdof_simp2, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   765
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   766
lemma pfdof_simp6: "proc_fd_of_file s f = {(p, fd)} \<Longrightarrow> file_of_proc_fd s p fd = Some f"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   767
apply (subgoal_tac "(p, fd) \<in> proc_fd_of_file s f")
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   768
by (simp add:pfdof_simp2, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   769
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   770
lemma cf2sfile_closefd:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   771
  "\<lbrakk>valid (CloseFd p fd # s); f \<in> current_files (CloseFd p fd # s)\<rbrakk>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   772
   \<Longrightarrow> cf2sfile (CloseFd p fd # s) f = cf2sfile s f"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   773
apply (frule vd_cons, frule vt_grant_os)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   774
apply (simp add:current_files_simps split:if_splits option.splits) 
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   775
(* costs too much time, but solved
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   776
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   777
apply (auto simp:cf2sfile_def sectxt_of_obj_simps get_parentfs_ctxts_simps is_file_simps is_dir_simps 
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   778
           split:if_splits option.splits  
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   779
            dest:init_file_dir_conflict pfdof_simp5 pfdof_simp6 file_of_pfd_is_file
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   780
                 not_deleted_init_file not_deleted_init_dir is_file_not_dir is_dir_not_file
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   781
            dest!:current_has_sec')
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   782
done
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   783
*)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   784
sorry
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   785
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   786
lemmas cf2sfile_simps = cf2sfile_open cf2sfile_mkdir cf2sfile_linkhard cf2sfile_other
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   787
  cf2sfile_unlink cf2sfile_rmdir cf2sfile_closefd
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   788
  
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   789
(*********** cfd2sfd simpset *********)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   790
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   791
lemma cfd2sfd_open1:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   792
  "valid (Open p f flags fd opt # s)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   793
   \<Longrightarrow> cfd2sfd (Open p f flags fd opt # s) p fd = 
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   794
     (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
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   795
        (Some sec, Some sf) \<Rightarrow> Some (sec, flags, sf)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   796
      | _ \<Rightarrow> None)"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   797
by (simp add:cfd2sfd_def sectxt_of_obj_simps split:if_splits)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   798
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   799
lemma cfd2sfd_open_some2:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   800
  "\<lbrakk>valid (Open p f flags fd (Some inum) # s); file_of_proc_fd s p' fd' = Some f'\<rbrakk>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   801
  \<Longrightarrow> cfd2sfd (Open p f flags fd (Some inum) # s) p' fd' = cfd2sfd s p' fd'"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   802
apply (frule vd_cons, frule vt_grant_os)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   803
apply (frule proc_fd_in_fds, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   804
apply (frule file_of_proc_fd_in_curf, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   805
apply (case_tac "f = f'", simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   806
apply (simp add:cfd2sfd_def sectxt_of_obj_simps cf2sfile_open_some1)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   807
apply (case_tac "p = p'", simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   808
apply (rule conjI, rule impI, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   809
apply (drule cf2sfile_open_some1, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   810
apply (auto split:option.splits)[1]
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   811
apply simp
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   812
apply (drule cf2sfile_open_some1, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   813
apply (auto split:option.splits)[1]
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   814
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   815
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   816
lemma cfd2sfd_open_none2:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   817
  "\<lbrakk>valid (Open p f flags fd None # s); file_of_proc_fd s p' fd' = Some f'\<rbrakk>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   818
  \<Longrightarrow> cfd2sfd (Open p f flags fd None # s) p' fd' = cfd2sfd s p' fd'"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   819
apply (frule vd_cons, frule vt_grant_os)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   820
apply (frule proc_fd_in_fds, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   821
apply (frule file_of_proc_fd_in_curf, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   822
apply (simp add:cfd2sfd_def sectxt_of_obj_simps cf2sfile_open_none)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   823
apply (case_tac "p = p'", simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   824
apply (rule conjI, rule impI, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   825
apply (drule cf2sfile_open_none)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   826
apply (auto split:option.splits)[1]
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   827
apply simp
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   828
apply (drule cf2sfile_open_none)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   829
apply (auto split:option.splits)[1]
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   830
done
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   831
  
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   832
lemma cfd2sfd_open2:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   833
  "\<lbrakk>valid (Open p f flags fd opt # s); file_of_proc_fd s p' fd' = Some f'\<rbrakk>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   834
  \<Longrightarrow> cfd2sfd (Open p f flags fd opt # s) p' fd' = cfd2sfd s p' fd'"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   835
apply (case_tac opt)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   836
apply (simp add:cfd2sfd_open_none2)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   837
apply (simp add:cfd2sfd_open_some2)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   838
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   839
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   840
lemma cfd2sfd_open:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   841
  "\<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>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   842
   \<Longrightarrow> cfd2sfd (Open p f flags fd opt # s) p' fd' = (if (p' = p \<and> fd' = fd) then 
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   843
     (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
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   844
        (Some sec, Some sf) \<Rightarrow> Some (sec, flags, sf)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   845
      | _ \<Rightarrow> None)         else cfd2sfd s p' fd')"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   846
apply (simp split:if_splits)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   847
apply (simp add:cfd2sfd_open1 split:option.splits)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   848
apply (simp add:cfd2sfd_open2)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   849
apply (rule impI, simp)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   850
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   851
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   852
lemma cfd2sfd_closefd:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   853
  "\<lbrakk>valid (CloseFd p fd # s); file_of_proc_fd (CloseFd p fd # s) p' fd' = Some f\<rbrakk>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   854
   \<Longrightarrow> cfd2sfd (CloseFd p fd # s) p' fd' = cfd2sfd  s p' fd'"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   855
apply (frule vd_cons, frule vt_grant_os)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   856
apply (frule proc_fd_in_fds, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   857
apply (frule file_of_proc_fd_in_curf, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   858
apply (frule cf2sfile_closefd, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   859
apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   860
apply (auto split:option.splits if_splits)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   861
done
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   862
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   863
lemma cfd2sfd_clone:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   864
  "\<lbrakk>valid (Clone p p' fds shms # s); file_of_proc_fd (Clone p p' fds shms # s) p'' fd' = Some f\<rbrakk>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   865
   \<Longrightarrow> cfd2sfd (Clone p p' fds shms # s) p'' fd' = (
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   866
     if (p'' = p') then cfd2sfd s p fd'
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   867
     else cfd2sfd s p'' fd')"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   868
apply (frule vd_cons, frule vt_grant_os)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   869
apply (frule proc_fd_in_fds, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   870
apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   871
apply (frule_tac cf2sfile_other', simp+)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   872
apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   873
apply (case_tac "p'' = p'", simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   874
apply (auto split:option.splits if_splits)[1]
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   875
apply (simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   876
apply (auto split:option.splits if_splits)[1]
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   877
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   878
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   879
lemma cfd2sfd_execve:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   880
  "\<lbrakk>valid (Execve p f fds # s); file_of_proc_fd (Execve p f fds # s) p' fd' = Some f'\<rbrakk>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   881
   \<Longrightarrow> cfd2sfd (Execve p f fds # s) p' fd' = cfd2sfd s p' fd'"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   882
apply (frule vd_cons, frule vt_grant_os)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   883
apply (frule proc_fd_in_fds, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   884
apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   885
apply (frule_tac cf2sfile_other', simp+)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   886
apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   887
apply (case_tac "p' = p", simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   888
apply (auto split:option.splits if_splits)[1]
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   889
apply (simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   890
apply (auto split:option.splits if_splits)[1]
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   891
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   892
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   893
lemma cfd2sfd_kill:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   894
  "\<lbrakk>valid (Kill p p'' # s); file_of_proc_fd (Kill p p'' # s) p' fd' = Some f'\<rbrakk>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   895
   \<Longrightarrow> cfd2sfd (Kill p p'' # s) p' fd' = cfd2sfd s p' fd'"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   896
apply (frule vd_cons, frule vt_grant_os)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   897
apply (frule proc_fd_in_fds, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   898
apply (frule proc_fd_in_procs, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   899
apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   900
apply (frule_tac cf2sfile_other', simp+)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   901
apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   902
apply (auto split:option.splits if_splits)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   903
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   904
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   905
lemma cfd2sfd_exit:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   906
  "\<lbrakk>valid (Exit p # s); file_of_proc_fd (Exit p # s) p' fd' = Some f'\<rbrakk>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   907
   \<Longrightarrow> cfd2sfd (Exit p # s) p' fd' = cfd2sfd s p' fd'"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   908
apply (frule vd_cons, frule vt_grant_os)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   909
apply (frule proc_fd_in_fds, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   910
apply (frule proc_fd_in_procs, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   911
apply (frule file_of_proc_fd_in_curf, simp, simp add:current_files_simps)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   912
apply (frule_tac cf2sfile_other', simp+)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   913
apply (simp add:cfd2sfd_def sectxt_of_obj_simps)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   914
apply (auto split:option.splits if_splits)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   915
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   916
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   917
lemma cfd2sfd_other:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   918
  "\<lbrakk>valid (e # s); file_of_proc_fd (e # s) p' fd' = Some f';
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   919
    \<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   920
    \<forall> p p'' fds shms. e \<noteq> Clone p p'' fds shms\<rbrakk>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   921
   \<Longrightarrow> cfd2sfd (e # s) p' fd' = cfd2sfd s p' fd'"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   922
apply (frule vd_cons, frule vt_grant_os)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   923
apply (frule proc_fd_in_fds, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   924
apply (frule proc_fd_in_procs, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   925
apply (frule file_of_proc_fd_in_curf, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   926
apply (case_tac e)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   927
apply (auto intro!:cfd2sfd_execve cfd2sfd_closefd cfd2sfd_kill cfd2sfd_exit)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   928
apply (auto simp:cfd2sfd_def sectxt_of_obj_simps current_files_simps cf2sfile_simps split:option.splits)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   929
apply (auto dest!:current_has_sec' dest:file_of_proc_fd_in_curf proc_fd_in_fds)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   930
done
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   931
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   932
lemmas cfd2sfd_simps = cfd2sfd_open cfd2sfd_clone cfd2sfd_other
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   933
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   934
(********** cpfd2sfds simpset **********)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   935
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   936
lemma current_filefd_has_flags:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   937
  "\<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"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   938
apply (induct s arbitrary:p)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   939
apply (simp only:flags_of_proc_fd.simps file_of_proc_fd.simps init_filefd_prop4)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   940
apply (frule vd_cons, frule vt_grant_os, case_tac a)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   941
apply (auto split:if_splits option.splits dest:proc_fd_in_fds)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   942
done
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   943
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   944
lemma current_filefd_has_flags':
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   945
  "\<lbrakk>flags_of_proc_fd s p fd = None; valid s\<rbrakk> \<Longrightarrow> file_of_proc_fd s p fd = None"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   946
apply (case_tac "file_of_proc_fd s p fd")
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   947
apply (simp, drule current_filefd_has_flags, simp+)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   948
done
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   949
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   950
lemma current_file_has_sfile':
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   951
  "\<lbrakk>cf2sfile s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   952
by (rule notI, drule current_file_has_sfile, simp+)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   953
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   954
lemma current_filefd_has_sfd:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   955
  "\<lbrakk>file_of_proc_fd s p fd = Some f; valid s\<rbrakk> \<Longrightarrow> \<exists>sfd. cfd2sfd s p fd = Some sfd"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   956
by (auto simp:cfd2sfd_def split:option.splits dest!:current_has_sec' current_file_has_sfile' 
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   957
         dest:file_of_proc_fd_in_curf proc_fd_in_fds current_filefd_has_flags)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   958
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   959
lemma current_filefd_has_sfd':
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   960
  "\<lbrakk>cfd2sfd s p fd = None; valid s\<rbrakk> \<Longrightarrow> file_of_proc_fd s p fd = None"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   961
by (case_tac "file_of_proc_fd s p fd", auto dest:current_filefd_has_sfd)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   962
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   963
lemma cpfd2sfds_open1:
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   964
  "valid (Open p f flags fd opt # s) \<Longrightarrow>
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   965
   cpfd2sfds (Open p f flags fd opt # s) p = (
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   966
    case (cfd2sfd (Open p f flags fd opt # s) p fd) of
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   967
        Some sfd \<Rightarrow> (cpfd2sfds s p) \<union> {sfd}
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   968
      | _ \<Rightarrow> cpfd2sfds s p)"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   969
apply (frule vd_cons, frule vt_grant_os)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   970
apply (split option.splits)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   971
apply (rule conjI, rule impI, drule current_filefd_has_sfd', simp, simp)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   972
apply (rule allI, rule impI)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   973
apply (rule set_eqI, rule iffI)
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   974
apply (case_tac "x = a", simp)
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   975
unfolding cpfd2sfds_def
13
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
   976
apply (erule CollectE, (erule conjE|erule bexE)+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
   977
apply (simp add:proc_file_fds_def split:if_splits)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
   978
apply (erule exE, rule_tac x = fda in exI)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
   979
apply (simp add:cfd2sfd_open2)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
   980
apply (case_tac "x = a", simp add:proc_file_fds_def)
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   981
apply (rule_tac x = fd in exI, simp+)
13
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
   982
apply (erule conjE|erule bexE)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
   983
apply (rule_tac x = fda in bexI)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
   984
apply (simp add:proc_file_fds_def, erule exE)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
   985
apply (simp add:cfd2sfd_open2)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
   986
apply (simp add:proc_file_fds_def)
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   987
done
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   988
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   989
lemma cpfd2sfds_open1':
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   990
  "valid (Open p f flags fd opt # s) \<Longrightarrow>
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   991
   cpfd2sfds (Open p f flags fd opt # s) p = (
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   992
    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
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   993
        (Some sec, Some sf) \<Rightarrow> (cpfd2sfds s p) \<union> {(sec, flags, sf)}
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   994
      | _ \<Rightarrow> cpfd2sfds s p)"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
   995
apply (frule cfd2sfd_open1)
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   996
apply (auto dest:cpfd2sfds_open1 split:option.splits)
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   997
done
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   998
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   999
lemma cpfd2sfds_open2:
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1000
  "\<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'"
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1001
apply (frule vt_grant_os, frule vd_cons)
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1002
unfolding cpfd2sfds_def
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1003
apply (rule set_eqI, rule iffI)
13
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1004
apply (simp add:proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1005
apply (erule exE|erule conjE)+
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1006
apply (simp only:file_of_proc_fd.simps cfd2sfd_open2 split:if_splits)
13
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1007
apply (rule_tac x = fda in exI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1008
apply (simp add:proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1009
apply (erule exE|erule conjE)+
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1010
apply (rule_tac x = fda in exI, simp add:cfd2sfd_open2)
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1011
done
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1012
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
  1013
lemma cpfd2sfds_open:
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
  1014
  "valid (Open p f flags fd opt # s)
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
  1015
   \<Longrightarrow> cpfd2sfds (Open p f flags fd opt # s) = (cpfd2sfds s) (p := (
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
  1016
    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
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
  1017
        (Some sec, Some sf) \<Rightarrow> (cpfd2sfds s p) \<union> {(sec, flags, sf)}
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
  1018
      | _ \<Rightarrow> cpfd2sfds s p))"
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
  1019
apply (rule ext)
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1020
apply (case_tac "x \<noteq> p")
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1021
apply (simp add:cpfd2sfds_open2)
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1022
apply (simp add:cpfd2sfds_open1')
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1023
done
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1024
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1025
lemma cpfd2sfds_execve:
13
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1026
  "valid (Execve p f fds # s) 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1027
   \<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})"
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1028
apply (frule vd_cons, frule vt_grant_os)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1029
apply (rule ext)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1030
apply (rule set_eqI, rule iffI)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1031
unfolding cpfd2sfds_def proc_file_fds_def
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1032
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1033
apply (simp split:if_splits)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1034
apply (frule_tac p' = p and fd' = fd in cfd2sfd_other, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1035
apply (rule_tac x = fd in bexI, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1036
apply (simp add:cpfd2sfds_def proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1037
apply (frule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1038
apply (rule_tac x = fd in exI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1039
apply (simp split:if_splits)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1040
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1041
apply (rule_tac x = fd in exI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1042
apply (frule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1043
apply (simp add:cpfd2sfds_def proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1044
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1045
apply (rule_tac x = fd in exI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1046
apply (frule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1047
done
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1048
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1049
lemma cpfd2sfds_clone:
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1050
  "valid (Clone p p' fds shms # s) 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1051
   \<Longrightarrow> cpfd2sfds (Clone p p' fds shms # 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})"
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1052
apply (frule vd_cons, frule vt_grant_os)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1053
apply (rule ext)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1054
apply (rule set_eqI, rule iffI)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1055
unfolding cpfd2sfds_def proc_file_fds_def
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1056
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1057
apply (simp split:if_splits)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1058
apply (frule_tac p'' = p' and fd' = fd in cfd2sfd_clone, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1059
apply (rule_tac x = fd in bexI, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1060
apply (simp add:cpfd2sfds_def proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1061
apply (frule_tac p'' = x and fd' = fd in cfd2sfd_clone, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1062
apply (rule_tac x = fd in exI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1063
apply (simp split:if_splits)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1064
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1065
apply (rule_tac x = fd in exI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1066
apply (frule_tac p'' = p' and fd' = fd in cfd2sfd_clone, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1067
apply (simp add:cpfd2sfds_def proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1068
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1069
apply (rule_tac x = fd in exI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1070
apply (frule_tac p'' = x and fd' = fd in cfd2sfd_clone, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1071
done
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1072
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1073
lemma cpfd2sfds_other:
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1074
  "\<lbrakk>valid (e # s);
13
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1075
    \<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1076
    \<forall> p f fds. e \<noteq> Execve p f fds;
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1077
    \<forall> p p'. e \<noteq> Kill p p';
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1078
    \<forall> p. e \<noteq> Exit p;
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1079
    \<forall> p fd. e \<noteq> CloseFd p fd;
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1080
    \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms\<rbrakk> \<Longrightarrow> cpfd2sfds (e # s) = cpfd2sfds s"
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1081
apply (frule vd_cons, frule vt_grant_os)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1082
apply (rule ext)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1083
unfolding cpfd2sfds_def proc_file_fds_def
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1084
apply (case_tac e)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1085
using cfd2sfd_other
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1086
by auto
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1087
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1088
lemma cpfd2sfds_kill:
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1089
  "valid (Kill p p' # s) \<Longrightarrow> cpfd2sfds (Kill p p' # s) = (cpfd2sfds s) (p' := {})"
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1090
apply (frule vd_cons, frule vt_grant_os)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1091
apply (rule ext, rule set_eqI)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1092
unfolding cpfd2sfds_def proc_file_fds_def
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1093
apply (rule iffI)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1094
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1095
apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1096
apply (rule_tac x = fd in exI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1097
apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1098
apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1099
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1100
apply (rule_tac x = fd in exI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1101
apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1102
done
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1103
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1104
lemma cpfd2sfds_exit:
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1105
  "valid (Exit p # s) \<Longrightarrow> cpfd2sfds (Exit p # s) = (cpfd2sfds s) (p := {})"
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1106
apply (frule vd_cons, frule vt_grant_os)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1107
apply (rule ext, rule set_eqI)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1108
unfolding cpfd2sfds_def proc_file_fds_def
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1109
apply (rule iffI)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1110
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1111
apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1112
apply (rule_tac x = fd in exI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1113
apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1114
apply (simp split:if_splits add: cpfd2sfds_def proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1115
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1116
apply (rule_tac x = fd in exI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1117
apply (drule_tac p' = x and fd' = fd in cfd2sfd_other, simp+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1118
done
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1119
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1120
lemma cpfd2sfds_closefd:
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1121
  "valid (CloseFd p fd # s) \<Longrightarrow> cpfd2sfds (CloseFd p fd # s) = (cpfd2sfds s) (p := 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1122
     if (fd \<in> proc_file_fds s p)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1123
     then (case (cfd2sfd s p fd) of 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1124
             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)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1125
                          then cpfd2sfds s p else cpfd2sfds s p - {sfd})
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1126
           | _        \<Rightarrow> cpfd2sfds s p)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1127
     else cpfd2sfds s p)"
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1128
apply (frule vd_cons)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1129
apply (rule ext, rule set_eqI, rule iffI)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1130
unfolding cpfd2sfds_def proc_file_fds_def
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1131
apply (erule CollectE| erule bexE| erule conjE| erule exE| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1132
apply (simp split:if_splits)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1133
apply (rule conjI, rule impI, rule conjI, rule impI, erule exE)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1134
apply (frule_tac p = p and fd = fd in current_filefd_has_sfd, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1135
apply (erule exE, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1136
apply (rule conjI, rule impI, (erule exE|erule conjE)+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1137
apply (rule_tac x = fda in exI, simp, simp add:cfd2sfd_closefd)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1138
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1139
apply (rule impI, rule conjI)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1140
apply (rule_tac x = fda in exI, simp, simp add:cfd2sfd_closefd)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1141
apply (rule notI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1142
apply (erule_tac x = fda in allE, simp add:cfd2sfd_closefd)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1143
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1144
apply (rule impI, simp add:cpfd2sfds_def proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1145
apply (erule exE, rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1146
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1147
apply (rule impI| rule conjI)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1148
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1149
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1150
apply (rule impI, simp add:cpfd2sfds_def proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1151
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
  1152
13
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1153
apply (simp split:if_splits)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1154
apply (frule_tac p = p and fd = fd in current_filefd_has_sfd, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1155
apply (erule exE, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1156
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")
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1157
apply simp
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1158
apply (case_tac "xa = sfd")
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1159
apply (erule exE|erule conjE)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1160
apply (rule_tac x = fd' in exI, simp add:cfd2sfd_closefd)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1161
apply (erule exE|erule conjE)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1162
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1163
apply (rule notI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1164
apply (simp, (erule exE|erule conjE)+)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1165
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1166
apply (rule notI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1167
apply (erule exE|erule conjE)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1168
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1169
apply (rule notI, simp)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1170
apply (simp add:cpfd2sfds_def proc_file_fds_def)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1171
apply (erule exE|erule conjE)+
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1172
apply (rule_tac x = fda in exI, simp add:cfd2sfd_closefd)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1173
done
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1174
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1175
lemmas cpfd2sfds_simps = cpfd2sfds_open cpfd2sfds_execve cpfd2sfds_clone cpfd2sfds_kill cpfd2sfds_exit
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1176
  cpfd2sfds_closefd cpfd2sfds_other
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1177
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1178
(********* ch2sshm simpset ********)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1179
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1180
lemma ch2sshm_createshm:
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1181
  "valid (CreateShM p h # s) 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1182
   \<Longrightarrow> ch2sshm (CreateShM p h # s) = (ch2sshm s) (h := 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1183
     (case (sectxt_of_obj (CreateShM p h # s) (O_shm h)) of
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1184
                    Some sec \<Rightarrow> 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1185
 Some (if (\<not> deleted (O_shm h) s \<and> h \<in> init_shms) then Init h else Created, sec)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1186
                  | _ \<Rightarrow> None))"
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1187
apply (frule vd_cons, frule vt_grant_os)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1188
apply (auto simp:ch2sshm_def sectxt_of_obj_simps dest!:current_has_sec' split:option.splits if_splits)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1189
done
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1190
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1191
lemma ch2sshm_other:
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1192
  "\<lbrakk>valid (e # s); 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1193
    \<forall> p h. e \<noteq> CreateShM p h; 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1194
    h' \<in> current_shms (e # s)\<rbrakk> \<Longrightarrow> ch2sshm (e # s) h' = ch2sshm s h'"
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1195
apply (frule vd_cons, frule vt_grant_os)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1196
apply (case_tac e)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1197
apply (auto simp:ch2sshm_def sectxt_of_obj_simps dest!:current_has_sec' split:option.splits if_splits)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1198
done
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1199
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1200
lemmas ch2sshm_simps = ch2sshm_createshm ch2sshm_other
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1201
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1202
lemma current_shm_has_sh:
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1203
  "\<lbrakk>h \<in> current_shms s; valid s\<rbrakk> \<Longrightarrow> \<exists> sh. ch2sshm s h = Some sh"
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1204
by (auto simp:ch2sshm_def split:option.splits dest!:current_has_sec')
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1205
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1206
lemma current_shm_has_sh':
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1207
  "\<lbrakk>ch2sshm s h = None; valid s\<rbrakk> \<Longrightarrow> h \<notin> current_shms s"
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1208
by (auto dest:current_shm_has_sh)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1209
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1210
(********** cph2spshs simpset **********)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1211
14
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1212
lemma cph2spshs_attach:
13
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1213
  "valid (Attach p h flag # s) \<Longrightarrow> 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1214
   cph2spshs (Attach p h flag # s) = (cph2spshs s) (p := 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1215
     (case (ch2sshm s h) of 
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1216
        Some sh \<Rightarrow> cph2spshs s p \<union> {(sh, flag)}
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1217
      | _       \<Rightarrow> cph2spshs s p) )"
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1218
apply (frule vd_cons, frule vt_grant_os, rule ext)
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1219
using ch2sshm_other[where e = "Attach p h flag" and s = s]
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1220
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
14
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1221
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 simp:cph2spshs_def)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1222
apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1223
apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1224
apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1225
apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1226
apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1227
apply (case_tac "ha = h", simp, frule procs_of_shm_prop1, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1228
apply (rule_tac x = ha in exI, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1229
apply (rule_tac x = ha in exI, simp, drule procs_of_shm_prop1, simp, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1230
apply (rule_tac x = ha in exI, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1231
apply (frule procs_of_shm_prop1, simp, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1232
apply (rule impI, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1233
done
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1234
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1235
lemma cph2spshs_detach: "valid (Detach p h # s) \<Longrightarrow> 
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1236
  cph2spshs (Detach p h # s) = (cph2spshs s) (p := 
15
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1237
    (case (ch2sshm s h, flag_of_proc_shm s p h) of 
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1238
       (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)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1239
                  then cph2spshs s p else cph2spshs s p - {(sh, flag)}
14
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1240
     | _       \<Rightarrow> cph2spshs s p)             )"
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1241
apply (frule vd_cons, frule vt_grant_os, rule ext)
15
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1242
apply (case_tac "x = p")  defer
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1243
using ch2sshm_other[where e = "Detach p h" and s = s] 
14
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1244
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
15
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1245
dest!:current_shm_has_sh' procs_of_shm_prop4' dest:procs_of_shm_prop1 procs_of_shm_prop3 simp:cph2spshs_def) [1]
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1246
apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp, simp)
14
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1247
apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp, simp)
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1248
apply (rule impI, simp)
15
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1249
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1250
apply (clarsimp)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1251
apply (frule current_shm_has_sh, simp, erule exE)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1252
apply (frule procs_of_shm_prop4, simp, simp)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1253
apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1254
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1255
apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1256
apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1257
apply (case_tac "ha = h", simp)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1258
apply (rule_tac x = sha in exI, rule_tac x = flaga in exI, rule_tac x = ha in exI, simp)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1259
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1260
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1261
apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1262
apply (case_tac "ha = h", simp)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1263
apply (rule_tac x = h' in exI, simp)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1264
apply (frule_tac flag = flag and flag' = flaga in procs_of_shm_prop3, simp+)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1265
apply (frule_tac flag = flaga in procs_of_shm_prop4, simp+)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1266
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1267
apply (frule_tac h = h' in procs_of_shm_prop1, simp, simp)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1268
apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1269
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1270
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1271
apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1272
apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1273
apply (erule CollectE | erule exE| erule conjE| rule DiffI |rule CollectI)+
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1274
apply (simp split:if_splits)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1275
apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1276
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1277
apply (rule notI, simp split:if_splits)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1278
apply (erule_tac x = ha in allE, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1279
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1280
apply (erule CollectE | erule exE| erule conjE| erule DiffE |rule CollectI)+
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1281
apply (simp split:if_splits)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1282
apply (erule_tac x = ha in allE, simp, rule_tac x = ha in exI, simp)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1283
apply (case_tac "ha = h", simp add:procs_of_shm_prop3, frule_tac h = ha in procs_of_shm_prop1, simp+)
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1284
using ch2sshm_other[where e = "Detach p h" and s = s] apply (simp add:procs_of_shm_prop1)
14
cc1e46225a81 shm attached at most once
chunhan
parents: 13
diff changeset
  1285
done
13
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1286
16
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1287
lemma cph2spshs_deleteshm:
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1288
  "valid (DeleteShM p h # s) \<Longrightarrow>     
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1289
   cph2spshs (DeleteShM p h # s) = (\<lambda> p'. 
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1290
    (case (ch2sshm s h, flag_of_proc_shm s p' h) of 
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1291
       (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)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1292
                  then cph2spshs s p' else cph2spshs s p' - {(sh, flag)}
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1293
     | _       \<Rightarrow> cph2spshs s p')   )"
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1294
apply (frule vd_cons, frule vt_grant_os, rule ext)
15
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1295
16
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1296
apply (clarsimp)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1297
apply (frule current_shm_has_sh, simp, erule exE, simp, simp split:option.splits)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1298
apply (rule conjI, rule impI)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1299
using ch2sshm_other[where e = "DeleteShM p h" and s = s] 
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1300
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1301
dest!:current_shm_has_sh' procs_of_shm_prop4' dest:procs_of_shm_prop1 procs_of_shm_prop3 simp:cph2spshs_def) [1]
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1302
apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp+)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1303
apply (rule_tac x = ha in exI, simp)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1304
apply (case_tac "ha = h", simp+, frule_tac h = ha in procs_of_shm_prop1, simp+)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1305
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1306
apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1307
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1308
apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1309
apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1310
apply (case_tac "ha = h", simp)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1311
apply (rule_tac x = sha in exI, rule_tac x = flag in exI, rule_tac x = ha in exI, simp)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1312
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1313
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1314
apply (erule CollectE | erule exE| erule conjE| rule CollectI)+
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1315
apply (case_tac "ha = h", simp)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1316
apply (rule_tac x = h' in exI, simp)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1317
apply (frule_tac flag = flag in procs_of_shm_prop4, simp+)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1318
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1319
apply (frule_tac h = h' in procs_of_shm_prop1, simp, simp)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1320
apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1321
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1322
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1323
apply (rule allI | rule conjI| erule conjE | erule exE | rule impI)+
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1324
apply (simp only:cph2spshs_def, rule set_eqI, rule iffI)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1325
apply (erule CollectE | erule exE| erule conjE| rule DiffI |rule CollectI)+
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1326
apply (simp split:if_splits)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1327
apply (rule_tac x = ha in exI, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1328
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1329
apply (rule notI, simp split:if_splits)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1330
apply (erule_tac x = ha in allE, simp, frule_tac h = ha in procs_of_shm_prop1, simp+)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1331
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1332
apply (erule CollectE | erule exE| erule conjE| erule DiffE |rule CollectI)+
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1333
apply (simp split:if_splits)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1334
apply (erule_tac x = ha in allE, simp, rule_tac x = ha in exI, simp)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1335
apply (case_tac "ha = h", simp add:procs_of_shm_prop4)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1336
apply (frule_tac h = ha in procs_of_shm_prop1, simp+)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1337
using ch2sshm_other[where e = "DeleteShM p h" and s = s] apply (simp add:procs_of_shm_prop1)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1338
done
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1339
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1340
lemma flag_of_proc_shm_prop1:
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1341
  "\<lbrakk>flag_of_proc_shm s p h = Some flag; valid s\<rbrakk> \<Longrightarrow> (p, flag) \<in> procs_of_shm s h"
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1342
apply (induct s arbitrary:p)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1343
apply (simp add:init_shmflag_has_proc)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1344
apply (frule vt_grant_os, frule vd_cons, case_tac a, auto split:if_splits option.splits)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1345
done
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1346
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1347
lemma cph2spshs_clone:
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1348
  "valid (Clone p p' fds shms # s) \<Longrightarrow> 
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1349
   cph2spshs (Clone p p' fds shms # s) = (cph2spshs s) (p' := 
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1350
     {(sh, flag) | h sh flag. h \<in> shms \<and> ch2sshm s h = Some sh \<and> flag_of_proc_shm s p h = Some flag} )"
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1351
apply (frule vd_cons, frule vt_grant_os, rule ext)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1352
using ch2sshm_other[where e = "Clone p p' fds shms" and s = s]
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1353
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1354
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1355
apply (erule_tac x = h in allE, frule procs_of_shm_prop1, simp, simp add:procs_of_shm_prop4)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1356
apply (rule_tac x = h in exI, simp, frule flag_of_proc_shm_prop1, simp+, simp add:procs_of_shm_prop1)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1357
apply (rule_tac x = h in exI, simp, frule procs_of_shm_prop1, simp+)+
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1358
done
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1359
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1360
lemma cph2spshs_execve:
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1361
  "valid (Execve p f fds # s) \<Longrightarrow>
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1362
  cph2spshs (Execve p f fds # s) = (cph2spshs s) (p := {})"
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1363
apply (frule vd_cons, frule vt_grant_os, rule ext)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1364
using ch2sshm_other[where e = "Execve p f fds" and s = s]
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1365
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1366
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1367
apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1368
done
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1369
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1370
lemma cph2spshs_kill:
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1371
  "valid (Kill p p' # s) \<Longrightarrow> cph2spshs (Kill p p' # s) = (cph2spshs s) (p' := {})"
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1372
apply (frule vd_cons, frule vt_grant_os, rule ext)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1373
using ch2sshm_other[where e = "Kill p p'" and s = s]
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1374
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1375
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1376
apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1377
done
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1378
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1379
lemma cph2spshs_exit:
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1380
  "valid (Exit p # s) \<Longrightarrow> cph2spshs (Exit p # s) = (cph2spshs s) (p := {})"
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1381
apply (frule vd_cons, frule vt_grant_os, rule ext)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1382
using ch2sshm_other[where e = "Exit p" and s = s]
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1383
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1384
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1385
apply (rule_tac x = h in exI, simp add:procs_of_shm_prop1)+
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1386
done
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1387
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1388
lemma cph2spshs_createshm:
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1389
  "valid (CreateShM p h # s) \<Longrightarrow> cph2spshs (CreateShM p h # s) = cph2spshs s"
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1390
apply (frule vt_grant_os, frule vd_cons, rule ext)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1391
apply (auto simp:cph2spshs_def)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1392
using ch2sshm_createshm
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1393
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1394
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 simp:cph2spshs_def)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1395
apply (rule_tac x = ha in exI, auto simp:procs_of_shm_prop1)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1396
done
15
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 14
diff changeset
  1397
13
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1398
lemma cph2spshs_other:
16
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1399
  "\<lbrakk>valid (e # s); 
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1400
    \<forall> p h flag. e \<noteq> Attach p h flag;
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1401
    \<forall> p h. e \<noteq> Detach p h;
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1402
    \<forall> p h. e \<noteq> DeleteShM p h;
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1403
    \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms;
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1404
    \<forall> p f fds. e \<noteq> Execve p f fds;
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1405
    \<forall> p p'. e \<noteq> Kill p p';
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1406
    \<forall> p. e \<noteq> Exit p\<rbrakk> \<Longrightarrow> cph2spshs (e # s) = cph2spshs s"
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1407
apply (frule vt_grant_os, frule vd_cons, rule ext)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1408
unfolding cph2spshs_def 
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1409
apply (rule set_eqI, rule iffI)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1410
apply (erule CollectE | erule conjE| erule exE| rule conjI| rule impI| rule CollectI)+
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1411
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)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1412
apply (case_tac e)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1413
using ch2sshm_other[where e = e and s = s] 
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1414
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1415
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1416
 simp:ch2sshm_createshm)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1417
apply (rule_tac x = h in exI, case_tac e)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1418
using ch2sshm_other[where e = e and s = s] 
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1419
apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1420
dest!:current_shm_has_sh' dest: procs_of_shm_prop1 procs_of_shm_prop2 procs_of_shm_prop3 
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1421
 simp:ch2sshm_createshm)
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1422
done
13
7b5e9fbeaf93 co2sobj simpset
chunhan
parents: 12
diff changeset
  1423
16
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1424
lemmas cph2spshs_simps = cph2spshs_attach cph2spshs_detach cph2spshs_deleteshm cph2spshs_clone cph2spshs_execve
c8b7c24f1db6 done with cph2spshs simpset
chunhan
parents: 15
diff changeset
  1425
  cph2spshs_exit cph2spshs_kill cph2spshs_other
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1426
10
ac66d8ba86d9 fix bugs and proofs
chunhan
parents: 8
diff changeset
  1427
(******** cp2sproc simpset *********)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1428
17
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1429
lemma cp2sproc_clone:
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1430
  "valid (Clone p p' fds shms # s) \<Longrightarrow> cp2sproc (Clone p p' fds shms # s) = (cp2sproc s) (p' := 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1431
     case (sectxt_of_obj s (O_proc p)) of 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1432
       Some sec \<Rightarrow> Some (Created, sec, 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1433
  {sfd. \<exists> fd \<in> fds. \<exists> f. file_of_proc_fd s p fd = Some f \<and> cfd2sfd s p fd = Some sfd},
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1434
  {(sh, flag) | h sh flag. h \<in> shms \<and> ch2sshm s h = Some sh \<and> flag_of_proc_shm s p h = Some flag})
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1435
     | _        \<Rightarrow> None)"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1436
apply (frule cph2spshs_clone, frule cpfd2sfds_clone)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1437
apply (frule vd_cons, frule vt_grant_os, simp only:os_grant.simps)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1438
apply ((erule exE| erule conjE)+, frule not_init_intro_proc, simp, rule ext, case_tac "x = p'", simp)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1439
apply (auto simp:cp2sproc_def sectxt_of_obj_simps dest!:current_has_sec' 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1440
            dest:current_proc_has_sec split:option.splits if_splits)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1441
done
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1442
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1443
lemma cp2sproc_other:
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1444
  "\<lbrakk>valid (e # s); 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1445
    \<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1446
    \<forall> p fd. e \<noteq> CloseFd p fd;
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1447
    \<forall> p h flag. e \<noteq> Attach p h flag;
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1448
    \<forall> p h. e \<noteq> Detach p h;
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1449
    \<forall> p h. e \<noteq> DeleteShM p h;
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1450
    \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms;
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1451
    \<forall> p f fds. e \<noteq> Execve p f fds;
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1452
    \<forall> p p'. e \<noteq> Kill p p';
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1453
    \<forall> p. e \<noteq> Exit p\<rbrakk> \<Longrightarrow> cp2sproc (e # s) = cp2sproc s"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1454
apply (frule cph2spshs_other, simp+, frule cpfd2sfds_other, simp+)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1455
apply (frule vt_grant_os, frule vd_cons, rule ext, case_tac e, simp_all) 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1456
apply (auto simp:cp2sproc_def sectxt_of_obj_simps dest!:current_has_sec' 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1457
            dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1458
done
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1459
17
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1460
lemma cp2sproc_open:
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1461
  "valid (Open p f flags fd opt # s) \<Longrightarrow> 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1462
   cp2sproc (Open p f flags fd opt # s) = (cp2sproc s) (p :=      
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1463
     case (sectxt_of_obj s (O_proc p), sectxt_of_obj (Open p f flags fd opt # s) (O_fd p fd), 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1464
           cf2sfile (Open p f flags fd opt # s) f) of 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1465
       (Some sec, Some fdsec, Some sf) \<Rightarrow> Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs) 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1466
                                               then Init p else Created, sec, 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1467
                                                (cpfd2sfds s p) \<union> {(fdsec, flags, sf)}, cph2spshs s p)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1468
     | _        \<Rightarrow> None)"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1469
apply (frule cph2spshs_other, simp, simp, simp, simp, simp, simp, simp)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1470
apply (frule cpfd2sfds_open, frule vt_grant_os, frule vd_cons, rule ext)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1471
apply (case_tac "x = p") 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1472
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1473
           dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1474
            dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1475
done
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1476
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1477
lemma cp2sproc_closefd:
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1478
  "valid (CloseFd p fd # s) \<Longrightarrow> 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1479
   cp2sproc (CloseFd p fd # s) = (cp2sproc s) (p := 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1480
     if (fd \<in> proc_file_fds s p)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1481
     then (case (cfd2sfd s p fd) of 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1482
             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)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1483
                          then cp2sproc s p 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1484
                          else (case (sectxt_of_obj s (O_proc p)) of
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1485
                                  Some sec \<Rightarrow> Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs) 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1486
                                                    then Init p else Created, 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1487
                                                    sec, cpfd2sfds s p - {sfd}, cph2spshs s p)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1488
                                | _        \<Rightarrow> None))
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1489
           | _        \<Rightarrow> cp2sproc s p)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1490
     else cp2sproc s p)"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1491
apply (frule cpfd2sfds_closefd, frule cph2spshs_other, simp, simp, simp, simp, simp, simp, simp)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1492
apply (frule vt_grant_os, frule vd_cons, rule ext)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1493
apply (case_tac "x = p") 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1494
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1495
           dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1496
            dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1497
done
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1498
17
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1499
lemma cp2sproc_attach:
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1500
  "valid (Attach p h flag # s) \<Longrightarrow> 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1501
   cp2sproc (Attach p h flag # s) = (cp2sproc s) (p := 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1502
     (case (ch2sshm s h) of 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1503
        Some sh \<Rightarrow> (case (sectxt_of_obj s (O_proc p)) of
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1504
                      Some sec \<Rightarrow> Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs) 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1505
                                        then Init p else Created, 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1506
                                        sec, cpfd2sfds s p, cph2spshs s p \<union> {(sh, flag)})
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1507
                    | _        \<Rightarrow> None)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1508
      | _       \<Rightarrow> cp2sproc s p) )"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1509
apply (frule cph2spshs_attach, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1510
apply (frule vt_grant_os, frule vd_cons, rule ext)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1511
apply (case_tac "x = p") 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1512
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1513
           dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1514
            dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1515
done
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1516
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1517
lemma cp2sproc_detach:
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1518
  "valid (Detach p h # s) \<Longrightarrow> 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1519
   cp2sproc (Detach p h # s) = (cp2sproc s) (p := 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1520
    (case (ch2sshm s h, flag_of_proc_shm s p h) of 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1521
       (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)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1522
                              then cp2sproc s p 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1523
                              else (case (sectxt_of_obj s (O_proc p)) of
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1524
                                      Some sec \<Rightarrow> Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs) 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1525
                                                       then Init p else Created, sec,
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1526
                                                       cpfd2sfds s p, cph2spshs s p - {(sh, flag)})
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1527
                                    | None     \<Rightarrow> None)                   
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1528
     | _       \<Rightarrow> cp2sproc s p)             )"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1529
apply (frule cph2spshs_detach, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1530
apply (frule vt_grant_os, frule vd_cons, rule ext)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1531
apply (case_tac "x = p") 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1532
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1533
           dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1534
            dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1535
done
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1536
17
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1537
lemma cp2sproc_deleteshm:
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1538
  "valid (DeleteShM p h # s) \<Longrightarrow> 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1539
   cp2sproc (DeleteShM p h # s) = (\<lambda> p'. 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1540
    (case (ch2sshm s h, flag_of_proc_shm s p' h) of 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1541
       (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)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1542
                              then cp2sproc s p' 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1543
                              else (case (sectxt_of_obj s (O_proc p')) of
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1544
                                      Some sec \<Rightarrow> Some (if (\<not> deleted (O_proc p') s \<and> p' \<in> init_procs) 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1545
                                                       then Init p' else Created, sec,
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1546
                                                       cpfd2sfds s p', cph2spshs s p' - {(sh, flag)})
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1547
                                    | None     \<Rightarrow> None)                   
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1548
     | _       \<Rightarrow> cp2sproc s p')             )"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1549
apply (frule cph2spshs_deleteshm, frule cpfd2sfds_other, simp, simp, simp, simp, simp, simp)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1550
apply (frule vt_grant_os, frule vd_cons, rule ext)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1551
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1552
           dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1553
            dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1554
done
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1555
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1556
lemma cp2sproc_execve:
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1557
  "valid (Execve p f fds # s) \<Longrightarrow> 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1558
   cp2sproc (Execve p f fds # s) = (cp2sproc s) (p := 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1559
     (case (sectxt_of_obj (Execve p f fds # s) (O_proc p)) of
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1560
        Some sec \<Rightarrow> Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs) then Init p else Created, sec, 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1561
                         {sfd. \<exists> fd \<in> fds. \<exists> f. file_of_proc_fd s p fd = Some f \<and> cfd2sfd s p fd = Some sfd}, {})
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1562
      | _        \<Rightarrow> None)                        )"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1563
apply (frule cph2spshs_execve, frule cpfd2sfds_execve)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1564
apply (frule vt_grant_os, frule vd_cons, rule ext)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1565
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1566
           dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1567
            dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1568
done
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1569
17
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1570
lemma cp2sproc_kill:
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1571
  "\<lbrakk>valid (Kill p p' # s); p'' \<noteq> p'\<rbrakk> \<Longrightarrow> 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1572
   cp2sproc (Kill p p' # s) p'' = (cp2sproc s) p''"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1573
apply (frule cph2spshs_kill, frule cpfd2sfds_kill)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1574
apply (frule vt_grant_os, frule vd_cons)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1575
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1576
           dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1577
            dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1578
done
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1579
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1580
lemma cp2sproc_kill':
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1581
  "\<lbrakk>valid (Kill p p' # s); p'' \<in> current_procs (Kill p p' # s)\<rbrakk> \<Longrightarrow> 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1582
   cp2sproc (Kill p p' # s) p'' = (cp2sproc s) p''"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1583
by (drule_tac p'' = p'' in cp2sproc_kill, simp+)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1584
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1585
lemma cp2sproc_exit:
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1586
  "\<lbrakk>valid (Exit p # s); p' \<noteq> p\<rbrakk> \<Longrightarrow> 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1587
   cp2sproc (Exit p # s) p' = (cp2sproc s) p'"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1588
apply (frule cph2spshs_exit, frule cpfd2sfds_exit)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1589
apply (frule vt_grant_os, frule vd_cons)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1590
apply (auto simp:cp2sproc_def sectxt_of_obj_simps current_files_simps 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1591
           dest!:current_has_sec' dest!:current_file_has_sfile' dest:is_file_in_current is_dir_in_current
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1592
            dest:current_proc_has_sec not_deleted_init_proc split:option.splits if_splits)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1593
done
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1594
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1595
lemma cp2sproc_exit':
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1596
  "\<lbrakk>valid (Exit p # s); p' \<in> current_procs (Exit p # s)\<rbrakk> \<Longrightarrow> 
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1597
   cp2sproc (Exit p # s) p' = (cp2sproc s) p'"
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1598
by (drule_tac p'= p' in cp2sproc_exit, simp+)
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1599
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1600
lemmas cp2sproc_simps = cp2sproc_open cp2sproc_closefd cp2sproc_attach cp2sproc_detach cp2sproc_deleteshm
570f90f175ee updated
chunhan
parents: 16
diff changeset
  1601
  cp2sproc_clone cp2sproc_execve cp2sproc_kill cp2sproc_exit cp2sproc_other 
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1602
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1603
end
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1604
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1605
(*<*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1606
end
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
  1607
(*>*)