obj2sobj_prop.thy
author chunhan
Fri, 12 Apr 2013 10:46:43 +0100
changeset 2 301f567e2a8e
parent 1 dcde836219bc
child 6 4294abb1f38c
permissions -rw-r--r--
add document
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
     1
theory obj2sobj_prop 
dcde836219bc add thy files
chunhan
parents:
diff changeset
     2
imports Main rc_theory os_rc deleted_prop
dcde836219bc add thy files
chunhan
parents:
diff changeset
     3
begin
dcde836219bc add thy files
chunhan
parents:
diff changeset
     4
dcde836219bc add thy files
chunhan
parents:
diff changeset
     5
context tainting_s_complete begin
dcde836219bc add thy files
chunhan
parents:
diff changeset
     6
dcde836219bc add thy files
chunhan
parents:
diff changeset
     7
(** file 2 sfile   **)
dcde836219bc add thy files
chunhan
parents:
diff changeset
     8
dcde836219bc add thy files
chunhan
parents:
diff changeset
     9
lemma init_son_deleted_D:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    10
  "\<lbrakk>deleted (File pf) s; f # pf \<in> init_files; valid s\<rbrakk> \<Longrightarrow> deleted (File (f # pf)) s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    11
apply (induct s, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    12
by (frule valid_cons, frule valid_os, case_tac a, auto dest:init_notin_curf_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    13
dcde836219bc add thy files
chunhan
parents:
diff changeset
    14
lemma init_parent_undeleted_I:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    15
  "\<lbrakk>\<not> deleted (File (f # pf)) s; f # pf \<in> init_files; valid s\<rbrakk> \<Longrightarrow> \<not> deleted (File pf) s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    16
by (rule notI, simp add:init_son_deleted_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    17
dcde836219bc add thy files
chunhan
parents:
diff changeset
    18
lemma source_dir_in_init:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    19
  "source_dir s f = Some sd \<Longrightarrow> sd \<in> init_files"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    20
by (induct f, auto split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    21
dcde836219bc add thy files
chunhan
parents:
diff changeset
    22
lemma source_dir_of_init: "\<lbrakk>source_dir [] f = Some sd; f \<in> init_files\<rbrakk> \<Longrightarrow> f = sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    23
by (induct f, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    24
dcde836219bc add thy files
chunhan
parents:
diff changeset
    25
lemma source_dir_of_init': "f \<in> init_files \<Longrightarrow> source_dir [] f = Some f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    26
by (induct f, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    27
dcde836219bc add thy files
chunhan
parents:
diff changeset
    28
lemma init_not_curf_imp_deleted:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    29
  "\<lbrakk>f \<in> init_files; f \<notin> current_files s; valid s\<rbrakk> \<Longrightarrow> deleted (File f) s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    30
apply (induct s, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    31
apply (frule valid_cons, frule valid_os, case_tac a, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    32
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    33
dcde836219bc add thy files
chunhan
parents:
diff changeset
    34
lemma source_dir_of_init'': 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    35
  "\<lbrakk>f \<in> init_files; \<not> deleted (File f) s; valid s\<rbrakk> \<Longrightarrow> source_dir s f = Some f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    36
by (induct f, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    37
dcde836219bc add thy files
chunhan
parents:
diff changeset
    38
dcde836219bc add thy files
chunhan
parents:
diff changeset
    39
lemma source_dir_createf:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    40
  "valid (CreateFile p (f#pf) # s) \<Longrightarrow> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    41
  source_dir (CreateFile p (f#pf) # s) = (source_dir s) ((f#pf)  := source_dir s pf)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    42
apply (frule valid_os, frule valid_cons)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    43
apply (rule ext, induct_tac x)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    44
apply (auto dest:init_not_curf_imp_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    45
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    46
dcde836219bc add thy files
chunhan
parents:
diff changeset
    47
lemma source_dir_createf':
dcde836219bc add thy files
chunhan
parents:
diff changeset
    48
  "valid (CreateFile p f # s) \<Longrightarrow> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    49
  source_dir (CreateFile p f # s) = (source_dir s) (f := (case (parent f) of
dcde836219bc add thy files
chunhan
parents:
diff changeset
    50
                                                            Some pf \<Rightarrow> source_dir s pf
dcde836219bc add thy files
chunhan
parents:
diff changeset
    51
                                                          | _       \<Rightarrow> None))"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    52
apply (frule valid_os, case_tac f, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    53
apply (drule source_dir_createf, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    54
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    55
dcde836219bc add thy files
chunhan
parents:
diff changeset
    56
lemma source_dir_other:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    57
  "\<lbrakk>valid (e # s); \<forall> p f. e \<noteq> CreateFile p f; \<forall> p f. e \<noteq> DeleteFile p f\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
    58
   \<Longrightarrow> source_dir (e#s) = source_dir s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    59
apply (rule ext, induct_tac x, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    60
apply (auto dest:not_deleted_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    61
apply (case_tac [!] e, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    62
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    63
dcde836219bc add thy files
chunhan
parents:
diff changeset
    64
lemma source_dir_deletef:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    65
  "valid (DeleteFile p f # s) \<Longrightarrow> source_dir (DeleteFile p f # s) f' =  
dcde836219bc add thy files
chunhan
parents:
diff changeset
    66
     (if (source_dir s f') = Some f then parent f else (source_dir s f'))"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    67
apply (frule valid_os, frule valid_cons)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    68
apply (case_tac "f \<in> init_files")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    69
apply (induct_tac f', simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    70
apply (auto dest!:init_parent_undeleted_I intro:parent_file_in_init'
dcde836219bc add thy files
chunhan
parents:
diff changeset
    71
            intro!: source_dir_of_init'')[1]
dcde836219bc add thy files
chunhan
parents:
diff changeset
    72
apply (induct_tac f', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    73
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    74
dcde836219bc add thy files
chunhan
parents:
diff changeset
    75
lemma source_dir_deletef':
dcde836219bc add thy files
chunhan
parents:
diff changeset
    76
  "valid (DeleteFile p f # s) \<Longrightarrow> source_dir (DeleteFile p f # s) = (\<lambda> f'.   
dcde836219bc add thy files
chunhan
parents:
diff changeset
    77
     (if (source_dir s f') = Some f then parent f else (source_dir s f')) )"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    78
by (auto dest:source_dir_deletef)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    79
dcde836219bc add thy files
chunhan
parents:
diff changeset
    80
lemmas source_dir_simps = source_dir_of_init' source_dir_of_init'' source_dir_createf' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    81
  source_dir_deletef' source_dir_other
dcde836219bc add thy files
chunhan
parents:
diff changeset
    82
dcde836219bc add thy files
chunhan
parents:
diff changeset
    83
declare source_dir.simps [simp del]
dcde836219bc add thy files
chunhan
parents:
diff changeset
    84
dcde836219bc add thy files
chunhan
parents:
diff changeset
    85
lemma source_dir_is_ancient:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    86
  "source_dir s f = Some sd ==> sd \<preceq> f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    87
apply (induct f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    88
by (auto simp:source_dir.simps no_junior_def split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    89
dcde836219bc add thy files
chunhan
parents:
diff changeset
    90
lemma no_junior_trans: "\<lbrakk>f \<preceq> f'; f' \<preceq> f''\<rbrakk> \<Longrightarrow> f \<preceq> f''"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    91
by (auto elim:no_juniorE)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    92
dcde836219bc add thy files
chunhan
parents:
diff changeset
    93
lemma ancient_has_parent:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    94
  "[| f \<preceq> f'; f \<noteq> f'|] ==> \<exists> sonf. parent sonf = Some f \<and> sonf \<preceq> f' "
dcde836219bc add thy files
chunhan
parents:
diff changeset
    95
apply (induct f')
dcde836219bc add thy files
chunhan
parents:
diff changeset
    96
apply (simp add:no_junior_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    97
apply (case_tac "f = f'")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    98
apply (rule_tac x = "a # f'" in exI, simp add:no_junior_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    99
apply (frule no_junior_noteq, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   100
apply clarsimp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   101
apply (rule_tac x = sonf in exI, simp add:no_junior_trans)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   102
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   103
dcde836219bc add thy files
chunhan
parents:
diff changeset
   104
lemma source_dir_prop:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   105
  "[|\<forall>fn. fn # f' \<notin> current_files s; source_dir s f = Some f'; f \<in> current_files s; valid s|]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   106
  ==> f = f'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   107
  apply (drule source_dir_is_ancient)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   108
  apply (case_tac "f = f'", simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   109
  apply (drule ancient_has_parent, simp, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   110
  apply (drule_tac ancient_file_in_current, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   111
  apply (case_tac sonf, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   112
  done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   113
dcde836219bc add thy files
chunhan
parents:
diff changeset
   114
lemma current_file_has_sd:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   115
  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> sd. source_dir s f = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   116
apply (induct s arbitrary:f, simp add:source_dir_of_init')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   117
apply (frule valid_cons, frule valid_os, case_tac a, auto simp:source_dir_simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   118
apply (case_tac list, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   119
apply (rule_tac f = f in cannot_del_root, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   120
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   121
dcde836219bc add thy files
chunhan
parents:
diff changeset
   122
lemma current_file_has_sd':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   123
  "\<lbrakk>source_dir s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   124
by (rule notI, auto dest:current_file_has_sd)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   125
dcde836219bc add thy files
chunhan
parents:
diff changeset
   126
lemma current_file_has_sfile:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   127
  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> et sd. cf2sfile s f = Some (et, sd)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   128
apply (frule current_file_has_sd, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   129
apply (frule current_file_has_etype, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   130
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   131
dcde836219bc add thy files
chunhan
parents:
diff changeset
   132
lemma current_file_has_sfile':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   133
  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> sf. cf2sfile s f = Some sf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   134
by (auto dest:current_file_has_sfile)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   135
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   136
(*
dcde836219bc add thy files
chunhan
parents:
diff changeset
   137
lemma not_deleted_sf_remains:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   138
  "\<lbrakk>f \<in> current_files s; \<not> deleted (File f) s; valid s\<rbrakk> \<Longrightarrow> "
dcde836219bc add thy files
chunhan
parents:
diff changeset
   139
*)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   140
dcde836219bc add thy files
chunhan
parents:
diff changeset
   141
lemma current_proc_has_sproc:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   142
  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> r fr pt u. cp2sproc s p = Some (r, fr, pt, u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   143
apply (frule current_proc_has_role, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   144
apply (frule current_proc_has_type, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   145
apply (frule current_proc_has_forcedrole, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   146
apply (frule current_proc_has_owner, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   147
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   148
dcde836219bc add thy files
chunhan
parents:
diff changeset
   149
lemma current_proc_has_sproc':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   150
  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sp. cp2sproc s p = Some sp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   151
by (auto dest!:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   152
dcde836219bc add thy files
chunhan
parents:
diff changeset
   153
lemma current_ipc_has_sipc: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   154
  "\<lbrakk>i \<in> current_ipcs s; valid s\<rbrakk> \<Longrightarrow> \<exists> t. ci2sipc s i = Some t"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   155
by (drule current_ipc_has_type, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   156
dcde836219bc add thy files
chunhan
parents:
diff changeset
   157
lemma init_file_has_sobj:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   158
  "f \<in> init_files \<Longrightarrow> \<exists> t sd. init_obj2sobj (File f) = SFile (t, sd) (Some f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   159
by (frule init_file_has_etype, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   160
dcde836219bc add thy files
chunhan
parents:
diff changeset
   161
lemma init_proc_has_sobj:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   162
  assumes pinit:"p \<in> init_processes"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   163
  shows "\<exists> r fr pt u. init_obj2sobj (Proc p) = SProc (r, fr, pt, u) (Some p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   164
proof -
dcde836219bc add thy files
chunhan
parents:
diff changeset
   165
  from pinit obtain r where "init_currentrole p = Some r" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   166
    using init_proc_has_role by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   167
  moreover from pinit obtain fr where "init_proc_forcedrole p = Some fr"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   168
    using init_proc_has_frole by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   169
  moreover from pinit obtain pt where "init_process_type p = Some pt"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   170
    using init_proc_has_type by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   171
  moreover from pinit obtain u where "init_owner p = Some u"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   172
    using init_proc_has_owner by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   173
  ultimately show ?thesis by auto 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   174
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   175
dcde836219bc add thy files
chunhan
parents:
diff changeset
   176
lemma init_ipc_has_sobj:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   177
  "i \<in> init_ipcs \<Longrightarrow> \<exists> t. init_obj2sobj (IPC i) = SIPC t (Some i)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   178
using init_ipc_has_type
dcde836219bc add thy files
chunhan
parents:
diff changeset
   179
by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   180
dcde836219bc add thy files
chunhan
parents:
diff changeset
   181
lemma init_obj_has_sobj:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   182
  "exists [] obj \<Longrightarrow> init_obj2sobj obj \<noteq> Unknown"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   183
apply (case_tac obj)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   184
apply (simp_all only:exists.simps current_procs.simps current_ipcs.simps current_files.simps) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   185
apply (auto dest!:init_proc_has_sobj init_file_has_sobj init_ipc_has_sobj)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   186
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   187
dcde836219bc add thy files
chunhan
parents:
diff changeset
   188
lemma exists_obj_has_sobj:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   189
  "\<lbrakk>exists s obj; valid s\<rbrakk> \<Longrightarrow> obj2sobj s obj \<noteq> Unknown"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   190
apply (case_tac obj)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   191
apply (auto dest!:current_ipc_has_sipc current_proc_has_sproc' current_file_has_sfile' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   192
            split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   193
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   194
dcde836219bc add thy files
chunhan
parents:
diff changeset
   195
lemma current_proc_has_srp:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   196
  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> srp. source_proc s p = Some srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   197
apply (induct s arbitrary:p, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   198
by (frule valid_cons, frule valid_os, case_tac a, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   199
dcde836219bc add thy files
chunhan
parents:
diff changeset
   200
lemma current_proc_has_sobj:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   201
  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> r fr t u srp. obj2sobj s (Proc p) = SProc (r,fr,t,u) (Some srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   202
apply (frule current_proc_has_sproc')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   203
apply (auto dest:current_proc_has_srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   204
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   205
dcde836219bc add thy files
chunhan
parents:
diff changeset
   206
lemma current_file_has_sobj:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   207
  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> t sd srf. obj2sobj s (File f) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   208
by (auto dest:current_file_has_sfile)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   209
dcde836219bc add thy files
chunhan
parents:
diff changeset
   210
lemma current_ipc_has_sobj:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   211
  "\<lbrakk>i \<in> current_ipcs s; valid s\<rbrakk> \<Longrightarrow> \<exists> t sri. obj2sobj s (IPC i) = SIPC t sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   212
by (auto dest:current_ipc_has_sipc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   213
dcde836219bc add thy files
chunhan
parents:
diff changeset
   214
lemma sobj_has_proc_role:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   215
  "obj2sobj s (Proc p) = SProc (r, fr, t, u) srp \<Longrightarrow> currentrole s p = Some r"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   216
by (auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   217
dcde836219bc add thy files
chunhan
parents:
diff changeset
   218
lemma chown_role_aux_valid:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   219
  "\<lbrakk>currentrole s p = Some r; proc_forcedrole s p = Some fr\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   220
  \<Longrightarrow> chown_role_aux r fr u = currentrole (ChangeOwner p u # s) p"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   221
by (auto split:t_role.splits simp:chown_role_aux_def dest:proc_forcedrole_valid)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   222
dcde836219bc add thy files
chunhan
parents:
diff changeset
   223
lemma chown_role_aux_valid':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   224
  "cp2sproc s p = Some (r, fr, t, u') \<Longrightarrow> chown_role_aux r fr u = currentrole (ChangeOwner p u # s) p"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   225
by (rule chown_role_aux_valid, auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   226
dcde836219bc add thy files
chunhan
parents:
diff changeset
   227
lemma chown_type_aux_valid:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   228
  "\<lbrakk>currentrole s p = Some r; currentrole (ChangeOwner p u # s) p = Some nr; type_of_process s p = Some t\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   229
  \<Longrightarrow> type_of_process (ChangeOwner p u # s) p = Some (chown_type_aux r nr t)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   230
apply (auto split:option.splits t_rc_proc_type.splits 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   231
             dest:default_process_create_type_valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   232
             simp:chown_type_aux_def pot_def pct_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   233
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   234
dcde836219bc add thy files
chunhan
parents:
diff changeset
   235
lemma chown_type_aux_valid':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   236
  "\<lbrakk>cp2sproc s p = Some (r, fr, t, u'); currentrole (ChangeOwner p u # s) p = Some nr\<rbrakk> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   237
   \<Longrightarrow> type_of_process (ChangeOwner p u # s) p = Some (chown_type_aux r nr t)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   238
by (rule chown_type_aux_valid, auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   239
dcde836219bc add thy files
chunhan
parents:
diff changeset
   240
lemma exec_type_aux_valid:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   241
  "\<lbrakk>currentrole s p = Some r; type_of_process s p = Some t\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   242
  \<Longrightarrow> type_of_process (Execute p f # s) p = Some (exec_type_aux r t)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   243
apply (auto split:option.splits t_rc_proc_type.splits 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   244
             dest:default_process_execute_type_valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   245
             simp:exec_type_aux_def pet_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   246
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   247
dcde836219bc add thy files
chunhan
parents:
diff changeset
   248
lemma exec_type_aux_valid':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   249
  "cp2sproc s p = Some (r, fr, t, u') \<Longrightarrow> type_of_process (Execute p f # s) p = Some (exec_type_aux r t)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   250
by (rule exec_type_aux_valid, auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   251
dcde836219bc add thy files
chunhan
parents:
diff changeset
   252
lemma non_initf_frole_inherit:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   253
  "\<lbrakk>f \<notin> init_files; f \<noteq> []\<rbrakk> \<Longrightarrow> forcedrole s f = Some InheritParentRole"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   254
apply (induct s) defer
dcde836219bc add thy files
chunhan
parents:
diff changeset
   255
apply (case_tac a, auto) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   256
apply (induct f, auto split:option.splits dest:init_frole_has_file)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   257
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   258
dcde836219bc add thy files
chunhan
parents:
diff changeset
   259
lemma non_initf_irole_inherit:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   260
  "\<lbrakk>f \<notin> init_files; f \<noteq> []\<rbrakk> \<Longrightarrow> initialrole s f = Some InheritParentRole"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   261
apply (induct s) defer
dcde836219bc add thy files
chunhan
parents:
diff changeset
   262
apply (case_tac a, auto) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   263
apply (induct f, auto split:option.splits dest:init_irole_has_file)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   264
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   265
dcde836219bc add thy files
chunhan
parents:
diff changeset
   266
lemma deleted_file_frole_inherit:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   267
  "\<lbrakk>deleted (File f) s; f \<in> current_files s\<rbrakk> \<Longrightarrow> forcedrole s f = Some InheritParentRole"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   268
apply (induct s, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   269
apply (case_tac a, auto) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   270
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   271
dcde836219bc add thy files
chunhan
parents:
diff changeset
   272
lemma deleted_file_irole_inherit:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   273
  "\<lbrakk>deleted (File f) s; f \<in> current_files s\<rbrakk> \<Longrightarrow> initialrole s f = Some InheritParentRole"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   274
apply (induct s, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   275
apply (case_tac a, auto) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   276
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   277
dcde836219bc add thy files
chunhan
parents:
diff changeset
   278
lemma sd_deter_efrole:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   279
  "\<lbrakk>source_dir s f = Some sd; valid s; f \<in> current_files s\<rbrakk> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   280
  \<Longrightarrow> effforcedrole s f = effforcedrole s sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   281
apply (induct f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   282
apply (drule source_dir_is_ancient, simp add:no_junior_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   283
apply (simp add:source_dir.simps split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   284
apply (frule parent_file_in_current', simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   285
apply (case_tac "a # f \<in> init_files", simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   286
apply (drule_tac deleted_file_frole_inherit, simp, simp add:effforcedrole_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   287
apply (drule_tac s = s in non_initf_frole_inherit, simp, simp add:effforcedrole_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   288
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   289
dcde836219bc add thy files
chunhan
parents:
diff changeset
   290
lemma sd_deter_eirole:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   291
  "\<lbrakk>source_dir s f = Some sd; valid s; f \<in> current_files s\<rbrakk> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   292
  \<Longrightarrow> effinitialrole s f = effinitialrole s sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   293
apply (induct f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   294
apply (drule source_dir_is_ancient, simp add:no_junior_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   295
apply (simp add:source_dir.simps split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   296
apply (frule parent_file_in_current', simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   297
apply (case_tac "a # f \<in> init_files", simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   298
apply (drule_tac deleted_file_irole_inherit, simp, simp add:effinitialrole_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   299
apply (drule_tac s = s in non_initf_irole_inherit, simp, simp add:effinitialrole_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   300
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   301
dcde836219bc add thy files
chunhan
parents:
diff changeset
   302
lemma undel_initf_keeps_frole:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   303
  "\<lbrakk>f \<in> init_files; \<not> deleted (File f) s; valid s\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   304
   \<Longrightarrow> forcedrole s f = init_file_forcedrole f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   305
apply (induct s, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   306
apply (frule valid_cons, frule valid_os, case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   307
apply (auto dest:init_notin_curf_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   308
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   309
dcde836219bc add thy files
chunhan
parents:
diff changeset
   310
lemma undel_initf_keeps_efrole:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   311
  "\<lbrakk>f \<in> init_files; \<not> deleted (File f) s; valid s\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   312
   \<Longrightarrow> effforcedrole s f = erole_functor init_file_forcedrole InheritUpMixed f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   313
apply (induct f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   314
apply (drule undel_initf_keeps_frole, simp, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   315
apply (simp add:effforcedrole_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   316
apply (frule parent_file_in_init', frule init_parent_undeleted_I, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   317
apply (drule undel_initf_keeps_frole, simp, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   318
apply (simp add:effforcedrole_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   319
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   320
dcde836219bc add thy files
chunhan
parents:
diff changeset
   321
lemma undel_initf_keeps_irole:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   322
  "\<lbrakk>f \<in> init_files; \<not> deleted (File f) s; valid s\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   323
   \<Longrightarrow> initialrole s f = init_file_initialrole f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   324
apply (induct s, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   325
apply (frule valid_cons, frule valid_os, case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   326
apply (auto dest:init_notin_curf_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   327
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   328
dcde836219bc add thy files
chunhan
parents:
diff changeset
   329
lemma undel_initf_keeps_eirole:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   330
  "\<lbrakk>f \<in> init_files; \<not> deleted (File f) s; valid s\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   331
   \<Longrightarrow> effinitialrole s f = erole_functor init_file_initialrole UseForcedRole f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   332
apply (induct f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   333
apply (drule undel_initf_keeps_irole, simp, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   334
apply (simp add:effinitialrole_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   335
apply (frule parent_file_in_init', frule init_parent_undeleted_I, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   336
apply (drule undel_initf_keeps_irole, simp, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   337
apply (simp add:effinitialrole_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   338
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   339
dcde836219bc add thy files
chunhan
parents:
diff changeset
   340
lemma source_dir_not_deleted:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   341
  "source_dir s f = Some sd \<Longrightarrow> \<not> deleted (File sd) s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   342
by (induct f, auto simp:source_dir.simps split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   343
dcde836219bc add thy files
chunhan
parents:
diff changeset
   344
lemma exec_role_aux_valid:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   345
  "\<lbrakk>currentrole s p = Some r; source_dir s f = Some sd; owner s p = Some u; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   346
  f \<in> current_files s; valid s\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   347
  \<Longrightarrow> exec_role_aux r sd u = currentrole (Execute p f # s) p"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   348
apply (frule sd_deter_eirole, simp+, frule sd_deter_efrole, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   349
apply (frule source_dir_in_init, drule source_dir_not_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   350
apply (simp add:undel_initf_keeps_eirole undel_initf_keeps_efrole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   351
apply (frule file_has_effinitialrole, simp, frule file_has_effforcedrole, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   352
apply (auto split:option.splits t_role.splits simp:map_comp_def exec_role_aux_def
dcde836219bc add thy files
chunhan
parents:
diff changeset
   353
             dest:effforcedrole_valid effinitialrole_valid)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   354
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   355
dcde836219bc add thy files
chunhan
parents:
diff changeset
   356
lemma exec_role_aux_valid':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   357
  "\<lbrakk>cp2sproc s p = Some (r, fr, t, u); source_dir s f = Some sd; f \<in> current_files s; valid s\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   358
  \<Longrightarrow> exec_role_aux r sd u = currentrole (Execute p f # s) p"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   359
by (rule exec_role_aux_valid, auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   360
dcde836219bc add thy files
chunhan
parents:
diff changeset
   361
lemma cp2sproc_nil_init:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   362
  "init_obj2sobj (Proc p) = (case (cp2sproc [] p) of 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   363
                               Some sp \<Rightarrow> SProc sp (Some p)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   364
                             | _       \<Rightarrow> Unknown)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   365
by (auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   366
dcde836219bc add thy files
chunhan
parents:
diff changeset
   367
lemma cf2sfile_nil_init:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   368
  "init_obj2sobj (File f) = (case (cf2sfile [] f) of 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   369
                               Some sf \<Rightarrow> SFile sf (Some f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   370
                             | _       \<Rightarrow> Unknown)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   371
apply (auto split:option.splits simp:etype_of_file_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   372
apply (case_tac "f \<in> init_files", simp add:source_dir_of_init')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   373
apply (induct f, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   374
apply (case_tac "f \<in> init_files", simp add:source_dir_of_init')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   375
apply (induct f, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   376
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   377
dcde836219bc add thy files
chunhan
parents:
diff changeset
   378
lemma ci2sipc_nil_init:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   379
  "init_obj2sobj (IPC i) = (case (ci2sipc [] i) of 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   380
                              Some si \<Rightarrow> SIPC si (Some i)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   381
                            | _       \<Rightarrow> Unknown)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   382
by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   383
dcde836219bc add thy files
chunhan
parents:
diff changeset
   384
lemma obj2sobj_nil_init:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   385
  "exists [] obj \<Longrightarrow> obj2sobj [] obj = init_obj2sobj obj" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   386
apply (case_tac obj) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   387
apply (auto simp:cf2sfile_nil_init cp2sproc_nil_init ci2sipc_nil_init 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   388
                 source_dir_of_init' etype_of_file_def
dcde836219bc add thy files
chunhan
parents:
diff changeset
   389
           split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   390
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   391
dcde836219bc add thy files
chunhan
parents:
diff changeset
   392
(**** cp2sproc simpset ****)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   393
dcde836219bc add thy files
chunhan
parents:
diff changeset
   394
lemma current_proc_has_role':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   395
  "\<lbrakk>currentrole s p = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   396
by (rule notI, auto dest:current_proc_has_role)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   397
dcde836219bc add thy files
chunhan
parents:
diff changeset
   398
lemma cp2sproc_chown:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   399
  assumes vs: "valid (ChangeOwner p u # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   400
  shows "cp2sproc (ChangeOwner p u # s) = (cp2sproc s) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   401
     (p := (case (cp2sproc s p) of
dcde836219bc add thy files
chunhan
parents:
diff changeset
   402
              Some (r,fr,pt,u') \<Rightarrow> (case (chown_role_aux r fr u) of 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   403
                                      Some nr \<Rightarrow> Some (nr,fr,chown_type_aux r nr pt,u)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   404
                                    | _       \<Rightarrow> None)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   405
            | _                 \<Rightarrow> None)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   406
      )" (is "?lhs = ?rhs")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   407
proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   408
  have os: "os_grant s (ChangeOwner p u)" and vs': "valid s" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   409
    by (auto dest:valid_cons valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   410
  have "\<And> x. x \<noteq> p \<Longrightarrow> ?lhs x = ?rhs x"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   411
    by (auto simp:type_of_process.simps split:option.splits t_role.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   412
  moreover have "?lhs p = ?rhs p"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   413
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   414
    from os have p_in: "p \<in> current_procs s" by (simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   415
    then obtain r fr t u' where csp: "cp2sproc s p = Some (r, fr, t, u')" using vs'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   416
      by (drule_tac current_proc_has_sproc, auto)  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   417
    from os have "u \<in> init_users" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   418
    hence "defrole u \<noteq> None" using init_user_has_role by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   419
    then obtain nr where nrole:"chown_role_aux r fr u = Some nr"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   420
      by (case_tac fr, auto simp:chown_role_aux_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   421
    have nr_eq: "currentrole (ChangeOwner p u # s) p = chown_role_aux r fr u" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   422
      using csp by (auto simp:chown_role_aux_valid'[where u = u])  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   423
    moreover have "type_of_process (ChangeOwner p u # s) p = Some (chown_type_aux r nr t)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   424
      using csp nrole nr_eq 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   425
      by (rule_tac fr = fr and u' = u' in chown_type_aux_valid', simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   426
    moreover have "proc_forcedrole (ChangeOwner p u # s) p = Some fr"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   427
      using csp by (auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   428
    moreover have "owner (ChangeOwner p u # s) p = Some u" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   429
    ultimately have "cp2sproc (ChangeOwner p u # s) p = Some (nr, fr, chown_type_aux r nr t, u)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   430
      using nrole by (simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   431
    thus ?thesis using csp nrole by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   432
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   433
  ultimately show ?thesis by (rule_tac ext, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   434
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   435
dcde836219bc add thy files
chunhan
parents:
diff changeset
   436
lemma cp2sproc_crole:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   437
  "valid (ChangeRole p r # s) \<Longrightarrow> cp2sproc (ChangeRole p r # s) = (cp2sproc s) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   438
     (p := (case (cp2sproc s p) of
dcde836219bc add thy files
chunhan
parents:
diff changeset
   439
              Some (r',fr,pt,u) \<Rightarrow> Some (r,fr,pt,u)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   440
            | _                 \<Rightarrow> None)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   441
      )"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   442
apply (frule valid_cons, frule valid_os, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   443
apply (frule current_proc_has_sproc, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   444
apply (rule ext, auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   445
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   446
dcde836219bc add thy files
chunhan
parents:
diff changeset
   447
lemma cp2sproc_exec:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   448
  assumes vs: "valid (Execute p f # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   449
  shows "cp2sproc (Execute p f # s) = (cp2sproc s) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   450
     (p := (case (cp2sproc s p, source_dir s f) of 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   451
              (Some (r,fr,pt,u), Some sd) \<Rightarrow> (
dcde836219bc add thy files
chunhan
parents:
diff changeset
   452
     case (exec_role_aux r sd u, erole_functor init_file_forcedrole InheritUpMixed sd) of
dcde836219bc add thy files
chunhan
parents:
diff changeset
   453
       (Some r', Some fr') \<Rightarrow> Some (r', fr', exec_type_aux r pt, u)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   454
     | _                   \<Rightarrow> None           )
dcde836219bc add thy files
chunhan
parents:
diff changeset
   455
            | _                \<Rightarrow> None))" (is "?lhs = ?rhs")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   456
proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   457
  have os: "os_grant s (Execute p f)" and vs': "valid s" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   458
    by (auto dest:valid_cons valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   459
  have "\<And> x. x \<noteq> p \<Longrightarrow> ?lhs x = ?rhs x"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   460
    by (auto simp:type_of_process.simps split:option.splits t_role.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   461
  moreover have "?lhs p = ?rhs p"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   462
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   463
    from os have p_in: "p \<in> current_procs s" by (simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   464
    then obtain r fr t u where csp: "cp2sproc s p = Some (r, fr, t, u)" using vs'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   465
      by (drule_tac current_proc_has_sproc, auto)  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   466
    from os have f_in: "f \<in> current_files s" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   467
    then obtain sd where sdir: "source_dir s f = Some sd" using vs'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   468
      by (drule_tac current_file_has_sd, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   469
    have "currentrole (Execute p f # s) p \<noteq> None" using vs p_in
dcde836219bc add thy files
chunhan
parents:
diff changeset
   470
      by (rule_tac notI, drule_tac current_proc_has_role', simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   471
    then obtain nr where nrole: "currentrole (Execute p f # s) p = Some nr" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   472
    have "proc_forcedrole (Execute p f # s) p \<noteq> None" using vs p_in
dcde836219bc add thy files
chunhan
parents:
diff changeset
   473
      by (rule_tac notI, drule_tac current_proc_has_forcedrole', simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   474
    then obtain nfr where nfrole: "proc_forcedrole (Execute p f # s) p = Some nfr" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   475
    have nr_eq: "currentrole (Execute p f # s) p = exec_role_aux r sd u" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   476
      using csp f_in sdir vs' by (simp only:exec_role_aux_valid')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   477
    moreover have "type_of_process (Execute p f # s) p = Some (exec_type_aux r t)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   478
      using csp by (simp only:exec_type_aux_valid')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   479
    moreover have nfr_eq: "proc_forcedrole (Execute p f # s) p = 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   480
                           erole_functor init_file_forcedrole InheritUpMixed sd" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   481
      using sdir vs' f_in
dcde836219bc add thy files
chunhan
parents:
diff changeset
   482
      apply (frule_tac source_dir_in_init, drule_tac source_dir_not_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   483
      by (simp add:undel_initf_keeps_efrole sd_deter_efrole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   484
    moreover have "owner (Execute p f # s) p = Some u" using csp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   485
      by (auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   486
    ultimately have "cp2sproc (Execute p f # s) p = Some (nr, nfr, exec_type_aux r t, u)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   487
      using nrole nfrole by (simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   488
    moreover have "exec_role_aux r sd u = Some nr" using nrole nr_eq by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   489
    moreover have "erole_functor init_file_forcedrole InheritUpMixed sd = Some nfr"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   490
      using nfrole nfr_eq by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   491
    ultimately show ?thesis using csp sdir by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   492
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   493
  ultimately show ?thesis by (rule_tac ext, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   494
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   495
dcde836219bc add thy files
chunhan
parents:
diff changeset
   496
lemma cp2sproc_clone:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   497
  "valid (Clone p p' # s) \<Longrightarrow> cp2sproc (Clone p p' # s) = (cp2sproc s) (p' := 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   498
      (case (cp2sproc s p) of
dcde836219bc add thy files
chunhan
parents:
diff changeset
   499
         Some (r, fr, pt, u) \<Rightarrow> Some (r, fr, clone_type_aux r pt, u)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   500
       | _                   \<Rightarrow> None))"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   501
apply (frule valid_cons, frule valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   502
apply (rule ext, auto split:option.splits t_rc_proc_type.splits 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   503
                       simp:pct_def clone_type_aux_def
dcde836219bc add thy files
chunhan
parents:
diff changeset
   504
                       dest:current_proc_has_type default_process_create_type_valid)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   505
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   506
dcde836219bc add thy files
chunhan
parents:
diff changeset
   507
lemma cp2sproc_other:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   508
  "\<lbrakk>valid (e # s); \<forall> p f. e \<noteq> Execute p f; \<forall> p p'. e \<noteq> Clone p p';
dcde836219bc add thy files
chunhan
parents:
diff changeset
   509
    \<forall> p r. e \<noteq> ChangeRole p r; \<forall> p u. e \<noteq> ChangeOwner p u\<rbrakk> \<Longrightarrow> cp2sproc (e#s) = cp2sproc s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   510
by (case_tac e, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   511
dcde836219bc add thy files
chunhan
parents:
diff changeset
   512
lemmas cp2sproc_simps = cp2sproc_exec cp2sproc_chown cp2sproc_crole cp2sproc_clone cp2sproc_other
dcde836219bc add thy files
chunhan
parents:
diff changeset
   513
dcde836219bc add thy files
chunhan
parents:
diff changeset
   514
lemma obj2sobj_file: "obj2sobj s obj = SFile sf fopt \<Longrightarrow> \<exists> f. obj = File f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   515
by (case_tac obj, case_tac [!] s, auto split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   516
dcde836219bc add thy files
chunhan
parents:
diff changeset
   517
lemma obj2sobj_proc: "obj2sobj s obj = SProc sp popt \<Longrightarrow> \<exists> p. obj = Proc p"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   518
by (case_tac obj, case_tac [!] s, auto split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   519
dcde836219bc add thy files
chunhan
parents:
diff changeset
   520
lemma obj2sobj_ipc: "obj2sobj s obj = SIPC si iopt \<Longrightarrow> \<exists> i. obj = IPC i"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   521
by (case_tac obj, case_tac [!] s, auto split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   522
dcde836219bc add thy files
chunhan
parents:
diff changeset
   523
lemma obj2sobj_file': 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   524
  "\<lbrakk>obj2sobj s (File f) = sobj; sobj \<noteq> Unknown\<rbrakk> \<Longrightarrow> \<exists> sf srf. sobj = SFile sf srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   525
by (case_tac sobj, case_tac [!] s, auto split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   526
dcde836219bc add thy files
chunhan
parents:
diff changeset
   527
lemma obj2sobj_proc': 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   528
  "\<lbrakk>obj2sobj s (Proc p) = sobj; sobj \<noteq> Unknown\<rbrakk> \<Longrightarrow> \<exists> sp srp. sobj = SProc sp srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   529
by (case_tac sobj, case_tac [!] s, auto split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   530
dcde836219bc add thy files
chunhan
parents:
diff changeset
   531
lemma obj2sobj_ipc': 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   532
  "\<lbrakk>obj2sobj s (IPC i) = sobj; sobj \<noteq> Unknown\<rbrakk> \<Longrightarrow> \<exists> si sri. sobj =  SIPC si sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   533
by (case_tac sobj, case_tac [!] s, auto split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   534
dcde836219bc add thy files
chunhan
parents:
diff changeset
   535
lemma obj2sobj_file_remains_cons:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   536
  assumes vs: "valid (e#s)" and exf: "f \<in> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   537
  and SF: "obj2sobj s (File f) = SFile sf srf" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   538
  and notdeled: "\<not> deleted (File f) (e#s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   539
  shows "obj2sobj (e#s) (File f) = SFile sf srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   540
proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   541
  from vs have os:"os_grant s e" and vs': "valid s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   542
    by (auto dest:valid_cons valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   543
  from notdeled exf have exf': "f \<in> current_files (e#s)" by (case_tac e, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   544
  have "etype_of_file (e # s) f = etype_of_file s f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   545
    using os vs vs' exf exf' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   546
    apply (case_tac e, auto simp:etype_of_file_def split:option.splits) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   547
    by (auto dest:ancient_file_in_current intro!:etype_aux_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   548
  moreover have "source_dir (e # s) f = source_dir s f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   549
    using os vs vs' exf exf'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   550
    by (case_tac e, auto simp:source_dir_simps dest:source_dir_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   551
  ultimately show ?thesis using vs SF notdeled 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   552
    by (auto split:if_splits option.splits dest:not_deleted_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   553
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   554
dcde836219bc add thy files
chunhan
parents:
diff changeset
   555
lemma obj2sobj_file_remains_cons':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   556
  "\<lbrakk>valid (e#s); f \<in> current_files s; obj2sobj s (File f) = SFile sf srf; no_del_event (e#s)\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   557
   \<Longrightarrow> obj2sobj (e#s) (File f) = SFile sf srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   558
by (auto intro!:obj2sobj_file_remains_cons nodel_imp_un_deleted
dcde836219bc add thy files
chunhan
parents:
diff changeset
   559
       simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   560
dcde836219bc add thy files
chunhan
parents:
diff changeset
   561
lemma obj2sobj_file_remains':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   562
  "\<lbrakk>obj2sobj s (File f) = sobj; sobj \<noteq> Unknown; valid (e#s); f \<in> current_files s;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   563
    no_del_event (e#s)\<rbrakk> \<Longrightarrow> obj2sobj (e#s) (File f) = sobj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   564
apply (frule obj2sobj_file', simp, (erule exE)+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   565
apply (simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   566
apply (erule obj2sobj_file_remains_cons', simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   567
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   568
dcde836219bc add thy files
chunhan
parents:
diff changeset
   569
lemma obj2sobj_file_remains_app:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   570
  "\<lbrakk>obj2sobj s (File f) = SFile sf srf; valid (s' @ s); f \<in> current_files s;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   571
    \<not> deleted (File f) (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (File f) = SFile sf srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   572
apply (induct s', simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   573
apply (simp only:cons_app_simp_aux)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   574
apply (frule valid_cons, frule not_deleted_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   575
apply (drule_tac s = "s'@s" in obj2sobj_file_remains_cons, auto simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   576
apply (drule_tac obj = "File f" in not_deleted_imp_exists', simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   577
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   578
dcde836219bc add thy files
chunhan
parents:
diff changeset
   579
lemma obj2sobj_file_remains_app':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   580
  "\<lbrakk>obj2sobj s (File f) = SFile sf srf; valid (s' @ s); f \<in> current_files s;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   581
    no_del_event (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (File f) = SFile sf srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   582
by (auto intro!:obj2sobj_file_remains_app nodel_imp_un_deleted
dcde836219bc add thy files
chunhan
parents:
diff changeset
   583
       simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   584
dcde836219bc add thy files
chunhan
parents:
diff changeset
   585
lemma obj2sobj_file_remains'':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   586
  "\<lbrakk>obj2sobj s (File f) = sobj; sobj \<noteq> Unknown; valid (s'@s); f \<in> current_files s;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   587
    no_del_event (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (File f) = sobj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   588
apply (frule obj2sobj_file', simp, (erule exE)+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   589
apply (simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   590
apply (erule obj2sobj_file_remains_app', simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   591
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   592
dcde836219bc add thy files
chunhan
parents:
diff changeset
   593
lemma obj2sobj_file_remains''':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   594
  "\<lbrakk>obj2sobj s (File f) = sobj; sobj \<noteq> Unknown; valid (s'@s); f \<in> current_files s;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   595
    \<not>deleted (File f) (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (File f) = sobj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   596
apply (frule obj2sobj_file', simp, (erule exE)+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   597
apply (simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   598
by (erule obj2sobj_file_remains_app, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   599
dcde836219bc add thy files
chunhan
parents:
diff changeset
   600
lemma obj2sobj_ipc_remains_cons:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   601
  "\<lbrakk>valid (e#s); i \<in> current_ipcs s; obj2sobj s (IPC i) = SIPC si sri; \<not> deleted (IPC i) (e#s)\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   602
  \<Longrightarrow> obj2sobj (e#s) (IPC i) = SIPC si sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   603
apply (frule valid_cons, frule valid_os, case_tac e)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   604
by (auto simp:ni_init_deled ni_notin_curi split:option.splits
dcde836219bc add thy files
chunhan
parents:
diff changeset
   605
        dest!:current_proc_has_role')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   606
dcde836219bc add thy files
chunhan
parents:
diff changeset
   607
lemma obj2sobj_ipc_remains_cons':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   608
  "\<lbrakk>valid (e#s); i \<in> current_ipcs s; obj2sobj s (IPC i) = SIPC si sri; no_del_event (e#s)\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   609
  \<Longrightarrow> obj2sobj (e#s) (IPC i) = SIPC si sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   610
by (auto intro!:obj2sobj_ipc_remains_cons nodel_imp_un_deleted
dcde836219bc add thy files
chunhan
parents:
diff changeset
   611
       simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   612
dcde836219bc add thy files
chunhan
parents:
diff changeset
   613
lemma obj2sobj_ipc_remains':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   614
  "\<lbrakk>obj2sobj s (IPC i) = sobj; sobj \<noteq> Unknown; valid (e#s); i \<in> current_ipcs s; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   615
    no_del_event (e#s)\<rbrakk> \<Longrightarrow> obj2sobj (e#s) (IPC i) = sobj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   616
apply (frule obj2sobj_ipc', simp, (erule exE)+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   617
apply (simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   618
apply (erule obj2sobj_ipc_remains_cons', simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   619
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   620
dcde836219bc add thy files
chunhan
parents:
diff changeset
   621
lemma obj2sobj_ipc_remains_app:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   622
  "\<lbrakk>obj2sobj s (IPC i) = SIPC si sri; valid (s'@s); i \<in> current_ipcs s; \<not> deleted (IPC i) (s'@s)\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   623
  \<Longrightarrow> obj2sobj (s'@s) (IPC i) = SIPC si sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   624
apply (induct s', simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   625
apply (simp only:cons_app_simp_aux)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   626
apply (frule valid_cons, frule not_deleted_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   627
apply (drule_tac s = "s'@s" in obj2sobj_ipc_remains_cons, auto simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   628
apply (drule_tac obj = "IPC i" in not_deleted_imp_exists', simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   629
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   630
dcde836219bc add thy files
chunhan
parents:
diff changeset
   631
lemma obj2sobj_ipc_remains_app':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   632
  "\<lbrakk>obj2sobj s (IPC i) = SIPC si sri; valid (s'@s); i \<in> current_ipcs s; no_del_event (s'@s)\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   633
  \<Longrightarrow> obj2sobj (s'@s) (IPC i) = SIPC si sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   634
by (auto intro!:obj2sobj_ipc_remains_app nodel_imp_un_deleted
dcde836219bc add thy files
chunhan
parents:
diff changeset
   635
       simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   636
dcde836219bc add thy files
chunhan
parents:
diff changeset
   637
lemma obj2sobj_ipc_remains'':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   638
  "\<lbrakk>obj2sobj s (IPC i) = sobj; sobj \<noteq> Unknown; valid (s'@s); i \<in> current_ipcs s; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   639
    no_del_event (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (IPC i) = sobj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   640
apply (frule obj2sobj_ipc', simp, (erule exE)+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   641
apply (simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   642
apply (erule obj2sobj_ipc_remains_app', simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   643
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   644
dcde836219bc add thy files
chunhan
parents:
diff changeset
   645
lemma obj2sobj_ipc_remains''':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   646
  "\<lbrakk>obj2sobj s (IPC i) = sobj; sobj \<noteq> Unknown; valid (s'@s); i \<in> current_ipcs s; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   647
    \<not> deleted (IPC i) (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (IPC i) = sobj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   648
apply (frule obj2sobj_ipc', simp, (erule exE)+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   649
apply (simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   650
apply (erule obj2sobj_ipc_remains_app, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   651
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   652
dcde836219bc add thy files
chunhan
parents:
diff changeset
   653
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
   654
dcde836219bc add thy files
chunhan
parents:
diff changeset
   655
context tainting_s_sound begin
dcde836219bc add thy files
chunhan
parents:
diff changeset
   656
dcde836219bc add thy files
chunhan
parents:
diff changeset
   657
lemma cp2sproc_clone':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   658
  "valid (Clone p p' # s) \<Longrightarrow> cp2sproc (Clone p p' # s) = (cp2sproc s) (p' := cp2sproc s p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   659
by (drule cp2sproc_clone, auto split:option.splits simp:clone_type_unchange clone_type_aux_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   660
dcde836219bc add thy files
chunhan
parents:
diff changeset
   661
lemmas cp2sproc_simps' =  cp2sproc_exec cp2sproc_chown cp2sproc_crole cp2sproc_clone' cp2sproc_other
dcde836219bc add thy files
chunhan
parents:
diff changeset
   662
dcde836219bc add thy files
chunhan
parents:
diff changeset
   663
lemma clone_sobj_keeps_same:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   664
  "valid (Clone p p' # s) \<Longrightarrow> obj2sobj (Clone p p' # s) (Proc p') = obj2sobj s (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   665
apply (frule valid_cons, frule valid_os, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   666
apply (auto split:option.splits t_rc_proc_type.splits
dcde836219bc add thy files
chunhan
parents:
diff changeset
   667
             dest:current_proc_has_role current_proc_has_forcedrole 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   668
                  current_proc_has_type current_proc_has_owner default_process_create_type_valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   669
             simp:pct_def clone_type_unchange)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   670
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   671
dcde836219bc add thy files
chunhan
parents:
diff changeset
   672
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
   673
dcde836219bc add thy files
chunhan
parents:
diff changeset
   674
end