os_rc.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Tue, 30 Apr 2013 14:46:18 +0100
changeset 4 8b6ba7168f2d
parent 1 dcde836219bc
child 6 4294abb1f38c
permissions -rw-r--r--
pictures
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
     1
theory os_rc 
dcde836219bc add thy files
chunhan
parents:
diff changeset
     2
imports Main rc_theory 
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
(****** below context is for lemmas of OS and RC ******)
dcde836219bc add thy files
chunhan
parents:
diff changeset
     6
context rc_basic begin
dcde836219bc add thy files
chunhan
parents:
diff changeset
     7
dcde836219bc add thy files
chunhan
parents:
diff changeset
     8
inductive_cases vs_step': "valid (e # \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
     9
dcde836219bc add thy files
chunhan
parents:
diff changeset
    10
lemma valid_cons:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    11
  "valid (e # \<tau>) \<Longrightarrow> valid \<tau>"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    12
by (drule vs_step', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    13
dcde836219bc add thy files
chunhan
parents:
diff changeset
    14
lemma valid_os: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    15
  "valid (e # \<tau>) \<Longrightarrow> os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    16
by (drule vs_step', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    17
dcde836219bc add thy files
chunhan
parents:
diff changeset
    18
lemma valid_rc:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    19
  "valid (e # \<tau>) \<Longrightarrow> rc_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    20
by (drule vs_step', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    21
dcde836219bc add thy files
chunhan
parents:
diff changeset
    22
lemma vs_history:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    23
  "\<lbrakk>s \<preceq> s'; valid s'\<rbrakk> \<Longrightarrow> valid s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    24
apply (induct s', simp add:no_junior_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    25
apply (case_tac "s = a # s'", simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    26
apply (drule no_junior_noteq, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    27
by (drule valid_cons)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    28
dcde836219bc add thy files
chunhan
parents:
diff changeset
    29
lemma parent_file_in_current: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    30
  "\<lbrakk>parent f = Some pf; f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> pf \<in> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    31
apply (induct s)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    32
apply (simp add:parent_file_in_init)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    33
apply (frule valid_cons, frule valid_rc, frule valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    34
apply (case_tac a, auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    35
apply (case_tac f, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    36
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    37
dcde836219bc add thy files
chunhan
parents:
diff changeset
    38
lemma parent_file_in_current': 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    39
  "\<lbrakk>fn # pf \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> pf \<in> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    40
by (auto intro!:parent_file_in_current[where pf = pf])
dcde836219bc add thy files
chunhan
parents:
diff changeset
    41
dcde836219bc add thy files
chunhan
parents:
diff changeset
    42
lemma parent_file_in_init': 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    43
  "fn # pf \<in> init_files \<Longrightarrow> pf \<in> init_files"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    44
by (auto intro!:parent_file_in_init[where pf = pf])
dcde836219bc add thy files
chunhan
parents:
diff changeset
    45
dcde836219bc add thy files
chunhan
parents:
diff changeset
    46
lemma ancient_file_in_current:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    47
  "\<lbrakk>f \<in> current_files s; valid s; af \<preceq> f\<rbrakk> \<Longrightarrow> af \<in> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    48
apply (induct f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    49
apply (simp add:no_junior_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    50
apply (case_tac "af = a # f", simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    51
apply (drule no_junior_noteq, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    52
apply (drule parent_file_in_current', simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    53
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    54
dcde836219bc add thy files
chunhan
parents:
diff changeset
    55
lemma cannot_del_root:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    56
  "\<lbrakk>valid (DeleteFile p [] # s); f \<noteq> []; f \<in> current_files s\<rbrakk> \<Longrightarrow> False"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    57
apply (frule valid_cons, frule valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    58
apply (case_tac f rule:rev_cases, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    59
apply (drule_tac af = "[y]" in ancient_file_in_current, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    60
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    61
dcde836219bc add thy files
chunhan
parents:
diff changeset
    62
lemma init_file_initialrole_imp_some: "\<exists> r. init_file_initialrole f = Some r"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    63
by (case_tac f, auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    64
dcde836219bc add thy files
chunhan
parents:
diff changeset
    65
lemma file_has_initialrole: "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> (\<exists> r. initialrole s f = Some r)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    66
apply (induct s arbitrary:f) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    67
apply (simp, rule init_file_initialrole_imp_some)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    68
apply (frule valid_cons, frule valid_os, case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    69
apply (auto split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    70
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    71
dcde836219bc add thy files
chunhan
parents:
diff changeset
    72
lemma file_has_initialrole':
dcde836219bc add thy files
chunhan
parents:
diff changeset
    73
  "\<lbrakk>initialrole s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    74
by (rule notI, auto dest:file_has_initialrole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    75
dcde836219bc add thy files
chunhan
parents:
diff changeset
    76
lemma file_has_effinitialrole: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    77
  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> r. effinitialrole s f = Some r"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    78
apply (induct f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    79
apply (auto simp:effinitialrole_def dest:file_has_initialrole parent_file_in_current')
dcde836219bc add thy files
chunhan
parents:
diff changeset
    80
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    81
dcde836219bc add thy files
chunhan
parents:
diff changeset
    82
lemma file_has_effinitialrole':
dcde836219bc add thy files
chunhan
parents:
diff changeset
    83
  "\<lbrakk>effinitialrole s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    84
by (rule notI, auto dest:file_has_effinitialrole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    85
dcde836219bc add thy files
chunhan
parents:
diff changeset
    86
lemma init_file_forcedrole_imp_some: "\<exists> r. init_file_forcedrole f = Some r"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    87
by (case_tac f, auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    88
dcde836219bc add thy files
chunhan
parents:
diff changeset
    89
lemma file_has_forcedrole: "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> (\<exists> r. forcedrole s f = Some r)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    90
apply (induct s arbitrary:f) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    91
apply (simp add:init_file_forcedrole_imp_some)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    92
apply (frule valid_cons, frule valid_os, case_tac a, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    93
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    94
dcde836219bc add thy files
chunhan
parents:
diff changeset
    95
lemma file_has_forcedrole':
dcde836219bc add thy files
chunhan
parents:
diff changeset
    96
  "\<lbrakk>forcedrole s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    97
by (rule notI, auto dest:file_has_forcedrole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    98
dcde836219bc add thy files
chunhan
parents:
diff changeset
    99
lemma file_has_effforcedrole: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   100
  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> r. effforcedrole s f = Some r"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   101
apply (induct f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   102
apply (auto simp:effforcedrole_def dest:file_has_forcedrole parent_file_in_current')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   103
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   104
dcde836219bc add thy files
chunhan
parents:
diff changeset
   105
lemma file_has_effforcedrole':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   106
  "\<lbrakk>effforcedrole s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   107
by (rule notI, auto dest:file_has_effforcedrole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   108
dcde836219bc add thy files
chunhan
parents:
diff changeset
   109
lemma current_proc_has_forcedrole:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   110
  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> r. proc_forcedrole s p = Some r"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   111
apply (induct s arbitrary:p) using init_proc_has_frole
dcde836219bc add thy files
chunhan
parents:
diff changeset
   112
apply (simp add:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   113
apply (frule valid_cons, frule valid_os, case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   114
apply (auto split:if_splits option.splits intro:file_has_effforcedrole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   115
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   116
dcde836219bc add thy files
chunhan
parents:
diff changeset
   117
lemma current_proc_has_forcedrole':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   118
  "\<lbrakk>proc_forcedrole s p = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   119
by (rule notI, auto dest:current_proc_has_forcedrole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   120
dcde836219bc add thy files
chunhan
parents:
diff changeset
   121
lemma current_proc_has_owner: "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> u. owner s p = Some u"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   122
apply (induct s arbitrary:p) using init_proc_has_owner
dcde836219bc add thy files
chunhan
parents:
diff changeset
   123
apply (simp add:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   124
apply (frule valid_cons, frule valid_os, case_tac a, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   125
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   126
dcde836219bc add thy files
chunhan
parents:
diff changeset
   127
lemma current_proc_has_owner':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   128
  "\<lbrakk>owner s p = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   129
by (rule notI, auto dest:current_proc_has_owner)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   130
dcde836219bc add thy files
chunhan
parents:
diff changeset
   131
(*
dcde836219bc add thy files
chunhan
parents:
diff changeset
   132
lemma effinitial_normal_intro: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   133
  "\<lbrakk>f \<in> current_files \<tau>; valid \<tau>; effinitialrole \<tau> f \<noteq> Some UseForcedRole\<rbrakk> \<Longrightarrow> \<exists>nr. effinitialrole \<tau> f = Some (NormalRole nr)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   134
apply (drule file_has_effinitialrole, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   135
apply (erule exE, frule effinitialrole_valid, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   136
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   137
dcde836219bc add thy files
chunhan
parents:
diff changeset
   138
lemma effforced_normal_intro: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   139
  "\<lbrakk>f \<in> current_files \<tau>; valid \<tau>; effforcedrole \<tau> f \<noteq> Some InheritUserRole; effforcedrole \<tau> f \<noteq> Some InheritProcessRole; effforcedrole \<tau> f \<noteq> Some InheritUpMixed\<rbrakk> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   140
  \<Longrightarrow> \<exists>nr. effforcedrole \<tau> f = Some (NormalRole nr)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   141
apply (drule file_has_effforcedrole, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   142
apply (erule exE, frule effforcedrole_valid, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   143
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   144
*)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   145
dcde836219bc add thy files
chunhan
parents:
diff changeset
   146
lemma owner_in_users: "\<lbrakk>owner s p = Some u; valid s\<rbrakk> \<Longrightarrow> u \<in> init_users"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   147
apply (induct s arbitrary:p) defer
dcde836219bc add thy files
chunhan
parents:
diff changeset
   148
apply (frule valid_cons, frule valid_os, case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   149
apply (auto split:if_splits option.splits intro!:init_owner_valid)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   150
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   151
dcde836219bc add thy files
chunhan
parents:
diff changeset
   152
lemma user_has_normalrole: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   153
  "u \<in> init_users \<Longrightarrow> \<exists> nr. defrole u = Some nr" using init_user_has_role
dcde836219bc add thy files
chunhan
parents:
diff changeset
   154
by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   155
dcde836219bc add thy files
chunhan
parents:
diff changeset
   156
lemma user_has_normalrole':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   157
  "defrole u = None \<Longrightarrow> u \<notin> init_users"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   158
by (rule notI, auto dest:user_has_normalrole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   159
dcde836219bc add thy files
chunhan
parents:
diff changeset
   160
lemma current_proc_has_role: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   161
  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> nr. currentrole s p = Some nr"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   162
apply (induct s arbitrary:p) using init_proc_has_role
dcde836219bc add thy files
chunhan
parents:
diff changeset
   163
apply (simp add:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   164
apply (frule valid_cons, frule valid_os, case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   165
apply (auto simp:map_comp_def split:if_splits option.splits t_role.splits
dcde836219bc add thy files
chunhan
parents:
diff changeset
   166
           dest!:current_proc_has_owner' user_has_normalrole' current_proc_has_forcedrole'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   167
                 file_has_forcedrole' file_has_effforcedrole' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   168
                 file_has_initialrole' file_has_effinitialrole'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   169
           intro:user_has_normalrole
dcde836219bc add thy files
chunhan
parents:
diff changeset
   170
            dest:owner_in_users effinitialrole_valid effforcedrole_valid proc_forcedrole_valid)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   171
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   172
dcde836219bc add thy files
chunhan
parents:
diff changeset
   173
lemma current_file_has_type: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   174
  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> t. type_of_file s f = Some t"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   175
apply (induct s)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   176
apply (simp split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   177
apply (frule valid_cons, frule valid_os, case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   178
apply (auto split:option.splits intro:current_proc_has_role)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   179
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   180
dcde836219bc add thy files
chunhan
parents:
diff changeset
   181
lemma init_file_has_etype: "f \<in> init_files \<Longrightarrow> \<exists> nt. etype_aux init_file_type_aux f = Some nt"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   182
apply (induct f) defer
dcde836219bc add thy files
chunhan
parents:
diff changeset
   183
apply (frule parent_file_in_init') 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   184
apply (auto split:option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   185
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   186
dcde836219bc add thy files
chunhan
parents:
diff changeset
   187
lemma current_file_has_etype[rule_format]:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   188
  "f \<in> current_files s \<Longrightarrow> valid s \<longrightarrow> (\<exists> nt. etype_of_file s f = Some nt)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   189
apply (induct f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   190
apply (auto simp:etype_of_file_def dest:current_file_has_type parent_file_in_current'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   191
           split:option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   192
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   193
dcde836219bc add thy files
chunhan
parents:
diff changeset
   194
lemma current_file_has_etype':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   195
  "\<lbrakk>etype_of_file s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   196
by (rule notI, auto dest:current_file_has_etype)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   197
dcde836219bc add thy files
chunhan
parents:
diff changeset
   198
(*** etype_of_file simpset ***)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   199
dcde836219bc add thy files
chunhan
parents:
diff changeset
   200
lemma etype_aux_prop:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   201
  "\<forall> x. x \<preceq> f \<longrightarrow> func' x = func x \<Longrightarrow> etype_aux func f = etype_aux func' f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   202
apply (induct f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   203
by (auto split:t_rc_file_type.splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   204
dcde836219bc add thy files
chunhan
parents:
diff changeset
   205
lemma etype_aux_prop1:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   206
  "func' = func ((a#f) := b) \<Longrightarrow> etype_aux func f = etype_aux func' f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   207
by (rule etype_aux_prop, auto simp:no_junior_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   208
dcde836219bc add thy files
chunhan
parents:
diff changeset
   209
lemma etype_aux_prop1':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   210
  "etype_aux func f = x \<Longrightarrow> etype_aux (func ((a#f) := b)) f = x"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   211
apply (subgoal_tac "etype_aux func f = etype_aux (func ((a#f) := b)) f")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   212
apply (simp, rule etype_aux_prop1, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   213
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   214
dcde836219bc add thy files
chunhan
parents:
diff changeset
   215
lemma etype_aux_prop2:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   216
  "\<lbrakk>f \<in> current_files s; f' \<notin> current_files s; valid s\<rbrakk> \<Longrightarrow>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   217
  etype_aux (func (f' := b)) f = etype_aux func f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   218
apply (rule etype_aux_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   219
by (auto dest:ancient_file_in_current)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   220
dcde836219bc add thy files
chunhan
parents:
diff changeset
   221
lemma etype_aux_prop3:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   222
  "parent f = Some pf 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   223
   \<Longrightarrow> etype_aux (func (f := Some InheritParent_file_type)) f = etype_aux func pf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   224
apply (case_tac f, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   225
by (rule etype_aux_prop, simp add:no_junior_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   226
dcde836219bc add thy files
chunhan
parents:
diff changeset
   227
lemma etype_aux_prop4:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   228
  "etype_aux (func (f := Some (NormalFile_type t))) f = Some t"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   229
by (case_tac f, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   230
dcde836219bc add thy files
chunhan
parents:
diff changeset
   231
lemma etype_of_file_delete:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   232
  "\<lbrakk>valid (DeleteFile p f # s); f' \<in> current_files s\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   233
   \<Longrightarrow> etype_of_file (DeleteFile p f # s) f' = etype_of_file s f'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   234
apply (frule valid_cons, frule valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   235
apply (simp add:etype_of_file_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   236
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   237
dcde836219bc add thy files
chunhan
parents:
diff changeset
   238
lemma current_proc_has_type: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   239
  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> nt. type_of_process s p = Some nt"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   240
apply (induct s arbitrary:p) using init_proc_has_type
dcde836219bc add thy files
chunhan
parents:
diff changeset
   241
apply (simp add:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   242
dcde836219bc add thy files
chunhan
parents:
diff changeset
   243
apply (frule valid_cons, frule valid_os, case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   244
dcde836219bc add thy files
chunhan
parents:
diff changeset
   245
apply (subgoal_tac "nat1 \<in> current_procs (a # s)") prefer 2 apply simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   246
apply (drule_tac p = nat1 in current_proc_has_role, simp, erule exE)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   247
dcde836219bc add thy files
chunhan
parents:
diff changeset
   248
apply (auto simp:pct_def pot_def pet_def dest:current_proc_has_role 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   249
           split:option.splits t_rc_proc_type.splits 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   250
           dest!:default_process_create_type_valid default_process_chown_type_valid 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   251
                 default_process_execute_type_valid)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   252
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   253
dcde836219bc add thy files
chunhan
parents:
diff changeset
   254
lemma current_ipc_has_type: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   255
  "\<lbrakk>i \<in> current_ipcs s; valid s\<rbrakk> \<Longrightarrow> \<exists> nt. type_of_ipc s i = Some nt"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   256
apply (induct s) using init_ipc_has_type
dcde836219bc add thy files
chunhan
parents:
diff changeset
   257
apply (simp add:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   258
dcde836219bc add thy files
chunhan
parents:
diff changeset
   259
apply (frule valid_cons, frule valid_os, case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   260
apply (auto dest:current_proc_has_role)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   261
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   262
dcde836219bc add thy files
chunhan
parents:
diff changeset
   263
(*** finite current_*  ***)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   264
dcde836219bc add thy files
chunhan
parents:
diff changeset
   265
lemma finite_cf: "finite (current_files s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   266
apply (induct s) defer apply (case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   267
using init_finite by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   268
dcde836219bc add thy files
chunhan
parents:
diff changeset
   269
lemma finite_cp: "finite (current_procs s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   270
apply (induct s) defer apply (case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   271
using init_finite by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   272
dcde836219bc add thy files
chunhan
parents:
diff changeset
   273
lemma finite_ci: "finite (current_ipcs s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   274
apply (induct s) defer apply (case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   275
using init_finite by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   276
dcde836219bc add thy files
chunhan
parents:
diff changeset
   277
(*** properties of new-proc new-ipc ... ***)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   278
dcde836219bc add thy files
chunhan
parents:
diff changeset
   279
lemma nn_notin_aux: "finite s \<Longrightarrow> \<forall> a \<in> s. Max s \<ge> a "
dcde836219bc add thy files
chunhan
parents:
diff changeset
   280
apply (erule finite.induct, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   281
apply (rule ballI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   282
apply (case_tac "aa = a", simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   283
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   284
dcde836219bc add thy files
chunhan
parents:
diff changeset
   285
lemma nn_notin: "finite s \<Longrightarrow> next_nat s \<notin> s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   286
apply (drule nn_notin_aux)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   287
apply (simp add:next_nat_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   288
by (auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   289
dcde836219bc add thy files
chunhan
parents:
diff changeset
   290
lemma np_notin_curp: "new_proc \<tau> \<notin> current_procs \<tau>" using finite_cp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   291
by (simp add:new_proc_def nn_notin)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   292
dcde836219bc add thy files
chunhan
parents:
diff changeset
   293
lemma np_notin_curp': "new_proc \<tau> \<in> current_procs \<tau> \<Longrightarrow> False"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   294
by (simp add:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   295
dcde836219bc add thy files
chunhan
parents:
diff changeset
   296
lemma ni_notin_curi: "new_ipc \<tau> \<notin> current_ipcs \<tau>" using finite_ci
dcde836219bc add thy files
chunhan
parents:
diff changeset
   297
by (simp add:new_ipc_def nn_notin)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   298
dcde836219bc add thy files
chunhan
parents:
diff changeset
   299
lemma ni_notin_curi': "new_ipc \<tau> \<in> current_ipcs \<tau> \<Longrightarrow> False"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   300
by (simp add:ni_notin_curi)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   301
dcde836219bc add thy files
chunhan
parents:
diff changeset
   302
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
   303
dcde836219bc add thy files
chunhan
parents:
diff changeset
   304
context tainting_s_complete begin
dcde836219bc add thy files
chunhan
parents:
diff changeset
   305
dcde836219bc add thy files
chunhan
parents:
diff changeset
   306
lemma init_notin_curf_deleted:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   307
  "\<lbrakk>f \<notin> current_files s; f \<in> init_files\<rbrakk> \<Longrightarrow> deleted (File f) s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   308
by (induct s, simp, case_tac a, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   309
dcde836219bc add thy files
chunhan
parents:
diff changeset
   310
lemma init_notin_curi_deleted:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   311
  "\<lbrakk>i \<notin> current_ipcs s; i \<in> init_ipcs\<rbrakk> \<Longrightarrow> deleted (IPC i) s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   312
by (induct s, simp, case_tac a, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   313
dcde836219bc add thy files
chunhan
parents:
diff changeset
   314
lemma init_notin_curp_deleted:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   315
  "\<lbrakk>p \<notin> current_procs s; p \<in> init_processes\<rbrakk> \<Longrightarrow> deleted (Proc p) s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   316
by (induct s, simp, case_tac a, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   317
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   318
lemma ni_init_deled: "new_ipc s \<in> init_ipcs \<Longrightarrow> deleted (IPC (new_ipc s)) s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   319
using ni_notin_curi[where \<tau> = s]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   320
by (drule_tac init_notin_curi_deleted, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   321
dcde836219bc add thy files
chunhan
parents:
diff changeset
   322
lemma np_init_deled: "new_proc s \<in> init_processes \<Longrightarrow> deleted (Proc (new_proc s)) s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   323
using np_notin_curp[where \<tau> = s]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   324
by (drule_tac init_notin_curp_deleted, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   325
dcde836219bc add thy files
chunhan
parents:
diff changeset
   326
lemma source_dir_in_init: "source_dir s f = Some sd \<Longrightarrow> sd \<in> init_files"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   327
by (induct f, auto split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   328
dcde836219bc add thy files
chunhan
parents:
diff changeset
   329
lemma source_proc_in_init:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   330
  "\<lbrakk>source_proc s p = Some p'; p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> p' \<in> init_processes"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   331
apply (induct s arbitrary:p, simp split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   332
apply (frule valid_os, frule valid_cons, case_tac a)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   333
by (auto simp:np_notin_curp split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   334
dcde836219bc add thy files
chunhan
parents:
diff changeset
   335
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
   336
dcde836219bc add thy files
chunhan
parents:
diff changeset
   337
context tainting_s_sound begin
dcde836219bc add thy files
chunhan
parents:
diff changeset
   338
dcde836219bc add thy files
chunhan
parents:
diff changeset
   339
lemma len_fname_all: "length (fname_all_a len) = len"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   340
by (induct len, auto simp:fname_all_a.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   341
dcde836219bc add thy files
chunhan
parents:
diff changeset
   342
lemma ncf_notin_curf: "new_childf f s \<notin> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   343
apply (simp add:new_childf_def next_fname_def all_fname_under_dir_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   344
apply (rule notI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   345
apply (subgoal_tac "(CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files s}))) \<in> {fn. fn # f \<in> current_files s}")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   346
defer apply simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   347
apply (subgoal_tac "length (CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files s}))) \<in> fname_length_set {fn. fn # f \<in> current_files s}")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   348
defer apply (auto simp:fname_length_set_def image_def)[1]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   349
apply (subgoal_tac "finite (fname_length_set {fn. fn # f \<in> current_files s})")  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   350
defer
dcde836219bc add thy files
chunhan
parents:
diff changeset
   351
apply (simp add:fname_length_set_def) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   352
apply (rule finite_imageI) using finite_cf[where s = s]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   353
apply (drule_tac h = "\<lambda> f'. case f' of [] \<Rightarrow> '''' | fn # pf' \<Rightarrow> if (pf' = f) then fn else ''''" in finite_imageI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   354
apply (rule_tac B = "(list_case [] (\<lambda>fn pf'. if pf' = f then fn else []) ` current_files s)" in finite_subset)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   355
unfolding image_def
dcde836219bc add thy files
chunhan
parents:
diff changeset
   356
apply(auto)[1]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   357
apply (rule_tac x = "x # f" in bexI, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   358
apply (drule_tac s = "(fname_length_set {fn. fn # f \<in> current_files s})" in nn_notin_aux)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   359
apply (erule_tac x = "length (CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files s})))" in ballE)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   360
apply (simp add:len_fname_all, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   361
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   362
dcde836219bc add thy files
chunhan
parents:
diff changeset
   363
lemma ncf_parent: "parent (new_childf f \<tau>) = Some f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   364
by (simp add:new_childf_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   365
dcde836219bc add thy files
chunhan
parents:
diff changeset
   366
lemma clone_event_no_limit: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   367
  "\<lbrakk>p \<in> current_procs \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> valid (Clone p (new_proc \<tau>) # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   368
apply (rule vs_step)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   369
apply (auto intro:clone_no_limit split:option.splits
dcde836219bc add thy files
chunhan
parents:
diff changeset
   370
            dest:current_proc_has_role current_proc_has_type)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   371
done 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   372
dcde836219bc add thy files
chunhan
parents:
diff changeset
   373
dcde836219bc add thy files
chunhan
parents:
diff changeset
   374
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
   375
dcde836219bc add thy files
chunhan
parents:
diff changeset
   376
end