Static.thy
author chunhan
Thu, 05 Sep 2013 13:23:03 +0800
changeset 42 021672ec28f5
parent 39 13bba99ca090
child 43 137358bd4921
permissions -rwxr-xr-x
update
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     1
theory Static
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     2
imports Static_type OS_type_def Flask_type Flask
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     3
begin
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     4
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     5
locale tainting_s = tainting 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     6
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     7
begin
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     8
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     9
(*
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    10
definition init_sectxt_of_file:: "t_file \<Rightarrow> security_context_t option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    11
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    12
  "init_sectxt_of_file f \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    13
     if (is_init_file f) then init_sectxt_of_obj (O_file f)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    14
     else if (is_init_dir f) then init_sectxt_of_obj (O_dir f)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    15
     else None"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    16
*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    17
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    18
definition sroot :: "t_sfile"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    19
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    20
  "sroot \<equiv> (Init [], sec_of_root, None, {})"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    21
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    22
definition init_cf2sfile :: "t_file \<Rightarrow> t_sfile option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    23
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    24
  "init_cf2sfile f \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    25
     case (parent f) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    26
       None \<Rightarrow> Some sroot
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    27
     | Some pf \<Rightarrow> if (is_init_file f) then 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    28
 (case (init_sectxt_of_obj (O_file f), init_sectxt_of_obj (O_dir pf), get_parentfs_ctxts [] pf) of
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    29
    (Some sec, Some psec, Some aseclist) \<Rightarrow> Some (Init f, sec, Some psec, set aseclist)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    30
  | _ \<Rightarrow> None)    else 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    31
 (case (init_sectxt_of_obj (O_dir f), init_sectxt_of_obj (O_dir pf), get_parentfs_ctxts [] pf) of
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    32
    (Some sec, Some psec, Some aseclist) \<Rightarrow> Some (Init f, sec, Some psec, set aseclist)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    33
  | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    34
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 2
diff changeset
    35
definition init_cf2sfiles :: "t_file \<Rightarrow> t_sfile set"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    36
where
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 2
diff changeset
    37
  "init_cf2sfiles f \<equiv> {sf. \<exists>f' \<in> init_same_inode_files f. init_cf2sfile f' = Some sf}"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    38
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    39
definition init_cfd2sfd :: "t_process \<Rightarrow> t_fd \<Rightarrow> (security_context_t \<times> t_open_flags \<times> t_sfile) option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    40
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    41
  "init_cfd2sfd p fd = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    42
     (case (init_file_of_proc_fd p fd, init_oflags_of_proc_fd p fd, init_sectxt_of_obj (O_fd p fd)) of
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    43
        (Some f, Some flags, Some sec) \<Rightarrow> (case (init_cf2sfile f) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    44
                                             Some sf \<Rightarrow> Some (sec, flags, sf)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    45
                                           | _ \<Rightarrow> None)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    46
      | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    47
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    48
definition init_cfds2sfds :: "t_process \<Rightarrow> (security_context_t \<times> t_open_flags \<times> t_sfile) set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    49
where
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
    50
  "init_cfds2sfds p \<equiv> { sfd. \<exists> fd \<in> init_proc_file_fds p. init_cfd2sfd p fd = Some sfd}"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    51
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    52
definition init_ch2sshm :: "t_shm \<Rightarrow> t_sshm option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    53
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    54
  "init_ch2sshm h \<equiv> (case (init_sectxt_of_obj (O_shm h)) of
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    55
                       Some sec \<Rightarrow> Some (Init h, sec)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    56
                     | _        \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    57
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    58
definition init_cph2spshs 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    59
  :: "t_process \<Rightarrow> (t_sshm \<times> t_shm_attach_flag) set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    60
where
15
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 12
diff changeset
    61
  "init_cph2spshs p \<equiv> {(sh, flag)| sh flag h. (p, flag) \<in> init_procs_of_shm h \<and> 
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    62
                                             init_ch2sshm h = Some sh}"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    63
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    64
definition init_cp2sproc :: "t_process \<Rightarrow> t_sproc option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    65
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    66
  "init_cp2sproc p \<equiv> (case (init_sectxt_of_obj (O_proc p)) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    67
       Some sec \<Rightarrow> Some (Init p, sec, (init_cfds2sfds p), (init_cph2spshs p))
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    68
     | None     \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    69
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    70
(* acient files of init-file 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    71
definition init_o2s_afs :: "t_file \<Rightarrow> security_context_t set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    72
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    73
  "init_o2s_afs f \<equiv> {sec. \<exists> f'. f' \<preceq> f \<and> init_sectxt_of_obj (O_dir f') = Some sec}" *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    74
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    75
definition init_cm2smsg :: "t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsg option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    76
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    77
  "init_cm2smsg q m \<equiv> (case (init_sectxt_of_obj (O_msg q m)) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    78
                         Some sec \<Rightarrow> Some (Init m, sec, (O_msg q m) \<in> seeds)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    79
                       | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    80
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    81
fun init_cqm2sms :: "t_msgq \<Rightarrow> t_msg list \<Rightarrow> (t_smsg list) option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    82
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    83
  "init_cqm2sms q []     = Some []"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    84
| "init_cqm2sms q (m#ms) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    85
     (case (init_cqm2sms q ms, init_cm2smsg q m) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    86
        (Some sms, Some sm) \<Rightarrow> Some (sm # sms)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    87
      | _        \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    88
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    89
definition init_cq2smsgq :: "t_msgq \<Rightarrow> t_smsgq option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    90
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    91
  "init_cq2smsgq q \<equiv> (case (init_sectxt_of_obj (O_msgq q), init_cqm2sms q (init_msgs_of_queue q)) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    92
       (Some sec, Some sms) \<Rightarrow> Some (Init q, sec, sms)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    93
     | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    94
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    95
fun init_obj2sobj :: "t_object \<Rightarrow> t_sobject option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    96
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    97
  "init_obj2sobj (O_proc p) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    98
     (case (init_cp2sproc p) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    99
       Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> seeds))
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   100
     | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   101
| "init_obj2sobj (O_file f) = 
37
029cccce84b4 remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
chunhan
parents: 34
diff changeset
   102
     Some (S_file (init_cf2sfiles f) 
029cccce84b4 remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
chunhan
parents: 34
diff changeset
   103
                  (\<exists> f'. f' \<in> (init_same_inode_files f) \<and> O_file f \<in> seeds))"
39
13bba99ca090 done with co2sobj_simps
chunhan
parents: 37
diff changeset
   104
| "init_obj2sobj (O_dir f) = 
37
029cccce84b4 remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
chunhan
parents: 34
diff changeset
   105
     (case (init_cf2sfile f) of
39
13bba99ca090 done with co2sobj_simps
chunhan
parents: 37
diff changeset
   106
           Some sf \<Rightarrow> Some (S_dir sf) 
13bba99ca090 done with co2sobj_simps
chunhan
parents: 37
diff changeset
   107
         | _ \<Rightarrow> None)"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   108
| "init_obj2sobj (O_msgq q) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   109
     (case (init_cq2smsgq q) of
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   110
       Some sq \<Rightarrow> Some (S_msgq sq)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   111
     | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   112
| "init_obj2sobj (O_shm h) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   113
     (case (init_ch2sshm h) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   114
       Some sh \<Rightarrow> Some (S_shm sh)
34
e7f850d1e08e remove O_msg from co2sobj/init_obj2sobj
chunhan
parents: 33
diff changeset
   115
     | _       \<Rightarrow> None)"  (*
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   116
| "init_obj2sobj (O_msg q m) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   117
     (case (init_cq2smsgq q, init_cm2smsg q m) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   118
        (Some sq, Some sm) \<Rightarrow> Some (S_msg sq sm)
34
e7f850d1e08e remove O_msg from co2sobj/init_obj2sobj
chunhan
parents: 33
diff changeset
   119
      | _                  \<Rightarrow> None)" *)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   120
| "init_obj2sobj _ = None"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   121
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   122
definition  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   123
  "init_static_state \<equiv> {sobj. \<exists> obj. init_alive obj \<and> init_obj2sobj obj = Some sobj}"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   124
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   125
fun is_init_sfile :: "t_sfile \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   126
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   127
  "is_init_sfile (Init _, sec, psec,asec) = True"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   128
| "is_init_sfile _ = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   129
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   130
fun is_many_sfile :: "t_sfile \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   131
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   132
  "is_many_sfile (Created, sec, psec, asec) = True"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   133
| "is_many_sfile _ = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   134
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   135
fun is_init_sproc :: "t_sproc \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   136
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   137
  "is_init_sproc (Init p, sec, fds, shms) = True"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   138
| "is_init_sproc _                        = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   139
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   140
fun is_many_sproc :: "t_sproc \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   141
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   142
  "is_many_sproc (Created, sec,fds,shms) = True"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   143
| "is_many_sproc _                       = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   144
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   145
fun is_many_smsg :: "t_smsg \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   146
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   147
  "is_many_smsg (Created,sec,tag) = True"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   148
| "is_many_smsg _                 = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   149
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   150
fun is_many_smsgq :: "t_smsgq \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   151
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   152
  "is_many_smsgq (Created,sec,sms) = (True \<and> (\<forall> sm \<in> (set sms). is_many_smsg sm))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   153
| "is_many_smsgq _                 = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   154
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   155
fun is_many_sshm :: "t_sshm \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   156
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   157
  "is_many_sshm (Created, sec) = True"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   158
| "is_many_sshm _              = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   159
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   160
(*
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   161
fun is_init_sobj :: "t_sobject \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   162
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   163
  "is_init_sobj (S_proc (popt, sec, fds, shms) tag) = (popt \<noteq> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   164
| "is_init_sobj (S_file sf tag) = (is_init_sfile sf)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   165
| "is_init_sobj (S_dir  sf tag) = (is_init_sfile sf)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   166
| "is_init_sobj (S_msg" *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   167
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   168
fun is_many :: "t_sobject \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   169
where
42
chunhan
parents: 39
diff changeset
   170
  "is_many (S_proc sp   tag) = is_many_sproc sp"
chunhan
parents: 39
diff changeset
   171
| "is_many (S_file sfs  tag) = (\<forall> sf \<in> sfs. is_many_sfile sf)"
chunhan
parents: 39
diff changeset
   172
| "is_many (S_dir  sf      ) = is_many_sfile sf"
chunhan
parents: 39
diff changeset
   173
| "is_many (S_msgq sq      ) = is_many_smsgq sq"
chunhan
parents: 39
diff changeset
   174
| "is_many (S_shm  sh      ) = is_many_sshm  sh"
chunhan
parents: 39
diff changeset
   175
| "is_many _                 = False"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   176
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   177
(*
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   178
fun update_ss_sp:: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   179
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   180
  "update_ss_sp ss (S_proc sp tag) (S_proc sp' tag') = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   181
     (if (is_many_sproc sp) then ss \<union> {S_proc sp' tag'} 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   182
      else (ss - {S_proc sp tag}) \<union> {S_proc sp' tag'})"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   183
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   184
fun update_ss_sd:: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   185
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   186
  "update_ss_sd ss (S_dir sf tag) (S_dir sf' tag') = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   187
     (if (is_many_sfile sf) then ss \<union> {S_dir sf' tag'} 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   188
      else (ss - {S_dir sf tag}) \<union> {S_dir sf' tag'})"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   189
*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   190
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   191
definition update_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   192
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   193
  "update_ss ss so so' \<equiv> if (is_many so) then ss \<union> {so'} else (ss - {so}) \<union> {so'}"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   194
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   195
definition add_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   196
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   197
  "add_ss ss so \<equiv> ss \<union> {so}"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   198
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   199
definition del_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   200
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   201
  "del_ss ss so \<equiv> if (is_many so) then ss else ss - {so}"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   202
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   203
(*
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   204
fun sparent :: "t_sfile \<Rightarrow> t_sfile option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   205
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   206
  "sparent (Sroot si sec) = None"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   207
| "sparent (Sfile si sec spf) = Some spf"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   208
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   209
inductive is_ancesf :: "t_sfile \<Rightarrow> t_sfile \<Rightarrow> bool" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   210
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   211
  "is_ancesf sf sf"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   212
| "sparent sf = Some spf \<Longrightarrow> is_ancesf spf sf"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   213
| "\<lbrakk>sparent sf = Some spf; is_ancesf saf spf\<rbrakk> \<Longrightarrow> is_ancesf saf sf"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   214
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   215
definition sfile_reparent :: "t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   216
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   217
  "sfile_reparent (Sroot)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   218
*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   219
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   220
(*
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   221
(* sfds rename aux definitions *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   222
definition sfds_rename_notrelated 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   223
  :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   224
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   225
  "sfds_rename_notrelated sfds from to sfds' \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   226
     (\<forall> sec flag sf. (sec,flag,sf) \<in> sfds \<and> (\<not> from \<preceq> sf) \<longrightarrow> (sec,flag,sf) \<in> sfds')"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   227
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   228
definition sfds_rename_renamed
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   229
  :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   230
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   231
  "sfds_rename_renamed sfds sf from to sfds' \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   232
     (\<forall> sec flag sf'. (sec,flag,sf) \<in> sfds \<and> (sf \<preceq> sf') \<longrightarrow>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   233
         (sec, flag, file_after_rename sf' from to) \<in> sfds' \<and> (sec,flag,sf) \<notin> sfds')"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   234
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   235
definition sfds_rename_remain
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   236
  :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   237
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   238
  "sfds_rename_remain sfds sf from to sfds' \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   239
     (\<forall> sec flag sf'. (sec,flag,sf') \<in> sfds \<and> (sf \<preceq> sf') \<longrightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   240
         (sec, flag, sf') \<in> sfds' \<and> (sec,flag, file_after_rename sf' from to) \<notin> sfds')"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   241
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   242
(* for not many, choose on renamed or not *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   243
definition sfds_rename_choices
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   244
  :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   245
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   246
  "sfds_rename_choices sfds sf from to sfds' \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   247
     sfds_rename_remain sfds sf from to sfds' \<or> sfds_rename_renamed sfds sf from to sfds'"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   248
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   249
(* for many, merge renamed with not renamed *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   250
definition sfds_rename_both
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   251
  :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   252
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   253
  "sfds_rename_both sfds sf from to sfds' \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   254
     (\<forall> sec flag sf'. (sec,flag,sf') \<in> sfds \<and> (sf \<preceq> sf') \<longrightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   255
         (sec, flag, sf') \<in> sfds' \<or> (sec,flag, file_after_rename sf' from to) \<in> sfds')"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   256
  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   257
(* added to the new sfds, are those only under the new sfile *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   258
definition sfds_rename_added
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   259
  :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   260
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   261
  "sfds_rename_added sfds from to sfds' \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   262
     (\<forall> sec' flag' sf' sec flag. (sec',flag',sf') \<in> sfds' \<and> (sec,flag,sf') \<notin> sfds \<longrightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   263
        (\<exists> sf. (sec,flag,sf) \<in> sfds \<and> sf' = file_after_rename from to sf))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   264
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   265
definition sproc_sfds_renamed
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   266
  :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   267
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   268
  "sproc_sfds_renamed ss sf from to ss' \<equiv>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   269
     (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   270
       (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_renamed sfds sf from to sfds'))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   271
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   272
definition sproc_sfds_remain
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   273
  :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   274
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   275
  "sproc_sfds_remain ss sf from to ss' \<equiv>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   276
     (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   277
       (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_remain sfds sf from to sfds'))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   278
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   279
(* for not many, choose on renamed or not *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   280
definition sproc_sfds_choices
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   281
  :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   282
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   283
  "sproc_sfds_choices ss sf from to ss' \<equiv>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   284
     (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   285
       (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_choices sfds sf from to sfds'))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   286
  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   287
(* for many, merge renamed with not renamed *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   288
definition sproc_sfds_both
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   289
  :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   290
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   291
  "sproc_sfds_both ss sf from to ss' \<equiv>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   292
     (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   293
       (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_both sfds sf from to sfds'))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   294
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   295
(* remove (\<forall> sp tag. S_proc sp tag \<in> ss \<longrightarrow> S_proc sp tag \<in> ss'),
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   296
 * cause sfds contains sfs informations *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   297
definition ss_rename_notrelated 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   298
  :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   299
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   300
  "ss_rename_notrelated ss sf sf' ss' \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   301
     (\<forall> sq. S_msgq sq \<in> ss \<longrightarrow> S_msgq sq \<in> ss') \<and> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   302
     (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow> (\<exists> sfds'.  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   303
         S_proc (pi,sec,sfds',sshms) tagp \<in> ss'\<and> sfds_rename_notrelated sfds sf sf' sfds'))  \<and>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   304
     (\<forall> sfs sf'' tag. S_file sfs tag \<in> ss \<and> sf'' \<in> sfs \<and> (\<not> sf \<preceq> sf'') \<longrightarrow> (\<exists> sfs'. 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   305
         S_file sfs tag \<in> ss' \<and> sf'' \<in> sfs')) \<and>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   306
     (\<forall> sf'' tag. S_dir sf'' tag \<in> ss \<and> (\<not> sf \<preceq> sf'') \<longrightarrow> S_dir sf'' tag \<in> ss')" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   307
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   308
(* rename from to, from should definited renamed if not many *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   309
definition all_descendant_sf_renamed 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   310
  :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   311
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   312
  "all_descendant_sf_renamed ss sf from to ss' \<equiv>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   313
     (\<forall> sfs sf' tagf. sf \<preceq> sf' \<and> S_file sfs tagf \<in> ss \<and> sf' \<in> sfs \<longrightarrow> (\<exists> sfs'. 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   314
       S_file sfs' tagf \<in> ss' \<and> file_after_rename from to sf' \<in> sfs' \<and> sf' \<notin> sfs')) \<and> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   315
     (\<forall> sf' tagf. sf \<preceq> sf' \<and> S_dir sf' tagf \<in> ss \<longrightarrow> S_dir (file_after_rename from to sf') tagf \<in> ss' \<and> S_dir sf' tagf \<notin> ss') \<and>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   316
     sproc_sfds_renamed ss sf from to ss'"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   317
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   318
(* not renamed *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   319
definition all_descendant_sf_remain
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   320
  :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   321
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   322
  "all_descendant_sf_remain ss sf from to ss' \<equiv>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   323
     (\<forall> sfs sf' tag'. sf \<preceq> sf' \<and> S_file sfs tag' \<in> ss \<and> sf' \<in> sfs \<longrightarrow> (\<exists> sfs'. 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   324
       S_file sfs' tag' \<in> ss \<and> file_after_rename from to sf' \<notin> sfs' \<and> sf' \<in> sfs')) \<and> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   325
     (\<forall> sf' tag'. sf \<preceq> sf' \<and> S_dir sf' tag' \<in> ss \<longrightarrow> S_dir (file_after_rename from to sf') tag' \<notin> ss' \<and> S_dir sf' tag' \<in> ss') \<and>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   326
     sproc_sfds_remain ss sf from to ss'"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   327
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   328
definition all_descendant_sf_choices
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   329
  :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   330
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   331
  "all_descendant_sf_choices ss sf from to ss' \<equiv>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   332
     all_descendant_sf_renamed ss sf from to ss' \<or> all_descendant_sf_remain ss sf from to ss'"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   333
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   334
definition all_descendant_sf_both
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   335
  :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   336
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   337
  "all_descendant_sf_both ss sf from to ss' \<equiv>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   338
     (\<forall> sfs sf' tag'. sf \<preceq> sf' \<and> S_file sfs tag' \<in> ss \<and> sf' \<in> sfs \<longrightarrow> (\<exists> sfs'. 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   339
        S_file sfs' tag' \<in> ss \<and> file_after_rename from to sf' \<in> sfs' \<or> sf' \<in> sfs')) \<and> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   340
     (\<forall> sf' tag'. sf \<preceq> sf' \<and> S_dir sf' tag' \<in> ss \<longrightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   341
        S_dir (file_after_rename from to sf') tag' \<in> ss' \<or> S_dir sf' tag' \<in> ss') \<and>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   342
     sproc_sfds_both ss sf from to ss'"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   343
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   344
definition ss_renamed_file 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   345
  :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   346
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   347
  "ss_renamed_file ss sf sf' ss' \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   348
     (if (is_many_sfile sf) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   349
        then all_descendant_sf_choices ss sf sf sf' ss'
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   350
        else all_descendant_sf_renamed ss sf sf sf' ss')"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   351
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   352
(* added to the new sfs, are those only under the new sfile *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   353
definition sfs_rename_added
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   354
  :: "t_sfile set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile set \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   355
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   356
  "sfs_rename_added sfs from to sfs' \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   357
     (\<forall> sf'. sf' \<in> sfs' \<and> sf' \<notin> sfs \<longrightarrow> (\<exists> sf. sf \<in> sfs \<and> sf' = file_after_rename from to sf))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   358
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   359
(* added to the new sfile, are those only under the new sfile *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   360
definition ss_rename_added
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   361
  :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   362
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   363
  "ss_rename_added ss from to ss' \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   364
     (\<forall> pi sec fds fds' shms tagp. S_proc (pi, sec, fds,shms) tagp \<in> ss \<and> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   365
        S_proc (pi,sec,fds',shms) tagp \<in> ss' \<longrightarrow> sfds_rename_added fds from to fds') \<and>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   366
     (\<forall> sq. S_msgq sq \<in> ss' \<longrightarrow> S_msgq sq \<in> ss) \<and>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   367
     (\<forall> sfs sfs' tagf. S_file sfs' tagf \<in> ss' \<and> S_file sfs tagf \<in> ss \<longrightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   368
        sfs_rename_added sfs from to sfs') \<and>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   369
     (\<forall> sf' tagf. S_dir sf' tagf \<in> ss' \<and> S_dir sf' tagf \<notin> ss \<longrightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   370
        (\<exists> sf. S_dir sf tagf \<in> ss \<and> sf' = file_after_rename from to sf))" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   371
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   372
definition sfile_alive :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   373
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   374
  "sfile_alive ss sf \<equiv> (\<exists> sfs tagf. sf \<in> sfs \<and> S_file sfs tagf \<in> ss)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   375
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   376
definition sf_alive :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   377
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   378
  "sf_alive ss sf \<equiv> (\<exists> tagd. S_dir sf tagd \<in> ss) \<or> sfile_alive ss sf"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   379
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   380
(* constrains that the new static state should satisfy *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   381
definition ss_rename:: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   382
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   383
  "ss_rename ss sf sf' ss' \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   384
     ss_rename_notrelated ss sf sf' ss' \<and>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   385
     ss_renamed_file ss sf sf' ss' \<and> ss_rename_added ss sf sf' ss' \<and>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   386
     (\<forall> sf''. sf_alive ss sf'' \<and> sf \<prec> sf'' \<longrightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   387
       (if (is_many_sfile sf'') 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   388
        then all_descendant_sf_choices ss sf'' sf sf' ss'
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   389
        else all_descendant_sf_both ss sf'' sf sf' ss'))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   390
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   391
(* two sfile, the last fname should not be equal *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   392
fun sfile_same_fname:: "t_sfile \<Rightarrow> t_sfile \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   393
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   394
  "sfile_same_fname ((Init n, sec)#spf) ((Init n', sec')#spf') = (n = n')"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   395
| "sfile_same_fname _                   _                      = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   396
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   397
(* no same init sfile/only sfile in the target-to spf, should be in static_state addmissble check *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   398
definition ss_rename_no_same_fname:: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   399
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   400
  "ss_rename_no_same_fname ss from spf \<equiv>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   401
    \<not> (\<exists> to. sfile_same_fname from to \<and> parent to = Some spf \<and> sf_alive ss to)" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   402
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   403
(* is not a function, is a relation (one 2 many)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   404
definition update_ss_rename :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   405
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   406
  "update_ss_rename ss sf sf' \<equiv> if (is_many_sfile sf) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   407
     then (ss \<union> {S_file (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_file sf'' tag \<in> ss}
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   408
              \<union> {S_dir  (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_dir sf'' tag \<in> ss})
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   409
     else (ss - {S_file sf'' tag | sf'' tag. sf \<preceq> sf'' \<and> S_file sf'' tag \<in> ss}
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   410
              - {S_dir  sf'' tag | sf'' tag. sf \<preceq> sf'' \<and> S_dir sf'' tag \<in> ss})
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   411
              \<union> {S_file (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_file sf'' tag \<in> ss}
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   412
              \<union> {S_dir  (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_dir sf'' tag \<in> ss}" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   413
*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   414
*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   415
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   416
fun sectxt_of_sproc :: "t_sproc \<Rightarrow> security_context_t"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   417
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   418
  "sectxt_of_sproc (pi,sec,fds,shms) = sec"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   419
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   420
fun sectxt_of_sfile :: "t_sfile \<Rightarrow> security_context_t"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   421
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   422
  "sectxt_of_sfile (fi,sec,psec,asecs) = sec"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   423
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   424
fun asecs_of_sfile :: "t_sfile \<Rightarrow> security_context_t set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   425
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   426
  "asecs_of_sfile (fi,sec,psec,asecs) = asecs"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   427
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   428
definition search_check_s :: "security_context_t \<Rightarrow> t_sfile \<Rightarrow> bool \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   429
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   430
  "search_check_s pctxt sf if_file = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   431
    (if if_file 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   432
      then search_check_file pctxt (sectxt_of_sfile sf) \<and> search_check_allp pctxt (asecs_of_sfile sf)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   433
      else search_check_dir  pctxt (sectxt_of_sfile sf) \<and> search_check_allp pctxt (asecs_of_sfile sf))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   434
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   435
definition sectxts_of_sfds :: "t_sfd set \<Rightarrow> security_context_t set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   436
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   437
  "sectxts_of_sfds sfds \<equiv> {ctxt. \<exists> flag sf. (ctxt, flag, sf) \<in> sfds}"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   438
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   439
definition inherit_fds_check_s :: "security_context_t \<Rightarrow> t_sfd set \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   440
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   441
  "inherit_fds_check_s pctxt sfds \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   442
     (\<forall> ctxt \<in> sectxts_of_sfds sfds. permission_check pctxt ctxt C_fd P_inherit)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   443
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   444
definition sectxts_of_sproc_sshms :: "t_sproc_sshm set \<Rightarrow> security_context_t set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   445
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   446
  "sectxts_of_sproc_sshms sshms \<equiv> {ctxt. \<exists> hi flag. ((hi, ctxt),flag) \<in> sshms}"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   447
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   448
definition inherit_shms_check_s :: "security_context_t \<Rightarrow> t_sproc_sshm set \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   449
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   450
  "inherit_shms_check_s pctxt sshms \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   451
     (\<forall> ctxt \<in> sectxts_of_sproc_sshms sshms. permission_check pctxt ctxt C_shm P_inherit)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   452
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   453
(*
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   454
fun info_flow_sshm :: "t_sproc \<Rightarrow> t_sproc \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   455
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   456
  "info_flow_sshm (pi,sec,fds,shms) (pi',sec',fds',shms') = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   457
     (\<exists> sh flag'. (sh, SHM_RDWR) \<in> shms \<and> (sh, flag') \<in> shms')" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   458
*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   459
definition info_flow_sproc_sshms :: "t_sproc_sshm set \<Rightarrow> t_sproc_sshm set \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   460
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   461
  "info_flow_sproc_sshms shms shms' \<equiv> (\<exists> sh flag'. (sh, SHM_RDWR) \<in> shms \<and> (sh, flag') \<in> shms')" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   462
25
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 20
diff changeset
   463
(*
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   464
fun info_flow_sshm :: "t_sproc \<Rightarrow> t_sproc \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   465
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   466
  "info_flow_sshm (pi,sec,fds,shms) (pi',sec',fds',shms') = info_flow_sproc_sshms shms shms'"
25
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 20
diff changeset
   467
*)
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 20
diff changeset
   468
inductive info_flow_sshm :: "t_sproc \<Rightarrow> t_sproc \<Rightarrow> bool"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 20
diff changeset
   469
where
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 20
diff changeset
   470
  "info_flow_sshm sp sp"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 20
diff changeset
   471
| "\<lbrakk>info_flow_sshm sp (pi,sec,fds,shms); info_flow_sproc_sshms shms shms'\<rbrakk>
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 20
diff changeset
   472
   \<Longrightarrow> info_flow_sshm sp (pi',sec',fds',shms')"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   473
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   474
definition update_ss_shms :: "t_static_state \<Rightarrow> t_sproc \<Rightarrow> bool \<Rightarrow> t_static_state"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   475
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   476
  "update_ss_shms ss spfrom tag \<equiv> {sobj. \<exists> sobj' \<in> ss. 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   477
     (case sobj' of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   478
        S_proc sp tagp \<Rightarrow> if (info_flow_sshm spfrom sp) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   479
                          then if (is_many_sproc sp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   480
                               then (sobj = S_proc sp tagp \<or> sobj = S_proc sp (tagp \<or> tag))
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   481
                               else (sobj = S_proc sp (tagp \<or> tag))
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   482
                          else (sobj = S_proc sp tag)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   483
     | _ \<Rightarrow> sobj = sobj')}"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   484
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   485
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   486
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   487
(* all reachable static states(sobjects set) *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   488
inductive_set static :: "t_static_state set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   489
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   490
  s_init:    "init_static_state \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   491
| s_execve:  "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; S_file sfs tagf \<in> ss;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   492
               (fi,fsec,pfsec,asecs) \<in> sfs; npctxt_execve pctxt fsec = Some pctxt';
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   493
               grant_execve pctxt fsec pctxt'; search_check_s pctxt (fi,fsec,pfsec,asecs) True; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   494
               inherit_fds_check_s pctxt' fds'; fds' \<subseteq> fds\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   495
  \<Longrightarrow> (update_ss ss (S_proc (pi, pctxt, fds, shms) tagp) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   496
                    (S_proc (pi, pctxt', fds', {}) (tagp \<or> tagf))) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   497
| s_clone:   "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   498
               permission_check pctxt pctxt C_process P_fork;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   499
               inherit_fds_check_s pctxt fds'; fds' \<subseteq> fds; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   500
               inherit_shms_check_s pctxt shms'; shms' \<subseteq> shms\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   501
  \<Longrightarrow> (add_ss ss (S_proc (Created, pctxt, fds', shms') tagp)) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   502
| s_kill:    "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   503
               S_proc (pi', pctxt', fds', shms') tagp' \<in> ss; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   504
               permission_check pctxt pctxt' C_process P_sigkill\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   505
  \<Longrightarrow> (del_ss ss (S_proc (pi', pctxt', fds', shms') tagp')) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   506
| s_ptrace:  "\<lbrakk>ss \<in> static; S_proc sp tagp \<in> ss; S_proc sp' tagp' \<in> ss; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   507
               permission_check (sextxt_of_sproc sp) (sectxt_of_sproc sp') C_process P_ptrace\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   508
  \<Longrightarrow> (update_ss_shms (update_ss_shms ss sp tagp) sp' tagp') \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   509
| s_exit:    "\<lbrakk>ss \<in> static; S_proc sp tagp \<in> ss\<rbrakk> \<Longrightarrow> (del_ss ss (S_proc sp tagp)) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   510
| s_open:    "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; S_file sfs tagf \<in> ss; sf \<in> sfs;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   511
              search_check_s pctxt sf True; \<not> is_creat_excl_flag flags; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   512
              oflags_check flags pctxt (sectxt_of_sfile sf); permission_check pctxt pctxt C_fd P_setattr\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   513
  \<Longrightarrow> (update_ss ss (S_proc (pi, pctxt, fds, shms) tagp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   514
                    (S_proc (pi, pctxt, fds \<union> {(pctxt,flags,sf)}, shms) tagp)) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   515
| s_open':   "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; is_creat_excl_flag flags;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   516
               S_dir (pfi,fsec,pfsec,asecs) \<in> ss; search_check_s pctxt (pfi,fsec,pfsec,asecs) False; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   517
               nfsec = nfctxt_create pctxt fsec C_file; oflags_check flags pctxt nfsec;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   518
               permission_check pctxt fsec C_dir P_add_rename; permission_check pctxt pctxt C_fd P_setattr\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   519
  \<Longrightarrow> (update_ss (add_ss ss (S_file {(Created, nfsec, Some fsec, asecs \<union> {fsec})} tagp))
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   520
         (S_proc (pi, pctxt, fds, shms) tagp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   521
         (S_proc (pi, pctxt, fds \<union> {(pctxt, flags, (Created, nfsec, Some fsec, asecs \<union> {fsec}))}, shms) tagp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   522
      ) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   523
| S_readf:   "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; (fdctxt,flags,sf) \<in> fds; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   524
               permission_check pctxt fdctxt C_fd P_setattr; S_file sfs tagf \<in> ss; sf \<in> sfs;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   525
               permission_check pctxt (sectxt_of_sfile sf) C_file P_read; is_read_flag flags\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   526
  \<Longrightarrow> (update_ss_shms ss (pi, pctxt,fds,shms) (tagp \<or> tagf)) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   527
| S_writef:  "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; (fdctxt,flags,sf) \<in> fds; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   528
               permission_check pctxt fdctxt C_fd P_setattr; sf \<in> sfs; S_file sfs tagf \<in> ss; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   529
               permission_check pctxt (sectxt_of_sfile sf) C_file P_write; is_write_flag flags\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   530
  \<Longrightarrow> (update_ss ss (S_file sfs tagf) (S_file sfs (tagp \<or> tagf))) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   531
| S_unlink:  "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_file sfs tagf \<in> ss;  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   532
               (Init f,fsec,Some pfsec,asecs) \<in> sfs; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   533
               search_check_s pctxt (Init f,fsec,Some pfsec,asecs) True; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   534
               permission_check pctxt fsec C_file P_unlink; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   535
               permission_check pctxt pfsec C_dir P_remove_name\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   536
  \<Longrightarrow> ((ss - {S_file sfs tagf}) \<union> {S_file (sfs - {(Init f,fsec,Some pfsec,asecs)}) tagf}) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   537
| S_rmdir:   "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   538
               S_dir (fi,fsec,Some pfsec,asecs) \<in> ss;  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   539
               search_check_s pctxt (fi,fsec,Some pfsec,asecs) False; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   540
               permission_check pctxt fsec C_dir P_rmdir;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   541
               permission_check pctxt pfsec C_dir P_remove_name\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   542
  \<Longrightarrow> (del_ss ss (S_dir (fi,fsec,Some pfsec,asecs))) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   543
| S_mkdir:   "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_dir (fi,fsec,pfsec,asecs) \<in> ss;  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   544
               search_check_s pctxt (fi,fsec,pfsec,asecs) False; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   545
               permission_check pctxt (nfctxt_create pctxt fsec C_dir) C_dir P_create;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   546
               permission_check pctxt fsec C_dir P_add_name\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   547
  \<Longrightarrow> (add_ss ss (S_dir (Created,nfctxt_create pctxt fsec C_dir,Some fsec,asecs \<union> {fsec}))) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   548
| s_link:    "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_dir (pfi,pfsec,ppfsec,asecs) \<in> ss;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   549
               S_file sfs tagf \<in> ss; sf \<in> sfs; nfsec = nfctxt_create pctxt pfsec C_file;  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   550
               search_check_s pctxt (pfi,pfsec,ppfsec,asecs) False; search_check_s pctxt sf True;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   551
               permission_check pctxt (sectxt_of_sfile sf) C_file P_link; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   552
               permission_check pctxt pfsec C_dir P_add_name\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   553
  \<Longrightarrow> (update_ss ss (S_file sfs tagf) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   554
                  (S_file (sfs \<union> {(Created,nfsec,Some pfsec, asecs \<union> {pfsec})}) tagf)) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   555
| s_trunc:   "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_file sfs tagf \<in> ss; sf \<in> sfs; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   556
               search_check_s pctxt sf True; permission_check pctxt (sectxt_of_sfile sf) C_file P_setattr\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   557
  \<Longrightarrow> (update_ss ss (S_file sfs tagf) (S_file sfs (tagf \<or> tagp))) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   558
(*
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   559
| s_rename:  "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_file sfs tagf \<in> ss;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   560
               (sf#spf') \<in> sfs; S_dir spf tagpf \<in> ss; \<not>((sf#spf') \<preceq> (sf#spf)); 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   561
               search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   562
               sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt;  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   563
               permission_check pctxt fctxt C_file P_rename;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   564
               permission_check pctxt pfctxt C_dir P_add_name;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   565
               ss_rename ss (sf#spf') (sf#spf) ss'; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   566
               ss_rename_no_same_fname ss (sf#spf') (sf#spf)\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   567
              \<Longrightarrow> ss' \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   568
| s_rename': "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_dir (sf#spf') tagf \<in> ss;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   569
               S_dir spf tagpf \<in> ss; \<not>((sf#spf') \<preceq> (sf#spf)); 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   570
               search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   571
               sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt;  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   572
               permission_check pctxt fctxt C_dir P_reparent;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   573
               permission_check pctxt pfctxt C_dir P_add_name;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   574
               ss_rename ss (sf#spf') (sf#spf) ss'; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   575
               ss_rename_no_same_fname ss (sf#spf') (sf#spf)\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   576
              \<Longrightarrow> ss' \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   577
*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   578
| s_createq: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   579
               permission_check pctxt pctxt C_msgq P_associate;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   580
               permission_check pctxt pctxt C_msgq P_create\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   581
  \<Longrightarrow> (add_ss ss (S_msgq (Created,pctxt,[]))) \<in> static" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   582
| s_sendmsg: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_msgq (qi,qctxt,sms) \<in> ss;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   583
               permission_check pctxt qctxt C_msgq P_enqueue;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   584
               permission_check pctxt qctxt C_msgq P_write; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   585
               permission_check pctxt pctxt C_msg  P_create\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   586
  \<Longrightarrow> (update_ss ss (S_msgq (qi,qctxt,sms)) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   587
                    (S_msgq (qi,qctxt,sms @ [(Created, pctxt, tagp)]))) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   588
| s_recvmsg: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   589
               S_msgq (qi,qctxt,(mi,mctxt,tagm)#sms) \<in> ss;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   590
               permission_check pctxt qctxt C_msgq P_read; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   591
               permission_check pctxt mctxt C_msg  P_receive\<rbrakk>
33
6884b3c9284b fix bug of static.thy for the static of recvmsg case
chunhan
parents: 25
diff changeset
   592
  \<Longrightarrow> (update_ss (update_ss_shms ss (pi,pctxt,fds,shms) (tagp \<or> tagm)) 
6884b3c9284b fix bug of static.thy for the static of recvmsg case
chunhan
parents: 25
diff changeset
   593
                 (S_msgq (qi, qctxt, (mi, mctxt, tagm)#sms))
6884b3c9284b fix bug of static.thy for the static of recvmsg case
chunhan
parents: 25
diff changeset
   594
                 (S_msgq (qi, qctxt, sms))) \<in> static"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   595
| s_removeq: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_msgq (qi,qctxt,sms) \<in> ss;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   596
               permission_check pctxt qctxt C_msgq P_destroy\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   597
  \<Longrightarrow> (del_ss ss (S_msgq (qi,qctxt,sms))) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   598
| s_createh: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   599
               permission_check pctxt pctxt C_shm P_associate; 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   600
               permission_check pctxt pctxt C_shm P_create\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   601
   \<Longrightarrow> (add_ss ss (S_shm (Created, pctxt))) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   602
| s_attach:  "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_shm (hi,hctxt) \<in> ss;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   603
               if flag = SHM_RDONLY then permission_check pctxt hctxt C_shm P_read
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   604
               else (permission_check pctxt hctxt C_shm P_read \<and>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   605
                     permission_check pctxt hctxt C_shm P_write)\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   606
   \<Longrightarrow> (update_ss ss (S_proc (pi,pctxt,fds,shms) tagp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   607
                    (S_proc (pi,pctxt,fds,shms \<union> {((hi,hctxt),flag)}) tagp)) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   608
| s_detach:  "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_shm sh \<in> ss;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   609
               (sh,flag) \<in> shms; \<not> is_many_sshm sh\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   610
   \<Longrightarrow> (update_ss ss (S_proc (pi,pctxt,fds,shms) tagp)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   611
                    (S_proc (pi,pctxt,fds,shms - {(sh,flag)}) tagp)) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   612
| s_deleteh: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_shm (hi,hctxt) \<in> ss;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   613
               permission_check pctxt hctxt C_shm P_destroy; \<not> is_many_sshm sh\<rbrakk>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   614
   \<Longrightarrow> (remove_sproc_sshm (del_ss ss (S_shm (hi,hctxt))) (hi,hctxt)) \<in> static"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   615
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   616
(*
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   617
fun smsg_related :: "t_msg \<Rightarrow> t_smsg list \<Rightarrow> bool" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   618
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   619
  "smsg_related m []                   = False" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   620
| "smsg_related m ((mi, sec, tag)#sms) =  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   621
    (if (mi = Init m) then True else smsg_related m sms)" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   622
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   623
fun smsgq_smsg_related :: "t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsgq \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   624
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   625
  "smsgq_smsg_related q m (qi, sec, smsgslist) = ((qi = Init q) \<and> (smsg_related m smsgslist))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   626
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   627
fun smsg_relatainted :: "t_msg \<Rightarrow> t_smsg list \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   628
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   629
  "smsg_relatainted m []                     = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   630
| "smsg_relatainted m ((mi, sec, tag)#sms) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   631
    (if (mi = Init m \<and> tag = True) then True else smsg_relatainted m sms)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   632
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   633
fun smsgq_smsg_relatainted :: "t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsgq \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   634
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   635
  "smsgq_smsg_relatainted q m (qi, sec, smsgslist) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   636
     ((qi = Init q) \<and> (smsg_relatainted m smsgslist))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   637
*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   638
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   639
fun sfile_related :: "t_sfile \<Rightarrow> t_file \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   640
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   641
  "sfile_related (fi,sec,psec,asecs) f = (fi = Init f)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   642
(*
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   643
fun sproc_related :: "t_process \<Rightarrow> t_sproc \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   644
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   645
  "sproc_related p (pi, sec, fds, shms) = (pi = Init p)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   646
*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   647
fun init_obj_related :: "t_sobject \<Rightarrow> t_object \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   648
where
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 15
diff changeset
   649
  "init_obj_related (S_proc (pi, sec, fds, shms) tag) (O_proc p') = (pi = Init p')"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   650
| "init_obj_related (S_file sfs tag) (O_file f) = (\<exists> sf \<in> sfs. sfile_related sf f)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   651
| "init_obj_related (S_dir sf) (O_dir f) = (sfile_related sf f)"
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 15
diff changeset
   652
| "init_obj_related (S_msgq (qi, sec, sms)) (O_msgq q') = (qi = Init q')"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 15
diff changeset
   653
| "init_obj_related (S_shm (hi, sec)) (O_shm h') = (hi = Init h')"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 15
diff changeset
   654
| "init_obj_related (S_msg (qi, sec, sms) (mi, secm, tagm)) (O_msg q' m') = (qi = Init q' \<and> mi = Init m')"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   655
| "init_obj_related _ _ = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   656
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   657
fun tainted_s :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   658
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   659
  "tainted_s ss (S_proc sp  tag) = (S_proc sp tag  \<in> ss \<and> tag = True)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   660
| "tainted_s ss (S_file sfs tag) = (S_file sfs tag \<in> ss \<and> tag = True)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   661
| "tainted_s ss (S_msg  (qi, sec, sms)  (mi, secm, tag)) = 
20
e2c6af3ccb0d fixed bug in tainted_s and update
chunhan
parents: 19
diff changeset
   662
     (S_msgq (qi, sec, sms) \<in> ss \<and> (mi,secm,tag) \<in> set sms \<and> tag = True)"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   663
| "tainted_s ss _ = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   664
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   665
(*
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   666
fun tainted_s :: "t_object \<Rightarrow> t_static_state \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   667
where 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   668
  "tainted_s (O_proc p)  ss = (\<exists> sp. S_proc sp True \<in> ss \<and> sproc_related p sp)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   669
| "tainted_s (O_file f)  ss = (\<exists> sfs sf. S_file sfs True \<in> ss \<and> sf \<in> sfs \<and> sfile_related f sf)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   670
| "tainted_s (O_msg q m) ss = (\<exists> sq. S_msgq sq \<in> ss \<and> smsgq_smsg_relatainted q m sq)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   671
| "tainted_s _           ss = False"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   672
*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   673
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   674
definition taintable_s :: "t_object \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   675
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   676
  "taintable_s obj \<equiv> \<exists> ss \<in> static. \<exists> sobj. tainted_s ss sobj \<and> init_obj_related sobj obj \<and> init_alive obj"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   677
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   678
definition deletable_s :: "t_object \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   679
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   680
  "deletable_s obj \<equiv> init_alive obj \<and> (\<exists> ss \<in> static. \<forall> sobj \<in> ss. \<not> init_obj_related sobj obj)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   681
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   682
definition undeletable_s :: "t_object \<Rightarrow> bool"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   683
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   684
  "undeletable_s obj \<equiv> init_alive obj \<and> (\<forall> ss \<in> static. \<exists> sobj \<in> ss. init_obj_related sobj obj)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   685
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   686
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   687
(**************** translation from dynamic to static *******************)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   688
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 2
diff changeset
   689
definition cf2sfile :: "t_state \<Rightarrow> t_file \<Rightarrow> t_sfile option"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   690
where
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 2
diff changeset
   691
  "cf2sfile s f \<equiv>
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   692
     case (parent f) of 
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 2
diff changeset
   693
       None \<Rightarrow> Some sroot
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 2
diff changeset
   694
     | Some pf \<Rightarrow> if (is_file s f) 
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   695
     then (case (sectxt_of_obj s (O_file f), sectxt_of_obj s (O_dir pf), get_parentfs_ctxts s pf) of
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   696
            (Some sec, Some psec, Some asecs) \<Rightarrow>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   697
 Some (if (\<not> deleted (O_file f) s \<and> is_init_file f) then Init f else Created, sec, Some psec, set asecs)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   698
          | _ \<Rightarrow> None) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   699
     else (case (sectxt_of_obj s (O_dir f), sectxt_of_obj s (O_dir pf), get_parentfs_ctxts s pf) of
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   700
            (Some sec, Some psec, Some asecs) \<Rightarrow>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   701
 Some (if (\<not> deleted (O_dir f) s \<and> is_init_dir f) then Init f else Created, sec, Some psec, set asecs)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   702
          | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   703
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 2
diff changeset
   704
definition cf2sfiles :: "t_state \<Rightarrow> t_file \<Rightarrow> t_sfile set"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   705
where
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 2
diff changeset
   706
  "cf2sfiles s f \<equiv> {sf. \<exists> f' \<in> (same_inode_files s f). cf2sfile s f' = Some sf}"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   707
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   708
(* here cf2sfile is passed with True, because, process' fds are only for files not dirs *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   709
definition cfd2sfd :: "t_state \<Rightarrow> t_process \<Rightarrow> t_fd \<Rightarrow> t_sfd option" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   710
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   711
  "cfd2sfd s p fd \<equiv> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   712
    (case (file_of_proc_fd s p fd, flags_of_proc_fd s p fd, sectxt_of_obj s (O_fd p fd)) of
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 2
diff changeset
   713
      (Some f, Some flags, Some sec) \<Rightarrow> (case (cf2sfile s f) of 
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   714
                                          Some sf \<Rightarrow> Some (sec, flags, sf)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   715
                                        | _       \<Rightarrow> None)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   716
    | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   717
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   718
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   719
definition cpfd2sfds :: "t_state \<Rightarrow> t_process \<Rightarrow> t_sfd set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   720
where
12
47a4b2ae0556 Modify definitions of cpfd2sfds
chunhan
parents: 11
diff changeset
   721
  "cpfd2sfds s p \<equiv> {sfd. \<exists> fd \<in> proc_file_fds s p. cfd2sfd s p fd = Some sfd}"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   722
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   723
definition ch2sshm :: "t_state \<Rightarrow> t_shm \<Rightarrow> t_sshm option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   724
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   725
  "ch2sshm s h \<equiv> (case (sectxt_of_obj s (O_shm h)) of
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   726
                    Some sec \<Rightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   727
 Some (if (\<not> deleted (O_shm h) s \<and> h \<in> init_shms) then Init h else Created, sec)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   728
                  | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   729
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   730
definition cph2spshs :: "t_state \<Rightarrow> t_process \<Rightarrow> t_sproc_sshm set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   731
where
15
4ca824cd0c59 finally get cph2spshs_attach done
chunhan
parents: 12
diff changeset
   732
  "cph2spshs s p \<equiv> {(sh, flag)| sh flag h. (p, flag) \<in> procs_of_shm s h \<and> ch2sshm s h = Some sh}"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   733
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   734
definition cp2sproc :: "t_state \<Rightarrow> t_process \<Rightarrow> t_sproc option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   735
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   736
  "cp2sproc s p \<equiv> (case (sectxt_of_obj s (O_proc p)) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   737
                     Some sec \<Rightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   738
 Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs) then Init p else Created, sec, 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   739
       cpfd2sfds s p, cph2spshs s p)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   740
                   | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   741
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   742
definition cm2smsg :: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsg option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   743
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   744
  "cm2smsg s q m \<equiv> (case (sectxt_of_obj s (O_msg q m)) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   745
                      Some sec \<Rightarrow>
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   746
 Some (if (\<not> deleted (O_msg q m) s \<and> m \<in> set (init_msgs_of_queue q)) then Init m else Created,
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   747
       sec, O_msg q m \<in> tainted s)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   748
                    | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   749
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   750
fun cqm2sms:: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_msg list \<Rightarrow> (t_smsg list) option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   751
where 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   752
  "cqm2sms s q [] = Some []"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   753
| "cqm2sms s q (m#ms) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   754
     (case (cqm2sms s q ms, cm2smsg s q m) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   755
       (Some sms, Some sm) \<Rightarrow> Some (sm#sms) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   756
     | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   757
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   758
definition cq2smsgq :: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_smsgq option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   759
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   760
  "cq2smsgq s q \<equiv> (case (sectxt_of_obj s (O_msgq q), cqm2sms s q (msgs_of_queue s q)) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   761
                     (Some sec, Some sms) \<Rightarrow> 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   762
 Some (if (\<not> deleted (O_msgq q) s \<and> q \<in> init_msgqs) then Init q else Created,
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   763
       sec, sms)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   764
                   | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   765
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   766
fun co2sobj :: "t_state \<Rightarrow> t_object \<Rightarrow> t_sobject option"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   767
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   768
  "co2sobj s (O_proc p) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   769
     (case (cp2sproc s p) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   770
        Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> tainted s))
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   771
      | _       \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   772
| "co2sobj s (O_file f) = 
37
029cccce84b4 remove in-current constrains from co2sobj, alive already does this job in co2sobj's properties
chunhan
parents: 34
diff changeset
   773
     Some (S_file (cf2sfiles s f) (O_file f \<in> tainted s))"
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   774
| "co2sobj s (O_dir f) = 
6
8779d321cc2e remove duplicated pre-condition of current_files in OS_grant, since it is property if is_file/is_dir.
chunhan
parents: 2
diff changeset
   775
     (case (cf2sfile s f) of
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   776
        Some sf  \<Rightarrow> Some (S_dir sf)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   777
      | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   778
| "co2sobj s (O_msgq q) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   779
     (case (cq2smsgq s q) of
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   780
        Some sq \<Rightarrow> Some (S_msgq sq)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   781
      | _ \<Rightarrow> None)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   782
| "co2sobj s (O_shm h) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   783
     (case (ch2sshm s h) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   784
        Some sh \<Rightarrow> Some (S_shm sh)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   785
      | _ \<Rightarrow> None)"
33
6884b3c9284b fix bug of static.thy for the static of recvmsg case
chunhan
parents: 25
diff changeset
   786
(*
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   787
| "co2sobj s (O_msg q m) = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   788
     (case (cq2smsgq s q, cm2smsg s q m) of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   789
       (Some sq, Some sm) \<Rightarrow> Some (S_msg sq sm)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   790
     | _ \<Rightarrow> None)"
33
6884b3c9284b fix bug of static.thy for the static of recvmsg case
chunhan
parents: 25
diff changeset
   791
*)
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   792
| "co2sobj s _ = None"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   793
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   794
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   795
(***************** for backward proof when Instancing static objects ******************)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   796
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   797
definition next_nat :: "nat set \<Rightarrow> nat"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   798
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   799
  "next_nat nset = (Max nset) + 1"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   800
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   801
definition new_proc :: "t_state \<Rightarrow> t_process"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   802
where 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   803
  "new_proc \<tau> = next_nat (current_procs \<tau>)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   804
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   805
definition new_inode_num :: "t_state \<Rightarrow> nat"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   806
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   807
  "new_inode_num \<tau> = next_nat (current_inode_nums \<tau>)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   808
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   809
definition new_msgq :: "t_state \<Rightarrow> t_msgq"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   810
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   811
  "new_msgq s = next_nat (current_msgqs s)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   812
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   813
definition new_msg :: "t_state \<Rightarrow> t_msgq \<Rightarrow> nat"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   814
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   815
  "new_msg s q = next_nat (set (msgs_of_queue s q))"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   816
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   817
definition new_shm :: "t_state \<Rightarrow> nat"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   818
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   819
  "new_shm \<tau> = next_nat (current_shms \<tau>)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   820
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   821
definition new_proc_fd :: "t_state \<Rightarrow> t_process \<Rightarrow> t_fd"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   822
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   823
  "new_proc_fd \<tau> p = next_nat (current_proc_fds \<tau> p)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   824
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   825
definition all_fname_under_dir:: "t_file \<Rightarrow> t_state \<Rightarrow> t_fname set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   826
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   827
  "all_fname_under_dir d \<tau> = {fn. \<exists> f. fn # d = f \<and> f \<in> current_files \<tau>}"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   828
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   829
fun fname_all_a:: "nat \<Rightarrow> t_fname"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   830
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   831
  "fname_all_a 0 = []" |
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   832
  "fname_all_a (Suc n) = ''a''@(fname_all_a n)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   833
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   834
definition fname_length_set :: "t_fname set \<Rightarrow> nat set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   835
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   836
  "fname_length_set fns = length`fns"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   837
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   838
definition next_fname:: "t_file \<Rightarrow> t_state \<Rightarrow> t_fname"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   839
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   840
  "next_fname pf \<tau> = fname_all_a ((Max (fname_length_set (all_fname_under_dir pf \<tau>))) + 1)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   841
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   842
definition new_childf:: "t_file \<Rightarrow> t_state \<Rightarrow> t_file"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   843
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   844
  "new_childf pf \<tau> = next_fname pf \<tau> # pf"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   845
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   846
definition s2ss :: "t_state \<Rightarrow> t_static_state"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   847
where
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   848
  "s2ss s \<equiv> {sobj. \<exists> obj. alive s obj \<and> co2sobj s obj = Some sobj}"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   849
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   850
end
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   851
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   852
end