all_sobj_prop.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Tue, 30 Apr 2013 14:46:18 +0100
changeset 4 8b6ba7168f2d
parent 1 dcde836219bc
child 6 4294abb1f38c
permissions -rw-r--r--
pictures
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
     1
theory all_sobj_prop
dcde836219bc add thy files
chunhan
parents:
diff changeset
     2
imports Main rc_theory os_rc obj2sobj_prop deleted_prop sound_defs_prop source_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 initf_has_effinitialrole:
dcde836219bc add thy files
chunhan
parents:
diff changeset
     8
  "f \<in> init_files ==> \<exists> r. effinitialrole [] f = Some r"
dcde836219bc add thy files
chunhan
parents:
diff changeset
     9
by (rule_tac f = f in file_has_effinitialrole, simp, simp add:vs_nil)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    10
dcde836219bc add thy files
chunhan
parents:
diff changeset
    11
lemma initf_has_effforcedrole:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    12
  "f \<in> init_files ==> \<exists> r. effforcedrole [] f = Some r"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    13
by (rule_tac f = f in file_has_effforcedrole, simp, simp add:vs_nil)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    14
dcde836219bc add thy files
chunhan
parents:
diff changeset
    15
lemma efffrole_sdir_some:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    16
  "sd \<in> init_files ==> \<exists> r. erole_functor init_file_forcedrole InheritUpMixed sd = Some r"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    17
