1
+ − 1
(*<*)
+ − 2
theory Co2sobj_prop
+ − 3
imports Main Flask Flask_type Static Static_type Sectxt_prop Init_prop Current_files_prop Current_sockets_prop
+ − 4
begin
+ − 5
(*<*)
+ − 6
+ − 7
context tainting_s begin
+ − 8
+ − 9
(************ simpset for cf2sfile ***************)
+ − 10
+ − 11
declare get_parentfs_ctxts.simps [simp del]
+ − 12
4
+ − 13
lemma cf2sfile_open_none1:
1
+ − 14
"valid (Open p f flag fd None # s) \<Longrightarrow> cf2sfile (Open p f flag fd None # s) f b = cf2sfile s f b"
+ − 15
apply (frule vd_cons, frule vt_grant_os)
4
+ − 16
apply (frule noroot_events, case_tac f, simp)
+ − 17
apply (auto split:if_splits option.splits dest!:current_has_sec' dest:parentf_in_current'
1
+ − 18
simp:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ − 19
get_parentfs_ctxts_simps)
4
+ − 20
done
+ − 21
+ − 22
lemma cf2sfile_open_none2:
+ − 23
"valid (Open p f flag fd None # s) \<Longrightarrow> cf2sfile (Open p f flag fd None # s) f' b = cf2sfile s f' b"
+ − 24
apply (frule vd_cons, frule vt_grant_os)
+ − 25
apply (induct f', simp add:cf2sfile_def)
+ − 26
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ − 27
get_parentfs_ctxts_simps)
+ − 28
done
+ − 29
+ − 30
lemma cf2sfile_open_none:
+ − 31
"valid (Open p f flag fd None # s) \<Longrightarrow> cf2sfile (Open p f flag fd None # s) = cf2sfile s"
+ − 32
apply (rule ext, rule ext)
+ − 33
apply (simp add:cf2sfile_open_none1 cf2sfile_open_none2)
+ − 34
done
1
+ − 35
3
+ − 36
lemma sroot_only:
+ − 37
"cf2sfile s [] = (\<lambda> b. if b then None else Some sroot)"
+ − 38
by (rule ext, simp add:cf2sfile_def)
+ − 39
1
+ − 40
lemma cf2sfile_open_some1:
+ − 41
"\<lbrakk>valid (Open p f flag fd (Some inum) # s); f' \<in> current_files s\<rbrakk>
+ − 42
\<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'"
+ − 43
apply (frule vd_cons, frule vt_grant_os, frule noroot_events, rule ext)
4
+ − 44
apply (case_tac "f = f'", simp)
+ − 45
apply (induct f', simp add:sroot_only, simp)
3
+ − 46
apply (frule parentf_in_current', simp+)
4
+ − 47
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
1
+ − 48
get_parentfs_ctxts_simps)
+ − 49
done
+ − 50
+ − 51
lemma cf2sfile_open_some2:
+ − 52
"\<lbrakk>valid (Open p f flag fd (Some inum) # s); is_file s f'\<rbrakk>
+ − 53
\<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' True = cf2sfile s f' True"
+ − 54
apply (frule vd_cons, drule is_file_in_current)
+ − 55
by (simp add:cf2sfile_open_some1)
+ − 56
+ − 57
lemma cf2sfile_open_some3:
+ − 58
"\<lbrakk>valid (Open p f flag fd (Some inum) # s); is_dir s f'\<rbrakk>
+ − 59
\<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f' False = cf2sfile s f' False"
+ − 60
apply (frule vd_cons, drule is_dir_in_current)
+ − 61
by (simp add:cf2sfile_open_some1)
+ − 62
+ − 63
lemma cf2sfile_open_some:
+ − 64
"valid (Open p f flag fd (Some inum) # s) \<Longrightarrow> cf2sfile (Open p f flag fd (Some inum) # s) f True = (
+ − 65
case (parent f) of
+ − 66
Some pf \<Rightarrow> (case (sectxt_of_obj (Open p f flag fd (Some inum) # s) (O_file f), sectxt_of_obj s (O_dir pf),
+ − 67
get_parentfs_ctxts s pf) of
+ − 68
(Some sec, Some psec, Some asecs) \<Rightarrow> Some (Created, sec, Some psec, set asecs)
+ − 69
| _ \<Rightarrow> None)
+ − 70
| None \<Rightarrow> None)"
+ − 71
apply (frule vd_cons, frule vt_grant_os, frule noroot_events)
4
+ − 72
apply (case_tac f, simp)
+ − 73
apply (simp add:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ − 74
get_parentfs_ctxts_simps)
+ − 75
apply (rule impI)
+ − 76
thm not_deleted_imp_exists
1
+ − 77
apply (auto split:if_splits option.splits dest!:current_has_sec'
+ − 78
simp:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps
+ − 79
get_parentfs_ctxts_simps)
+ − 80
+ − 81
+ − 82
+ − 83
lemma cf2sfile_open:
+ − 84
"valid (Open p f flag fd opt # s) \<Longrightarrow> cf2sfile (Open p f flag fd opt # s) = (\<lambda> f' b.
+ − 85
if (opt = None) then cf2sfile s f' b
+ − 86
else if (f' = f \<and> b = True)
+ − 87
then (case (parent f) of
+ − 88
Some pf \<Rightarrow> (case (sectxt_of_obj s (O_file f), sectxt_of_obj s (O_dir pf),
+ − 89
get_parentfs_ctxts s pf) of
+ − 90
(Some sec, Some psec, Some asecs) \<Rightarrow>
+ − 91
Some (if (\<not> deleted (O_file f) s \<and> f \<in> init_files)
+ − 92
then Init f else Created, sec, Some psec, set asecs)
+ − 93
| _ \<Rightarrow> None)
+ − 94
| None \<Rightarrow> None)
+ − 95
else cf2sfile s f' b)"
+ − 96
apply (frule vd_cons, frule vt_grant_os, rule ext, rule ext)
+ − 97
apply (auto split:if_splits option.splits
+ − 98
simp:sectxt_of_obj_simps)
+ − 99
+ − 100
+ − 101
+ − 102
+ − 103
+ − 104
+ − 105
+ − 106
+ − 107
+ − 108
+ − 109
+ − 110
+ − 111
lemma cf2sfile_keep_path: "\<lbrakk>f \<preceq> f'; \<tau> \<in> vt rc_cs\<rbrakk> \<Longrightarrow> cf2sfile \<tau> f \<preceq> cf2sfile \<tau> f'"
+ − 112
apply (induct f', simp add:no_junior_def)
+ − 113
apply (case_tac "f = a # f'", simp add:cf2sfile.simps)
+ − 114
apply (drule no_junior_noteq, simp, simp add:cf2sfile.simps)
+ − 115
done
+ − 116
+ − 117
lemma ckp'_aux: "\<not> f \<preceq> a # f' \<Longrightarrow> \<not> f \<preceq> f'"
+ − 118
by (auto simp:no_junior_def)
+ − 119
+ − 120
lemma cf2sfile_keep_path': "\<lbrakk>\<not> f \<preceq> f'; \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>; f' \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> \<not> cf2sfile \<tau> f \<preceq> cf2sfile \<tau> f'"
+ − 121
apply (induct f', simp add:no_junior_def cf2sfile.simps, rule notI, drule cf2sfile_root_file, simp)
+ − 122
apply (case_tac "f = a # f'", simp add:cf2sfile.simps)
+ − 123
apply (rule notI)
+ − 124
apply (frule ckp'_aux, simp, frule parentf_in_current', simp+)
+ − 125
apply (case_tac "cf2sfile \<tau> f = cf2sfile \<tau> (a # f')", drule cf2sfile_inj, simp+)
+ − 126
apply (simp add:cf2sfile.simps)
+ − 127
by (drule no_junior_noteq, simp+)
+ − 128
+ − 129
lemma cf2sfile_keep_path'': "\<lbrakk>cf2sfile \<tau> f \<preceq> cf2sfile \<tau> f'; \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>; f' \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> f \<preceq> f'"
+ − 130
using cf2sfile_keep_path'
+ − 131
by (auto)
+ − 132
+ − 133
lemma cf2sfile_open_some': "\<lbrakk>f \<in> current_files \<tau>; Open p f' flags fd (Some inum) # \<tau> \<in> vt rc_cs\<rbrakk> \<Longrightarrow> cf2sfile (Open p f' flags fd (Some inum) # \<tau>) f = cf2sfile \<tau> f"
+ − 134
apply (frule vt_cons')
+ − 135
apply (drule vt_grant_os)
+ − 136
apply (induct f)
+ − 137
apply (simp add:cf2sfile.simps)+
+ − 138
apply (frule parentf_in_current', simp)
+ − 139
apply (auto simp:os_grant.simps index_of_file.simps split:option.splits)
+ − 140
done
+ − 141
+ − 142
lemma cf2sfile_open_some: "\<lbrakk>Open p f flags fd (Some inum) # \<tau> \<in> vt rc_cs; parent f = Some pf\<rbrakk>
+ − 143
\<Longrightarrow> cf2sfile (Open p f flags fd (Some inum) # \<tau>) f = (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf))"
+ − 144
apply (case_tac f, simp)
+ − 145
apply (frule vt_grant_os)
+ − 146
apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_open_some')
+ − 147
done
+ − 148
+ − 149
lemma cf2sfile_open_none: "cf2sfile (Open p f flags fd None # \<tau>) f' = cf2sfile \<tau> f'"
+ − 150
apply (induct f')
+ − 151
by (simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits)+
+ − 152
+ − 153
lemma cf2sfile_open: "\<lbrakk>Open p f flags fd opt # \<tau> \<in> vt rc_cs; f' \<in> current_files (Open p f flags fd opt # \<tau>)\<rbrakk> \<Longrightarrow>
+ − 154
cf2sfile (Open p f flags fd opt # \<tau>) f' = (
+ − 155
if (opt = None) then cf2sfile \<tau> f'
+ − 156
else if (f' = f) then (case (parent f) of
+ − 157
Some pf \<Rightarrow> (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf))
+ − 158
| _ \<Rightarrow> [])
+ − 159
else cf2sfile \<tau> f' )"
+ − 160
apply (frule vt_grant_os)
+ − 161
by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_open_none cf2sfile_open_some cf2sfile_open_some' split:if_splits option.splits)
+ − 162
+ − 163
lemma cf2sfile_mkdir_some': "\<lbrakk>Mkdir p f' inum # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> cf2sfile (Mkdir p f' inum # \<tau>) f = cf2sfile \<tau> f"
+ − 164
apply (frule vt_cons', drule vt_grant_os)
+ − 165
apply (induct f, (simp add:cf2sfile.simps)+)
+ − 166
apply (frule parentf_in_current', simp)
+ − 167
apply (auto simp:os_grant.simps index_of_file.simps split:option.splits)
+ − 168
done
+ − 169
+ − 170
lemma cf2sfile_mkdir_some: "\<lbrakk>Mkdir p f inum # \<tau> \<in> vt rc_cs; parent f = Some pf\<rbrakk>
+ − 171
\<Longrightarrow> cf2sfile (Mkdir p f inum # \<tau>) f = (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf))"
+ − 172
apply (case_tac f, simp)
+ − 173
apply (frule vt_grant_os)
+ − 174
apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_mkdir_some')
+ − 175
done
+ − 176
+ − 177
lemma cf2sfile_mkdir: "\<lbrakk>Mkdir p f inum # \<tau> \<in> vt rc_cs; f' \<in> current_files (Mkdir p f inum # \<tau>)\<rbrakk> \<Longrightarrow>
+ − 178
cf2sfile (Mkdir p f inum # \<tau>) f' = (
+ − 179
if (f' = f) then (case (parent f) of
+ − 180
Some pf \<Rightarrow> (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf))
+ − 181
| _ \<Rightarrow> [])
+ − 182
else cf2sfile \<tau> f' ) "
+ − 183
apply (frule vt_grant_os)
+ − 184
by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_mkdir_some cf2sfile_mkdir_some' split:if_splits option.splits)
+ − 185
+ − 186
lemma cf2sfile_linkhard_none: "\<lbrakk>LinkHard p f\<^isub>1 f\<^isub>2 # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>) f = cf2sfile \<tau> f"
+ − 187
apply (frule vt_cons')
+ − 188
apply (drule vt_grant_os)
+ − 189
apply (induct f)
+ − 190
apply (simp add:cf2sfile.simps)+
+ − 191
apply (frule parentf_in_current', simp)
+ − 192
apply (auto simp:os_grant.simps index_of_file.simps split:option.splits)
+ − 193
done
+ − 194
+ − 195
lemma cf2sfile_linkhard_some:
+ − 196
"\<lbrakk>LinkHard p f\<^isub>1 f\<^isub>2 # \<tau> \<in> vt rc_cs; parent f\<^isub>2 = Some pf\<^isub>2\<rbrakk> \<Longrightarrow> cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>) f\<^isub>2 = (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>2))"
+ − 197
apply (case_tac f\<^isub>2, simp)
+ − 198
apply (frule vt_grant_os)
+ − 199
apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_linkhard_none)
+ − 200
done
+ − 201
+ − 202
lemma cf2sfile_linkhard: "\<lbrakk>LinkHard p f\<^isub>1 f\<^isub>2 # \<tau> \<in> vt rc_cs; f \<in> current_files (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>)\<rbrakk> \<Longrightarrow>
+ − 203
cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>) f = (
+ − 204
if (f = f\<^isub>2) then (case (parent f\<^isub>2) of
+ − 205
Some pf\<^isub>2 \<Rightarrow> SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>2)
+ − 206
| _ \<Rightarrow> [])
+ − 207
else cf2sfile \<tau> f )"
+ − 208
apply (frule vt_grant_os)
+ − 209
by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_linkhard_none cf2sfile_linkhard_some split:if_splits option.splits)
+ − 210
+ − 211
+ − 212
lemma no_junior_aux: "\<not> f\<^isub>2 \<preceq> a # f \<Longrightarrow> \<not> f\<^isub>2 \<preceq> f"
+ − 213
by (auto simp:no_junior_def)
+ − 214
+ − 215
lemma cf2sfile_rename_aux: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>\<rbrakk> \<Longrightarrow> f \<noteq> f\<^isub>3 \<and> \<not> f\<^isub>3 \<preceq> f"
+ − 216
apply (frule vt_grant_os, simp add:os_grant.simps, (erule exE|erule conjE)+)
+ − 217
apply (rule conjI, rule notI, simp)
+ − 218
apply (rule notI, drule vt_cons', simp add:ancenf_in_current)
+ − 219
done
+ − 220
+ − 221
lemma cf2sfile_rename'1:
+ − 222
"\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>; \<not> (f\<^isub>2 \<preceq> f)\<rbrakk> \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = cf2sfile \<tau> f"
+ − 223
apply (frule vt_cons',frule vt_grant_os, induct f)
+ − 224
apply (simp add:cf2sfile.simps)
+ − 225
apply (frule parentf_in_current', simp, simp)
+ − 226
apply (frule no_junior_aux, simp)
+ − 227
apply (simp add:os_grant.simps, (erule exE|erule conjE)+)
+ − 228
apply (drule cf2sfile_rename_aux, simp, erule conjE)
+ − 229
apply (auto simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits)
+ − 230
done
+ − 231
+ − 232
lemma cf2sfile_rename'2:
+ − 233
"\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; \<not> (f\<^isub>3 \<preceq> f)\<rbrakk> \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = cf2sfile \<tau> f"
+ − 234
apply (frule vt_cons', induct f)
+ − 235
by (auto simp add:index_of_file.simps cf2sfile.simps no_junior_def split:option.splits nat.splits)
+ − 236
+ − 237
lemma cf2sfile_rename1:
+ − 238
"\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; parent f\<^isub>3 = Some pf\<^isub>3\<rbrakk> \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f\<^isub>3 = SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)"
+ − 239
apply (case_tac f\<^isub>3, simp add:cf2sfile.simps)
+ − 240
apply (auto dest!:cf2sfile_rename'2 simp add:no_junior_def cf2sfile.simps index_of_file.simps split:option.splits)
+ − 241
done
+ − 242
+ − 243
lemma index_of_file_rename1: "\<lbrakk>f\<^isub>2 \<preceq> f; f \<noteq> f\<^isub>2\<rbrakk> \<Longrightarrow> index_of_file (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) (file_after_rename f\<^isub>2 f\<^isub>3 f) = index_of_file \<tau> f"
+ − 244
apply (clarsimp simp add:index_of_file.simps split:option.splits)
+ − 245
by (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop3, simp, erule conjE, simp add:file_renaming_prop5)
+ − 246
+ − 247
lemma cf2sfile_rename2: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; (file_before_rename f\<^isub>2 f\<^isub>3 f) \<in> current_files \<tau>; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>3 \<preceq> f\<rbrakk>
+ − 248
\<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> (file_before_rename f\<^isub>2 f\<^isub>3 f))"
+ − 249
apply (induct f, simp add:no_junior_def)
+ − 250
apply (case_tac "a # f = f\<^isub>3", simp)
+ − 251
apply (drule cf2sfile_rename1, simp, simp add:file_renaming_prop0' file_renaming_prop0)
+ − 252
+ − 253
apply (frule no_junior_noteq, simp, simp)
+ − 254
apply (frule_tac file_renaming_prop1')
+ − 255
apply (frule_tac f = f\<^isub>2 and \<tau> = \<tau> in cf2sfile_keep_path, simp add:vt_cons')
+ − 256
apply (frule_tac f\<^isub>2 = f\<^isub>2 and a = a in file_renaming_prop8')
+ − 257
apply (frule_tac f\<^isub>2 = f\<^isub>2 and a = a in file_renaming_prop6')
+ − 258
apply (frule_tac f = "file_before_rename f\<^isub>2 f\<^isub>3 (a # f)" and \<tau> = \<tau> and f\<^isub>2 = f\<^isub>2 and f\<^isub>3 = f\<^isub>3 and p = p in index_of_file_rename1, simp add:file_before_rename_def)
+ − 259
apply (drule_tac f\<^isub>3 = f\<^isub>3 and f\<^isub>1 = f and a = a and f\<^isub>2 = f\<^isub>2 in file_renaming_prop9', simp)
+ − 260
apply (drule parentf_in_current', simp add:vt_cons')
+ − 261
apply (simp add:cf2sfile.simps)
+ − 262
apply (erule file_renaming_prop6[THEN sym])
+ − 263
done
+ − 264
+ − 265
lemma cf2sfile_rename2': "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>2 \<preceq> f\<rbrakk>
+ − 266
\<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> f)"
+ − 267
apply (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop5)
+ − 268
apply (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop1)
+ − 269
apply (drule_tac f = "file_after_rename f\<^isub>2 f\<^isub>3 f" in cf2sfile_rename2, simp+)
+ − 270
done
+ − 271
+ − 272
lemma cf2sfile_rename3: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; f \<in> current_files \<tau>; parent f\<^isub>3 = Some pf\<^isub>3\<rbrakk>
+ − 273
\<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> f)"
+ − 274
apply (case_tac "f\<^isub>2 \<preceq> f")
+ − 275
apply (rule cf2sfile_rename2', simp+)
+ − 276
apply (frule vt_grant_os)
+ − 277
apply (frule_tac \<tau> = \<tau> in cf2sfile_keep_path', simp add:vt_cons', simp add:os_grant.simps, simp)
+ − 278
apply (simp add:file_after_rename_def)
+ − 279
by (rule cf2sfile_rename'1, simp+)
+ − 280
+ − 281
lemma cf2sfile_rename: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; f \<in> current_files (Rename p f\<^isub>2 f\<^isub>3 # \<tau>)\<rbrakk> \<Longrightarrow>
+ − 282
cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = (
+ − 283
if (f\<^isub>3 \<preceq> f) then (case (parent f\<^isub>3) of
+ − 284
Some pf\<^isub>3 \<Rightarrow> file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> (file_before_rename f\<^isub>2 f\<^isub>3 f))
+ − 285
| _ \<Rightarrow> [])
+ − 286
else cf2sfile \<tau> f )"
+ − 287
apply (frule vt_grant_os)
+ − 288
apply (case_tac "f = f\<^isub>3", clarsimp simp:os_grant.simps, drule cf2sfile_rename1, simp+, simp add:file_renaming_prop0' file_renaming_prop0)
+ − 289
+ − 290
apply (auto simp:os_grant.simps current_files_simps intro:cf2sfile_rename'2 cf2sfile_rename1 split:if_splits option.splits)
+ − 291
apply (rule cf2sfile_rename2, simp, drule rename_renaming_decom, simp+)
+ − 292
apply (drule_tac f\<^isub>2 = f\<^isub>2 and f\<^isub>1 = f\<^isub>1 and f\<^isub>3 = f\<^isub>3 in file_renaming_prop5, simp+)
+ − 293
done
+ − 294
+ − 295
lemma cf2sfile_other: "\<lbrakk>
+ − 296
\<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ − 297
\<forall> p f im. e \<noteq> Mkdir p f im;
+ − 298
\<forall> p f\<^isub>1 f\<^isub>2. e \<noteq> LinkHard p f\<^isub>1 f\<^isub>2;
+ − 299
\<forall> p f\<^isub>2 f\<^isub>3. e \<noteq> Rename p f\<^isub>2 f\<^isub>3\<rbrakk> \<Longrightarrow> cf2sfile (e # \<tau>) f' = cf2sfile \<tau> f'"
+ − 300
apply (induct f', simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits)
+ − 301
apply (case_tac e, auto simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits)
+ − 302
done
+ − 303
+ − 304
lemmas cf2sfile_nil = init_cf2sfile
+ − 305
+ − 306
lemma cf2sfile_nil': "f \<in> current_files [] \<Longrightarrow> cf2sfile [] f = map SInit f"
+ − 307
by (simp add:cf2sfile_nil current_files_simps)
+ − 308
+ − 309
lemmas cf2sfile_simps = cf2sfile_nil cf2sfile_nil' cf2sfile_open cf2sfile_mkdir cf2sfile_linkhard cf2sfile_rename cf2sfile_other
+ − 310
+ − 311
lemmas cf2sfile_simpss = cf2sfile_nil cf2sfile_nil' cf2sfile_open_some' cf2sfile_open_some cf2sfile_open_none cf2sfile_open cf2sfile_mkdir_some' cf2sfile_mkdir_some
+ − 312
cf2sfile_mkdir cf2sfile_linkhard_none cf2sfile_linkhard_some cf2sfile_linkhard cf2sfile_rename'1 cf2sfile_rename'2 cf2sfile_rename1
+ − 313
cf2sfile_rename2 cf2sfile_rename2' cf2sfile_rename3 cf2sfile_rename cf2sfile_other
+ − 314
+ − 315
lemma cf2sfile_open_some'_inum: "\<lbrakk>Open p f' flags fd (Some inum) # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im\<rbrakk> \<Longrightarrow> cf2sfile (Open p f' flags fd (Some inum) # \<tau>) f = cf2sfile \<tau> f"
+ − 316
by (simp add:cf2sfile_open_some' current_files_def)
+ − 317
+ − 318
lemma cf2sfile_mkdir_some'_inum: "\<lbrakk>Mkdir p f' inum # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im\<rbrakk> \<Longrightarrow> cf2sfile (Mkdir p f' inum # \<tau>) f = cf2sfile \<tau> f"
+ − 319
by (simp add:cf2sfile_mkdir_some' current_files_def)
+ − 320
+ − 321
lemma cf2sfile_linkhard_none_inum: "\<lbrakk>LinkHard p f\<^isub>1 f\<^isub>2 # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im\<rbrakk> \<Longrightarrow> cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \<tau>) f = cf2sfile \<tau> f"
+ − 322
by (simp add:cf2sfile_linkhard_none current_files_def)
+ − 323
+ − 324
lemma cf2sfile_rename'1_inum:
+ − 325
"\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im ; \<not> (f\<^isub>2 \<preceq> f)\<rbrakk> \<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = cf2sfile \<tau> f"
+ − 326
by (simp add:cf2sfile_rename'1 current_files_def)
+ − 327
+ − 328
lemma cf2sfile_rename2_inum: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; inum_of_file \<tau> (file_before_rename f\<^isub>2 f\<^isub>3 f) = Some im; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>3 \<preceq> f\<rbrakk>
+ − 329
\<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) f = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> (file_before_rename f\<^isub>2 f\<^isub>3 f))"
+ − 330
by (simp add:cf2sfile_rename2 current_files_def)
+ − 331
+ − 332
lemma cf2sfile_rename2'_inum: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>2 \<preceq> f\<rbrakk>
+ − 333
\<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> f)"
+ − 334
by (simp add:cf2sfile_rename2' current_files_def)
+ − 335
+ − 336
lemma cf2sfile_rename3_inum: "\<lbrakk>Rename p f\<^isub>2 f\<^isub>3 # \<tau> \<in> vt rc_cs; inum_of_file \<tau> f = Some im; parent f\<^isub>3 = Some pf\<^isub>3\<rbrakk>
+ − 337
\<Longrightarrow> cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \<tau>) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \<tau> f\<^isub>2) (SCrea (Suc (length \<tau>)) # (cf2sfile \<tau> pf\<^isub>3)) (cf2sfile \<tau> f)"
+ − 338
by (simp add:cf2sfile_rename3 current_files_def)
+ − 339
+ − 340
lemma cf2sfile_nil'_inum: "inum_of_file [] f = Some im \<Longrightarrow> cf2sfile [] f = map SInit f"
+ − 341
by (simp add:cf2sfile_nil' current_files_def)
+ − 342
+ − 343
lemmas cf2sfile_simps_inum = cf2sfile_nil cf2sfile_nil'_inum cf2sfile_open_some'_inum cf2sfile_open_some cf2sfile_open_none cf2sfile_open cf2sfile_mkdir_some'_inum cf2sfile_mkdir_some
+ − 344
cf2sfile_mkdir cf2sfile_linkhard_none_inum cf2sfile_linkhard_some cf2sfile_linkhard cf2sfile_rename'1_inum cf2sfile_rename'2 cf2sfile_rename1
+ − 345
cf2sfile_rename2_inum cf2sfile_rename2'_inum cf2sfile_rename3_inum cf2sfile_rename cf2sfile_other
+ − 346
+ − 347
(*************** cp2sproc simpset *********************)
+ − 348
+ − 349
lemma cp2sproc_nil: "p \<in> init_processes \<Longrightarrow> cp2sproc [] p = SInit p"
+ − 350
by (simp add:cp2sproc_def index_of_proc.simps)
+ − 351
+ − 352
lemma cp2sproc_nil': "p \<in> current_procs [] \<Longrightarrow> cp2sproc [] p = SInit p"
+ − 353
by (simp add:cp2sproc_nil current_procs.simps)
+ − 354
+ − 355
lemma cp2sproc_clone: "cp2sproc (Clone p p' # \<tau>) p'' = (
+ − 356
if (p'' = p') then SCrea (Suc (length \<tau>))
+ − 357
else cp2sproc \<tau> p'' )"
+ − 358
by (auto simp:cp2sproc_def index_of_proc.simps d2s_aux.simps)
+ − 359
+ − 360
lemma cp2sproc_other: "\<forall> p p'. e \<noteq> Clone p p' \<Longrightarrow> cp2sproc (e # \<tau>) p'' = cp2sproc \<tau> p''"
+ − 361
apply (case_tac e)
+ − 362
by (auto simp:cp2sproc_def index_of_proc.simps d2s_aux.simps)
+ − 363
+ − 364
lemmas cp2sproc_simps = cp2sproc_nil cp2sproc_nil' cp2sproc_clone cp2sproc_other
+ − 365
+ − 366
(******************** ch2sshm simpset **************************)
+ − 367
+ − 368
lemma ch2sshm_nil: "h \<in> init_shms \<Longrightarrow> ch2sshm [] h = SInit h"
+ − 369
by (simp add:ch2sshm_def index_of_shm.simps)
+ − 370
+ − 371
lemma ch2sshm_nil': "h \<in> current_shms [] \<Longrightarrow> ch2sshm [] h = SInit h"
+ − 372
by (simp add:ch2sshm_nil current_shms.simps)
+ − 373
+ − 374
lemma ch2sshm_createshm: "ch2sshm (CreateShM p h # \<tau>) h' = (if (h' = h) then SCrea (Suc (length \<tau>)) else ch2sshm \<tau> h')"
+ − 375
by (simp add:ch2sshm_def index_of_shm.simps d2s_aux.simps)
+ − 376
+ − 377
lemma ch2sshm_other: "\<forall> p h. e \<noteq> CreateShM p h \<Longrightarrow> ch2sshm (e # \<tau>) h' = ch2sshm \<tau> h'"
+ − 378
apply (case_tac e)
+ − 379
by (auto simp add:ch2sshm_def index_of_shm.simps d2s_aux.simps)
+ − 380
+ − 381
lemmas ch2sshm_simps = ch2sshm_nil ch2sshm_nil' ch2sshm_createshm ch2sshm_other
+ − 382
+ − 383
(********************* cm2smsg simpset ***********************)
+ − 384
+ − 385
lemma cm2smsg_nil: "m \<in> init_msgs \<Longrightarrow> cm2smsg [] m = SInit m"
+ − 386
by (simp add:cm2smsg_def index_of_msg.simps)
+ − 387
+ − 388
lemma cm2smsg_nil': "m \<in> current_msgs [] \<Longrightarrow> cm2smsg [] m = SInit m"
+ − 389
by (simp add:cm2smsg_nil current_msgs.simps)
+ − 390
+ − 391
lemma cm2smsg_createmsg: "cm2smsg (CreateMsg p m # \<tau>) m' = (if (m' = m) then SCrea (Suc (length \<tau>)) else cm2smsg \<tau> m')"
+ − 392
by (simp add:cm2smsg_def index_of_msg.simps d2s_aux.simps)
+ − 393
+ − 394
lemma cm2smsg_other: "\<forall> p m. e \<noteq> CreateMsg p m \<Longrightarrow> cm2smsg (e # \<tau>) m' = cm2smsg \<tau> m'"
+ − 395
apply (case_tac e)
+ − 396
by (auto simp:cm2smsg_def index_of_msg.simps d2s_aux.simps)
+ − 397
+ − 398
lemmas cm2smsg_simps = cm2smsg_nil cm2smsg_nil' cm2smsg_createmsg cm2smsg_other
+ − 399
+ − 400
(********************** cfd2fd_s simpset ******************************)
+ − 401
+ − 402
lemma cfd2fd_s_nil: "fd \<in> init_fds_of_proc p \<Longrightarrow> cfd2fd_s [] p fd = SInit fd"
+ − 403
by (simp add:cfd2fd_s_def index_of_fd.simps)
+ − 404
+ − 405
lemma cfd2fd_s_nil': "fd \<in> current_proc_fds [] p \<Longrightarrow> cfd2fd_s [] p fd = SInit fd"
+ − 406
by (simp add:cfd2fd_s_nil current_proc_fds.simps)
+ − 407
+ − 408
lemma cfd2fd_s_open: "cfd2fd_s (Open p f flags fd opt # \<tau>) p' fd' = (
+ − 409
if (p = p') then (if (fd = fd') then SCrea (Suc (length \<tau>))
+ − 410
else cfd2fd_s \<tau> p' fd')
+ − 411
else cfd2fd_s \<tau> p' fd' )"
+ − 412
by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps)
+ − 413
+ − 414
lemma cfd2fd_s_createsock: "cfd2fd_s (CreateSock p af st fd im # \<tau>) p' fd' = (
+ − 415
if (p = p') then (if (fd = fd') then SCrea (Suc (length \<tau>))
+ − 416
else cfd2fd_s \<tau> p' fd')
+ − 417
else cfd2fd_s \<tau> p' fd' )"
+ − 418
by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps)
+ − 419
+ − 420
lemma cfd2fd_s_accept: "cfd2fd_s (Accept p fd addr port fd' im # \<tau>) p' fd'' = (
+ − 421
if (p' = p) then (if (fd'' = fd') then SCrea (Suc (length \<tau>))
+ − 422
else cfd2fd_s \<tau> p' fd'')
+ − 423
else cfd2fd_s \<tau> p' fd'' )"
+ − 424
by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps)
+ − 425
+ − 426
lemma cfd2fd_s_clone: "cfd2fd_s (Clone p p' # \<tau>) p'' fd = (if (p'' = p') then cfd2fd_s \<tau> p fd else cfd2fd_s \<tau> p'' fd)"
+ − 427
by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps)
+ − 428
+ − 429
lemma cfd2fd_s_other: "\<lbrakk>\<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ − 430
\<forall> p af st fd im. e \<noteq> CreateSock p af st fd im;
+ − 431
\<forall> p fd addr port fd' im. e \<noteq> Accept p fd addr port fd' im;
+ − 432
\<forall> p p'. e \<noteq> Clone p p'\<rbrakk> \<Longrightarrow> cfd2fd_s (e # \<tau>) p'' fd'' = cfd2fd_s \<tau> p'' fd''"
+ − 433
by (case_tac e, auto simp:cfd2fd_s_def index_of_fd.simps d2s_aux.simps)
+ − 434
+ − 435
lemmas cfd2fd_s_simps = cfd2fd_s_nil cfd2fd_s_nil' cfd2fd_s_open cfd2fd_s_createsock cfd2fd_s_accept cfd2fd_s_clone cfd2fd_s_other
+ − 436
+ − 437
(************* cim2im_s simpset **************************)
+ − 438
+ − 439
(* no such lemma
+ − 440
lemma cim2im_s_nil: "init_itag_of_inum im = Some tag \<Longrightarrow> cim2im_s [] im = SInit im"
+ − 441
by (simp add:cim2im_s_def)
+ − 442
*)
+ − 443
+ − 444
lemma cim2im_s_open: "cim2im_s (Open p f flags fd (Some im) # \<tau>) im' = (if (im' = im) then SCrea (Suc (length \<tau>)) else cim2im_s \<tau> im')"
+ − 445
by (simp add:cim2im_s_def)
+ − 446
+ − 447
lemma cim2im_s_open': "cim2im_s (Open p f flags fd None # \<tau>) im = cim2im_s \<tau> im"
+ − 448
by (simp add:cim2im_s_def)
+ − 449
+ − 450
lemma cim2im_s_mkdir: "cim2im_s (Mkdir p f im # \<tau>) im' = (if (im' = im) then SCrea (Suc (length \<tau>)) else cim2im_s \<tau> im')"
+ − 451
by (simp add:cim2im_s_def)
+ − 452
+ − 453
lemma cim2im_s_createsock: "cim2im_s (CreateSock p sf st fd im # \<tau>) im' = (if (im' = im) then SCrea (Suc (length \<tau>)) else cim2im_s \<tau> im')"
+ − 454
by (simp add:cim2im_s_def)
+ − 455
+ − 456
lemma cim2im_s_accept: "cim2im_s (Accept p fd addr port fd' im # \<tau>) im' = (if (im' = im) then SCrea (Suc (length \<tau>)) else cim2im_s \<tau> im')"
+ − 457
by (simp add:cim2im_s_def)
+ − 458
+ − 459
lemma cim2im_s_other: "\<lbrakk>\<forall> p f flags fd opt. e \<noteq> Open p f flags fd opt;
+ − 460
\<forall> p f im. e \<noteq> Mkdir p f im;
+ − 461
\<forall> p sf st fd im. e \<noteq> CreateSock p sf st fd im;
+ − 462
\<forall> p fd addr port fd' im. e \<noteq> Accept p fd addr port fd' im\<rbrakk> \<Longrightarrow> cim2im_s (e # \<tau>) im = cim2im_s \<tau> im"
+ − 463
by (case_tac e, auto simp:cim2im_s_def)
+ − 464
+ − 465
lemmas cim2im_s_simps = cim2im_s_open cim2im_s_open' cim2im_s_mkdir cim2im_s_createsock cim2im_s_accept cim2im_s_other
+ − 466
+ − 467
+ − 468
lemma cig2ig_s_simp: "cig2ig_s (e # \<tau>) tag = cig2ig_s \<tau> tag"
+ − 469
apply (case_tac tag)
+ − 470
by auto
+ − 471
+ − 472
(******************* cobj2sobj no Suc (length \<tau>) ***********************)
+ − 473
+ − 474
lemma cf2sfile_le_len: "\<lbrakk>cf2sfile \<tau> f = SCrea (Suc (length \<tau>)) # spf; f \<in> current_files \<tau>; \<tau> \<in> vt rc_cs\<rbrakk> \<Longrightarrow> False"
+ − 475
apply (case_tac f, (simp add:cf2sfile.simps d2s_aux.simps)+)
+ − 476
apply (case_tac "index_of_file \<tau> (a # list)", (simp add:d2s_aux.simps)+)
+ − 477
by (drule index_of_file_le_length', simp+)
+ − 478
+ − 479
lemma cf2sfile_le_len': "\<lbrakk>SCrea (Suc (length \<tau>)) # spf \<preceq> cf2sfile \<tau> f; f \<in> current_files \<tau>; \<tau> \<in> vt rc_cs\<rbrakk> \<Longrightarrow> False"
+ − 480
apply (induct f)
+ − 481
apply (simp add:no_junior_def cf2sfile.simps d2s_aux.simps)
+ − 482
apply (case_tac "cf2sfile \<tau> (a # f) = SCrea (Suc (length \<tau>)) # spf")
+ − 483
apply (drule_tac f = "a # f" in cf2sfile_le_len, simp+)
+ − 484
apply (simp only:cf2sfile.simps d2s_aux.simps)
+ − 485
apply (drule_tac no_junior_noteq, simp+)
+ − 486
apply (rule impI, erule impE, simp+)
+ − 487
apply (drule parentf_in_current', simp+)
+ − 488
done
+ − 489
+ − 490
lemma cp2sproc_le_len: "cp2sproc \<tau> p = SCrea (Suc (length \<tau>)) \<Longrightarrow> False"
+ − 491
apply (simp add:cp2sproc_def, case_tac "index_of_proc \<tau> p")
+ − 492
apply (simp add:d2s_aux.simps)+
+ − 493
by (drule index_of_proc_le_length', simp)
+ − 494
+ − 495
lemma ch2sshm_le_len: "ch2sshm \<tau> h = SCrea (Suc (length \<tau>)) \<Longrightarrow> False"
+ − 496
apply (simp add:ch2sshm_def, case_tac "index_of_shm \<tau> h")
+ − 497
apply (simp add:d2s_aux.simps)+
+ − 498
by (drule index_of_shm_le_length', simp)
+ − 499
+ − 500
lemma cm2smsg_le_len: "cm2smsg \<tau> m = SCrea (Suc (length \<tau>)) \<Longrightarrow> False"
+ − 501
apply (simp add:cm2smsg_def, case_tac "index_of_msg \<tau> m")
+ − 502
apply (simp add:d2s_aux.simps)+
+ − 503
by (drule index_of_msg_le_length', simp)
+ − 504
+ − 505
lemma cim2im_s_le_len: "cim2im_s \<tau> im = SCrea (Suc (length \<tau>)) \<Longrightarrow> False"
+ − 506
apply (simp add:cim2im_s_def, case_tac "inum2ind \<tau> im")
+ − 507
apply (simp add:d2s_aux.simps)+
+ − 508
by (drule inum2ind_le_length', simp)
+ − 509
+ − 510
lemma cfd2fd_s_le_len: "cfd2fd_s \<tau> p fd = SCrea (Suc (length \<tau>)) \<Longrightarrow> False"
+ − 511
apply (simp add:cfd2fd_s_def, case_tac "index_of_fd \<tau> p fd")
+ − 512
apply (simp add:d2s_aux.simps)+
+ − 513
by (drule index_of_fd_le_length', simp)
+ − 514
+ − 515
end
+ − 516
+ − 517
(*<*)
+ − 518
end
+ − 519
(*>*)