tainted_vs_tainted_s.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Sun, 16 Jun 2013 20:42:07 -0400
changeset 10 569222a42cf5
parent 6 4294abb1f38c
permissions -rw-r--r--
updated the paper for submission
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
     1
theory tainted_vs_tainted_s
dcde836219bc add thy files
chunhan
parents:
diff changeset
     2
imports Main rc_theory os_rc deleted_prop tainted obj2sobj_prop source_prop all_sobj_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
lemma t2ts[rule_format]: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
     8
  "obj \<in> tainted s \<Longrightarrow> obj2sobj s obj \<in> tainted_s "  
dcde836219bc add thy files
chunhan
parents:
diff changeset
     9
proof (induct rule:tainted.induct)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    10
  case (t_init obj) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    11
  assume seed: "obj \<in> seeds" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    12
  hence"exists [] obj" by (drule_tac seeds_in_init, case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    13
  thus ?case using seed by (simp add:ts_init obj2sobj_nil_init)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    14
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
    15
  case (t_clone p s p') 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    16
  assume p1: "Proc p \<in> tainted s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    17
  and p2: "obj2sobj s (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    18
  and p3: "valid (Clone p p' # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    19
  from p3 have os: "os_grant s (Clone p p')" and rc: "rc_grant s (Clone p p')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    20
    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    21
  from os have exp: "exists s (Proc p)" by (simp add:os_grant.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    22
  from exp obtain r fr pt u sp where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    23
    and srp: "source_proc s p = Some sp" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
    24
    apply (simp del:cp2sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    25
    by (frule current_proc_has_sproc, simp, frule current_proc_has_srp, simp, blast)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    26
  with exp all_sobjs_I[where obj = "Proc p" and s = s] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
    27
dcde836219bc add thy files
chunhan
parents:
diff changeset
    28
  have "obj2sobj (Clone p p' # s) (Proc p') = SProc (r,fr,clone_type_aux r pt, u) (Some sp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    29
    using sproc srp p3
dcde836219bc add thy files
chunhan
parents:
diff changeset
    30
    by (simp add:obj2sobj.simps cp2sproc_clone del:cp2sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    31
  moreover have "(r, Proc_type pt, CREATE) \<in> compatible" using rc sproc
dcde836219bc add thy files
chunhan
parents:
diff changeset
    32
    by (auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    33
  moreover have "SProc (r, fr, pt, u) (Some sp) \<in> tainted_s" using p2 sproc srp
dcde836219bc add thy files
chunhan
parents:
diff changeset
    34
    by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
    35
  ultimately show ?case by (auto intro:ts_clone) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    36
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
    37
  case (t_exec f s p)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    38
  assume p1: "File f \<in> tainted s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    39
  and p2: "obj2sobj s (File f) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    40
  and p3: "valid (Execute p f # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    41
  from p3 have os: "os_grant s (Execute p f)" and rc: "rc_grant s (Execute p f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    42
    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    43
  from os have exp: "exists s (Proc p)" and exf: "exists s (File f)" by auto 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    44
  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    45
    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    46
  with exp all_sobjs_I[where obj = "Proc p" and s = s] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
    47
  have SP: "SProc (r,fr,pt,u) (source_proc s p) \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    48
    by (auto simp:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    49
  from exf obtain ft sd where etype: "etype_of_file s f = Some ft"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    50
    and srdir: "source_dir s f = Some sd" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
    51
    by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    52
  with p2 obtain srf where TF: "SFile (ft, sd) srf \<in> tainted_s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    53
    by (auto simp:obj2sobj.simps split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    54
  from sproc srdir have "u \<in> init_users" and "sd \<in> init_files" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
    55
    by (auto simp:obj2sobj.simps cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    56
            intro:source_dir_in_init owner_in_users 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    57
            split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    58
  then obtain nr nfr where nrole: "exec_role_aux r sd u = Some nr" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    59
    and nfrole: "erole_functor init_file_forcedrole InheritUpMixed sd = Some nfr"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    60
    by (frule_tac r = r in exec_role_some, simp, frule_tac efffrole_sdir_some, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    61
  hence "cp2sproc (Execute p f # s) p = Some (nr, nfr, exec_type_aux r pt, u)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    62
    using p3 srdir sproc by (simp add:cp2sproc_exec)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    63
  with nrole nfrole TF SP rc sproc etype
dcde836219bc add thy files
chunhan
parents:
diff changeset
    64
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    65
    by (auto simp:obj2sobj.simps cp2sproc.simps intro!:ts_exec1 split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    66
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
    67
  case (t_cfile p s f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    68
  assume p1: "Proc p \<in> tainted s" and p2: "obj2sobj s (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    69
    and p3: "valid (CreateFile p f # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    70
  from p3 have os: "os_grant s (CreateFile p f)" and rc: "rc_grant s (CreateFile p f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    71
    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    72
  from os obtain pf where parent: "parent f = Some pf" and exp: "exists s (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    73
    and expf: "exists s (File pf)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
    74
  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    75
    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    76
  with p2 have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    77
    by (auto simp:obj2sobj.simps cp2sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    78
  from expf obtain pft sd where etype: "etype_of_file s pf = Some pft"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    79
    and srdir: "source_dir s pf = Some sd" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
    80
    by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    81
  with expf all_sobjs_I[where obj = "File pf" and s = s] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
    82
  obtain srpf where SF: "SFile (pft, sd) srpf \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    83
    by (auto simp:obj2sobj.simps cp2sproc.simps split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    84
  show ?case using etype srdir p3 parent os
dcde836219bc add thy files
chunhan
parents:
diff changeset
    85
    apply (auto simp:source_dir_simps init_notin_curf_deleted obj2sobj.simps split:option.splits
dcde836219bc add thy files
chunhan
parents:
diff changeset
    86
               dest!:current_file_has_etype')
dcde836219bc add thy files
chunhan
parents:
diff changeset
    87
    apply (case_tac "default_fd_create_type r")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    88
    using SF TP rc sproc
dcde836219bc add thy files
chunhan
parents:
diff changeset
    89
    apply (rule_tac sf = srpf in ts_cfd',
dcde836219bc add thy files
chunhan
parents:
diff changeset
    90
           auto simp:etype_of_file_def etype_aux_prop3 obj2sobj.simps cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
    91
               split:option.splits) [1]
dcde836219bc add thy files
chunhan
parents:
diff changeset
    92
    using SF TP rc sproc
dcde836219bc add thy files
chunhan
parents:
diff changeset
    93
    apply (rule_tac sf = srpf in ts_cfd)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    94
    apply (auto simp:etype_of_file_def etype_aux_prop4 cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    95
    done
dcde836219bc add thy files
chunhan
parents:
diff changeset
    96
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
    97
  case (t_cipc p s i)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    98
  assume p1: "Proc p \<in> tainted s" and p2: "obj2sobj s (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    99
    and p3: "valid (CreateIPC p i # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   100
  from p3 have os: "os_grant s (CreateIPC p i)" and rc: "rc_grant s (CreateIPC p i)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   101
    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   102
  from os have exp: "exists s (Proc p)" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   103
  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   104
    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   105
  with p2 have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   106
    by (auto simp:obj2sobj.simps cp2sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   107
  show ?case using p3 sproc os rc TP 
6
4294abb1f38c update scripts with no new_* in admissable-check
chunhan
parents: 1
diff changeset
   108
    by (auto simp:obj2sobj.simps cp2sproc.simps dest:not_deleted_imp_exists
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
   109
            split:option.splits intro!:ts_cipc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   110
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   111
  case (t_read f s p)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   112
  assume p1: "File f \<in> tainted s" and p2: "obj2sobj s (File f) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   113
    and p3: "valid (ReadFile p f # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   114
  from p3 have os: "os_grant s (ReadFile p f)" and rc: "rc_grant s (ReadFile p f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   115
    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   116
  from os have exp: "exists s (Proc p)" and exf: "exists s (File f)" by auto 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   117
  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   118
    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   119
  with exp all_sobjs_I[where obj = "Proc p" and s = s] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   120
  have SP: "SProc (r,fr,pt,u) (source_proc s p) \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   121
    by (auto simp:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   122
  from exf obtain ft sd where etype: "etype_of_file s f = Some ft"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   123
    and srdir: "source_dir s f = Some sd" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   124
    by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   125
  with p2 obtain srf where TF: "SFile (ft, sd) srf \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   126
    by (auto simp:obj2sobj.simps cp2sproc.simps split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   127
  show ?case using sproc SP TF rc etype 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   128
    by (auto simp:obj2sobj.simps cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   129
            split:option.splits intro!:ts_read) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   130
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   131
  case (t_write p s f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   132
  assume p1: "Proc p \<in> tainted s" and p2: "obj2sobj s (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   133
    and p3: "valid (WriteFile p f # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   134
  from p3 have os: "os_grant s (WriteFile p f)" and rc: "rc_grant s (WriteFile p f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   135
    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   136
  from os have exp: "exists s (Proc p)" and exf: "exists s (File f)" by auto 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   137
  from exf obtain ft sd where etype: "etype_of_file s f = Some ft"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   138
    and srdir: "source_dir s f = Some sd" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   139
    by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   140
  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   141
    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   142
  with exp p2 have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   143
    by (auto simp:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   144
  from etype p3 have etype':"etype_of_file (WriteFile p f # s) f = Some ft"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   145
    by (case_tac f, auto simp:etype_of_file_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   146
  have SF: "obj2sobj s (File f) \<in> all_sobjs" using exf vs 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   147
    by (rule_tac all_sobjs_I, simp+) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   148
  show ?case using sproc TP rc etype p3 srdir etype' SF
dcde836219bc add thy files
chunhan
parents:
diff changeset
   149
    by (auto simp:source_dir_simps obj2sobj.simps cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   150
            split:option.splits intro!:ts_write) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   151
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   152
  case (t_send p s i)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   153
  assume p1: "Proc p \<in> tainted s" and p2: "obj2sobj s (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   154
    and p3: "valid (Send p i # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   155
  from p3 have os: "os_grant s (Send p i)" and rc: "rc_grant s (Send p i)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   156
    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   157
  from os have exp: "exists s (Proc p)" and exi: "exists s (IPC i)" by auto 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   158
  from exi obtain t where etype: "type_of_ipc s i = Some t" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   159
    by (simp, drule_tac current_ipc_has_type, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   160
  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   161
    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   162
  with exp p2 have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   163
    by (auto simp:obj2sobj.simps cp2sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   164
  have SI: "obj2sobj s (IPC i) \<in> all_sobjs" using exi vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   165
    by (simp add:all_sobjs_I del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   166
  show ?case using sproc TP rc etype p3 vs SI
dcde836219bc add thy files
chunhan
parents:
diff changeset
   167
    by (auto simp:obj2sobj.simps cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   168
            split:option.splits intro!:ts_send)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   169
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   170
  case (t_recv i s p)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   171
  assume p1: "IPC i \<in> tainted s" and p2: "obj2sobj s (IPC i) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   172
    and p3: "valid (Recv p i # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   173
  from p3 have os: "os_grant s (Recv p i)" and rc: "rc_grant s (Recv p i)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   174
    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   175
  from os have exp: "exists s (Proc p)" and exi: "exists s (IPC i)" by auto 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   176
  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   177
    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   178
  with exp all_sobjs_I[where obj = "Proc p" and s = s] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   179
  have SP: "SProc (r,fr,pt,u) (source_proc s p) \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   180
    by (auto simp:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   181
  from exi obtain t where etype: "type_of_ipc s i = Some t"using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   182
    by (simp, drule_tac current_ipc_has_type, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   183
  with p2 obtain sri where TI: "SIPC t sri \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   184
    by (auto simp:obj2sobj.simps split:if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   185
  show ?case using sproc SP TI rc etype 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   186
    by (auto simp:obj2sobj.simps cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   187
            split:option.splits intro!:ts_recv) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   188
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   189
  case (t_remain obj s e)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   190
  from t_remain(1) have p5: "exists s obj" by (rule tainted_is_current)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   191
  from t_remain(3) have os: "os_grant s e" and vs: "valid s" and rc: "rc_grant s e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   192
    by (auto dest:valid_os valid_cons valid_rc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   193
  show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   194
  proof (cases obj)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   195
    case (File f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   196
    have "obj2sobj (e # s) (File f) = obj2sobj s (File f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   197
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   198
      have "etype_of_file (e # s) f = etype_of_file s f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   199
        using p5 os vs File t_remain(3,4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   200
        apply (case_tac e, auto simp:etype_of_file_def split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   201
        by (auto dest:ancient_file_in_current intro!:etype_aux_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   202
      moreover have "source_dir (e # s) f = source_dir s f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   203
        using p5 os vs File t_remain(3,4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   204
        by (case_tac e, auto simp:source_dir_simps dest:source_dir_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   205
      ultimately show ?thesis using vs t_remain(4) File
dcde836219bc add thy files
chunhan
parents:
diff changeset
   206
        apply (auto simp:obj2sobj.simps cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   207
                   split:if_splits option.splits dest:not_deleted_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   208
        by (case_tac e, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   209
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   210
    with File t_remain(2) show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   211
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   212
    case (IPC i) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   213
    have "obj2sobj (e # s) (IPC i) = obj2sobj s (IPC i)" using p5 t_remain(3,4) os IPC
6
4294abb1f38c update scripts with no new_* in admissable-check
chunhan
parents: 1
diff changeset
   214
      by (case_tac e, auto simp:obj2sobj.simps dest:not_deleted_imp_exists
4294abb1f38c update scripts with no new_* in admissable-check
chunhan
parents: 1
diff changeset
   215
                          split:option.splits  dest!:current_proc_has_role')
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
   216
    with IPC t_remain(2) show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   217
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   218
    case (Proc p) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   219
    show ?thesis 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   220
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   221
      have "\<And> f. e = Execute p f \<Longrightarrow> obj2sobj (e # s) (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   222
      proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   223
        fix f
dcde836219bc add thy files
chunhan
parents:
diff changeset
   224
        assume ev: "e = Execute p f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   225
        show "obj2sobj (e # s) (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   226
        proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   227
          from os ev have exp: "exists s (Proc p)" and exf: "exists s (File f)" by auto 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   228
          from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   229
            using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   230
          with Proc t_remain(2) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   231
          have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   232
            by (auto simp:obj2sobj.simps cp2sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   233
          from exf obtain ft sd where etype: "etype_of_file s f = Some ft"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   234
            and srdir: "source_dir s f = Some sd" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   235
            by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   236
          with exf all_sobjs_I[where obj = "File f" and s = s] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   237
          obtain srf where SF: "SFile (ft, sd) srf \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   238
            by (auto simp:obj2sobj.simps split:if_splits) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   239
          from sproc srdir have "u \<in> init_users" and "sd \<in> init_files" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   240
            by (auto simp:obj2sobj.simps cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   241
                    intro:source_dir_in_init owner_in_users split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   242
          then obtain nr nfr where nrole: "exec_role_aux r sd u = Some nr" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   243
            and nfrole: "erole_functor init_file_forcedrole InheritUpMixed sd = Some nfr"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   244
            by (frule_tac r = r in exec_role_some, simp, frule_tac efffrole_sdir_some, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   245
          hence "cp2sproc (Execute p f # s) p = Some (nr, nfr, exec_type_aux r pt, u)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   246
            using t_remain(3) srdir sproc ev by (simp add:cp2sproc_exec)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   247
          with nrole nfrole SF TP rc sproc etype ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   248
          show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
   249
            by (auto simp:obj2sobj.simps cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   250
                   intro!:ts_exec2 split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   251
        qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   252
      qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   253
      have "\<And> r'. e = ChangeRole p r' \<Longrightarrow> obj2sobj (e # s) (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   254
      proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   255
        fix r' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   256
        assume ev: "e = ChangeRole p r'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   257
        show "obj2sobj (e # s) (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   258
        proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   259
          from os ev have exp: "exists s (Proc p)" by auto 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   260
          from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   261
            using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   262
          with Proc t_remain(2)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   263
          have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   264
            by (auto simp:obj2sobj.simps cp2sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   265
          with rc sproc ev show ?thesis 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   266
            apply (simp add:obj2sobj.simps cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   267
            by (rule_tac ts_crole, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   268
        qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   269
      qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   270
      have "\<And> u'. e = ChangeOwner p u' \<Longrightarrow> obj2sobj (e # s) (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   271
      proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   272
        fix u'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   273
        assume ev: "e = ChangeOwner p u'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   274
        show "obj2sobj (e # s) (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   275
        proof-        
dcde836219bc add thy files
chunhan
parents:
diff changeset
   276
          from os ev have exp: "exists s (Proc p)" by auto 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   277
          from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   278
            using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   279
          with Proc t_remain(2)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   280
          have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   281
            by (auto simp:obj2sobj.simps cp2sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   282
          from os ev have uinit: "u' \<in> init_users" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   283
          then obtain nr where chown: "chown_role_aux r fr u' = Some nr" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   284
            by (auto dest:chown_role_some)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   285
          hence nsproc:"cp2sproc (e#s) p = Some (nr,fr,chown_type_aux r nr pt, u')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   286
            using sproc ev os
dcde836219bc add thy files
chunhan
parents:
diff changeset
   287
            by (auto split:option.splits t_role.splits 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   288
                  simp del:currentrole.simps type_of_process.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   289
                  simp add:obj2sobj.simps cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   290
                           chown_role_aux_valid chown_type_aux_valid)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   291
          with rc sproc ev TP uinit chown
dcde836219bc add thy files
chunhan
parents:
diff changeset
   292
          show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
   293
            by (auto simp:obj2sobj.simps cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   294
                   intro!:ts_chown split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   295
        qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   296
      qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   297
      have "\<lbrakk>\<forall> f. e \<noteq> Execute p f; \<forall> r. e \<noteq> ChangeRole p r; \<forall> u. e \<noteq> ChangeOwner p u\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   298
        \<Longrightarrow> obj2sobj (e # s) (Proc p) \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   299
        using t_remain(2,3,4) os p5 Proc
6
4294abb1f38c update scripts with no new_* in admissable-check
chunhan
parents: 1
diff changeset
   300
        by (case_tac e, auto simp add:obj2sobj.simps cp2sproc_simps dest:not_deleted_imp_exists 
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
   301
                             simp del:cp2sproc.simps  split:option.splits) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   302
      ultimately show ?thesis using Proc
dcde836219bc add thy files
chunhan
parents:
diff changeset
   303
        by (case_tac e, auto simp del:obj2sobj.simps) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   304
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   305
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   306
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   307
dcde836219bc add thy files
chunhan
parents:
diff changeset
   308
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
   309
dcde836219bc add thy files
chunhan
parents:
diff changeset
   310
context tainting_s_sound begin
dcde836219bc add thy files
chunhan
parents:
diff changeset
   311
dcde836219bc add thy files
chunhan
parents:
diff changeset
   312
lemma tainted_s'_eq1: "sobj \<in> tainted_s \<Longrightarrow> sobj \<in> tainted_s'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   313
apply (erule tainted_s.induct)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   314
apply (auto elim:ts'_init ts'_exec1 ts'_exec2 ts'_cfd ts'_cfd' ts'_cipc ts'_read ts'_write ts'_recv ts'_send ts'_crole ts'_chown simp:all_sobjs'_eq)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   315
by (simp add:clone_type_aux_def clone_type_unchange)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   316
dcde836219bc add thy files
chunhan
parents:
diff changeset
   317
lemma tainted_s'_eq2: "sobj \<in> tainted_s' \<Longrightarrow> sobj \<in> tainted_s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   318
apply (erule tainted_s'.induct)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   319
by (auto intro:ts_init ts_exec1 ts_exec2 ts_cfd ts_cfd' ts_cipc ts_read ts_write ts_recv ts_send ts_crole ts_chown ts_clone simp:all_sobjs'_eq)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   320
dcde836219bc add thy files
chunhan
parents:
diff changeset
   321
lemma tainted_s'_eq: "(sobj \<in> tainted_s) = (sobj \<in> tainted_s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   322
by (auto intro:iffI tainted_s'_eq1 tainted_s'_eq2)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   323
dcde836219bc add thy files
chunhan
parents:
diff changeset
   324
(* cause sobj_source_eq_sobj may conflict with initp_intact, so remove it too. *) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   325
lemma ts2t_intact:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   326
  "sobj \<in> tainted_s' \<Longrightarrow> \<exists> obj s. obj2sobj s obj = sobj \<and> obj \<in> tainted s \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   327
                                 no_del_event s \<and> initp_intact s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   328
dcde836219bc add thy files
chunhan
parents:
diff changeset
   329
proof (induct rule:tainted_s'.induct)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   330
  case (ts'_init obj)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   331
  hence ex: "exists [] obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   332
    apply (drule_tac seeds_in_init) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   333
    by (case_tac obj, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   334
  have "obj2sobj [] obj = init_obj2sobj obj" using ex 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   335
    by (simp add:obj2sobj_nil_init)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   336
  moreover have "obj \<in> tainted []" using ts'_init by (simp add:t_init)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   337
  moreover have "initp_intact []"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   338
    by (auto simp:initp_intact_def obj2sobj.simps cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   339
  ultimately show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   340
    by (rule_tac x = obj in exI, rule_tac x = "[]" in exI, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   341
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   342
  case (ts'_exec1 t sd srf r fr pt u srp r' fr')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   343
  then obtain f s  where TF: "(File f) \<in> tainted s" and vds: "valid s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   344
    and "f \<in> current_files s \<and> obj2sobj s (File f) = SFile (t, sd) srf \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   345
         no_del_event s \<and> initp_intact s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   346
    apply (clarsimp, frule_tac obj2sobj_file)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   347
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   348
  with ts'_exec1(3) obtain p s' where vds's: "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   349
    and etype: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   350
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   351
    and intacts': "initp_intact (s'@s)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   352
    and ISP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   353
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   354
    and exf: "f \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   355
    apply (drule_tac s' =s and obj' = "File f" in all_sobjs_E0, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   356
    apply (frule_tac obj = "File f" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   357
    by (frule obj2sobj_proc, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   358
  obtain ev \<tau> where ev: "ev = Execute (new_proc (s'@s)) f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   359
    and tau: "\<tau> = Clone p (new_proc (s'@s)) # (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   360
  hence vs_tau:"valid \<tau>" using exp vds's by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   361
dcde836219bc add thy files
chunhan
parents:
diff changeset
   362
  have valid: "valid (ev # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   363
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   364
    have "os_grant \<tau> ev"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   365
      using ev tau by (simp add:exf exp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   366
    moreover have "rc_grant \<tau> ev" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   367
      using ev tau ts'_exec1 ISP etype
dcde836219bc add thy files
chunhan
parents:
diff changeset
   368
      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   369
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   370
    ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   371
      by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   372
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   373
  have "obj2sobj (ev # \<tau>) (Proc (new_proc (s'@s))) = 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   374
        SProc (r', fr', exec_type_aux r pt, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   375
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   376
    have "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp" using ISP tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   377
      by (auto split:option.splits simp:cp2sproc_simps obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   378
    hence "obj2sobj \<tau> (Proc (new_proc (s'@s))) = SProc (r,fr,pt,u) srp" using tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   379
      by (auto simp:obj2sobj.simps cp2sproc.simps pct_def clone_type_unchange 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   380
              split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   381
    moreover have "source_dir \<tau> f = Some sd" using vs_tau etype tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   382
      by (auto simp:obj2sobj.simps source_dir_simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   383
    ultimately show ?thesis using valid ts'_exec1(5) ts'_exec1(6) ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   384
      by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   385
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   386
  have "Proc (new_proc (s'@s)) \<in> tainted (ev # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   387
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   388
    have "(File f) \<in> tainted \<tau>" using TF nodels' tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   389
      by (drule_tac s = "Clone p (new_proc (s' @ s)) # s'" in t_remain_app,auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   390
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   391
      by (drule_tac t_exec, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   392
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   393
  have "no_del_event (ev # \<tau>)" using ev tau nodels' by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   394
  moreover have "initp_intact (ev#\<tau>)" using ev tau intacts' valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   395
    by (simp add:initp_intact_I_exec)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   396
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   397
    apply (rule_tac x = "Proc (new_proc (s'@s))" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   398
    by (rule_tac x = "ev#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   399
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   400
  case (ts'_exec2 r fr pt u srp t sd srf r' fr')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   401
  then obtain p s where TP: "(Proc p) \<in> tainted s" and vds: "valid s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   402
    and "p \<in> current_procs s \<and> obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   403
         no_del_event s \<and> initp_intact s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   404
    apply (clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   405
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   406
  with ts'_exec2(3) obtain f s' where vds's: "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   407
    and etype: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   408
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   409
    and intacts': "initp_intact (s'@s)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   410
    and ISP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   411
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   412
    and exf: "f \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   413
    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   414
    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   415
    by (frule obj2sobj_file, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   416
  obtain ev \<tau> where ev: "ev = Execute (new_proc (s'@s)) f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   417
    and tau: "\<tau> = Clone p (new_proc (s'@s)) # (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   418
  hence vs_tau:"valid \<tau>" using exp vds's by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   419
dcde836219bc add thy files
chunhan
parents:
diff changeset
   420
  have valid: "valid (ev # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   421
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   422
    have "os_grant \<tau> ev"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   423
      using ev tau by (simp add:exf exp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   424
    moreover have "rc_grant \<tau> ev" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   425
      using ev tau ts'_exec2(4) ISP etype
dcde836219bc add thy files
chunhan
parents:
diff changeset
   426
      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   427
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   428
    ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   429
      by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   430
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   431
  have "obj2sobj (ev # \<tau>) (Proc (new_proc (s'@s))) = 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   432
        SProc (r', fr', exec_type_aux r pt, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   433
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   434
    have "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp" using ISP tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   435
      by (auto split:option.splits simp:cp2sproc_simps obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   436
    hence "obj2sobj \<tau> (Proc (new_proc (s'@s))) = SProc (r,fr,pt,u) srp" using tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   437
      by (auto simp:obj2sobj.simps cp2sproc.simps pct_def clone_type_unchange 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   438
              split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   439
    moreover have "source_dir \<tau> f = Some sd" using vs_tau etype tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   440
      by (auto simp:source_dir_simps obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   441
    ultimately show ?thesis using valid ts'_exec2(5) ts'_exec2(6) ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   442
      by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   443
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   444
  have "Proc (new_proc (s'@s)) \<in> tainted (ev # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   445
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   446
    have "Proc p \<in> tainted (s' @ s)" using TP vds's nodels'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   447
      by (drule_tac s = s' in t_remain_app, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   448
    hence "Proc (new_proc (s'@s)) \<in> tainted \<tau>" using TP tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   449
      by (auto intro:t_clone)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   450
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   451
      by (drule_tac t_remain, auto dest:valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   452
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   453
  have "no_del_event (ev # \<tau>)" using ev tau nodels' by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   454
  moreover have "initp_intact (ev#\<tau>)" using ev tau intacts' valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   455
    by (simp add:initp_intact_I_exec)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   456
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   457
    by (rule_tac x = "Proc (new_proc (s'@s))" in exI, rule_tac x = "ev#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   458
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   459
  case (ts'_cfd r fr pt u srp t sd srf t')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   460
  then obtain p s where TP: "(Proc p) \<in> tainted s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   461
    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   462
         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   463
    apply (clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   464
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   465
  with ts'_cfd(3) obtain pf s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   466
    and SF: "obj2sobj (s'@s) (File pf) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   467
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   468
    and intacts': "initp_intact (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   469
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   470
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   471
    and exf: "pf \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   472
    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   473
    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   474
    by (frule obj2sobj_file, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   475
  obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   476
    and tau: "\<tau> = (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   477
dcde836219bc add thy files
chunhan
parents:
diff changeset
   478
  have valid: "valid (e# \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   479
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   480
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   481
      using ev tau exp exf by (simp add:ncf_notin_curf ncf_parent)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   482
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   483
      using ev tau ts'_cfd(4,5,6) SP SF
dcde836219bc add thy files
chunhan
parents:
diff changeset
   484
      by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   485
        split:if_splits option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   486
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   487
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   488
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   489
  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   490
  have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t', sd) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   491
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   492
    have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   493
      using ev tau SF SP ts'_cfd(4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   494
      by (auto simp:obj2sobj.simps etype_of_file_def cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   495
              split:option.splits if_splits  intro!:etype_aux_prop4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   496
    moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   497
      using ev tau SF SP valid ncf_parent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   498
      by (auto simp:source_dir_simps obj2sobj.simps cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   499
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   500
    ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   501
      nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   502
      by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   503
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   504
  have "initp_intact (e#\<tau>)" using ev tau intacts' valid nodels'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   505
    by (auto intro!:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   506
  have "File (new_childf pf \<tau>) \<in> tainted (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   507
  proof- 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   508
    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   509
      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   510
    thus ?thesis using ev tau valid by (auto intro:t_cfile)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   511
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   512
  ultimately show ?case using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   513
    apply (rule_tac x = "File (new_childf pf \<tau>)" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   514
    by (rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   515
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   516
  case (ts'_cfd' r fr pt u srp t sd srf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   517
  then obtain p s where TP: "(Proc p) \<in> tainted s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   518
    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   519
         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   520
    apply (clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   521
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   522
  with ts'_cfd'(3) obtain pf s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   523
    and SF: "obj2sobj (s'@s) (File pf) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   524
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   525
    and intacts': "initp_intact (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   526
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   527
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   528
    and exf: "pf \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   529
    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   530
    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   531
    by (frule obj2sobj_file, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   532
  obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   533
    and tau: "\<tau> = (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   534
dcde836219bc add thy files
chunhan
parents:
diff changeset
   535
  have valid: "valid (e# \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   536
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   537
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   538
      using ev tau exp exf by (simp add:ncf_notin_curf ncf_parent)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   539
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   540
      using ev tau ts'_cfd'(4,5) SP SF
dcde836219bc add thy files
chunhan
parents:
diff changeset
   541
      by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   542
        split:if_splits option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   543
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   544
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   545
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   546
  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   547
  have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t, sd) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   548
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   549
      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   550
      proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   551
        have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = etype_of_file \<tau> pf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   552
          using ev tau SP ts'_cfd'(4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   553
          by (auto simp:obj2sobj.simps ncf_parent etype_of_file_def cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   554
            split:option.splits   intro!:etype_aux_prop3)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   555
        thus ?thesis using SF tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   556
          by (auto simp:obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   557
      qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   558
    moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   559
      using ev tau SF SP valid ncf_parent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   560
      by (auto simp:source_dir_simps obj2sobj.simps cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   561
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   562
    ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   563
      nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   564
      by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   565
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   566
  have "initp_intact (e#\<tau>)" using ev tau intacts' valid nodels'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   567
    by (auto intro!:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   568
  have "File (new_childf pf \<tau>) \<in> tainted (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   569
  proof- 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   570
    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   571
      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   572
    thus ?thesis using ev tau valid by (auto intro:t_cfile)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   573
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   574
  ultimately show ?case using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   575
    apply (rule_tac x = "File (new_childf pf \<tau>)" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   576
    by (rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   577
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   578
  case (ts'_cipc r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   579
  then obtain p s where TP: "(Proc p) \<in> tainted s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   580
    and vds: "valid s" and exp: "p \<in> current_procs s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   581
    and nodels: "no_del_event s" and intacts: "initp_intact s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   582
    and SP: "obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   583
    apply (clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   584
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   585
  obtain e where ev: "e = CreateIPC p (new_ipc s)" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   586
      
dcde836219bc add thy files
chunhan
parents:
diff changeset
   587
  have valid: "valid (e # s)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   588
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   589
    have "os_grant s e"
6
4294abb1f38c update scripts with no new_* in admissable-check
chunhan
parents: 1
diff changeset
   590
      using ev exp by (simp add:ni_notin_curi)
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
   591
    moreover have "rc_grant s e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   592
      using ev ts'_cipc(3) SP
dcde836219bc add thy files
chunhan
parents:
diff changeset
   593
      by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   594
    ultimately show ?thesis using vds
dcde836219bc add thy files
chunhan
parents:
diff changeset
   595
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   596
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   597
  have nodel: "no_del_event (e#s)" using nodels ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   598
  have "initp_intact (e#s)" using ev intacts valid nodels
dcde836219bc add thy files
chunhan
parents:
diff changeset
   599
    by (auto intro!:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   600
  have "IPC (new_ipc s) \<in> tainted (e#s)" using TP ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   601
    by (auto intro:t_cipc)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   602
  have "obj2sobj (e#s) (IPC (new_ipc s)) = SIPC (default_ipc_create_type r) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   603
    using ev SP nodel nodel_imp_exists[where obj = "IPC (new_ipc s)" and s=s]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   604
    by (auto simp:obj2sobj.simps ni_notin_curi cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   605
            split:option.splits dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   606
  ultimately show ?case using ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   607
    apply (rule_tac x = "IPC (new_ipc s)" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   608
    by (rule_tac x = "e # s" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   609
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   610
  case (ts'_read t sd srf r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   611
  then obtain f s  where TF: "(File f) \<in> tainted s" and vds: "valid s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   612
    and "f \<in> current_files s \<and> obj2sobj s (File f) = SFile (t, sd) srf \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   613
         no_del_event s \<and> initp_intact s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   614
    apply (clarsimp, frule_tac obj2sobj_file)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   615
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   616
  with ts'_read(3) obtain p s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   617
    and SF: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   618
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   619
    and intacts': "initp_intact (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   620
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   621
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   622
    and exf: "f \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   623
    apply (drule_tac s' =s and obj' = "File f" in all_sobjs_E0, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   624
    apply (frule_tac obj = "File f" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   625
    by (frule obj2sobj_proc, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   626
  obtain e \<tau> where ev: "e = ReadFile p f" and tau: "\<tau> = s' @ s" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   627
dcde836219bc add thy files
chunhan
parents:
diff changeset
   628
  have valid: "valid (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   629
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   630
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   631
      using ev tau by (simp add:exf exp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   632
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   633
      using ev tau ts'_read(4) SP SF
dcde836219bc add thy files
chunhan
parents:
diff changeset
   634
      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   635
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   636
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   637
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   638
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   639
  have "obj2sobj (e # \<tau>) (Proc p) = SProc (r, fr, pt, u) srp" using valid tau ev SP
dcde836219bc add thy files
chunhan
parents:
diff changeset
   640
    by (auto simp:obj2sobj.simps cp2sproc_simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   641
  moreover  have "Proc p \<in> tainted (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   642
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   643
    have "(File f) \<in> tainted \<tau>" using TF nodels' tau vds'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   644
      by (drule_tac s = "s'" in t_remain_app,auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   645
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   646
      by (drule_tac t_read, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   647
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   648
  have "no_del_event (e # \<tau>)" using ev tau nodels' by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   649
  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   650
    by (auto intro!:initp_intact_I_others) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   651
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   652
    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   653
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   654
  case (ts'_write r fr pt u srp t sd srf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   655
  then obtain p s where TP: "(Proc p) \<in> tainted s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   656
    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   657
         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   658
    apply (clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   659
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   660
  with ts'_write(3) obtain f s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   661
    and SF: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   662
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   663
    and intacts': "initp_intact (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   664
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   665
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   666
    and exf: "f \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   667
    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   668
    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   669
    by (frule obj2sobj_file, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   670
  obtain e \<tau> where ev: "e = WriteFile p f" and tau: "\<tau> = (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   671
dcde836219bc add thy files
chunhan
parents:
diff changeset
   672
  have valid: "valid (e# \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   673
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   674
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   675
      using ev tau exp exf by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   676
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   677
      using ev tau ts'_write(4) SP SF
dcde836219bc add thy files
chunhan
parents:
diff changeset
   678
      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   679
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   680
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   681
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   682
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   683
  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   684
  have "obj2sobj (e#\<tau>) (File f) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   685
    using tau ev SF valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   686
    by (auto simp:obj2sobj.simps source_dir_simps etype_of_file_def
dcde836219bc add thy files
chunhan
parents:
diff changeset
   687
            split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   688
  moreover  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   689
    by (auto intro!:initp_intact_I_others)   moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   690
  have "File f \<in> tainted (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   691
  proof- 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   692
    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   693
      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   694
    thus ?thesis using ev tau valid by (auto intro:t_write)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   695
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   696
  ultimately show ?case using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   697
    by (rule_tac x = "File f" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   698
next 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   699
  case (ts'_recv t sri r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   700
  then obtain i s  where TI: "(IPC i) \<in> tainted s" and vds: "valid s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   701
    and "i \<in> current_ipcs s \<and> obj2sobj s (IPC i) = SIPC t sri \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   702
         no_del_event s \<and> initp_intact s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   703
    apply (clarsimp, frule_tac obj2sobj_ipc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   704
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   705
  with ts'_recv(3) obtain p s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   706
    and SI: "obj2sobj (s'@s) (IPC i) = SIPC t sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   707
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   708
    and intacts': "initp_intact (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   709
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   710
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   711
    and exi: "i \<in> current_ipcs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   712
    apply (drule_tac s' =s and obj' = "IPC i" in all_sobjs_E0, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   713
    apply (frule_tac obj = "IPC i" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   714
    by (frule obj2sobj_proc, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   715
  obtain e \<tau> where ev: "e = Recv p i" and tau: "\<tau> = s' @ s" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   716
dcde836219bc add thy files
chunhan
parents:
diff changeset
   717
  have valid: "valid (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   718
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   719
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   720
      using ev tau by (simp add:exi exp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   721
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   722
      using ev tau ts'_recv(4) SP SI
dcde836219bc add thy files
chunhan
parents:
diff changeset
   723
      by (auto simp:cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   724
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   725
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   726
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   727
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   728
  have "obj2sobj (e # \<tau>) (Proc p) = SProc (r, fr, pt, u) srp" using valid tau ev SP
dcde836219bc add thy files
chunhan
parents:
diff changeset
   729
    by (auto simp:obj2sobj.simps cp2sproc_simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   730
  moreover  have "Proc p \<in> tainted (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   731
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   732
    have "(IPC i) \<in> tainted \<tau>" using TI nodels' tau vds'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   733
      by (drule_tac s = "s'" in t_remain_app,auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   734
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   735
      by (drule_tac t_recv, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   736
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   737
  have "no_del_event (e # \<tau>)" using ev tau nodels' by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   738
  moreover have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   739
    by (auto intro!:initp_intact_I_others)   
dcde836219bc add thy files
chunhan
parents:
diff changeset
   740
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   741
    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   742
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   743
  case (ts'_send r fr pt u srp t sri)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   744
  then obtain p s where TP: "(Proc p) \<in> tainted s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   745
    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   746
         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   747
    apply (clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   748
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   749
  with ts'_send(3) obtain i s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   750
    and SI: "obj2sobj (s'@s) (IPC i) = SIPC t sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   751
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   752
    and intacts': "initp_intact (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   753
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   754
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   755
    and exi: "i \<in> current_ipcs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   756
    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   757
    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   758
    by (frule obj2sobj_ipc, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   759
  obtain e \<tau> where ev: "e = Send p i" and tau: "\<tau> = (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   760
dcde836219bc add thy files
chunhan
parents:
diff changeset
   761
  have valid: "valid (e# \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   762
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   763
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   764
      using ev tau exp exi by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   765
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   766
      using ev tau ts'_send(4) SP SI
dcde836219bc add thy files
chunhan
parents:
diff changeset
   767
      by (auto simp:cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   768
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   769
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   770
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   771
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   772
  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   773
  have "obj2sobj (e#\<tau>) (IPC i) = SIPC t sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   774
    using tau ev SI valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   775
    by (auto simp:obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   776
  moreover have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   777
    by (auto intro!:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   778
  have "IPC i \<in> tainted (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   779
  proof- 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   780
    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   781
      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   782
    thus ?thesis using ev tau valid by (auto intro:t_send)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   783
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   784
  ultimately show ?case using tau ev 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   785
    by (rule_tac x = "IPC i" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   786
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   787
  case (ts'_crole r fr pt u srp r')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   788
  then obtain p s where exp: "p \<in> current_procs s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   789
    and TP: "(Proc p) \<in> tainted s" and vds: "valid s" and nodels: "no_del_event s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   790
    and SP: "obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   791
    and intacts: "initp_intact s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   792
    apply (clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   793
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   794
  obtain e \<tau> where ev: "e = ChangeRole (new_proc s) r'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   795
    and tau: "\<tau> = Clone p (new_proc s) # s" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   796
  hence vs_tau:"valid \<tau>" using exp vds by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   797
dcde836219bc add thy files
chunhan
parents:
diff changeset
   798
  have valid: "valid (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   799
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   800
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   801
      using ev tau exp by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   802
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   803
      using ev tau ts'_crole(3) SP
dcde836219bc add thy files
chunhan
parents:
diff changeset
   804
      by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   805
    ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   806
      by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   807
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   808
  have "obj2sobj (e # \<tau>) (Proc (new_proc s)) = SProc (r',fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   809
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   810
    have "obj2sobj \<tau> (Proc (new_proc s)) = SProc (r,fr,pt,u) srp" using SP tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   811
      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   812
    thus ?thesis using valid ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   813
      by (auto simp:cp2sproc_crole obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   814
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   815
  have "Proc (new_proc s) \<in> tainted (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   816
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   817
    have "(Proc (new_proc s)) \<in> tainted \<tau>" using TP tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   818
      by (auto intro!:t_clone)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   819
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   820
      by (drule_tac t_remain, auto dest:valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   821
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   822
  have "no_del_event (e # \<tau>)" using ev tau nodels by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   823
  moreover have "initp_intact (e#\<tau>)" using ev intacts valid nodels tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   824
    by (simp add:initp_intact_I_crole)  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   825
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   826
    by (rule_tac x = "Proc (new_proc s)" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   827
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   828
  case (ts'_chown r fr pt u srp u' nr)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   829
  then obtain p s where exp: "p \<in> current_procs s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   830
    and TP: "(Proc p) \<in> tainted s" and vds: "valid s" and nodels: "no_del_event s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   831
    and SP: "obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   832
    and intacts: "initp_intact s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   833
    apply (clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   834
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   835
  obtain e \<tau> where ev: "e = ChangeOwner (new_proc s) u'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   836
    and tau: "\<tau> = Clone p (new_proc s) # s" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   837
  hence vs_tau:"valid \<tau>" using exp vds by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   838
dcde836219bc add thy files
chunhan
parents:
diff changeset
   839
  have valid: "valid (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   840
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   841
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   842
      using ev tau exp ts'_chown(3) by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   843
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   844
      using ev tau ts'_chown(5) SP
dcde836219bc add thy files
chunhan
parents:
diff changeset
   845
      by (auto simp:cp2sproc.simps obj2sobj.simps pct_def clone_type_unchange
dcde836219bc add thy files
chunhan
parents:
diff changeset
   846
              split:option.splits t_rc_proc_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   847
        (* here is another place of no_limit of clone event assumption *)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   848
    ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   849
      by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   850
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   851
  have "obj2sobj (e # \<tau>) (Proc (new_proc s)) = SProc (nr, fr, chown_type_aux r nr pt, u') srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   852
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   853
    have "obj2sobj \<tau> (Proc (new_proc s)) = SProc (r,fr,pt,u) srp" using SP tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   854
      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   855
    thus ?thesis using valid ev ts'_chown(4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   856
      by (auto simp:cp2sproc_chown obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   857
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   858
  have "Proc (new_proc s) \<in> tainted (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   859
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   860
    have "Proc (new_proc s) \<in> tainted \<tau>" using TP tau vs_tau exp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   861
      by (auto intro!:t_clone)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   862
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   863
      by (drule_tac t_remain, auto dest:valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   864
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   865
  have "no_del_event (e # \<tau>)" using ev tau nodels by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   866
  moreover have "initp_intact (e#\<tau>)"  using ev intacts valid nodels tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   867
    by (simp add:initp_intact_I_chown)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   868
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   869
    by (rule_tac x = "Proc (new_proc s)" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   870
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   871
dcde836219bc add thy files
chunhan
parents:
diff changeset
   872
lemma ts2t:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   873
  "sobj \<in> tainted_s' \<Longrightarrow> \<exists> obj s. obj2sobj s obj = sobj \<and> obj \<in> tainted s \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   874
                                 sobj_source_eq_obj sobj obj \<and> no_del_event s \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   875
                                 initp_intact_but s sobj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   876
proof (induct rule:tainted_s'.induct)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   877
  case (ts'_init obj)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   878
  hence ex: "exists [] obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   879
    apply (drule_tac seeds_in_init) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   880
    by (case_tac obj, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   881
  have "obj2sobj [] obj = init_obj2sobj obj" using ex 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   882
    by (simp add:obj2sobj_nil_init)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   883
  moreover have "obj \<in> tainted []" using ts'_init by (simp add:t_init)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   884
  moreover have "sobj_source_eq_obj (init_obj2sobj obj) obj" using ex
dcde836219bc add thy files
chunhan
parents:
diff changeset
   885
    apply (frule_tac init_obj_has_sobj)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   886
    apply (case_tac "init_obj2sobj obj", case_tac[!] obj)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   887
    by (auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   888
  ultimately show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   889
    apply (rule_tac x = obj in exI, rule_tac x = "[]" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   890
    by (auto simp:initp_intact_but_nil)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   891
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   892
  case (ts'_exec1 t sd srf r fr pt u srp r' fr')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   893
  then obtain f s  where "sobj_source_eq_obj (SFile (t, sd) srf) (File f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   894
    and TF: "(File f) \<in> tainted s" and vds: "valid s" and "f \<in> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   895
    and "obj2sobj s (File f) = SFile (t, sd) srf \<and> no_del_event s \<and> initp_intact s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   896
    apply (clarsimp, frule_tac obj2sobj_file)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   897
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   898
  with ts'_exec1(3) obtain p s' where vds's: "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   899
    and etype: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   900
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   901
    and intacts': "initp_intact_but (s'@s) (SProc (r,fr,pt,u) srp)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   902
    and ISP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   903
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   904
    and exf: "f \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   905
    and sreq: "sobj_source_eq_obj (SProc (r,fr,pt,u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   906
    apply (drule_tac s' =s and obj' = "File f" in all_sobjs_E2, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   907
    apply (frule_tac obj = "File f" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   908
    by (frule obj2sobj_proc, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   909
  from vds's ISP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   910
  with exp vds's ISP have initp: "p \<in> init_processes" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   911
    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   912
  obtain ev \<tau> where ev: "ev = Execute p f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   913
    and tau: "\<tau> = Clone p (new_proc (s'@s)) # (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   914
  hence vs_tau:"valid \<tau>" using exp vds's by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   915
dcde836219bc add thy files
chunhan
parents:
diff changeset
   916
  have valid: "valid (ev # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   917
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   918
    have "os_grant \<tau> ev"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   919
      using ev tau by (simp add:exf exp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   920
    moreover have "rc_grant \<tau> ev" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   921
      using ev tau ts'_exec1 ISP etype
dcde836219bc add thy files
chunhan
parents:
diff changeset
   922
      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   923
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   924
    ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   925
      by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   926
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   927
  have "obj2sobj (ev # \<tau>) (Proc p) = SProc (r', fr', exec_type_aux r pt, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   928
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   929
    have "obj2sobj \<tau> (Proc p) = SProc (r,fr,pt,u) srp" using ISP tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   930
      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   931
    moreover have "source_dir \<tau> f = Some sd" using vs_tau etype tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   932
      by (auto simp:obj2sobj.simps source_dir_simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   933
    ultimately show ?thesis using valid ts'_exec1(5) ts'_exec1(6) ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   934
      by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   935
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   936
  have "Proc p \<in> tainted (ev # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   937
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   938
    have "(File f) \<in> tainted \<tau>" using TF nodels' tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   939
      by (drule_tac s = "Clone p (new_proc (s' @ s)) # s'" in t_remain_app,auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   940
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   941
      by (drule_tac t_exec, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   942
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   943
  have "no_del_event (ev # \<tau>)" using ev tau nodels' by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   944
  moreover have "initp_intact_but (ev#\<tau>) (SProc (r', fr', exec_type_aux r pt, u) srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   945
    using ev tau nodels' intacts' srpeq valid initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   946
    by (simp, rule_tac initp_intact_butp_I_exec, simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   947
  moreover have "sobj_source_eq_obj (SProc (r', fr', exec_type_aux r pt, u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   948
    using sreq by (case_tac srp, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   949
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   950
    by (rule_tac x = "Proc p" in exI, rule_tac x = "ev#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   951
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   952
  case (ts'_exec2 r fr pt u srp t sd srf r' fr')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   953
  then obtain p s  where sreq:"sobj_source_eq_obj (SProc (r, fr, pt, u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   954
    and TP: "(Proc p) \<in> tainted s" and vds: "valid s" and "p \<in> current_procs s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   955
    and "obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp \<and> no_del_event s \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   956
         initp_intact_but s (SProc (r, fr, pt, u) srp)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   957
    apply (clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   958
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   959
  with ts'_exec2(3) obtain f s' where vds's: "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   960
    and etype: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   961
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   962
    and intacts': "initp_intact_but (s'@s) (SProc (r,fr,pt,u) srp)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   963
    and ISP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   964
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   965
    and exf: "f \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   966
    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E1, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   967
    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   968
    by (frule obj2sobj_file, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   969
  from vds's ISP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   970
  with exp vds's ISP have initp: "p \<in> init_processes" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   971
    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   972
  obtain ev \<tau> where ev: "ev = Execute p f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   973
    and tau: "\<tau> = Clone p (new_proc (s'@s)) # (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   974
  hence vs_tau:"valid \<tau>" using exp vds's by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   975
dcde836219bc add thy files
chunhan
parents:
diff changeset
   976
  have valid: "valid (ev # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   977
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   978
    have "os_grant \<tau> ev"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   979
      using ev tau by (simp add:exf exp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   980
    moreover have "rc_grant \<tau> ev" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   981
      using ev tau ts'_exec2 ISP etype
dcde836219bc add thy files
chunhan
parents:
diff changeset
   982
      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   983
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   984
    ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   985
      by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   986
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   987
  have "obj2sobj (ev # \<tau>) (Proc p) = SProc (r', fr', exec_type_aux r pt, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   988
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   989
    have "obj2sobj \<tau> (Proc p) = SProc (r,fr,pt,u) srp" using ISP tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   990
      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   991
    moreover have "source_dir \<tau> f = Some sd" using vs_tau etype tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   992
      by (auto simp:source_dir_simps obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   993
    ultimately show ?thesis using valid ts'_exec2(5) ts'_exec2(6) ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   994
      by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   995
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   996
  have "Proc p \<in> tainted (ev # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   997
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   998
    have "(Proc p) \<in> tainted \<tau>" using TP nodels' tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   999
      by (drule_tac s = "Clone p (new_proc (s' @ s)) # s'" in t_remain_app,auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1000
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1001
      by (drule_tac t_remain, auto dest:valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1002
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1003
  have "no_del_event (ev # \<tau>)" using ev tau nodels' by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1004
  moreover have "initp_intact_but (ev#\<tau>) (SProc (r', fr', exec_type_aux r pt, u) srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1005
    using ev tau nodels' intacts' srpeq valid initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1006
    by (simp, rule_tac initp_intact_butp_I_exec, simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1007
  moreover have "sobj_source_eq_obj (SProc (r', fr', exec_type_aux r pt, u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1008
    using sreq by (case_tac srp, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1009
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1010
    by (rule_tac x = "Proc p" in exI, rule_tac x = "ev#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1011
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1012
  case (ts'_cfd r fr pt u srp t sd srf t')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1013
  from ts'_cfd(1) obtain p s where TP: "(Proc p) \<in> tainted s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1014
    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1015
         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1016
    apply (drule_tac ts2t_intact, clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1017
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1018
  with ts'_cfd(3) obtain pf s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1019
    and SF: "obj2sobj (s'@s) (File pf) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1020
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1021
    and intacts': "initp_intact (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1022
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1023
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1024
    and exf: "pf \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1025
    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1026
    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1027
    by (frule obj2sobj_file, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1028
  obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1029
    and tau: "\<tau> = (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1030
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1031
  have valid: "valid (e# \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1032
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1033
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1034
      using ev tau exp exf by (simp add:ncf_notin_curf ncf_parent)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1035
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1036
      using ev tau ts'_cfd(4,5,6) SP SF
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1037
      by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1038
        split:if_splits option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1039
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1040
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1041
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1042
  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1043
  have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t', sd) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1044
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1045
    have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1046
      using ev tau SF SP ts'_cfd(4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1047
      by (auto simp:obj2sobj.simps etype_of_file_def cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1048
              split:option.splits if_splits  intro!:etype_aux_prop4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1049
    moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1050
      using ev tau SF SP valid ncf_parent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1051
      by (auto simp:source_dir_simps obj2sobj.simps cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1052
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1053
    ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1054
      nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1055
      by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1056
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1057
  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1058
    by (auto intro!:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1059
  have "File (new_childf pf \<tau>) \<in> tainted (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1060
  proof- 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1061
    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1062
      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1063
    thus ?thesis using ev tau valid by (auto intro:t_cfile)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1064
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1065
  ultimately show ?case using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1066
    apply (rule_tac x = "File (new_childf pf \<tau>)" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1067
    by (rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1068
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1069
  case (ts'_cfd' r fr pt u srp t sd srf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1070
  from ts'_cfd'(1) obtain p s where TP: "(Proc p) \<in> tainted s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1071
    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1072
         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1073
    apply (drule_tac ts2t_intact, clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1074
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1075
  with ts'_cfd'(3) obtain pf s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1076
    and SF: "obj2sobj (s'@s) (File pf) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1077
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1078
    and intacts': "initp_intact (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1079
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1080
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1081
    and exf: "pf \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1082
    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1083
    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1084
    by (frule obj2sobj_file, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1085
  obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1086
    and tau: "\<tau> = (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1087
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1088
  have valid: "valid (e# \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1089
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1090
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1091
      using ev tau exp exf by (simp add:ncf_notin_curf ncf_parent)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1092
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1093
      using ev tau ts'_cfd'(4,5) SP SF
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1094
      by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1095
        split:if_splits option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1096
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1097
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1098
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1099
  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1100
  have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t, sd) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1101
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1102
      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1103
      proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1104
        have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = etype_of_file \<tau> pf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1105
          using ev tau SP ts'_cfd'(4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1106
          by (auto simp:obj2sobj.simps ncf_parent etype_of_file_def cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1107
            split:option.splits   intro!:etype_aux_prop3)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1108
        thus ?thesis using SF tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1109
          by (auto simp:obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1110
      qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1111
    moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1112
      using ev tau SF SP valid ncf_parent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1113
      by (auto simp:source_dir_simps obj2sobj.simps cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1114
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1115
    ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1116
      nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1117
      by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1118
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1119
  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1120
    by (auto intro!:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1121
  have "File (new_childf pf \<tau>) \<in> tainted (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1122
  proof- 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1123
    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1124
      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1125
    thus ?thesis using ev tau valid by (auto intro:t_cfile)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1126
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1127
  ultimately show ?case using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1128
    apply (rule_tac x = "File (new_childf pf \<tau>)" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1129
    by (rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1130
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1131
  case (ts'_cipc r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1132
  from ts'_cipc(1) obtain p s where TP: "(Proc p) \<in> tainted s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1133
    and vds: "valid s" and exp: "p \<in> current_procs s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1134
    and nodels: "no_del_event s" and intacts: "initp_intact s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1135
    and SP: "obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1136
    apply (drule_tac ts2t_intact, clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1137
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1138
  obtain e where ev: "e = CreateIPC p (new_ipc s)" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1139
      
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1140
  have valid: "valid (e # s)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1141
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1142
    have "os_grant s e"
6
4294abb1f38c update scripts with no new_* in admissable-check
chunhan
parents: 1
diff changeset
  1143
      using ev exp by (simp add:ni_notin_curi)
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1144
    moreover have "rc_grant s e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1145
      using ev ts'_cipc(3) SP
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1146
      by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1147
    ultimately show ?thesis using vds
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1148
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1149
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1150
  have nodel: "no_del_event (e#s)" using nodels ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1151
  have "initp_intact (e#s)" using ev intacts valid nodels
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1152
    by (auto intro!:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1153
  have "IPC (new_ipc s) \<in> tainted (e#s)" using TP ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1154
    by (auto intro:t_cipc)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1155
  have "obj2sobj (e#s) (IPC (new_ipc s)) = SIPC (default_ipc_create_type r) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1156
    using ev SP nodel nodel_imp_exists[where obj = "IPC (new_ipc s)" and s=s]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1157
    by (auto simp:obj2sobj.simps ni_notin_curi cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1158
            split:option.splits dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1159
  ultimately show ?case using ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1160
    apply (rule_tac x = "IPC (new_ipc s)" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1161
    by (rule_tac x = "e # s" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1162
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1163
  case (ts'_read t sd srf r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1164
  then obtain f s  where "sobj_source_eq_obj (SFile (t, sd) srf) (File f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1165
    and TF: "(File f) \<in> tainted s" and vds: "valid s" and "f \<in> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1166
    and "obj2sobj s (File f) = SFile (t, sd) srf \<and> no_del_event s \<and> initp_intact s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1167
    apply (clarsimp, frule_tac obj2sobj_file)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1168
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1169
  with ts'_read(3) obtain p s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1170
    and SF: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1171
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1172
    and intacts': "initp_intact_but (s'@s) (SProc (r,fr,pt,u) srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1173
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1174
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1175
    and exf: "f \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1176
    and sreq: "sobj_source_eq_obj (SProc (r,fr,pt,u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1177
    apply (drule_tac s' =s and obj' = "File f" in all_sobjs_E2, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1178
    apply (frule_tac obj = "File f" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1179
    by (frule obj2sobj_proc, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1180
  from vds' SP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1181
  with exp vds' SP have initp: "p \<in> init_processes" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1182
    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1183
  obtain e \<tau> where ev: "e = ReadFile p f" and tau: "\<tau> = s' @ s" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1184
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1185
  have valid: "valid (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1186
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1187
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1188
      using ev tau by (simp add:exf exp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1189
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1190
      using ev tau ts'_read(4) SP SF
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1191
      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1192
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1193
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1194
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1195
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1196
  have "obj2sobj (e # \<tau>) (Proc p) = SProc (r, fr, pt, u) srp" using valid tau ev SP
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1197
    by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1198
  moreover  have "Proc p \<in> tainted (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1199
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1200
    have "(File f) \<in> tainted \<tau>" using TF nodels' tau vds'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1201
      by (drule_tac s = "s'" in t_remain_app,auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1202
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1203
      by (drule_tac t_read, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1204
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1205
  have "no_del_event (e # \<tau>)" using ev tau nodels' by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1206
  moreover  have "initp_intact_but (e#\<tau>) (SProc (r, fr, pt, u) srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1207
    using ev tau nodels' intacts' srpeq valid initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1208
    by (simp, rule_tac initp_intact_butp_I_others, simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1209
  moreover have "sobj_source_eq_obj (SProc (r,fr,pt,u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1210
    using sreq by (case_tac srp, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1211
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1212
    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1213
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1214
  case (ts'_write r fr pt u srp t sd srf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1215
  from ts'_write(1) obtain p s where TP: "(Proc p) \<in> tainted s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1216
    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1217
         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1218
    apply (drule_tac ts2t_intact, clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1219
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1220
  with ts'_write(3) obtain f s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1221
    and SF: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1222
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1223
    and intacts': "initp_intact (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1224
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1225
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1226
    and exf: "f \<in> current_files (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1227
    and sreq: "sobj_source_eq_obj (SFile (t, sd) srf) (File f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1228
    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E2, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1229
    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1230
    by (frule obj2sobj_file, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1231
  obtain e \<tau> where ev: "e = WriteFile p f" and tau: "\<tau> = (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1232
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1233
  have valid: "valid (e# \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1234
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1235
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1236
      using ev tau exp exf by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1237
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1238
      using ev tau ts'_write(4) SP SF
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1239
      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1240
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1241
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1242
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1243
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1244
  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1245
  have "obj2sobj (e#\<tau>) (File f) = SFile (t, sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1246
    using tau ev SF valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1247
    by (auto simp:obj2sobj.simps source_dir_simps etype_of_file_def
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1248
            split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1249
  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1250
  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1251
    by (auto intro!:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1252
  have "File f \<in> tainted (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1253
  proof- 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1254
    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1255
      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1256
    thus ?thesis using ev tau valid by (auto intro:t_write)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1257
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1258
  ultimately show ?case using tau ev sreq
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1259
    by (rule_tac x = "File f" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1260
next 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1261
  case (ts'_recv t sri r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1262
  then obtain i s  where "sobj_source_eq_obj (SIPC t sri) (IPC i)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1263
    and TI: "(IPC i) \<in> tainted s" and vds: "valid s" and "i \<in> current_ipcs s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1264
    and "obj2sobj s (IPC i) = SIPC t sri \<and> no_del_event s \<and> initp_intact s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1265
    apply (clarsimp, frule_tac obj2sobj_ipc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1266
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1267
  with ts'_recv(3) obtain p s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1268
    and SI: "obj2sobj (s'@s) (IPC i) = SIPC t sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1269
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1270
    and intacts': "initp_intact_but (s'@s) (SProc (r,fr,pt,u) srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1271
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1272
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1273
    and exi: "i \<in> current_ipcs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1274
    and sreq: "sobj_source_eq_obj (SProc (r,fr,pt,u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1275
    apply (drule_tac s' =s and obj' = "IPC i" in all_sobjs_E2, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1276
    apply (frule_tac obj = "IPC i" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1277
    by (frule obj2sobj_proc, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1278
  from vds' SP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1279
  with exp vds' SP have initp: "p \<in> init_processes" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1280
    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1281
  obtain e \<tau> where ev: "e = Recv p i" and tau: "\<tau> = s' @ s" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1282
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1283
  have valid: "valid (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1284
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1285
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1286
      using ev tau by (simp add:exi exp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1287
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1288
      using ev tau ts'_recv(4) SP SI
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1289
      by (auto simp:cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1290
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1291
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1292
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1293
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1294
  have "obj2sobj (e # \<tau>) (Proc p) = SProc (r, fr, pt, u) srp" using valid tau ev SP
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1295
    by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1296
  moreover  have "Proc p \<in> tainted (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1297
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1298
    have "(IPC i) \<in> tainted \<tau>" using TI nodels' tau vds'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1299
      by (drule_tac s = "s'" in t_remain_app,auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1300
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1301
      by (drule_tac t_recv, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1302
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1303
  have "no_del_event (e # \<tau>)" using ev tau nodels' by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1304
  moreover  have "initp_intact_but (e#\<tau>) (SProc (r, fr, pt, u) srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1305
    using ev tau nodels' intacts' srpeq valid initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1306
    by (simp, rule_tac initp_intact_butp_I_others, simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1307
  moreover have "sobj_source_eq_obj (SProc (r,fr,pt,u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1308
    using sreq by (case_tac srp, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1309
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1310
    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1311
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1312
  case (ts'_send r fr pt u srp t sri)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1313
  from ts'_send(1) obtain p s where TP: "(Proc p) \<in> tainted s" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1314
    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1315
         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1316
    apply (drule_tac ts2t_intact, clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1317
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1318
  with ts'_send(3) obtain i s' where vds': "valid (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1319
    and SI: "obj2sobj (s'@s) (IPC i) = SIPC t sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1320
    and nodels': "no_del_event (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1321
    and intacts': "initp_intact (s'@s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1322
    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1323
    and exp: "p \<in> current_procs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1324
    and exi: "i \<in> current_ipcs (s' @ s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1325
    and sreq: "sobj_source_eq_obj (SIPC t sri) (IPC i)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1326
    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E2, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1327
    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1328
    by (frule obj2sobj_ipc, clarsimp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1329
  obtain e \<tau> where ev: "e = Send p i" and tau: "\<tau> = (s' @ s)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1330
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1331
  have valid: "valid (e# \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1332
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1333
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1334
      using ev tau exp exi by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1335
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1336
      using ev tau ts'_send(4) SP SI
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1337
      by (auto simp:cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1338
              split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1339
    ultimately show ?thesis using vds' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1340
      by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1341
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1342
  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1343
  have "obj2sobj (e#\<tau>) (IPC i) = SIPC t sri"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1344
    using tau ev SI valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1345
    by (auto simp:obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1346
  moreover  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1347
    by (auto intro!:initp_intact_I_others)  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1348
  moreover  have "IPC i \<in> tainted (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1349
  proof- 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1350
    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1351
      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1352
    thus ?thesis using ev tau valid by (auto intro:t_send)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1353
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1354
  ultimately show ?case using tau ev sreq
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1355
    by (rule_tac x = "IPC i" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1356
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1357
  case (ts'_crole r fr pt u srp r')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1358
  then obtain p s where exp: "p \<in> current_procs s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1359
    and sreq:"sobj_source_eq_obj (SProc (r, fr, pt, u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1360
    and TP: "(Proc p) \<in> tainted s" and vds: "valid s" and nodels: "no_del_event s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1361
    and SP: "obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1362
    and intacts: "initp_intact_but s (SProc (r, fr, pt, u) srp)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1363
    apply (clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1364
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1365
  from vds SP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1366
  with exp vds SP have initp: "p \<in> init_processes" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1367
    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1368
  obtain e \<tau> where ev: "e = ChangeRole p r'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1369
    and tau: "\<tau> = Clone p (new_proc s) # s" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1370
  hence vs_tau:"valid \<tau>" using exp vds by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1371
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1372
  have valid: "valid (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1373
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1374
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1375
      using ev tau exp by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1376
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1377
      using ev tau ts'_crole(3) SP
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1378
      by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1379
    ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1380
      by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1381
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1382
  have "obj2sobj (e # \<tau>) (Proc p) = SProc (r', fr, pt, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1383
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1384
    have "obj2sobj \<tau> (Proc p) = SProc (r,fr,pt,u) srp" using SP tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1385
      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1386
    thus ?thesis using valid ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1387
      by (auto simp:cp2sproc_crole obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1388
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1389
  have "Proc p \<in> tainted (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1390
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1391
    have "(Proc p) \<in> tainted \<tau>" using TP tau vs_tau exp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1392
      by (auto intro!:t_remain)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1393
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1394
      by (drule_tac t_remain, auto dest:valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1395
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1396
  have "no_del_event (e # \<tau>)" using ev tau nodels by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1397
  moreover have "initp_intact_but (e#\<tau>) (SProc (r', fr, pt, u) srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1398
    using ev tau nodels intacts srpeq valid initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1399
    by (simp, rule_tac initp_intact_butp_I_crole, simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1400
  moreover have "sobj_source_eq_obj (SProc (r',fr,pt,u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1401
    using sreq by (case_tac srp, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1402
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1403
    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1404
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1405
  case (ts'_chown r fr pt u srp u' nr)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1406
  then obtain p s where exp: "p \<in> current_procs s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1407
    and sreq:"sobj_source_eq_obj (SProc (r, fr, pt, u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1408
    and TP: "(Proc p) \<in> tainted s" and vds: "valid s" and nodels: "no_del_event s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1409
    and SP: "obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1410
    and intacts: "initp_intact_but s (SProc (r, fr, pt, u) srp)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1411
    apply (clarsimp, frule_tac obj2sobj_proc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1412
    by (frule tainted_is_valid, frule tainted_is_current, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1413
  from vds SP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1414
  with exp vds SP have initp: "p \<in> init_processes" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1415
    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1416
  obtain e \<tau> where ev: "e = ChangeOwner p u'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1417
    and tau: "\<tau> = Clone p (new_proc s) # s" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1418
  hence vs_tau:"valid \<tau>" using exp vds by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1419
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1420
  have valid: "valid (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1421
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1422
    have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1423
      using ev tau exp ts'_chown(3) by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1424
    moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1425
      using ev tau ts'_chown(5) SP
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1426
      by (auto simp:cp2sproc.simps obj2sobj.simps pct_def clone_type_unchange
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1427
              split:option.splits t_rc_proc_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1428
        (* here is another place of no_limit of clone event assumption *)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1429
    ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1430
      by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1431
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1432
  have "obj2sobj (e # \<tau>) (Proc p) = SProc (nr, fr, chown_type_aux r nr pt, u') srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1433
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1434
    have "obj2sobj \<tau> (Proc p) = SProc (r,fr,pt,u) srp" using SP tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1435
      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1436
    thus ?thesis using valid ev ts'_chown(4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1437
      by (auto simp:cp2sproc_chown obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1438
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1439
  have "Proc p \<in> tainted (e # \<tau>)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1440
  proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1441
    have "(Proc p) \<in> tainted \<tau>" using TP tau vs_tau exp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1442
      by (auto intro!:t_remain)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1443
    thus ?thesis using ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1444
      by (drule_tac t_remain, auto dest:valid_os)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1445
  qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1446
  have "no_del_event (e # \<tau>)" using ev tau nodels by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1447
  moreover have "initp_intact_but (e#\<tau>) (SProc (nr, fr, chown_type_aux r nr pt, u') srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1448
    using ev tau nodels intacts srpeq valid initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1449
    by (simp, rule_tac initp_intact_butp_I_chown, simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1450
  moreover have "sobj_source_eq_obj (SProc (nr, fr, chown_type_aux r nr pt, u') srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1451
    using sreq by (case_tac srp, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1452
  ultimately show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1453
    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1454
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1455
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1456
lemma tainted_s2tainted:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1457
  "sobj \<in> tainted_s \<Longrightarrow> \<exists> obj s. obj2sobj s obj = sobj \<and> obj \<in> tainted s \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1458
                                 sobj_source_eq_obj sobj obj \<and> no_del_event s \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1459
                                 initp_intact_but s sobj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1460
apply (simp add:tainted_s'_eq)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1461
by (erule ts2t)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1462
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1463
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1464
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1465
end