apply (frule_tac s = "[]" in undel_initf_keeps_efrole, simp, simp add:vs_nil)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    18
by (drule initf_has_effforcedrole, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    19
dcde836219bc add thy files
chunhan
parents:
diff changeset
    20
lemma efffrole_sdir_some':
dcde836219bc add thy files
chunhan
parents:
diff changeset
    21
  "erole_functor init_file_forcedrole InheritUpMixed sd = None ==> sd \<notin> init_files"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    22
by (rule notI, auto dest!:efffrole_sdir_some)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    23
dcde836219bc add thy files
chunhan
parents:
diff changeset
    24
lemma effirole_sdir_some:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    25
  "sd \<in> init_files ==> \<exists> r. erole_functor init_file_initialrole UseForcedRole sd = Some r"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    26
apply (frule_tac s = "[]" in undel_initf_keeps_eirole, simp, simp add:vs_nil)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    27
by (drule initf_has_effinitialrole, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    28
dcde836219bc add thy files
chunhan
parents:
diff changeset
    29
lemma effirole_sdir_some':
dcde836219bc add thy files
chunhan
parents:
diff changeset
    30
  "erole_functor init_file_initialrole UseForcedRole sd = None ==> sd \<notin> init_files"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    31
by (rule notI, auto dest:effirole_sdir_some)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    32
dcde836219bc add thy files
chunhan
parents:
diff changeset
    33
lemma erole_func_irole_simp:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    34
  "erole_functor init_file_initialrole UseForcedRole sd = effinitialrole [] sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    35
by (simp add:effinitialrole_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    36
dcde836219bc add thy files
chunhan
parents:
diff changeset
    37
lemma erole_func_frole_simp:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    38
  "erole_functor init_file_forcedrole InheritUpMixed sd = effforcedrole [] sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    39
by (simp add:effforcedrole_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    40
dcde836219bc add thy files
chunhan
parents:
diff changeset
    41
lemma init_effforcedrole_valid: "erole_functor init_file_forcedrole InheritUpMixed sd = Some r \<Longrightarrow> r = InheritUserRole \<or> r = InheritProcessRole  \<or> r = InheritUpMixed \<or> (\<exists> nr. r = NormalRole nr)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    42
by (simp add:erole_func_frole_simp, erule effforcedrole_valid)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    43
dcde836219bc add thy files
chunhan
parents:
diff changeset
    44
lemma init_effinitialrole_valid: "erole_functor init_file_initialrole UseForcedRole sd = Some r \<Longrightarrow> r = UseForcedRole \<or> (\<exists> nr. r = NormalRole nr)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    45
by (simp add:erole_func_irole_simp, erule effinitialrole_valid)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    46
dcde836219bc add thy files
chunhan
parents:
diff changeset
    47
lemma exec_role_some:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    48
  "[|sd \<in> init_files; u \<in> init_users|] ==> \<exists> r'. exec_role_aux r sd u = Some r'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    49
by (auto simp:exec_role_aux_def split:option.splits t_role.splits 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    50
        dest!:effirole_sdir_some' efffrole_sdir_some'
dcde836219bc add thy files
chunhan
parents:
diff changeset
    51
        dest:init_effforcedrole_valid init_effinitialrole_valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
    52
       intro:effirole_sdir_some efffrole_sdir_some user_has_normalrole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    53
dcde836219bc add thy files
chunhan
parents:
diff changeset
    54
lemma chown_role_some:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    55
  "u \<in> init_users ==> \<exists> r'. chown_role_aux r fr u = Some r'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    56
by (auto simp:chown_role_aux_def split:option.splits t_role.splits 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    57
        dest!:effirole_sdir_some' efffrole_sdir_some'
dcde836219bc add thy files
chunhan
parents:
diff changeset
    58
        dest:init_effforcedrole_valid init_effinitialrole_valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
    59
       intro:effirole_sdir_some efffrole_sdir_some user_has_normalrole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    60
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
    61
declare obj2sobj.simps [simp del]
dcde836219bc add thy files
chunhan
parents:
diff changeset
    62
dcde836219bc add thy files
chunhan
parents:
diff changeset
    63
lemma all_sobjs_I:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    64
  assumes ex: "exists s obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    65
  and vd: "valid s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    66
  shows "obj2sobj s obj \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    67
using ex vd
dcde836219bc add thy files
chunhan
parents:
diff changeset
    68
proof (induct s arbitrary:obj)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    69
  case Nil
dcde836219bc add thy files
chunhan
parents:
diff changeset
    70
  assume ex:"exists [] obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    71
  show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
    72
  proof (cases obj)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    73
    case (Proc p) assume prem: "obj = Proc p" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    74
    with ex have initp:"p \<in> init_processes" by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    75
    from initp obtain r where "init_currentrole p = Some r" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    76
      using init_proc_has_role by (auto simp:bidirect_in_init_def)  
dcde836219bc add thy files
chunhan
parents:
diff changeset
    77
    moreover from initp obtain t where "init_process_type p = Some t" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    78
      using init_proc_has_type by (auto simp:bidirect_in_init_def)  
dcde836219bc add thy files
chunhan
parents:
diff changeset
    79
    moreover from initp obtain fr where "init_proc_forcedrole p = Some fr" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    80
      using init_proc_has_frole by (auto simp:bidirect_in_init_def)  
dcde836219bc add thy files
chunhan
parents:
diff changeset
    81
    moreover from initp obtain u where "init_owner p = Some u" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    82
      using init_proc_has_owner by (auto simp:bidirect_in_init_def)  
dcde836219bc add thy files
chunhan
parents:
diff changeset
    83
    ultimately have "obj2sobj [] (Proc p) \<in> all_sobjs" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    84
      using initp by (auto intro!:ap_init simp:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    85
    with prem show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
    86
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
    87
    case (File f) assume prem: "obj = File f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    88
    with ex have initf: "f \<in> init_files" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
    89
    from initf obtain t where "etype_aux init_file_type_aux f = Some t" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    90
      using init_file_has_etype by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
    91
    moreover from initf have "source_dir [] f = Some f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    92
      by (simp add:source_dir_of_init')
dcde836219bc add thy files
chunhan
parents:
diff changeset
    93
    ultimately have "obj2sobj [] (File f) \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    94
      using initf by (auto simp:etype_of_file_def obj2sobj.simps intro!:af_init)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    95
    with prem show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
    96
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
    97
    case (IPC i) assume prem: "obj = IPC i"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    98
    with ex have initi: "i \<in> init_ipcs" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
    99
    from initi obtain t where "init_ipc_type i = Some t" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   100
      using init_ipc_has_type by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   101
    hence "obj2sobj [] (IPC i) \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   102
      using initi by (auto intro!:ai_init simp:obj2sobj.simps) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   103
    with prem show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   104
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   105
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   106
  case (Cons e s)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   107
  assume prem: "\<And> obj. \<lbrakk>exists s obj; valid s\<rbrakk> \<Longrightarrow> obj2sobj s obj \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   108
    and ex_cons: "exists (e # s) obj" and vs_cons: "valid (e # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   109
  have vs: "valid s" and rc: "rc_grant s e" and os: "os_grant s e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   110
    using vs_cons by (auto intro:valid_cons valid_os valid_rc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   111
  from prem and vs have prem': "\<And> obj. exists s obj \<Longrightarrow> obj2sobj s obj \<in> all_sobjs" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   112
  show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   113
  proof (cases e)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   114
    case (ChangeOwner p u)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   115
    assume ev: "e = ChangeOwner p u"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   116
    show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
   117
    proof (cases "obj = Proc p")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   118
      case True 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   119
      have curp: "p \<in> current_procs s" and exp: "exists s (Proc p)" using os ev by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   120
      from curp obtain r fr t u' srp where sp: "obj2sobj s (Proc p) = SProc (r,fr,t,u') (Some srp)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   121
        using vs apply (drule_tac current_proc_has_sobj, simp) by blast
dcde836219bc add thy files
chunhan
parents:
diff changeset
   122
      hence sp_in: "SProc (r,fr,t,u') (Some srp) \<in> all_sobjs" using prem' exp by metis
dcde836219bc add thy files
chunhan
parents:
diff changeset
   123
      have comp: "(r, Proc_type t, CHANGE_OWNER) \<in> compatible" using sp ev rc
dcde836219bc add thy files
chunhan
parents:
diff changeset
   124
        by (auto simp:obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   125
      from os ev have uinit: "u \<in> init_users" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   126
      then obtain nr where chown: "chown_role_aux r fr u = Some nr" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   127
        by (auto dest:chown_role_some)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   128
      hence nsp_in:"obj2sobj (e#s) (Proc p) = SProc (nr,fr,chown_type_aux r nr t, u) (Some srp)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   129
        using sp ev os
dcde836219bc add thy files
chunhan
parents:
diff changeset
   130
        by (auto split:option.splits t_role.splits 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   131
          simp del:currentrole.simps type_of_process.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   132
          simp add:chown_role_aux_valid chown_type_aux_valid obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   133
      thus ?thesis using True ev os rc sp_in sp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   134
        by (auto simp:chown comp intro!:ap_chown[where u'=u']) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   135
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   136
      case False
dcde836219bc add thy files
chunhan
parents:
diff changeset
   137
      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons
dcde836219bc add thy files
chunhan
parents:
diff changeset
   138
        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   139
                              split:option.splits t_role.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   140
      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   141
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   142
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   143
    case (Clone p p')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   144
    assume ev: "e = Clone p p'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   145
    show ?thesis 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   146
    proof (cases "obj = Proc p'")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   147
      case True
dcde836219bc add thy files
chunhan
parents:
diff changeset
   148
      from ev os have exp: "exists s (Proc p)" by (simp add:os_grant.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   149
      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
   150
        and srp: "source_proc s p = Some sp" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   151
        apply (simp del:cp2sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   152
        by (frule current_proc_has_sproc, simp, frule current_proc_has_srp, simp, blast)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   153
      hence SP: "SProc (r,fr,pt,u) (Some sp) \<in> all_sobjs" using exp prem'[where obj = "Proc p"] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   154
        by (auto split:option.splits simp add:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   155
      have "obj2sobj (e # s) (Proc p') = SProc (r,fr,clone_type_aux r pt, u) (Some sp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   156
        using ev sproc srp vs_cons 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   157
        by (simp add:obj2sobj.simps cp2sproc_clone del:cp2sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   158
      thus ?thesis using True SP by (simp add:ap_clone)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   159
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   160
      case False
dcde836219bc add thy files
chunhan
parents:
diff changeset
   161
      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons
dcde836219bc add thy files
chunhan
parents:
diff changeset
   162
        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   163
                              split:option.splits t_role.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   164
      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   165
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   166
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   167
    case (Execute p f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   168
    assume ev: "e = Execute p f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   169
    show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
   170
    proof (cases "obj = Proc p")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   171
      case True
dcde836219bc add thy files
chunhan
parents:
diff changeset
   172
      from ev os have exp: "exists s (Proc p)" and exf: "exists s (File f)" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   173
      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
   174
        and srp: "source_proc s p = Some sp" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   175
        apply (simp del:cp2sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   176
        by (frule current_proc_has_sproc, simp, frule current_proc_has_srp, simp, blast)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   177
      hence SP: "SProc (r,fr,pt,u) (Some sp) \<in> all_sobjs" using exp prem'[where obj = "Proc p"] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   178
        by (auto split:option.splits simp add:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   179
      from exf obtain sd t where srdir: "source_dir s f = Some sd" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   180
        etype: "etype_of_file s f = Some t" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   181
        by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   182
      then obtain srf where SF: "SFile (t, sd) srf \<in> all_sobjs" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   183
        using exf prem'[where obj = "File f"] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   184
        by (auto split:option.splits if_splits simp:obj2sobj.simps dest:current_file_has_etype)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   185
      from sproc srdir have "u \<in> init_users" and "sd \<in> init_files" using vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   186
        by (auto intro:source_dir_in_init owner_in_users split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   187
      then obtain nr where "exec_role_aux r sd u = Some nr" by (auto dest:exec_role_some) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   188
      
dcde836219bc add thy files
chunhan
parents:
diff changeset
   189
      hence "obj2sobj (e # s) (Proc p) \<in> all_sobjs" using ev vs_cons srdir sproc srp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   190
        apply (auto simp:obj2sobj.simps cp2sproc_simps source_proc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   191
                   intro:source_dir_in_init simp del:cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   192
                   split:option.splits dest!:efffrole_sdir_some')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   193
        apply (rule ap_exec) using SF SP rc ev etype by (auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   194
      with True show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   195
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   196
      case False
dcde836219bc add thy files
chunhan
parents:
diff changeset
   197
      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons
dcde836219bc add thy files
chunhan
parents:
diff changeset
   198
        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   199
                              split:option.splits t_role.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   200
      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   201
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   202
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   203
    case (CreateFile p f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   204
    assume ev: "e = CreateFile p f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   205
    show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
   206
    proof (cases "obj = File f")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   207
      case True
dcde836219bc add thy files
chunhan
parents:
diff changeset
   208
      from os ev obtain pf where expf: "exists s (File pf)" and parent:"parent f = Some pf" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   209
      from expf obtain pft sd srpf where SF: "SFile (pft, sd) srpf \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   210
        and eptype: "etype_of_file s pf = Some pft" and srpf: "source_dir s pf = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   211
        using prem'[where obj = "File pf"] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   212
        by (auto split:option.splits if_splits simp:obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   213
                  dest:current_file_has_etype current_file_has_sd)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   214
      from os ev have exp: "exists s (Proc p)" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   215
      then obtain r pt fr u srp where SP: "SProc (r, fr, pt, u) srp \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   216
        and sproc: "cp2sproc s p = Some (r, fr, pt, u)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   217
        using prem'[where obj = "Proc p"] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   218
        by (auto split:option.splits if_splits simp:obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   219
                  dest:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   220
      have "obj2sobj (e # s) (File f) \<in> all_sobjs" using ev vs_cons sproc srpf parent os
dcde836219bc add thy files
chunhan
parents:
diff changeset
   221
        apply (auto simp:obj2sobj.simps source_dir_simps init_notin_curf_deleted  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   222
                   split:option.splits dest!:current_file_has_etype')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   223
        apply (case_tac "default_fd_create_type r")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   224
        using SF SP rc ev eptype sproc
dcde836219bc add thy files
chunhan
parents:
diff changeset
   225
        apply (rule_tac sf = srpf in af_cfd', auto simp:etype_of_file_def etype_aux_prop3) [1]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   226
        using SF SP rc ev eptype sproc
dcde836219bc add thy files
chunhan
parents:
diff changeset
   227
        apply (rule_tac sf = srpf in af_cfd)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   228
        apply (auto simp:etype_of_file_def etype_aux_prop4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   229
        done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   230
      with True show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   231
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   232
      case False 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   233
      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   234
        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps etype_aux_prop2
dcde836219bc add thy files
chunhan
parents:
diff changeset
   235
                              split:option.splits t_role.splits )
dcde836219bc add thy files
chunhan
parents:
diff changeset
   236
      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   237
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   238
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   239
    case (CreateIPC p i)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   240
    assume ev: "e = CreateIPC p i"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   241
    show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
   242
    proof (cases "obj = IPC i")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   243
      case True
dcde836219bc add thy files
chunhan
parents:
diff changeset
   244
      from os ev have exp: "exists s (Proc p)" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   245
      then obtain r pt fr u srp where SP: "SProc (r, fr, pt, u) srp \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   246
        and sproc: "cp2sproc s p = Some (r, fr, pt, u)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   247
        using prem'[where obj = "Proc p"] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   248
        by (auto split:option.splits if_splits simp:obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   249
                  dest:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   250
      have "obj2sobj (e # s) (IPC i) \<in> all_sobjs" using ev vs_cons sproc os
dcde836219bc add thy files
chunhan
parents:
diff changeset
   251
        apply (auto simp:obj2sobj.simps ni_init_deled split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   252
        apply (rule ai_cipc) using SP sproc rc ev by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   253
      with True show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   254
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   255
      case False 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   256
      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   257
        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   258
                              split:option.splits t_role.splits )
dcde836219bc add thy files
chunhan
parents:
diff changeset
   259
      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   260
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   261
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   262
    case (ChangeRole p r')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   263
    assume ev: "e = ChangeRole p r'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   264
    show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
   265
    proof (cases "obj = Proc p")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   266
      case True
dcde836219bc add thy files
chunhan
parents:
diff changeset
   267
      from os ev have exp: "exists s (Proc p)" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   268
      then obtain r pt fr u srp where SP: "SProc (r, fr, pt, u) srp \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   269
        and sproc: "cp2sproc s p = Some (r, fr, pt, u)" and srproc: "source_proc s p = srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   270
        using prem'[where obj = "Proc p"] vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   271
        by (auto split:option.splits if_splits simp:obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   272
                  dest:current_proc_has_sproc)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   273
      have "obj2sobj (e # s) (Proc p) \<in> all_sobjs" using ev vs_cons sproc os 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   274
        apply (auto simp:obj2sobj.simps ni_init_deled split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   275
        apply (rule ap_crole) using SP sproc rc ev srproc by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   276
      with True show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   277
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   278
      case False 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   279
      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   280
        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   281
                              split:option.splits t_role.splits )
dcde836219bc add thy files
chunhan
parents:
diff changeset
   282
      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   283
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   284
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   285
    case (ReadFile p f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   286
    assume ev: "e = ReadFile p f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   287
    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   288
      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   289
                            split:option.splits t_role.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   290
    moreover have "exists s obj" using ev ex_cons
dcde836219bc add thy files
chunhan
parents:
diff changeset
   291
      by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   292
    ultimately show ?thesis using prem[where obj = obj] vs by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   293
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   294
    case (WriteFile p f)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   295
    assume ev: "e = WriteFile p f"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   296
    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   297
      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   298
                            split:option.splits t_role.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   299
    moreover have "exists s obj" using ev ex_cons
dcde836219bc add thy files
chunhan
parents:
diff changeset
   300
      by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   301
    ultimately show ?thesis using prem[where obj = obj] vs by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   302
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   303
    case (Send p i)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   304
    assume ev: "e = Send p i"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   305
    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   306
      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   307
                            split:option.splits t_role.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   308
    moreover have "exists s obj" using ev ex_cons
dcde836219bc add thy files
chunhan
parents:
diff changeset
   309
      by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   310
    ultimately show ?thesis using prem[where obj = obj] vs by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   311
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   312
    case (Recv p i)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   313
    assume ev: "e = Recv p i"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   314
    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   315
      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   316
                            split:option.splits t_role.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   317
    moreover have "exists s obj" using ev ex_cons
dcde836219bc add thy files
chunhan
parents:
diff changeset
   318
      by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   319
    ultimately show ?thesis using prem[where obj = obj] vs by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   320
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   321
    case (Kill p p')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   322
    assume ev: "e = Kill p p'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   323
    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   324
      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   325
                            split:option.splits t_role.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   326
    thus ?thesis using prem[where obj = obj] vs ex_cons ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   327
      by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   328
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   329
    case (DeleteFile p f')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   330
    assume ev: "e = DeleteFile p f'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   331
    have "obj2sobj (e#s) obj = obj2sobj s obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   332
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   333
      have "\<And> f. obj = File f ==> obj2sobj (e#s) (File f) = obj2sobj s (File f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   334
        using ev vs os ex_cons vs_cons
dcde836219bc add thy files
chunhan
parents:
diff changeset
   335
        by (auto simp:obj2sobj.simps etype_of_file_delete source_dir_simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   336
                split:option.splits t_role.splits if_splits
dcde836219bc add thy files
chunhan
parents:
diff changeset
   337
                dest!:current_file_has_etype' current_file_has_sd'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   338
                 dest:source_dir_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   339
      moreover have "\<forall> f. obj \<noteq> File f ==> obj2sobj (e#s) obj = obj2sobj s obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   340
        using ev vs_cons ex_cons os vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   341
        by (case_tac obj, auto simp:obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   342
      ultimately show ?thesis by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   343
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   344
    thus ?thesis using prem[where obj = obj] vs ex_cons ev 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   345
      by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   346
  next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   347
    case (DeleteIPC p i)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   348
    assume ev: "e = DeleteIPC p i"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   349
    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   350
      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   351
                            split:option.splits t_role.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   352
    thus ?thesis using prem[where obj = obj] vs ex_cons ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   353
      by (case_tac obj, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   354
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   355
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   356
dcde836219bc add thy files
chunhan
parents:
diff changeset
   357
declare obj2sobj.simps [simp add]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   358
dcde836219bc add thy files
chunhan
parents:
diff changeset
   359
lemma seeds_in_all_sobjs: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   360
  assumes seed: "obj \<in> seeds" shows "init_obj2sobj obj \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   361
proof (cases obj)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   362
  case (Proc p)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   363
  assume p0: "obj = Proc p" (*?*)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   364
  from seed p0 have pinit: "p \<in> init_processes" by (drule_tac seeds_in_init, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   365
  from pinit obtain r where "init_currentrole p = Some r" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   366
    using init_proc_has_role by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   367
  moreover from pinit obtain fr where "init_proc_forcedrole p = Some fr"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   368
    using init_proc_has_frole by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   369
  moreover from pinit obtain pt where "init_process_type p = Some pt"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   370
    using init_proc_has_type by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   371
  moreover from pinit obtain u where "init_owner p = Some u"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   372
    using init_proc_has_owner by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   373
  ultimately show ?thesis using p0 by (auto intro:ap_init)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   374
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   375
  case (File f) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   376
  assume p0: "obj = File f" (*?*)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   377
  from seed p0 have finit: "f \<in> init_files" by (drule_tac seeds_in_init, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   378
  then obtain t where "etype_aux init_file_type_aux f = Some t"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   379
    by (auto dest:init_file_has_etype)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   380
  with finit p0 show ?thesis by (auto intro:af_init)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   381
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   382
  case (IPC i)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   383
  assume p0: "obj = IPC i" (*?*)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   384
  from seed p0 have iinit: "i \<in> init_ipcs" by (drule_tac seeds_in_init, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   385
  then obtain t where "init_ipc_type i = Some t" using init_ipc_has_type
dcde836219bc add thy files
chunhan
parents:
diff changeset
   386
    by (auto simp:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   387
  with iinit p0 show ?thesis by (auto intro:ai_init)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   388
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   389
dcde836219bc add thy files
chunhan
parents:
diff changeset
   390
lemma tainted_s_in_all_sobjs:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   391
  "sobj \<in> tainted_s \<Longrightarrow> sobj \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   392
apply (erule tainted_s.induct)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   393
apply (erule seeds_in_all_sobjs)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   394
apply (auto intro:ap_crole ap_exec ap_chown ai_cipc af_cfd af_cfd' ap_clone)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   395
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   396
dcde836219bc add thy files
chunhan
parents:
diff changeset
   397
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
   398
dcde836219bc add thy files
chunhan
parents:
diff changeset
   399
context tainting_s_sound begin
dcde836219bc add thy files
chunhan
parents:
diff changeset
   400
dcde836219bc add thy files
chunhan
parents:
diff changeset
   401
(*** all_sobjs' equal with all_sobjs in the view of soundness ***)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   402
dcde836219bc add thy files
chunhan
parents:
diff changeset
   403
lemma all_sobjs'_eq1: "sobj \<in> all_sobjs \<Longrightarrow> sobj \<in> all_sobjs'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   404
apply (erule all_sobjs.induct)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   405
apply (auto intro:af'_init af'_cfd af'_cfd' ai'_init ai'_cipc ap'_init ap'_crole ap'_exec ap'_chown)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   406
by (simp add:clone_type_aux_def clone_type_unchange)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   407
dcde836219bc add thy files
chunhan
parents:
diff changeset
   408
lemma all_sobjs'_eq2: "sobj \<in> all_sobjs' \<Longrightarrow> sobj \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   409
apply (erule all_sobjs'.induct)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   410
by (auto intro:af_init af_cfd af_cfd' ai_init ai_cipc ap_init ap_crole ap_exec ap_chown)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   411
dcde836219bc add thy files
chunhan
parents:
diff changeset
   412
lemma all_sobjs'_eq: "(sobj \<in> all_sobjs) = (sobj \<in> all_sobjs')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   413
by (auto intro:iffI all_sobjs'_eq1 all_sobjs'_eq2)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   414
dcde836219bc add thy files
chunhan
parents:
diff changeset
   415
(************************ all_sobjs Elim Rules ********************)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   416
dcde836219bc add thy files
chunhan
parents:
diff changeset
   417
declare obj2sobj.simps [simp del]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   418
declare cp2sproc.simps [simp del]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   419
dcde836219bc add thy files
chunhan
parents:
diff changeset
   420
lemma all_sobjs_E0_aux[rule_format]:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   421
  "sobj \<in> all_sobjs' \<Longrightarrow> (\<forall> s' obj' sobj'. valid s' \<and> obj2sobj s' obj' = sobj' \<and> exists s' obj' \<and> sobj' \<noteq> Unknown \<and> no_del_event s' \<and> initp_intact s' \<longrightarrow> (\<exists> s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact (s @ s') \<and> obj2sobj (s @ s') obj = sobj \<and> exists (s @ s') obj))"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   422
proof (induct rule:all_sobjs'.induct)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   423
  case (af'_init f t) show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   424
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
   425
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   426
    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   427
      and nodels': "no_del_event s'"and intacts':"initp_intact s'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   428
      and exso': "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   429
    from nodels' af'_init(1) have exf: "f \<in> current_files s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   430
      by (drule_tac obj = "File f" in nodel_imp_exists, simp+)    
dcde836219bc add thy files
chunhan
parents:
diff changeset
   431
    have "obj2sobj s' (File f) = SFile (t, f) (Some f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   432
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   433
      have "obj2sobj [] (File f) = SFile (t, f) (Some f)"  using af'_init 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   434
        by (auto simp:etype_of_file_def source_dir_of_init' obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   435
                split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   436
      thus ?thesis using vss' exf nodels' af'_init(1) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   437
        by (drule_tac obj2sobj_file_remains_app', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   438
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   439
    thus "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   440
             exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact (s @ s') \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   441
             obj2sobj (s @ s') obj = SFile (t, f) (Some f) \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   442
      apply (rule_tac x = "[]" in exI, rule_tac x = "File f" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   443
      by (simp add:vss' sobjs' nodels' intacts' exf exso')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   444
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   445
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   446
  case (af'_cfd t sd srf r fr pt u srp t') 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   447
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   448
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
   449
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   450
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   451
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and exobj':"exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   452
    with af'_cfd(1,2) obtain sa pf where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   453
      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   454
      "exists (sa@s') obj'" and "initp_intact (sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   455
      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   456
      exfa: "pf \<in> current_files (sa@s')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   457
      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   458
      by (frule obj2sobj_file, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   459
    with af'_cfd(3,4) notUkn obtain sb p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   460
      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   461
      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   462
      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   463
      exsoab: "exists (sb@sa@s') obj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   464
      intactab: "initp_intact (sb@sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   465
      nodelab: "no_del_event (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   466
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   467
    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   468
      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   469
      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   470
    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   471
      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   472
    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   473
      and tau: "\<tau>=sb@sa@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   474
    
dcde836219bc add thy files
chunhan
parents:
diff changeset
   475
    have valid: "valid (e # \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   476
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   477
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   478
        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   479
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   480
        using ev tau af'_cfd(5,6,7) SPab SFab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   481
        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   482
                split:if_splits option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   483
      ultimately show ?thesis using vsab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   484
        by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   485
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   486
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   487
    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   488
        by (case_tac obj', simp+)   moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   489
    have "obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   490
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   491
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   492
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   493
        by (auto intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   494
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   495
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   496
        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   497
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   498
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   499
        by (auto simp:obj2sobj.simps cp2sproc_simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   500
             simp del:cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   501
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   502
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   503
    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
   504
      by (simp_all add:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   505
    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t', sd) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   506
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   507
      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   508
        using ev tau SFab SPab af'_cfd(5)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   509
        by (auto simp:obj2sobj.simps cp2sproc.simps etype_of_file_def
dcde836219bc add thy files
chunhan
parents:
diff changeset
   510
          split:option.splits if_splits  intro!:etype_aux_prop4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   511
      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   512
        using ev tau SFab SPab valid ncf_parent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   513
        by (auto simp:source_dir_simps obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   514
                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   515
      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   516
        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   517
        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   518
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   519
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   520
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact (s @ s') \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   521
      obj2sobj (s @ s') obj = SFile (t', sd) None \<and> exists (s @ s') obj "
dcde836219bc add thy files
chunhan
parents:
diff changeset
   522
      using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   523
      by (rule_tac x = "e#sb@sa" in exI, rule_tac x = "File (new_childf pf \<tau>)" in exI, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   524
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   525
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   526
  case (af'_cfd' t sd srf r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   527
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   528
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
   529
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   530
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   531
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and exobj':"exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   532
    with af'_cfd'(1,2) obtain sa pf where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   533
      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   534
      "exists (sa@s') obj'" and "initp_intact (sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   535
      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   536
      exfa: "pf \<in> current_files (sa@s')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   537
      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   538
      by (frule obj2sobj_file, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   539
    with af'_cfd'(3,4) notUkn obtain sb p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   540
      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   541
      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   542
      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   543
      exsoab: "exists (sb@sa@s') obj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   544
      intactab: "initp_intact (sb@sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   545
      nodelab: "no_del_event (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   546
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   547
    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   548
      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   549
      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   550
    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   551
      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   552
    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   553
      and tau: "\<tau>=sb@sa@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   554
    
dcde836219bc add thy files
chunhan
parents:
diff changeset
   555
    have valid: "valid (e # \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   556
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   557
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   558
        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   559
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   560
        using ev tau af'_cfd'(5,6) SPab SFab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   561
        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   562
                split:if_splits option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   563
      ultimately show ?thesis using vsab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   564
        by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   565
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   566
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   567
    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   568
        by (case_tac obj', simp+)   moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   569
    have "obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   570
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   571
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   572
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   573
        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   574
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   575
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   576
        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   577
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   578
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   579
        by (auto simp:obj2sobj.simps cp2sproc_simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   580
             simp del:cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   581
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   582
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   583
    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
   584
      by (simp add:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   585
    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t, sd) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   586
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   587
      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   588
      proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   589
        have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = etype_of_file \<tau> pf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   590
          using ev tau SPab af'_cfd'(5) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   591
          by (auto simp:obj2sobj.simps ncf_parent etype_of_file_def cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   592
            split:option.splits   intro!:etype_aux_prop3)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   593
        thus ?thesis using SFab tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   594
          by (auto simp:obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   595
      qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   596
      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   597
        using ev tau SFab SPab valid ncf_parent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   598
        by (auto simp:source_dir_simps obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   599
                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   600
      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   601
        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   602
        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   603
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   604
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   605
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact (s @ s') \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   606
      obj2sobj (s @ s') obj = SFile (t, sd) None \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   607
      using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   608
      by (rule_tac x = "e#sb@sa" in exI, rule_tac x = "File (new_childf pf \<tau>)" in exI, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   609
  qed 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   610
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   611
  case (ai'_init i t) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   612
  hence initi: "i \<in> init_ipcs" using init_ipc_has_type
dcde836219bc add thy files
chunhan
parents:
diff changeset
   613
    by (simp add:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   614
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   615
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
   616
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   617
    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   618
      and nodels': "no_del_event s'"and intacts':"initp_intact s'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   619
      and exso': "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   620
    from nodels' initi have exi: "i \<in> current_ipcs s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   621
      by (drule_tac obj = "IPC i" in nodel_imp_exists, simp+)    
dcde836219bc add thy files
chunhan
parents:
diff changeset
   622
    have "obj2sobj s' (IPC i) = SIPC t (Some i)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   623
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   624
      have "obj2sobj [] (IPC i) = SIPC t (Some i)"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   625
        using ai'_init initi by (auto simp:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   626
      thus ?thesis using vss' exi nodels' initi 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   627
        by (drule_tac obj2sobj_ipc_remains'', auto simp:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   628
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   629
    thus "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   630
             exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact (s @ s') \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   631
             obj2sobj (s @ s') obj = SIPC t (Some i) \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   632
      apply (rule_tac x = "[]" in exI, rule_tac x = "IPC i" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   633
      by (simp add:vss' sobjs' nodels' exi exso' intacts' del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   634
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   635
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   636
  case (ai'_cipc r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   637
  show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   638
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
   639
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   640
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   641
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and exobj':"exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   642
    with ai'_cipc(1,2) notUkn obtain s p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   643
      SPab: "obj2sobj (s@s') (Proc p) = SProc (r,fr,pt,u) srp" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   644
      expab: "exists (s@s') (Proc p)" and vsab: "valid (s@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   645
      soab: "obj2sobj (s@s') obj' = sobj'" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   646
      exsoab: "exists (s@s') obj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   647
      intactab: "initp_intact (s@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   648
      nodelab: "no_del_event (s@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   649
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   650
    obtain e \<tau> where ev: "e = CreateIPC p (new_ipc \<tau>)" and tau: "\<tau>=s@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   651
    
dcde836219bc add thy files
chunhan
parents:
diff changeset
   652
    have valid: "valid (e # \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   653
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   654
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   655
        using ev tau expab by (simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   656
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   657
        using ev tau ai'_cipc(3) SPab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   658
        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   659
      ultimately show ?thesis using vsab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   660
        by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   661
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   662
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   663
    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   664
        by (case_tac obj', simp+)   moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   665
    have "obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   666
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   667
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   668
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   669
        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   670
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   671
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   672
        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   673
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   674
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   675
        by (auto simp:obj2sobj.simps cp2sproc_simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   676
             simp del:cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   677
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   678
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   679
    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
   680
      by (simp add:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   681
    have "obj2sobj (e#\<tau>) (IPC (new_ipc \<tau>)) = SIPC (default_ipc_create_type r) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   682
      using ev tau SPab nodel 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   683
        nodel_imp_exists[where obj = "IPC (new_ipc \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   684
      by (auto simp:obj2sobj.simps ni_notin_curi cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   685
              split:option.splits  dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   686
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   687
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact (s @ s') \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   688
      obj2sobj (s @ s') obj = SIPC (default_ipc_create_type r) None \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   689
      using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
   690
      by (rule_tac x = "e#s" in exI, rule_tac x = "IPC (new_ipc \<tau>)" in exI, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   691
  qed 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   692
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   693
  case (ap'_init p r fr t u)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   694
  hence initp: "p \<in> init_processes" using init_proc_has_role
dcde836219bc add thy files
chunhan
parents:
diff changeset
   695
    by (simp add:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   696
  show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   697
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
   698
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   699
    assume VSs': "valid s'" and SOs': "obj2sobj s' obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   700
      and Nodels': "no_del_event s'"and Intacts':"initp_intact s'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   701
      and exso': "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   702
    from Nodels' initp have exp: "p \<in> current_procs s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   703
      apply (drule_tac obj = "Proc p" in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   704
      by (drule not_deleted_imp_exists, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   705
    with Intacts' initp ap'_init have "obj2sobj s' (Proc p) = SProc (r, fr, t, u) (Some p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   706
      by (auto simp:initp_intact_def split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   707
    thus "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   708
             exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact (s @ s') \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   709
             obj2sobj (s @ s') obj = SProc (r, fr, t, u) (Some p) \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   710
      apply (rule_tac x = "[]" in exI, rule_tac x = "Proc p" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   711
      by (simp add:VSs' SOs' Nodels' exp exso' initp intact_imp_butp Intacts'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   712
               del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   713
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   714
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   715
  case (ap'_crole r fr t u srp r')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   716
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   717
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
   718
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   719
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   720
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   721
    with ap'_crole(1,2) obtain s p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   722
      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   723
      and nodelab: "no_del_event (s@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   724
      and intactab: "initp_intact (s@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   725
      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   726
      and exp:"exists (s@s') (Proc p)" and exobj'ab:"exists (s@s') obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   727
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   728
    obtain e \<tau> where ev: "e = ChangeRole (new_proc (s@s')) r'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   729
      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   730
    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   731
dcde836219bc add thy files
chunhan
parents:
diff changeset
   732
    have valid: "valid (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   733
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   734
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   735
        using ev tau exp by (simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   736
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   737
        using ev tau ap'_crole(3) SPab 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   738
        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   739
      ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   740
        by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   741
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   742
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   743
    have "initp_intact (e#\<tau>)" using tau ev intactab valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   744
      by (simp add:initp_intact_I_crole)   moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   745
    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   746
      by (case_tac obj', simp+)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   747
    have "obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   748
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   749
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   750
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   751
        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   752
        by (auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   753
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   754
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   755
        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   756
        by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   757
      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   758
        apply (case_tac "p' = new_proc (s @ s')") 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   759
        using vs_tau exobj'ab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   760
        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   761
        using tau ev SOab' valid notUkn vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   762
        by (auto simp:obj2sobj.simps cp2sproc_simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   763
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   764
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   765
    have "obj2sobj (e#\<tau>) (Proc (new_proc (s@s'))) = SProc (r', fr, t, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   766
      using SPab tau vs_tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   767
      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   768
              split:option.splits)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   769
    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   770
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   771
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact (s @ s') \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   772
      obj2sobj (s @ s') obj = SProc (r', fr, t, u) srp \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   773
      using ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   774
      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   775
      by (rule_tac x = "Proc (new_proc (s@s'))" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   776
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   777
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   778
  case (ap'_chown r fr t u srp u' nr)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   779
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   780
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
   781
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   782
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   783
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   784
    with ap'_chown(1,2) obtain s p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   785
      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   786
      and nodelab: "no_del_event (s@s')" and intactab: "initp_intact (s@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   787
      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   788
      and exp:"exists (s@s') (Proc p)" and exobj'ab:"exists (s@s') obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   789
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   790
    obtain e \<tau> where ev: "e = ChangeOwner (new_proc (s@s')) u'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   791
      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   792
    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   793
dcde836219bc add thy files
chunhan
parents:
diff changeset
   794
    have valid: "valid (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   795
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   796
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   797
        using ev tau exp ap'_chown(3) by (simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   798
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   799
        using ev tau ap'_chown(5) SPab 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   800
        by (auto simp:cp2sproc.simps obj2sobj.simps pct_def clone_type_unchange
dcde836219bc add thy files
chunhan
parents:
diff changeset
   801
                split:option.splits t_rc_proc_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   802
          (* here is another place of no_limit of clone event assumption *)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   803
      ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   804
        by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   805
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   806
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   807
    have "initp_intact (e#\<tau>)" using intactab tau ev valid 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   808
      by (simp add:initp_intact_I_chown)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   809
    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   810
      by (case_tac obj', simp+)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   811
    have "obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   812
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   813
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   814
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   815
        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   816
        by (auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   817
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   818
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   819
        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   820
        by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   821
      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   822
        apply (case_tac "p' = new_proc (s @ s')") 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   823
        using vs_tau exobj'ab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   824
        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   825
        using tau ev SOab' valid notUkn vs_tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   826
        by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   827
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   828
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   829
    have "obj2sobj (e#\<tau>) (Proc (new_proc (s@s'))) = 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   830
          SProc (nr,fr,chown_type_aux r nr t,u') srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   831
      using SPab tau vs_tau ev valid ap'_chown(4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   832
      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   833
              split:option.splits)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   834
    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp   moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   835
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   836
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact (s @ s') \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   837
      obj2sobj (s @ s') obj = SProc (nr,fr,chown_type_aux r nr t,u') srp \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   838
      exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   839
      using ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   840
      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   841
      by (rule_tac x = "Proc (new_proc (s@s'))" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   842
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   843
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   844
  case (ap'_exec r fr pt u sp t sd sf r' fr')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   845
  show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   846
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
   847
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   848
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   849
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   850
    with ap'_exec(3,4) obtain sa f where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   851
      SFa: "obj2sobj (sa @ s') (File f) = SFile (t, sd) sf" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   852
      Exfa: "exists (sa @ s') (File f)" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   853
      butsa: "initp_intact (sa @ s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   854
      othersa:"valid (sa @ s') \<and> obj2sobj (sa @ s') obj' = sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   855
               exists (sa @s') obj' \<and> no_del_event (sa @ s')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   856
      by (blast dest:obj2sobj_file intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   857
    with ap'_exec(1,2) notUkn obtain sb p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   858
      VSab: "valid (sb@sa@s')" and SOab': "obj2sobj (sb@sa@s') obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   859
      and nodelab: "no_del_event (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   860
      and intactab: "initp_intact (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   861
      and SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r, fr, pt, u) sp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   862
      and exp:"exists (sb@sa@s') (Proc p)" and exobj'ab:"exists (sb@sa@s') obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   863
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   864
    obtain e \<tau> where ev: "e = Execute (new_proc (sb@sa@s')) f" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   865
      and tau: "\<tau> = Clone p (new_proc (sb@sa@s'))#sb@sa@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   866
    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   867
    from Exfa nodelab have exf:"f \<in> current_files (sb@sa@s')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   868
      apply (drule_tac obj = "File f" in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   869
      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   870
    from SFa VSab Exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File f) = SFile (t,sd) sf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   871
      by (rule_tac s = "sa@s'" in obj2sobj_file_remains'', simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   872
dcde836219bc add thy files
chunhan
parents:
diff changeset
   873
    have valid: "valid (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   874
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   875
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   876
        using ev tau exp by (simp add:exf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   877
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   878
        using ev tau ap'_exec(5) SPab SFab
dcde836219bc add thy files
chunhan
parents:
diff changeset
   879
        by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   880
                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   881
      ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   882
        by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   883
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   884
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   885
    have "initp_intact (e#\<tau>)" using tau ev intactab valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   886
      by (simp add:initp_intact_I_exec)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   887
    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   888
      by (case_tac obj', simp+) moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
   889
    have "obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   890
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   891
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   892
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   893
        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   894
        by (auto simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   895
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   896
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   897
        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   898
        by (auto simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   899
      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   900
        apply (case_tac "p' = new_proc (sb @ sa @ s')") 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   901
        using vs_tau exobj'ab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   902
        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   903
        using tau ev SOab' valid notUkn vs_tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   904
        by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   905
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   906
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   907
    have "obj2sobj (e#\<tau>) (Proc (new_proc (sb @ sa @ s'))) = 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   908
          SProc (r',fr',exec_type_aux r pt, u) sp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   909
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   910
      have "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) sp" using SPab tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   911
        by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   912
      hence "obj2sobj \<tau> (Proc (new_proc (sb@sa@s'))) = SProc (r,fr,pt,u) sp" using tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   913
      by (auto simp:obj2sobj.simps cp2sproc.simps pct_def clone_type_unchange 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   914
              split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   915
      moreover have "source_dir \<tau> f = Some sd" using vs_tau SFab tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   916
        by (auto simp:source_dir_simps obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   917
      ultimately show ?thesis using valid ev ap'_exec(6,7) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   918
        by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   919
    qed 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   920
    ultimately 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   921
    show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   922
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact (s @ s') \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   923
      obj2sobj (s @ s') obj = SProc (r', fr', exec_type_aux r pt, u) sp \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   924
      exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   925
      using ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
   926
      apply (rule_tac x = "e#Clone p (new_proc (sb @ sa @ s')) #sb@sa" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   927
      by (rule_tac x = "Proc (new_proc (sb @ sa @ s'))" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   928
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   929
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   930
dcde836219bc add thy files
chunhan
parents:
diff changeset
   931
(* this is for ts2t createfile case ... *)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   932
lemma all_sobjs_E0:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   933
  "\<lbrakk>sobj \<in> all_sobjs'; valid s'; obj2sobj s' obj' = sobj'; exists s' obj'; sobj' \<noteq> Unknown; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   934
    no_del_event s'; initp_intact s'\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   935
   \<Longrightarrow> \<exists> s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> exists (s@s') obj \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   936
                no_del_event (s @ s') \<and> initp_intact (s @ s') \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   937
                obj2sobj (s @ s') obj = sobj \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   938
by (drule all_sobjs_E0_aux, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   939
dcde836219bc add thy files
chunhan
parents:
diff changeset
   940
lemma all_sobjs_E1_aux[rule_format]:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   941
  "sobj \<in> all_sobjs' \<Longrightarrow> (\<forall> s' obj' sobj'. valid s' \<and> obj2sobj s' obj' = sobj' \<and> exists s' obj' \<and> sobj' \<noteq> Unknown \<and> no_del_event s' \<and> initp_intact_but s' sobj' \<longrightarrow> (\<exists> s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj' \<and> obj2sobj (s @ s') obj = sobj \<and> exists (s @ s') obj))"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   942
proof (induct rule:all_sobjs'.induct)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   943
  case (af'_init f t) show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
   944
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
   945
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   946
    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   947
      and nodels': "no_del_event s'"and intacts':"initp_intact_but s' sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   948
      and exso': "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   949
    from nodels' af'_init(1) have exf: "f \<in> current_files s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   950
      by (drule_tac obj = "File f" in nodel_imp_exists, simp+)    
dcde836219bc add thy files
chunhan
parents:
diff changeset
   951
    have "obj2sobj s' (File f) = SFile (t, f) (Some f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   952
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   953
      have "obj2sobj [] (File f) = SFile (t, f) (Some f)"  using af'_init 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   954
        by (auto simp:etype_of_file_def source_dir_of_init' obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
   955
                split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   956
      thus ?thesis using vss' exf nodels' af'_init(1) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   957
        by (drule_tac obj2sobj_file_remains_app', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   958
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   959
    thus "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   960
             exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   961
             obj2sobj (s @ s') obj = SFile (t, f) (Some f) \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   962
      apply (rule_tac x = "[]" in exI, rule_tac x = "File f" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   963
      by (simp add:vss' sobjs' nodels' intacts' exf exso')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   964
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   965
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
   966
  case (af'_cfd t sd srf r fr pt u srp t') 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   967
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   968
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
   969
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
   970
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   971
      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   972
      and exobj':"exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   973
    with af'_cfd(1,2) obtain sa pf where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   974
      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   975
      "exists (sa@s') obj'" and "initp_intact_but (sa@s') sobj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   976
      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   977
      exfa: "pf \<in> current_files (sa@s')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   978
      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   979
      by (frule obj2sobj_file, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   980
    with af'_cfd(3,4) notUkn obtain sb p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   981
      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   982
      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   983
      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   984
      exsoab: "exists (sb@sa@s') obj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   985
      intactab: "initp_intact_but (sb@sa@s') sobj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   986
      nodelab: "no_del_event (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   987
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   988
    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   989
      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   990
      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   991
    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   992
      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   993
    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   994
      and tau: "\<tau>=sb@sa@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
   995
    
dcde836219bc add thy files
chunhan
parents:
diff changeset
   996
    have valid: "valid (e # \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   997
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
   998
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   999
        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1000
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1001
        using ev tau af'_cfd(5,6,7) SPab SFab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1002
        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1003
                split:if_splits option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1004
      ultimately show ?thesis using vsab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1005
        by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1006
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1007
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1008
    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1009
        by (case_tac obj', simp+)   moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1010
    have "obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1011
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1012
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1013
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1014
        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1015
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1016
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1017
        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1018
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1019
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1020
        by (auto simp:obj2sobj.simps cp2sproc_simps' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1021
             simp del:cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1022
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1023
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1024
    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t', sd) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1025
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1026
      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1027
        using ev tau SFab SPab af'_cfd(5)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1028
        by (auto simp:obj2sobj.simps cp2sproc.simps etype_of_file_def
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1029
          split:option.splits if_splits  intro!:etype_aux_prop4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1030
      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1031
        using ev tau SFab SPab valid ncf_parent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1032
        by (auto simp:source_dir_simps obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1033
                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1034
      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1035
        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1036
        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1037
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1038
    have "initp_intact_but (e#\<tau>) sobj'" using intactab tau ev valid nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1039
        apply (case_tac sobj', case_tac option)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1040
        by (simp_all add:initp_intact_butp_I_others initp_intact_I_others)  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1041
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1042
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1043
      obj2sobj (s @ s') obj = SFile (t', sd) None \<and> exists (s @ s') obj "
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1044
      using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1045
      apply (rule_tac x = "e#sb@sa" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1046
      by (rule_tac x = "File (new_childf pf \<tau>)" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1047
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1048
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1049
  case (af'_cfd' t sd srf r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1050
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1051
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1052
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1053
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1054
      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1055
      and exobj':"exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1056
    with af'_cfd'(1,2) obtain sa pf where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1057
      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1058
      "exists (sa@s') obj'" and "initp_intact_but (sa@s') sobj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1059
      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1060
      exfa: "pf \<in> current_files (sa@s')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1061
      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1062
      by (frule obj2sobj_file, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1063
    with af'_cfd'(3,4) notUkn obtain sb p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1064
      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1065
      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1066
      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1067
      exsoab: "exists (sb@sa@s') obj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1068
      intactab: "initp_intact_but (sb@sa@s') sobj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1069
      nodelab: "no_del_event (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1070
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1071
    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1072
      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1073
      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1074
    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1075
      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1076
    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1077
      and tau: "\<tau>=sb@sa@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1078
    
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1079
    have valid: "valid (e # \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1080
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1081
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1082
        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1083
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1084
        using ev tau af'_cfd'(5,6) SPab SFab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1085
        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1086
                split:if_splits option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1087
      ultimately show ?thesis using vsab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1088
        by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1089
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1090
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1091
    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1092
        by (case_tac obj', simp+)   moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1093
    have "obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1094
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1095
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1096
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1097
        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1098
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1099
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1100
        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1101
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1102
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1103
        by (auto simp:obj2sobj.simps cp2sproc_simps' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1104
             simp del:cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1105
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1106
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1107
    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t, sd) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1108
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1109
      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1110
      proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1111
        have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = etype_of_file \<tau> pf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1112
          using ev tau SPab af'_cfd'(5) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1113
          by (auto simp:obj2sobj.simps cp2sproc.simps ncf_parent etype_of_file_def
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1114
            split:option.splits   intro!:etype_aux_prop3)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1115
        thus ?thesis using SFab tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1116
          by (auto simp:obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1117
      qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1118
      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1119
        using ev tau SFab SPab valid ncf_parent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1120
        by (auto simp:source_dir_simps obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1121
                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1122
      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1123
        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1124
        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1125
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1126
    have "initp_intact_but (e#\<tau>) sobj'" using intactab tau ev valid nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1127
        apply (case_tac sobj', case_tac option)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1128
        by (simp_all add:initp_intact_butp_I_others initp_intact_I_others)  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1129
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1130
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1131
      obj2sobj (s @ s') obj = SFile (t, sd) None \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1132
      using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1133
      apply (rule_tac x = "e#sb@sa" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1134
      by (rule_tac x = "File (new_childf pf \<tau>)" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1135
  qed 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1136
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1137
  case (ai'_init i t) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1138
  hence initi: "i \<in> init_ipcs" using init_ipc_has_type
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1139
    by (simp add:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1140
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1141
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1142
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1143
    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1144
      and nodels': "no_del_event s'"and intacts':"initp_intact_but s' sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1145
      and exso': "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1146
    from nodels' initi have exi: "i \<in> current_ipcs s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1147
      by (drule_tac obj = "IPC i" in nodel_imp_exists, simp+)    
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1148
    have "obj2sobj s' (IPC i) = SIPC t (Some i)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1149
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1150
      have "obj2sobj [] (IPC i) = SIPC t (Some i)"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1151
        using ai'_init initi by (auto simp:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1152
      thus ?thesis using vss' exi nodels' initi 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1153
        by (drule_tac obj2sobj_ipc_remains'', auto simp:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1154
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1155
    thus "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1156
             exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1157
             obj2sobj (s @ s') obj = SIPC t (Some i) \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1158
      apply (rule_tac x = "[]" in exI, rule_tac x = "IPC i" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1159
      by (simp add:vss' sobjs' nodels' exi exso' intacts' del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1160
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1161
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1162
  case (ai'_cipc r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1163
  show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1164
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1165
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1166
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1167
      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1168
      and exobj':"exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1169
    with ai'_cipc(1,2) notUkn obtain s p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1170
      SPab: "obj2sobj (s@s') (Proc p) = SProc (r,fr,pt,u) srp" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1171
      expab: "exists (s@s') (Proc p)" and vsab: "valid (s@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1172
      soab: "obj2sobj (s@s') obj' = sobj'" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1173
      exsoab: "exists (s@s') obj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1174
      intactab: "initp_intact_but (s@s') sobj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1175
      nodelab: "no_del_event (s@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1176
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1177
    obtain e \<tau> where ev: "e = CreateIPC p (new_ipc \<tau>)" and tau: "\<tau>=s@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1178
    
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1179
    have valid: "valid (e # \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1180
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1181
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1182
        using ev tau expab by (simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1183
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1184
        using ev tau ai'_cipc(3) SPab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1185
        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1186
      ultimately show ?thesis using vsab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1187
        by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1188
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1189
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1190
    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1191
        by (case_tac obj', simp+)   moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1192
    have "obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1193
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1194
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1195
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1196
        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1197
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1198
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1199
        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1200
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1201
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1202
        by (auto simp:obj2sobj.simps cp2sproc_simps' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1203
             simp del:cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1204
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1205
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1206
    have "obj2sobj (e#\<tau>) (IPC (new_ipc \<tau>)) = SIPC (default_ipc_create_type r) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1207
      using ev tau SPab nodel 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1208
        nodel_imp_exists[where obj = "IPC (new_ipc \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1209
      by (auto simp:obj2sobj.simps ni_notin_curi cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1210
              split:option.splits dest:no_del_event_cons_D)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1211
    have "initp_intact_but (e#\<tau>) sobj'" using intactab tau ev valid nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1212
        apply (case_tac sobj', case_tac option)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1213
        by (simp_all add:initp_intact_butp_I_others initp_intact_I_others)  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1214
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1215
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1216
      obj2sobj (s @ s') obj = SIPC (default_ipc_create_type r) None \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1217
      using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1218
      by (rule_tac x = "e#s" in exI, rule_tac x = "IPC (new_ipc \<tau>)" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1219
  qed 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1220
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1221
  case (ap'_init p r fr t u)  (* the big difference from other elims is in this case *)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1222
  hence initp: "p \<in> init_processes" using init_proc_has_role
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1223
    by (simp add:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1224
  show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1225
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1226
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1227
    assume VSs': "valid s'" and SOs': "obj2sobj s' obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1228
      and Nodels': "no_del_event s'"and Intacts':"initp_intact_but s' sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1229
      and exso': "exists s' obj'" and notUkn: "sobj' \<noteq> Unknown"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1230
    from Nodels' initp have exp: "p \<in> current_procs s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1231
      by (drule_tac obj = "Proc p" in nodel_imp_exists, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1232
    have "\<exists> p'. obj2sobj s' (Proc p') = SProc (r,fr,t,u) (Some p) \<and> p' \<in> current_procs s'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1233
    proof (cases sobj')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1234
      case (SProc sp srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1235
      show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1236
      proof (cases srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1237
        case None
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1238
        with SProc Intacts' have "initp_intact s'" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1239
        thus ?thesis using initp ap'_init
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1240
          apply (rule_tac x = p in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1241
          by (auto simp:initp_intact_def exp split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1242
      next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1243
        case (Some p')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1244
        show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1245
        proof (cases "p' = p")
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1246
          case True
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1247
          with Intacts' SProc Some have "initp_alter s' p" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1248
            by (simp add:initp_intact_butp_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1249
          then obtain pa where "pa \<in> current_procs s'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1250
            and "obj2sobj s' (Proc pa) = init_obj2sobj (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1251
            by (auto simp only:initp_alter_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1252
          thus ?thesis using ap'_init initp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1253
            by (rule_tac x = pa in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1254
        next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1255
          case False
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1256
          with Intacts' SProc Some initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1257
          have "obj2sobj s' (Proc p) = init_obj2sobj (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1258
            apply (simp only:initp_intact_butp_def initp_intact_but.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1259
            by (erule conjE, erule_tac x = p in allE, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1260
          thus ?thesis using ap'_init exp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1261
            by (rule_tac x = p in exI, auto split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1262
        qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1263
      qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1264
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1265
      case (SFile sf srf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1266
      thus ?thesis using ap'_init exp Intacts' initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1267
        by (rule_tac x = p in exI, auto split:option.splits simp:initp_intact_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1268
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1269
      case (SIPC si sri)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1270
      thus ?thesis using ap'_init exp Intacts' initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1271
        by (rule_tac x = p in exI, auto split:option.splits simp:initp_intact_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1272
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1273
      case Unknown
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1274
      thus ?thesis using notUkn by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1275
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1276
    then obtain p' where "obj2sobj s' (Proc p') = SProc (r, fr, t, u) (Some p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1277
      and "p' \<in> current_procs s'" by blast
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1278
    thus "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1279
             exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1280
             obj2sobj (s @ s') obj = SProc (r, fr, t, u) (Some p) \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1281
      apply (rule_tac x = "[]" in exI, rule_tac x = "Proc p'" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1282
      by (simp add:VSs' SOs' Nodels' exp exso' Intacts') 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1283
  qed 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1284
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1285
  case (ap'_crole r fr t u srp r')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1286
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1287
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1288
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1289
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1290
      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1291
    with ap'_crole(1,2) obtain s p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1292
      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1293
      and nodelab: "no_del_event (s@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1294
      and intactab: "initp_intact_but (s@s') sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1295
      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1296
      and exp:"exists (s@s') (Proc p)" and exobj'ab:"exists (s@s') obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1297
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1298
    obtain e \<tau> where ev: "e = ChangeRole (new_proc (s@s')) r'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1299
      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1300
    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1301
    have np_not_initp: "new_proc (s@s') \<notin> init_processes" using nodelab 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1302
      apply (rule_tac notI, drule_tac obj = "Proc (new_proc (s@s'))" in nodel_imp_exists)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1303
      by (auto simp:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1304
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1305
    have valid: "valid (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1306
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1307
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1308
        using ev tau exp by (simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1309
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1310
        using ev tau ap'_crole(3) SPab 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1311
        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1312
      ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1313
        by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1314
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1315
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1316
    have "initp_intact_but (e#\<tau>) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1317
    proof (cases sobj')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1318
      case (SProc sp srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1319
      show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1320
      proof (cases srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1321
        case (Some p')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1322
        with SOab' exobj'ab VSab intactab notUkn SProc
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1323
        have butp: "p' \<in> init_processes \<and> initp_intact_butp (s@s') p'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1324
          by (case_tac obj', auto intro:source_proc_in_init simp:obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1325
                                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1326
        then obtain p'' where exp': "p'' \<in> current_procs (s@s')" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1327
          SP': "obj2sobj (s@s') (Proc p'') = init_obj2sobj (Proc p')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1328
          by (auto simp:initp_alter_def initp_intact_butp_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1329
        hence "initp_alter (e#\<tau>) p'" using tau ev notUkn nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1330
          apply (simp add:initp_alter_def del:init_obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1331
          apply (rule_tac x = p'' in exI, rule conjI, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1332
          apply (subgoal_tac "p'' \<noteq> new_proc (s @s')")
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1333
          apply (auto simp:obj2sobj.simps cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1334
                  simp del:init_obj2sobj.simps  split:option.splits)[1]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1335
          by (rule notI, simp add:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1336
        thus ?thesis using SProc Some intactab tau ev valid vs_tau np_not_initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1337
          apply (simp add:initp_intact_butp_def del:init_obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1338
          apply (rule impI|rule allI|rule conjI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1339
          apply (erule_tac x = pa in allE)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1340
          by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:init_obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1341
                  split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1342
      next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1343
        case None
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1344
        with intactab SProc 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1345
        have "initp_intact (s@s')" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1346
        hence "initp_intact (e#\<tau>)" using tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1347
          by (simp add:initp_intact_I_crole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1348
        thus ?thesis using SProc None by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1349
      qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1350
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1351
      case (SFile sf srf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1352
      with intactab SFile
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1353
      have "initp_intact (s@s')" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1354
      hence "initp_intact (e#\<tau>)" using tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1355
        by (simp add:initp_intact_I_crole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1356
      thus ?thesis using SFile by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1357
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1358
      case (SIPC si sri)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1359
      with intactab SIPC
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1360
      have "initp_intact (s@s')" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1361
      hence "initp_intact (e#\<tau>)" using tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1362
        by (simp add:initp_intact_I_crole)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1363
      thus ?thesis using SIPC by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1364
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1365
      case Unknown 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1366
      with notUkn show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1367
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1368
    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1369
      by (case_tac obj', simp+)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1370
    have "obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1371
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1372
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1373
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1374
        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1375
        by (auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1376
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1377
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1378
        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1379
        by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1380
      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1381
        apply (case_tac "p' = new_proc (s @ s')") 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1382
        using vs_tau exobj'ab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1383
        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1384
        using tau ev SOab' valid notUkn vs_tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1385
        by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1386
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1387
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1388
    have "obj2sobj (e#\<tau>) (Proc (new_proc (s@s'))) = SProc (r', fr, t, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1389
      using SPab tau vs_tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1390
      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1391
              split:option.splits)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1392
    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1393
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1394
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1395
      obj2sobj (s @ s') obj = SProc (r', fr, t, u) srp \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1396
      using ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1397
      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1398
      by (rule_tac x = "Proc (new_proc (s@s'))" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1399
  qed 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1400
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1401
  case (ap'_chown r fr t u srp u' nr)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1402
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1403
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1404
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1405
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1406
      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1407
    with ap'_chown(1,2) obtain s p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1408
      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1409
      and nodelab: "no_del_event (s@s')" and intactab: "initp_intact_but (s@s') sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1410
      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1411
      and exp:"exists (s@s') (Proc p)" and exobj'ab:"exists (s@s') obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1412
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1413
    obtain e \<tau> where ev: "e = ChangeOwner (new_proc (s@s')) u'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1414
      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1415
    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1416
    have np_not_initp: "new_proc (s@s') \<notin> init_processes" using nodelab 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1417
      apply (rule_tac notI, drule_tac obj = "Proc (new_proc (s@s'))" in nodel_imp_exists)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1418
      by (auto simp:np_notin_curp)
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 ap'_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 ap'_chown(5) SPab 
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 nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1433
    have "initp_intact_but (e#\<tau>) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1434
    proof (cases sobj')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1435
      case (SProc sp srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1436
      show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1437
      proof (cases srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1438
        case (Some p')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1439
        with SOab' exobj'ab VSab intactab notUkn SProc
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1440
        have butp: "p' \<in> init_processes \<and> initp_intact_butp (s@s') p'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1441
          by (case_tac obj', auto intro:source_proc_in_init simp:obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1442
                                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1443
        then obtain p'' where exp': "p'' \<in> current_procs (s@s')" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1444
          SP': "obj2sobj (s@s') (Proc p'') = init_obj2sobj (Proc p')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1445
          by (auto simp:initp_alter_def initp_intact_butp_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1446
        hence "initp_alter (e#\<tau>) p'" using tau ev notUkn nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1447
          apply (simp add:initp_alter_def del:init_obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1448
          apply (rule_tac x = p'' in exI, rule conjI, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1449
          apply (subgoal_tac "p'' \<noteq> new_proc (s @s')")
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1450
          apply (auto simp:obj2sobj.simps cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1451
                  simp del:init_obj2sobj.simps  split:option.splits)[1]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1452
          by (rule notI, simp add:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1453
        thus ?thesis using SProc Some intactab tau ev valid vs_tau np_not_initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1454
          apply (simp add:initp_intact_butp_def del:init_obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1455
          apply (rule impI|rule allI|rule conjI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1456
          apply (erule_tac x = pa in allE)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1457
          by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:init_obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1458
                  split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1459
      next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1460
        case None
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1461
        with intactab SProc 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1462
        have "initp_intact (s@s')" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1463
        hence "initp_intact (e#\<tau>)" using tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1464
          by (simp add:initp_intact_I_chown)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1465
        thus ?thesis using SProc None by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1466
      qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1467
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1468
      case (SFile sf srf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1469
      with intactab SFile
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1470
      have "initp_intact (s@s')" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1471
      hence "initp_intact (e#\<tau>)" using tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1472
        by (simp add:initp_intact_I_chown)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1473
      thus ?thesis using SFile by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1474
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1475
      case (SIPC si sri)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1476
      with intactab SIPC
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1477
      have "initp_intact (s@s')" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1478
      hence "initp_intact (e#\<tau>)" using tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1479
        by (simp add:initp_intact_I_chown)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1480
      thus ?thesis using SIPC by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1481
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1482
      case Unknown 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1483
      with notUkn show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1484
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1485
    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1486
      by (case_tac obj', simp+)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1487
    have "obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1488
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1489
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1490
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1491
        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1492
        by (auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1493
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1494
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1495
        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1496
        by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1497
      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1498
        apply (case_tac "p' = new_proc (s @ s')") 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1499
        using vs_tau exobj'ab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1500
        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1501
        using tau ev SOab' valid notUkn vs_tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1502
        by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1503
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1504
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1505
    have "obj2sobj (e#\<tau>) (Proc (new_proc (s@s'))) = 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1506
          SProc (nr,fr,chown_type_aux r nr t,u') srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1507
      using SPab tau vs_tau ev valid ap'_chown(4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1508
      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1509
              split:option.splits)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1510
    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp   moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1511
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1512
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1513
      obj2sobj (s @ s') obj = SProc (nr,fr,chown_type_aux r nr t,u') srp \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1514
      exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1515
      using ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1516
      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1517
      by (rule_tac x = "Proc (new_proc (s@s'))" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1518
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1519
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1520
  case (ap'_exec r fr pt u sp t sd sf r' fr')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1521
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1522
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1523
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1524
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1525
      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1526
    with ap'_exec(3,4) obtain sa f where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1527
      SFa: "obj2sobj (sa @ s') (File f) = SFile (t, sd) sf" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1528
      Exfa: "exists (sa @ s') (File f)" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1529
      butsa: "initp_intact_but (sa @ s') sobj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1530
      othersa:"valid (sa @ s') \<and> obj2sobj (sa @ s') obj' = sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1531
               exists (sa @s') obj' \<and> no_del_event (sa @ s')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1532
      by (blast dest:obj2sobj_file intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1533
    with ap'_exec(1,2) notUkn obtain sb p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1534
      VSab: "valid (sb@sa@s')" and SOab': "obj2sobj (sb@sa@s') obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1535
      and nodelab: "no_del_event (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1536
      and intactab: "initp_intact_but (sb@sa@s') sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1537
      and SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r, fr, pt, u) sp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1538
      and exp:"exists (sb@sa@s') (Proc p)" and exobj'ab:"exists (sb@sa@s') obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1539
      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1540
    obtain e \<tau> where ev: "e = Execute (new_proc (sb@sa@s')) f" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1541
      and tau: "\<tau> = Clone p (new_proc (sb@sa@s'))#sb@sa@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1542
    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1543
    from Exfa nodelab have exf:"f \<in> current_files (sb@sa@s')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1544
      apply (drule_tac obj = "File f" in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1545
      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1546
    from SFa VSab Exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File f) = SFile (t,sd) sf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1547
      by (rule_tac s = "sa@s'" in obj2sobj_file_remains'', simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1548
    have np_not_initp: "new_proc (sb@sa@s') \<notin> init_processes" using nodelab 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1549
      apply (rule_tac notI, drule_tac obj = "Proc (new_proc (sb@sa@s'))" in nodel_imp_exists)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1550
      by (auto simp:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1551
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1552
    have valid: "valid (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1553
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1554
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1555
        using ev tau exp by (simp add:exf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1556
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1557
        using ev tau ap'_exec(5) SPab SFab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1558
        by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1559
                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1560
      ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1561
        by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1562
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1563
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1564
    have "initp_intact_but (e#\<tau>) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1565
    proof (cases sobj')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1566
      case (SProc sp srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1567
      show ?thesis
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1568
      proof (cases srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1569
        case (Some p')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1570
        with SOab' exobj'ab VSab intactab notUkn SProc
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1571
        have butp: "p' \<in> init_processes \<and> initp_intact_butp (sb@sa@s') p'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1572
          by (case_tac obj', auto intro:source_proc_in_init simp:obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1573
                                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1574
        then obtain p'' where exp': "p'' \<in> current_procs (sb@sa@s')" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1575
          SP': "obj2sobj (sb@sa@s') (Proc p'') = init_obj2sobj (Proc p')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1576
          by (auto simp:initp_alter_def initp_intact_butp_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1577
        hence "initp_alter (e#\<tau>) p'" using tau ev notUkn nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1578
          apply (simp add:initp_alter_def del:init_obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1579
          apply (rule_tac x = p'' in exI, rule conjI, simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1580
          apply (subgoal_tac "p'' \<noteq> new_proc (sb@sa@s')")
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1581
          apply (auto simp:obj2sobj.simps cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1582
                  simp del:init_obj2sobj.simps  split:option.splits)[1]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1583
          by (rule notI, simp add:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1584
        thus ?thesis using SProc Some intactab tau ev valid vs_tau np_not_initp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1585
          apply (simp add:initp_intact_butp_def del:init_obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1586
          apply (rule impI|rule allI|rule conjI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1587
          apply (erule_tac x = pa in allE)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1588
          by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:init_obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1589
                  split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1590
      next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1591
        case None
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1592
        with intactab SProc 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1593
        have "initp_intact (sb@sa@s')" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1594
        hence "initp_intact (e#\<tau>)" using tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1595
          by (simp add:initp_intact_I_exec)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1596
        thus ?thesis using SProc None by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1597
      qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1598
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1599
      case (SFile sf srf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1600
      with intactab SFile
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1601
      have "initp_intact (sb@sa@s')" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1602
      hence "initp_intact (e#\<tau>)" using tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1603
        by (simp add:initp_intact_I_exec)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1604
      thus ?thesis using SFile by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1605
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1606
      case (SIPC si sri)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1607
      with intactab SIPC
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1608
      have "initp_intact (sb@sa@s')" by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1609
      hence "initp_intact (e#\<tau>)" using tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1610
        by (simp add:initp_intact_I_exec)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1611
      thus ?thesis using SIPC by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1612
    next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1613
      case Unknown 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1614
      with notUkn show ?thesis by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1615
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1616
    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1617
      by (case_tac obj', simp+) moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1618
    have "obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1619
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1620
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1621
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1622
        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1623
        by (auto simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1624
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1625
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1626
        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1627
        by (auto simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1628
      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1629
        apply (case_tac "p' = new_proc (sb @ sa @ s')") 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1630
        using vs_tau exobj'ab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1631
        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1632
        using tau ev SOab' valid notUkn vs_tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1633
        by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1634
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1635
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1636
    have "obj2sobj (e#\<tau>) (Proc (new_proc (sb @ sa @ s'))) = 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1637
          SProc (r',fr',exec_type_aux r pt, u) sp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1638
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1639
      have "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) sp" using SPab tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1640
        by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1641
      hence "obj2sobj \<tau> (Proc (new_proc (sb@sa@s'))) = SProc (r,fr,pt,u) sp" using tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1642
      by (auto simp:obj2sobj.simps cp2sproc.simps pct_def clone_type_unchange 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1643
              split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1644
      moreover have "source_dir \<tau> f = Some sd" using vs_tau SFab tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1645
        by (auto simp:source_dir_simps obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1646
      ultimately show ?thesis using valid ev ap'_exec(6,7) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1647
        by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1648
    qed 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1649
    ultimately  show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1650
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1651
      obj2sobj (s @ s') obj = SProc (r', fr', exec_type_aux r pt, u) sp \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1652
      exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1653
      using ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1654
      apply (rule_tac x = "e#Clone p (new_proc (sb @ sa @ s')) #sb@sa" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1655
      by (rule_tac x = "Proc (new_proc (sb @ sa @ s'))" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1656
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1657
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1658
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1659
(* this is for all_sobjs_E2 *)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1660
lemma all_sobjs_E1:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1661
  "\<lbrakk>sobj \<in> all_sobjs'; valid s'; obj2sobj s' obj' = sobj'; exists s' obj'; sobj' \<noteq> Unknown; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1662
    no_del_event s'; initp_intact_but s' sobj'\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1663
   \<Longrightarrow> \<exists> s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>  exists (s@s') obj \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1664
                no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1665
                obj2sobj (s @ s') obj = sobj \<and> exists (s @ s') obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1666
by (drule all_sobjs_E1_aux, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1667
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1668
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1669
lemma all_sobjs_E2_aux[rule_format]:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1670
  "sobj \<in> all_sobjs' \<Longrightarrow> (\<forall> s' obj' sobj'. valid s' \<and> obj2sobj s' obj' = sobj' \<and> exists s' obj' \<and> sobj' \<noteq> Unknown \<and> not_both_sproc sobj sobj' \<and> no_del_event s' \<and> initp_intact s' \<longrightarrow> (\<exists> s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj \<and> obj2sobj (s @ s') obj = sobj \<and> exists (s @ s') obj \<and> sobj_source_eq_obj sobj obj))"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1671
proof (induct rule:all_sobjs'.induct)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1672
  case (af'_init f t) show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1673
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1674
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1675
    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1676
      and nodels': "no_del_event s'"and intacts':"initp_intact s'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1677
      and notboth: "not_both_sproc (SFile (t, f) (Some f)) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1678
      and exso': "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1679
    from nodels' af'_init(1) have exf: "f \<in> current_files s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1680
      by (drule_tac obj = "File f" in nodel_imp_exists, simp+)    
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1681
    have "obj2sobj s' (File f) = SFile (t, f) (Some f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1682
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1683
      have "obj2sobj [] (File f) = SFile (t, f) (Some f)"  using af'_init 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1684
        by (auto simp:etype_of_file_def source_dir_of_init' obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1685
                split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1686
      thus ?thesis using vss' exf nodels' af'_init(1) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1687
        by (drule_tac obj2sobj_file_remains_app', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1688
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1689
    thus "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1690
             exists (s @ s') obj' \<and> no_del_event (s @ s') \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1691
             initp_intact_but (s @ s') (SFile (t, f) (Some f)) \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1692
             obj2sobj (s @ s') obj = SFile (t, f) (Some f) \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1693
             exists (s @ s') obj \<and> sobj_source_eq_obj (SFile (t, f) (Some f)) obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1694
      apply (rule_tac x = "[]" in exI, rule_tac x = "File f" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1695
      by (simp add:vss' sobjs' nodels' intacts' exf exso')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1696
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1697
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1698
  case (af'_cfd t sd srf r fr pt u srp t') 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1699
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1700
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1701
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1702
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1703
      and Both:"not_both_sproc (SFile (t', sd) None) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1704
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1705
      and exobj':"exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1706
    with af'_cfd(1,2) obtain sa pf where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1707
      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1708
      "exists (sa@s') obj'" and "initp_intact (sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1709
      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1710
      exfa: "pf \<in> current_files (sa@s')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1711
      apply (drule_tac sf' = "(t, sd)" and srf' = srf in not_both_I_file)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1712
      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1713
      by (frule obj2sobj_file, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1714
    with af'_cfd(3) notUkn obtain sb p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1715
      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1716
      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1717
      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1718
      exsoab: "exists (sb@sa@s') obj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1719
      intactab: "initp_intact (sb@sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1720
      nodelab: "no_del_event (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1721
      apply (drule_tac s'= "sa@s'" and obj' = obj' in all_sobjs_E0, auto) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1722
      apply (frule obj2sobj_proc, erule exE)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1723
      by (auto intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1724
    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1725
      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1726
      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1727
    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1728
      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1729
    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1730
      and tau: "\<tau>=sb@sa@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1731
    
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1732
    have valid: "valid (e # \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1733
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1734
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1735
        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1736
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1737
        using ev tau af'_cfd(5,6,7) SPab SFab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1738
        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1739
                split:if_splits option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1740
      ultimately show ?thesis using vsab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1741
        by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1742
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1743
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1744
    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1745
        by (case_tac obj', simp+)   moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1746
    have "obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1747
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1748
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1749
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1750
        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1751
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1752
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1753
        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1754
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1755
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1756
        by (auto simp:obj2sobj.simps cp2sproc_simps' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1757
             simp del:cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1758
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1759
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1760
    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1761
      by (simp add:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1762
    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t', sd) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1763
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1764
      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1765
        using ev tau SFab SPab af'_cfd(5)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1766
        by (auto simp:obj2sobj.simps etype_of_file_def cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1767
          split:option.splits if_splits  intro!:etype_aux_prop4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1768
      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1769
        using ev tau SFab SPab valid ncf_parent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1770
        by (auto simp:source_dir_simps obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1771
                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1772
      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1773
        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1774
        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1775
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1776
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1777
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1778
      initp_intact_but (s @ s') (SFile (t', sd) None) \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1779
      obj2sobj (s @ s') obj = SFile (t', sd) None \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1780
      exists (s @ s') obj \<and> sobj_source_eq_obj (SFile (t', sd) None) obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1781
      using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1782
      apply (rule_tac x = "e#sb@sa" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1783
      by (rule_tac x = "File (new_childf pf \<tau>)" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1784
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1785
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1786
  case (af'_cfd' t sd srf r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1787
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1788
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1789
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1790
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1791
      and Both:"not_both_sproc (SFile (t, sd) None) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1792
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1793
      and exobj':"exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1794
    with af'_cfd'(1,2) obtain sa pf where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1795
      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1796
      "exists (sa@s') obj'" and "initp_intact (sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1797
      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1798
      exfa: "pf \<in> current_files (sa@s')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1799
      apply (drule_tac sf' = "(t, sd)" and srf' = srf in not_both_I_file)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1800
      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1801
      by (frule obj2sobj_file, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1802
    with af'_cfd'(3) notUkn obtain sb p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1803
      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1804
      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1805
      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1806
      exsoab: "exists (sb@sa@s') obj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1807
      intactab: "initp_intact (sb@sa@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1808
      nodelab: "no_del_event (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1809
      apply (drule_tac s'= "sa@s'" and obj' = obj' in all_sobjs_E0, auto) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1810
      apply (frule obj2sobj_proc, erule exE)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1811
      by (auto intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1812
    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1813
      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1814
      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1815
    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1816
      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1817
    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1818
      and tau: "\<tau>=sb@sa@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1819
    
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1820
    have valid: "valid (e # \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1821
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1822
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1823
        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1824
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1825
        using ev tau af'_cfd'(5,6) SPab SFab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1826
        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1827
                split:if_splits option.splits t_rc_file_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1828
      ultimately show ?thesis using vsab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1829
        by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1830
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1831
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1832
    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1833
        by (case_tac obj', simp+)   moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1834
    have "obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1835
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1836
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1837
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1838
        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1839
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1840
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1841
        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1842
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1843
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1844
        by (auto simp:obj2sobj.simps cp2sproc_simps' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1845
             simp del:cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1846
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1847
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1848
    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1849
      by (simp add:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1850
    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t, sd) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1851
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1852
      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1853
      proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1854
        have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = etype_of_file \<tau> pf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1855
          using ev tau SPab af'_cfd'(5) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1856
          by (auto simp:obj2sobj.simps cp2sproc.simps ncf_parent etype_of_file_def
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1857
            split:option.splits   intro!:etype_aux_prop3)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1858
        thus ?thesis using SFab tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1859
          by (auto simp:obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1860
      qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1861
      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1862
        using ev tau SFab SPab valid ncf_parent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1863
        by (auto simp:source_dir_simps obj2sobj.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1864
                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1865
      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1866
        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1867
        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1868
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1869
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1870
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1871
      initp_intact_but (s @ s') (SFile (t, sd) None) \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1872
      obj2sobj (s @ s') obj = SFile (t, sd) None \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1873
      exists (s @ s') obj \<and> sobj_source_eq_obj (SFile (t, sd) None) obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1874
      using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1875
      apply (rule_tac x = "e#sb@sa" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1876
      by (rule_tac x = "File (new_childf pf \<tau>)" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1877
  qed 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1878
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1879
  case (ai'_init i t) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1880
  hence initi: "i \<in> init_ipcs" using init_ipc_has_type
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1881
    by (simp add:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1882
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1883
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1884
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1885
    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1886
      and nodels': "no_del_event s'"and intacts':"initp_intact s'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1887
      and notboth: "not_both_sproc (SIPC t (Some i)) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1888
      and exso': "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1889
    from nodels' initi have exi: "i \<in> current_ipcs s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1890
      by (drule_tac obj = "IPC i" in nodel_imp_exists, simp+)    
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1891
    have "obj2sobj s' (IPC i) = SIPC t (Some i)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1892
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1893
      have "obj2sobj [] (IPC i) = SIPC t (Some i)"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1894
        using ai'_init initi by (auto simp:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1895
      thus ?thesis using vss' exi nodels' initi
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1896
        by (drule_tac obj2sobj_ipc_remains'', auto simp:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1897
    qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1898
    thus "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1899
             exists (s @ s') obj' \<and> no_del_event (s @ s') \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1900
             initp_intact_but (s @ s') (SIPC t (Some i)) \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1901
             obj2sobj (s @ s') obj = SIPC t (Some i) \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1902
             exists (s @ s') obj \<and> sobj_source_eq_obj (SIPC t (Some i)) obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1903
      apply (rule_tac x = "[]" in exI, rule_tac x = "IPC i" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1904
      by (simp add:vss' sobjs' nodels' intacts' exi exso' del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1905
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1906
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1907
  case (ai'_cipc r fr pt u srp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1908
  show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1909
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1910
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1911
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1912
      and Both:"not_both_sproc (SIPC (default_ipc_create_type r) None) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1913
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1914
      and exobj':"exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1915
    with ai'_cipc(1) notUkn obtain s p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1916
      SPab: "obj2sobj (s@s') (Proc p) = SProc (r,fr,pt,u) srp" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1917
      expab: "exists (s@s') (Proc p)" and vsab: "valid (s@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1918
      soab: "obj2sobj (s@s') obj' = sobj'" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1919
      exsoab: "exists (s@s') obj'" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1920
      intactab: "initp_intact (s@s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1921
      nodelab: "no_del_event (s@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1922
      apply (drule_tac s'= "s'" and obj' = obj' in all_sobjs_E0, auto) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1923
      apply (frule obj2sobj_proc, erule exE)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1924
      by (auto intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1925
    obtain e \<tau> where ev: "e = CreateIPC p (new_ipc \<tau>)" and tau: "\<tau>=s@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1926
    
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1927
    have valid: "valid (e # \<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1928
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1929
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1930
        using ev tau expab by (simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1931
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1932
        using ev tau ai'_cipc(3) SPab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1933
        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1934
      ultimately show ?thesis using vsab tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1935
        by (rule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1936
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1937
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1938
    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1939
        by (case_tac obj', simp+)   moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1940
    have "obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1941
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1942
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1943
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1944
        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1945
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1946
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1947
        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1948
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1949
        using soab tau valid notUkn nodel ev exsoab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1950
        by (auto simp:obj2sobj.simps cp2sproc_simps' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1951
             simp del:cp2sproc.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1952
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1953
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1954
    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1955
      by (simp add:initp_intact_I_others)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1956
    have "obj2sobj (e#\<tau>) (IPC (new_ipc \<tau>)) = SIPC (default_ipc_create_type r) None"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1957
      using ev tau SPab nodel 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1958
        nodel_imp_exists[where obj = "IPC (new_ipc \<tau>)" and s =\<tau>]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1959
      by (auto simp:obj2sobj.simps ni_notin_curi cp2sproc.simps 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1960
              split:option.splits  dest:no_del_event_cons_D)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1961
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1962
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1963
      initp_intact_but (s @ s') (SIPC (default_ipc_create_type r) None) \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1964
      obj2sobj (s @ s') obj = SIPC (default_ipc_create_type r) None \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1965
      exists (s @ s') obj \<and> sobj_source_eq_obj (SIPC (default_ipc_create_type r) None) obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1966
      using tau ev
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1967
      by (rule_tac x = "e#s" in exI, rule_tac x = "IPC (new_ipc \<tau>)" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1968
  qed 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1969
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1970
  case (ap'_init p r fr t u)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1971
  hence initp: "p \<in> init_processes" using init_proc_has_role
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1972
    by (simp add:bidirect_in_init_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1973
  show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1974
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1975
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1976
    assume VSs': "valid s'" and SOs': "obj2sobj s' obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1977
      and Nodels': "no_del_event s'"and Intacts':"initp_intact s'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1978
      and notboth: "not_both_sproc (SProc (r,fr,t,u) (Some p)) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1979
      and exso': "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1980
    from Nodels' initp have exp: "p \<in> current_procs s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1981
      apply (drule_tac obj = "Proc p" in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1982
      by (drule not_deleted_imp_exists, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1983
    with Intacts' initp ap'_init 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1984
    have "obj2sobj s' (Proc p) = SProc (r, fr, t, u) (Some p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1985
      by (auto simp:initp_intact_def split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1986
    thus "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1987
             exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1988
             initp_intact_but (s @ s') (SProc (r, fr, t, u) (Some p)) \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1989
             obj2sobj (s @ s') obj = SProc (r, fr, t, u) (Some p) \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1990
             exists (s @ s') obj \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1991
             sobj_source_eq_obj (SProc (r, fr, t, u) (Some p)) obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1992
      apply (rule_tac x = "[]" in exI, rule_tac x = "Proc p" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1993
      by (simp add:VSs' SOs' Nodels' Intacts' initp intact_imp_butp exp exso' 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1994
               del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1995
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1996
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1997
  case (ap'_crole r fr t u srp r')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1998
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1999
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2000
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2001
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2002
      and Both:"not_both_sproc (SProc (r', fr, t, u) srp) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2003
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2004
    with ap'_crole(1,2) obtain s p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2005
      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2006
      and nodelab: "no_del_event (s@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2007
      and butab: "initp_intact_but (s@s') (SProc (r, fr, t, u) srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2008
      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2009
      and exp:"exists (s@s') (Proc p)" and exobj':"exists (s@s') obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2010
      and sreq: "sobj_source_eq_obj (SProc (r, fr, t, u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2011
      by (blast dest:not_both_I obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2012
    from VSab SPab sreq exp have srpeq: "srp = Some p" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2013
      by (simp add:proc_source_eq_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2014
    with exp VSab SPab have initp: "p \<in> init_processes"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2015
      by (auto dest:source_proc_in_init simp:obj2sobj.simps split:option.splits) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2016
    obtain e \<tau> where ev: "e = ChangeRole p r'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2017
      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2018
    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2019
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2020
    have valid: "valid (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2021
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2022
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2023
        using ev tau exp by (simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2024
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2025
        using ev tau ap'_crole(3) SPab 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2026
        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2027
      ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2028
        by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2029
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2030
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2031
    have "initp_intact_but (e#\<tau>) (SProc (r', fr, t, u) srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2032
      using butab tau ev valid initp srpeq nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2033
      by (simp add:initp_intact_butp_I_crole)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2034
    have exobj': "exists (e#\<tau>) obj'" using exobj' valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2035
      by (case_tac obj', simp+)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2036
    have "obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2037
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2038
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2039
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2040
        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2041
        by (auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2042
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2043
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2044
        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2045
        by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2046
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow>False"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2047
        using Both SOab' notUkn 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2048
        by (simp del:obj2sobj.simps, drule_tac obj2sobj_proc', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2049
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2050
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2051
    have "obj2sobj (e#\<tau>) (Proc p) = SProc (r', fr, t, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2052
      using SPab tau vs_tau ev valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2053
      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2054
              split:option.splits)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2055
    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp   moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2056
    have "sobj_source_eq_obj (SProc (r', fr, t, u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2057
      using sreq by (case_tac srp, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2058
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2059
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2060
      initp_intact_but (s @ s') (SProc (r', fr, t, u) srp) \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2061
      obj2sobj (s @ s') obj = SProc (r', fr, t, u) srp \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2062
      exists (s @ s') obj \<and> sobj_source_eq_obj (SProc (r', fr, t, u) srp) obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2063
      using ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2064
      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2065
      by (rule_tac x = "Proc p" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2066
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2067
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2068
  case (ap'_chown r fr t u srp u' nr)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2069
  show ?case 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2070
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2071
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2072
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2073
      and Both:"not_both_sproc (SProc (nr,fr,chown_type_aux r nr t,u') srp) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2074
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2075
    with ap'_chown(1,2) obtain s p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2076
      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2077
      and nodelab: "no_del_event (s@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2078
      and butab: "initp_intact_but (s@s') (SProc (r, fr, t, u) srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2079
      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2080
      and exp:"exists (s@s') (Proc p)" and exobj':"exists (s@s') obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2081
      and sreq: "sobj_source_eq_obj (SProc (r, fr, t, u) srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2082
      by (blast dest:not_both_I obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2083
    from VSab SPab sreq exp have srpeq: "srp = Some p" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2084
      by (simp add:proc_source_eq_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2085
    with exp VSab SPab have initp: "p \<in> init_processes"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2086
      by (auto dest:source_proc_in_init simp:obj2sobj.simps split:option.splits) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2087
    obtain e \<tau> where ev: "e = ChangeOwner p u'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2088
      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2089
    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2090
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2091
    have valid: "valid (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2092
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2093
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2094
        using ev tau exp ap'_chown(3) by (simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2095
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2096
        using ev tau ap'_chown(5) SPab 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2097
        by (auto simp:cp2sproc.simps obj2sobj.simps pct_def clone_type_unchange
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2098
                split:option.splits t_rc_proc_type.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2099
          (* here is another place of no_limit of clone event assumption *)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2100
      ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2101
        by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2102
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2103
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2104
    have "initp_intact_but (e#\<tau>) (SProc (nr,fr,chown_type_aux r nr t,u') srp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2105
      using butab tau ev valid initp srpeq nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2106
      by (simp add:initp_intact_butp_I_chown)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2107
    have exobj': "exists (e#\<tau>) obj'" using exobj' valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2108
      by (case_tac obj', simp+)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2109
    have "obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2110
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2111
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2112
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2113
        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2114
        by (auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2115
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2116
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2117
        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2118
        by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2119
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow>False"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2120
        using Both SOab' notUkn 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2121
        by (simp del:obj2sobj.simps, drule_tac obj2sobj_proc', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2122
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2123
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2124
    have "obj2sobj (e#\<tau>) (Proc p) = SProc (nr,fr,chown_type_aux r nr t,u') srp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2125
      using SPab tau vs_tau ev valid ap'_chown(4)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2126
      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2127
              split:option.splits)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2128
    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp   moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2129
    have "sobj_source_eq_obj (SProc (nr,fr,chown_type_aux r nr t,u') srp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2130
      using sreq by (case_tac srp, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2131
    ultimately show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2132
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2133
      initp_intact_but (s @ s') (SProc (nr,fr,chown_type_aux r nr t,u') srp) \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2134
      obj2sobj (s @ s') obj = SProc (nr,fr,chown_type_aux r nr t,u') srp \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2135
      exists (s @ s') obj \<and> sobj_source_eq_obj (SProc (nr,fr,chown_type_aux r nr t,u') srp) obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2136
      using ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2137
      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2138
      by (rule_tac x = "Proc p" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2139
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2140
next
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2141
  case (ap'_exec r fr pt u sp t sd sf r' fr')
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2142
  show ?case
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2143
  proof (rule allI|rule impI|erule conjE)+
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2144
    fix s' obj' sobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2145
    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2146
      and Both:"not_both_sproc (SProc (r', fr', exec_type_aux r pt, u) sp) sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2147
      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2148
    with ap'_exec(3,4) obtain sa f where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2149
      SFa: "obj2sobj (sa @ s') (File f) = SFile (t, sd) sf" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2150
      Exfa: "exists (sa @ s') (File f)" and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2151
      buta: "initp_intact (sa @ s')" and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2152
      othersa:"valid (sa @ s') \<and> obj2sobj (sa @ s') obj' = sobj' \<and> exists (sa @s') obj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2153
      no_del_event (sa @ s') \<and> sobj_source_eq_obj (SFile (t, sd) sf) (File f)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2154
      apply (simp only:not_both_sproc.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2155
      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2156
      apply (erule_tac x = sobj' in allE, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2157
      by (frule obj2sobj_file, auto intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2158
    with SFa Exfa othersa ap'_exec(1,2) Both notUkn obtain sb p where
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2159
      VSab: "valid (sb@sa@s')" and SOab': "obj2sobj (sb@sa@s') obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2160
      and nodelab: "no_del_event (sb@sa@s')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2161
      and butab: "initp_intact_but (sb@sa@s') (SProc (r, fr, pt, u) sp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2162
      and SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r, fr, pt, u) sp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2163
      and exp:"exists (sb@sa@s') (Proc p)" and exobj':"exists (sb@sa@s') obj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2164
      and sreq: "sobj_source_eq_obj (SProc (r, fr, pt, u) sp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2165
      by (blast dest:not_both_I obj2sobj_proc intro:nodel_exists_remains)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2166
    from VSab SPab sreq exp have srpeq: "sp = Some p" by (simp add:proc_source_eq_prop)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2167
    with exp VSab SPab have initp: "p \<in> init_processes"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2168
      by (auto dest:source_proc_in_init simp:obj2sobj.simps split:option.splits) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2169
    obtain e \<tau> where ev: "e = Execute p f" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2170
      and tau: "\<tau> = Clone p (new_proc (sb@sa@s'))#sb@sa@s'" by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2171
    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2172
    from Exfa nodelab have exf:"f \<in> current_files (sb@sa@s')" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2173
      apply (drule_tac obj = "File f" in nodel_imp_un_deleted)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2174
      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2175
    from SFa VSab Exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File f) = SFile (t,sd) sf"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2176
      by (rule_tac s = "sa@s'" in obj2sobj_file_remains'', simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2177
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2178
    have valid: "valid (e#\<tau>)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2179
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2180
      have "os_grant \<tau> e"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2181
        using ev tau exp by (simp add:exf)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2182
      moreover have "rc_grant \<tau> e" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2183
        using ev tau ap'_exec(5) SPab SFab
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2184
        by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2185
                split:if_splits option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2186
      ultimately show ?thesis using vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2187
        by (erule_tac vs_step, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2188
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2189
    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2190
    have "initp_intact_but (e#\<tau>) (SProc (r', fr', exec_type_aux r pt, u) sp)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2191
      using butab tau ev valid initp srpeq nodel
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2192
      by (simp add:initp_intact_butp_I_exec)  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2193
    have exobj': "exists (e#\<tau>) obj'" using exobj' valid ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2194
      by (case_tac obj', simp+) moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2195
    have "obj2sobj (e#\<tau>) obj' = sobj'"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2196
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2197
      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2198
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2199
        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2200
        by (auto simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2201
      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2202
        using SOab' tau ev valid notUkn nodel exobj'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2203
        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2204
        by (auto simp del:obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2205
      moreover have "\<And> p. obj' = Proc p \<Longrightarrow>False"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2206
        using Both SOab' notUkn 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2207
        by (simp del:obj2sobj.simps, drule_tac obj2sobj_proc', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2208
      ultimately show ?thesis by (case_tac obj', auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2209
    qed  moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2210
    have "obj2sobj (e#\<tau>) (Proc p) = SProc (r',fr',exec_type_aux r pt, u) sp"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2211
    proof-
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2212
      have "obj2sobj \<tau> (Proc p) = SProc (r,fr,pt,u) sp" using SPab tau vs_tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2213
        by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2214
      moreover have "source_dir \<tau> f = Some sd" using vs_tau SFab tau 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2215
        by (auto simp:source_dir_simps obj2sobj.simps split:option.splits if_splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2216
      ultimately show ?thesis using valid ev ap'_exec(6,7) 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2217
        by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2218
    qed  moreover
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2219
    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp   moreover 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2220
    have "sobj_source_eq_obj (SProc (r',fr',exec_type_aux r pt,u) sp) (Proc p)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2221
      using sreq by (case_tac sp, simp+)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2222
    ultimately 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2223
    show "\<exists>s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2224
      exists (s @ s') obj' \<and> no_del_event (s @ s') \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2225
      initp_intact_but (s @ s') (SProc (r', fr', exec_type_aux r pt, u) sp) \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2226
      obj2sobj (s @ s') obj = SProc (r', fr', exec_type_aux r pt, u) sp \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2227
      exists (s @ s') obj \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2228
      sobj_source_eq_obj (SProc (r', fr', exec_type_aux r pt, u) sp) obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2229
      using ev tau
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2230
      apply (rule_tac x = "e#Clone p (new_proc (sb @ sa @ s')) #sb@sa" in exI)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2231
      by (rule_tac x = "Proc p" in exI, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2232
  qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2233
qed
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2234
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2235
lemma all_sobjs_E2: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2236
  "\<lbrakk>sobj \<in> all_sobjs'; valid s'; obj2sobj s' obj' = sobj'; exists s' obj'; sobj' \<noteq> Unknown; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2237
    not_both_sproc sobj sobj'; no_del_event s'; initp_intact s'\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2238
   \<Longrightarrow> \<exists> s obj. valid (s @ s') \<and> obj2sobj (s @ s') obj' = sobj' \<and> exists (s@s') obj \<and>
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2239
                no_del_event (s @ s') \<and> initp_intact_but (s @ s') sobj \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2240
                obj2sobj (s @ s') obj = sobj \<and> exists (s @ s') obj \<and> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2241
                sobj_source_eq_obj sobj obj"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2242
by (drule all_sobjs_E2_aux, auto)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2243
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2244
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2245
dcde836219bc add thy files
chunhan
parents:
diff changeset
  2246
end