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