add thy files
authorchunhan
Fri, 12 Apr 2013 10:43:11 +0100
changeset 1 dcde836219bc
parent 0 b992684e9ff6
child 2 301f567e2a8e
add thy files
List_Prefix.thy
Paper.thy
all_sobj_prop.thy
del_vs_del_s.thy
deleted_prop.thy
final_theorems.thy
finite_static.thy
my_list_prefix.thy
obj2sobj_prop.thy
os_rc.thy
rc_theory.thy
sound_defs_prop.thy
source_prop.thy
tainted.thy
tainted_vs_tainted_s.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/List_Prefix.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,382 @@
+(*  Title:      HOL/Library/List_Prefix.thy
+    Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
+*)
+
+header {* List prefixes and postfixes *}
+
+theory List_Prefix
+imports List Main
+begin
+
+subsection {* Prefix order on lists *}
+
+instantiation list :: (type) "{order, bot}"
+begin
+
+definition
+  prefix_def: "xs \<le> ys \<longleftrightarrow> (\<exists>zs. ys = xs @ zs)"
+
+definition
+  strict_prefix_def: "xs < ys \<longleftrightarrow> xs \<le> ys \<and> xs \<noteq> (ys::'a list)"
+
+definition
+  "bot = []"
+
+instance proof
+qed (auto simp add: prefix_def strict_prefix_def bot_list_def)
+
+end
+
+lemma prefixI [intro?]: "ys = xs @ zs ==> xs \<le> ys"
+  unfolding prefix_def by blast
+
+lemma prefixE [elim?]:
+  assumes "xs \<le> ys"
+  obtains zs where "ys = xs @ zs"
+  using assms unfolding prefix_def by blast
+
+lemma strict_prefixI' [intro?]: "ys = xs @ z # zs ==> xs < ys"
+  unfolding strict_prefix_def prefix_def by blast
+
+lemma strict_prefixE' [elim?]:
+  assumes "xs < ys"
+  obtains z zs where "ys = xs @ z # zs"
+proof -
+  from `xs < ys` obtain us where "ys = xs @ us" and "xs \<noteq> ys"
+    unfolding strict_prefix_def prefix_def by blast
+  with that show ?thesis by (auto simp add: neq_Nil_conv)
+qed
+
+lemma strict_prefixI [intro?]: "xs \<le> ys ==> xs \<noteq> ys ==> xs < (ys::'a list)"
+  unfolding strict_prefix_def by blast
+
+lemma strict_prefixE [elim?]:
+  fixes xs ys :: "'a list"
+  assumes "xs < ys"
+  obtains "xs \<le> ys" and "xs \<noteq> ys"
+  using assms unfolding strict_prefix_def by blast
+
+
+subsection {* Basic properties of prefixes *}
+
+theorem Nil_prefix [iff]: "[] \<le> xs"
+  by (simp add: prefix_def)
+
+theorem prefix_Nil [simp]: "(xs \<le> []) = (xs = [])"
+  by (induct xs) (simp_all add: prefix_def)
+
+lemma prefix_snoc [simp]: "(xs \<le> ys @ [y]) = (xs = ys @ [y] \<or> xs \<le> ys)"
+proof
+  assume "xs \<le> ys @ [y]"
+  then obtain zs where zs: "ys @ [y] = xs @ zs" ..
+  show "xs = ys @ [y] \<or> xs \<le> ys"
+    by (metis append_Nil2 butlast_append butlast_snoc prefixI zs)
+next
+  assume "xs = ys @ [y] \<or> xs \<le> ys"
+  then show "xs \<le> ys @ [y]"
+    by (metis order_eq_iff order_trans prefixI)
+qed
+
+lemma Cons_prefix_Cons [simp]: "(x # xs \<le> y # ys) = (x = y \<and> xs \<le> ys)"
+  by (auto simp add: prefix_def)
+
+lemma less_eq_list_code [code]:
+  "([]\<Colon>'a\<Colon>{equal, ord} list) \<le> xs \<longleftrightarrow> True"
+  "(x\<Colon>'a\<Colon>{equal, ord}) # xs \<le> [] \<longleftrightarrow> False"
+  "(x\<Colon>'a\<Colon>{equal, ord}) # xs \<le> y # ys \<longleftrightarrow> x = y \<and> xs \<le> ys"
+  by simp_all
+
+lemma same_prefix_prefix [simp]: "(xs @ ys \<le> xs @ zs) = (ys \<le> zs)"
+  by (induct xs) simp_all
+
+lemma same_prefix_nil [iff]: "(xs @ ys \<le> xs) = (ys = [])"
+  by (metis append_Nil2 append_self_conv order_eq_iff prefixI)
+
+lemma prefix_prefix [simp]: "xs \<le> ys ==> xs \<le> ys @ zs"
+  by (metis order_le_less_trans prefixI strict_prefixE strict_prefixI)
+
+lemma append_prefixD: "xs @ ys \<le> zs \<Longrightarrow> xs \<le> zs"
+  by (auto simp add: prefix_def)
+
+theorem prefix_Cons: "(xs \<le> y # ys) = (xs = [] \<or> (\<exists>zs. xs = y # zs \<and> zs \<le> ys))"
+  by (cases xs) (auto simp add: prefix_def)
+
+theorem prefix_append:
+  "(xs \<le> ys @ zs) = (xs \<le> ys \<or> (\<exists>us. xs = ys @ us \<and> us \<le> zs))"
+  apply (induct zs rule: rev_induct)
+   apply force
+  apply (simp del: append_assoc add: append_assoc [symmetric])
+  apply (metis append_eq_appendI)
+  done
+
+lemma append_one_prefix:
+  "xs \<le> ys ==> length xs < length ys ==> xs @ [ys ! length xs] \<le> ys"
+  unfolding prefix_def
+  by (metis Cons_eq_appendI append_eq_appendI append_eq_conv_conj
+    eq_Nil_appendI nth_drop')
+
+theorem prefix_length_le: "xs \<le> ys ==> length xs \<le> length ys"
+  by (auto simp add: prefix_def)
+
+lemma prefix_same_cases:
+  "(xs\<^isub>1::'a list) \<le> ys \<Longrightarrow> xs\<^isub>2 \<le> ys \<Longrightarrow> xs\<^isub>1 \<le> xs\<^isub>2 \<or> xs\<^isub>2 \<le> xs\<^isub>1"
+  unfolding prefix_def by (metis append_eq_append_conv2)
+
+lemma set_mono_prefix: "xs \<le> ys \<Longrightarrow> set xs \<subseteq> set ys"
+  by (auto simp add: prefix_def)
+
+lemma take_is_prefix: "take n xs \<le> xs"
+  unfolding prefix_def by (metis append_take_drop_id)
+
+lemma map_prefixI: "xs \<le> ys \<Longrightarrow> map f xs \<le> map f ys"
+  by (auto simp: prefix_def)
+
+lemma prefix_length_less: "xs < ys \<Longrightarrow> length xs < length ys"
+  by (auto simp: strict_prefix_def prefix_def)
+
+lemma strict_prefix_simps [simp, code]:
+  "xs < [] \<longleftrightarrow> False"
+  "[] < x # xs \<longleftrightarrow> True"
+  "x # xs < y # ys \<longleftrightarrow> x = y \<and> xs < ys"
+  by (simp_all add: strict_prefix_def cong: conj_cong)
+
+lemma take_strict_prefix: "xs < ys \<Longrightarrow> take n xs < ys"
+  apply (induct n arbitrary: xs ys)
+   apply (case_tac ys, simp_all)[1]
+  apply (metis order_less_trans strict_prefixI take_is_prefix)
+  done
+
+lemma not_prefix_cases:
+  assumes pfx: "\<not> ps \<le> ls"
+  obtains
+    (c1) "ps \<noteq> []" and "ls = []"
+  | (c2) a as x xs where "ps = a#as" and "ls = x#xs" and "x = a" and "\<not> as \<le> xs"
+  | (c3) a as x xs where "ps = a#as" and "ls = x#xs" and "x \<noteq> a"
+proof (cases ps)
+  case Nil then show ?thesis using pfx by simp
+next
+  case (Cons a as)
+  note c = `ps = a#as`
+  show ?thesis
+  proof (cases ls)
+    case Nil then show ?thesis by (metis append_Nil2 pfx c1 same_prefix_nil)
+  next
+    case (Cons x xs)
+    show ?thesis
+    proof (cases "x = a")
+      case True
+      have "\<not> as \<le> xs" using pfx c Cons True by simp
+      with c Cons True show ?thesis by (rule c2)
+    next
+      case False
+      with c Cons show ?thesis by (rule c3)
+    qed
+  qed
+qed
+
+lemma not_prefix_induct [consumes 1, case_names Nil Neq Eq]:
+  assumes np: "\<not> ps \<le> ls"
+    and base: "\<And>x xs. P (x#xs) []"
+    and r1: "\<And>x xs y ys. x \<noteq> y \<Longrightarrow> P (x#xs) (y#ys)"
+    and r2: "\<And>x xs y ys. \<lbrakk> x = y; \<not> xs \<le> ys; P xs ys \<rbrakk> \<Longrightarrow> P (x#xs) (y#ys)"
+  shows "P ps ls" using np
+proof (induct ls arbitrary: ps)
+  case Nil then show ?case
+    by (auto simp: neq_Nil_conv elim!: not_prefix_cases intro!: base)
+next
+  case (Cons y ys)
+  then have npfx: "\<not> ps \<le> (y # ys)" by simp
+  then obtain x xs where pv: "ps = x # xs"
+    by (rule not_prefix_cases) auto
+  show ?case by (metis Cons.hyps Cons_prefix_Cons npfx pv r1 r2)
+qed
+
+
+subsection {* Parallel lists *}
+
+definition
+  parallel :: "'a list => 'a list => bool"  (infixl "\<parallel>" 50) where
+  "(xs \<parallel> ys) = (\<not> xs \<le> ys \<and> \<not> ys \<le> xs)"
+
+lemma parallelI [intro]: "\<not> xs \<le> ys ==> \<not> ys \<le> xs ==> xs \<parallel> ys"
+  unfolding parallel_def by blast
+
+lemma parallelE [elim]:
+  assumes "xs \<parallel> ys"
+  obtains "\<not> xs \<le> ys \<and> \<not> ys \<le> xs"
+  using assms unfolding parallel_def by blast
+
+theorem prefix_cases:
+  obtains "xs \<le> ys" | "ys < xs" | "xs \<parallel> ys"
+  unfolding parallel_def strict_prefix_def by blast
+
+theorem parallel_decomp:
+  "xs \<parallel> ys ==> \<exists>as b bs c cs. b \<noteq> c \<and> xs = as @ b # bs \<and> ys = as @ c # cs"
+proof (induct xs rule: rev_induct)
+  case Nil
+  then have False by auto
+  then show ?case ..
+next
+  case (snoc x xs)
+  show ?case
+  proof (rule prefix_cases)
+    assume le: "xs \<le> ys"
+    then obtain ys' where ys: "ys = xs @ ys'" ..
+    show ?thesis
+    proof (cases ys')
+      assume "ys' = []"
+      then show ?thesis by (metis append_Nil2 parallelE prefixI snoc.prems ys)
+    next
+      fix c cs assume ys': "ys' = c # cs"
+      then show ?thesis
+        by (metis Cons_eq_appendI eq_Nil_appendI parallelE prefixI
+          same_prefix_prefix snoc.prems ys)
+    qed
+  next
+    assume "ys < xs" then have "ys \<le> xs @ [x]" by (simp add: strict_prefix_def)
+    with snoc have False by blast
+    then show ?thesis ..
+  next
+    assume "xs \<parallel> ys"
+    with snoc obtain as b bs c cs where neq: "(b::'a) \<noteq> c"
+      and xs: "xs = as @ b # bs" and ys: "ys = as @ c # cs"
+      by blast
+    from xs have "xs @ [x] = as @ b # (bs @ [x])" by simp
+    with neq ys show ?thesis by blast
+  qed
+qed
+
+lemma parallel_append: "a \<parallel> b \<Longrightarrow> a @ c \<parallel> b @ d"
+  apply (rule parallelI)
+    apply (erule parallelE, erule conjE,
+      induct rule: not_prefix_induct, simp+)+
+  done
+
+lemma parallel_appendI: "xs \<parallel> ys \<Longrightarrow> x = xs @ xs' \<Longrightarrow> y = ys @ ys' \<Longrightarrow> x \<parallel> y"
+  by (simp add: parallel_append)
+
+lemma parallel_commute: "a \<parallel> b \<longleftrightarrow> b \<parallel> a"
+  unfolding parallel_def by auto
+
+
+subsection {* Postfix order on lists *}
+
+definition
+  postfix :: "'a list => 'a list => bool"  ("(_/ >>= _)" [51, 50] 50) where
+  "(xs >>= ys) = (\<exists>zs. xs = zs @ ys)"
+
+lemma postfixI [intro?]: "xs = zs @ ys ==> xs >>= ys"
+  unfolding postfix_def by blast
+
+lemma postfixE [elim?]:
+  assumes "xs >>= ys"
+  obtains zs where "xs = zs @ ys"
+  using assms unfolding postfix_def by blast
+
+lemma postfix_refl [iff]: "xs >>= xs"
+  by (auto simp add: postfix_def)
+lemma postfix_trans: "\<lbrakk>xs >>= ys; ys >>= zs\<rbrakk> \<Longrightarrow> xs >>= zs"
+  by (auto simp add: postfix_def)
+lemma postfix_antisym: "\<lbrakk>xs >>= ys; ys >>= xs\<rbrakk> \<Longrightarrow> xs = ys"
+  by (auto simp add: postfix_def)
+
+lemma Nil_postfix [iff]: "xs >>= []"
+  by (simp add: postfix_def)
+lemma postfix_Nil [simp]: "([] >>= xs) = (xs = [])"
+  by (auto simp add: postfix_def)
+
+lemma postfix_ConsI: "xs >>= ys \<Longrightarrow> x#xs >>= ys"
+  by (auto simp add: postfix_def)
+lemma postfix_ConsD: "xs >>= y#ys \<Longrightarrow> xs >>= ys"
+  by (auto simp add: postfix_def)
+
+lemma postfix_appendI: "xs >>= ys \<Longrightarrow> zs @ xs >>= ys"
+  by (auto simp add: postfix_def)
+lemma postfix_appendD: "xs >>= zs @ ys \<Longrightarrow> xs >>= ys"
+  by (auto simp add: postfix_def)
+
+lemma postfix_is_subset: "xs >>= ys ==> set ys \<subseteq> set xs"
+proof -
+  assume "xs >>= ys"
+  then obtain zs where "xs = zs @ ys" ..
+  then show ?thesis by (induct zs) auto
+qed
+
+lemma postfix_ConsD2: "x#xs >>= y#ys ==> xs >>= ys"
+proof -
+  assume "x#xs >>= y#ys"
+  then obtain zs where "x#xs = zs @ y#ys" ..
+  then show ?thesis
+    by (induct zs) (auto intro!: postfix_appendI postfix_ConsI)
+qed
+
+lemma postfix_to_prefix [code]: "xs >>= ys \<longleftrightarrow> rev ys \<le> rev xs"
+proof
+  assume "xs >>= ys"
+  then obtain zs where "xs = zs @ ys" ..
+  then have "rev xs = rev ys @ rev zs" by simp
+  then show "rev ys <= rev xs" ..
+next
+  assume "rev ys <= rev xs"
+  then obtain zs where "rev xs = rev ys @ zs" ..
+  then have "rev (rev xs) = rev zs @ rev (rev ys)" by simp
+  then have "xs = rev zs @ ys" by simp
+  then show "xs >>= ys" ..
+qed
+
+lemma distinct_postfix: "distinct xs \<Longrightarrow> xs >>= ys \<Longrightarrow> distinct ys"
+  by (clarsimp elim!: postfixE)
+
+lemma postfix_map: "xs >>= ys \<Longrightarrow> map f xs >>= map f ys"
+  by (auto elim!: postfixE intro: postfixI)
+
+lemma postfix_drop: "as >>= drop n as"
+  unfolding postfix_def
+  apply (rule exI [where x = "take n as"])
+  apply simp
+  done
+
+lemma postfix_take: "xs >>= ys \<Longrightarrow> xs = take (length xs - length ys) xs @ ys"
+  by (clarsimp elim!: postfixE)
+
+lemma parallelD1: "x \<parallel> y \<Longrightarrow> \<not> x \<le> y"
+  by blast
+
+lemma parallelD2: "x \<parallel> y \<Longrightarrow> \<not> y \<le> x"
+  by blast
+
+lemma parallel_Nil1 [simp]: "\<not> x \<parallel> []"
+  unfolding parallel_def by simp
+
+lemma parallel_Nil2 [simp]: "\<not> [] \<parallel> x"
+  unfolding parallel_def by simp
+
+lemma Cons_parallelI1: "a \<noteq> b \<Longrightarrow> a # as \<parallel> b # bs"
+  by auto
+
+lemma Cons_parallelI2: "\<lbrakk> a = b; as \<parallel> bs \<rbrakk> \<Longrightarrow> a # as \<parallel> b # bs"
+  by (metis Cons_prefix_Cons parallelE parallelI)
+
+lemma not_equal_is_parallel:
+  assumes neq: "xs \<noteq> ys"
+    and len: "length xs = length ys"
+  shows "xs \<parallel> ys"
+  using len neq
+proof (induct rule: list_induct2)
+  case Nil
+  then show ?case by simp
+next
+  case (Cons a as b bs)
+  have ih: "as \<noteq> bs \<Longrightarrow> as \<parallel> bs" by fact
+  show ?case
+  proof (cases "a = b")
+    case True
+    then have "as \<noteq> bs" using Cons by simp
+    then show ?thesis by (rule Cons_parallelI2 [OF True ih])
+  next
+    case False
+    then show ?thesis by (rule Cons_parallelI1)
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Paper.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,1517 @@
+(*<*)
+theory Paper
+imports rc_theory final_theorems rc_theory os_rc 
+begin
+
+(* THEOREMS *)
+
+
+notation (Rule output)
+  "==>"  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
+
+syntax (Rule output)
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:\mbox{}\inferrule{>_\<^raw:}>\<^raw:{\mbox{>_\<^raw:}}>")
+
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" 
+  ("\<^raw:\mbox{>_\<^raw:}\\>/ _")
+
+  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
+
+notation (Axiom output)
+  "Trueprop"  ("\<^raw:\mbox{}\inferrule{\mbox{}}{\mbox{>_\<^raw:}}>")
+
+notation (IfThen output)
+  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+syntax (IfThen output)
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
+  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
+
+notation (IfThenNoBox output)
+  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+syntax (IfThenNoBox output)
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
+  "_asm" :: "prop \<Rightarrow> asms" ("_")
+
+(* insert *)
+notation (latex)
+  "Set.empty" ("\<emptyset>")
+
+translations 
+  "{x} \<union> A" <= "CONST insert x A"
+  "{x,y}" <= "{x} \<union> {y}"
+  "{x,y} \<union> A" <= "{x} \<union> ({y} \<union> A)"
+  "{x}" <= "{x} \<union> \<emptyset>"
+
+lemma impeq:
+  "A = B \<Longrightarrow> (B \<Longrightarrow> A)"
+by auto
+
+
+
+
+consts DUMMY::'a
+
+abbreviation
+  "is_parent f pf \<equiv> (parent f = Some pf)"
+
+context tainting_s_sound begin 
+
+notation (latex output)
+  source_dir ("anchor") and
+  SProc ("P_\<^bsup>_\<^esup>") and
+  SFile ("F_\<^bsup>_\<^esup>") and
+  SIPC ("I'(_')\<^bsup>_\<^esup>") and
+  READ ("Read") and 
+  WRITE ("Write") and
+  EXECUTE ("Execute") and
+  CHANGE_OWNER ("ChangeOwner") and
+  CREATE ("Create") and
+  SEND ("Send") and
+  RECEIVE ("Receive") and
+  DELETE ("Delete") and
+  compatible ("permissions") and
+  comproles ("compatible") and
+  DUMMY  ("\<^raw:\mbox{$\_$}>") and
+  Cons ("_::_" [78,77] 79) and 
+  Proc ("") and
+  File ("") and
+  File_type ("") and
+  Proc_type ("") and
+  IPC ("") and
+  init_processes ("init'_procs") and
+  os_grant ("admissible") and
+  rc_grant ("granted") and
+  exists ("alive") and
+  default_fd_create_type ("default'_type") and
+  InheritParent_file_type ("InheritPatentType") and
+  NormalFile_type ("NormalFileType") and
+  deleted ("deleted _ _" [50, 100] 100) and
+  taintable_s ("taintable\<^isup>s") and
+  tainted_s ("tainted\<^isup>s") and
+  all_sobjs ("reachable\<^isup>s") and
+  init_obj2sobj ("\<lbrakk>_\<rbrakk>") and
+  erole_functor ("erole'_aux") --"I have a erole_functor and etype_aux to handle 
+                                 efficient, but their name not same, so ..., but don't work"
+
+
+abbreviation
+  "is_process_type s p t \<equiv> (type_of_process s p = Some t)"
+
+abbreviation
+  "is_current_role s p r \<equiv> (currentrole s p = Some r)"
+
+abbreviation
+  "is_file_type s f t \<equiv> (etype_of_file s f = Some t)"
+
+lemma osgrant2:
+  "\<lbrakk>p \<in> current_procs \<tau>; f \<notin> current_files \<tau>; parent f = Some pf; pf \<in> current_files \<tau>\<rbrakk> \<Longrightarrow>
+   os_grant \<tau> (CreateFile p f)"
+by simp
+
+lemma osgrant6:
+ "\<lbrakk>p \<in> current_procs \<tau>; u \<in> init_users\<rbrakk> \<Longrightarrow> os_grant \<tau> (ChangeOwner p u)"
+by simp
+
+lemma osgrant10:
+  "\<lbrakk>p \<in> current_procs \<tau>; p' = new_proc \<tau>\<rbrakk> \<Longrightarrow> os_grant \<tau> (Clone p p')"
+by simp 
+
+
+lemma rcgrant1:
+  "\<lbrakk>is_parent f pf; is_file_type s pf t; is_current_role s p r;
+    default_fd_create_type r = InheritParent_file_type; 
+    (r, File_type t, WRITE) \<in> compatible\<rbrakk>
+   \<Longrightarrow> rc_grant s (CreateFile p f)"
+by simp
+
+lemma rcgrant1':
+  "\<lbrakk>is_parent f pf; is_file_type s pf t; is_current_role s p r; 
+    default_fd_create_type r = NormalFile_type t'; 
+    (r, File_type t, WRITE) \<in> compatible;
+    (r, File_type t', CREATE) \<in> compatible\<rbrakk>
+   \<Longrightarrow> rc_grant s (CreateFile p f)"
+by simp
+
+lemma rcgrant4:
+  "\<lbrakk>is_current_role s p r; is_file_type s f t; (r, File_type t, EXECUTE) \<in> compatible\<rbrakk>
+   \<Longrightarrow> rc_grant s (Execute p f)"
+by simp
+
+lemma rcgrant7: 
+  "\<lbrakk>is_current_role s p r; r' \<in> comproles r\<rbrakk> \<Longrightarrow> rc_grant s (ChangeRole p r')"
+by simp
+
+lemma rcgrant_CHO:
+"\<lbrakk>is_current_role s p r; 
+  type_of_process s p = Some t;
+  (r, Proc_type t, CHANGE_OWNER) \<in> compatible\<rbrakk> \<Longrightarrow> rc_grant s (ChangeOwner p u)"
+by(simp)
+
+lemma pf_in_current_paper:
+  "\<lbrakk>is_parent f pf; f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> pf \<in> current_files s"
+by (simp add:parent_file_in_current)
+
+lemma dels:
+  shows "deleted (Proc p') ((Kill p p')#s)" 
+  and   "deleted (File f) ((DeleteFile p f)#s)" 
+  and   "deleted (IPC i) ((DeleteIPC p i)#s)"
+  and   "deleted obj s \<Longrightarrow> deleted obj (e#s)"
+apply simp_all
+apply(case_tac e)
+apply(simp_all)
+done
+
+lemma tainted_10:
+  "\<lbrakk>(File f) \<in> tainted s; valid (e # s); f \<in> current_files (e # s)\<rbrakk> 
+  \<Longrightarrow> (File f) \<in> tainted (e # s)"
+apply(rule tainted.intros)
+apply(assumption)
+apply(assumption)
+apply(simp only: exists.simps)
+done
+
+definition
+  Init ("init _")
+where
+  "Init obj \<equiv> exists [] obj" 
+
+lemma Init_rhs:
+  shows "Init (File f) = (f \<in> init_files)"
+  and   "Init (Proc p) = (p \<in> init_processes)"
+  and   "Init (IPC i) = (i \<in> init_ipcs)"
+unfolding Init_def
+by(simp_all)
+
+notation (latex output)
+  Init ("_ \<in> init")
+
+lemma af_init':
+  "\<lbrakk>f \<in> init_files; is_file_type [] f t\<rbrakk>
+  \<Longrightarrow> SFile (t, f) (Some f) \<in> all_sobjs"
+apply(rule af_init)
+apply(simp)
+by (simp add:etype_of_file_def)
+
+declare [[show_question_marks = false]]
+
+
+(*>*)
+
+section {* Introduction *}
+
+text {*
+  Role-based access control models are used in many operating systems
+  for enforcing security properties.  The
+  \emph{Role-Compatibility Model} (RC-Model), introduced by Ott
+  \cite{ottrc,ottthesis}, is one such role-based access control
+  model. It defines \emph{roles}, which are associated with processes,
+  and defines \emph{types}, which are associated with system
+  resources, such as files and directories. The RC-Model also includes
+  types for interprocess communication, that is message queues,
+  sockets and shared memory. A policy in the RC-Model gives every user
+  a default role, and also specifies how roles can be
+  changed. Moreover, it specifies which types of resources a role has
+  permission to access, and also the \emph{mode} with which the role
+  can access the resources, for example read, write, send, receive and
+  so on.
+   
+  The RC-Model is built on top of a collection of system calls
+  provided by the operating system, for instance system calls for
+  reading and writing files, cloning and killing of processes, and
+  sending and receiving messages. The purpose of the RC-Model is to
+  restrict access to these system calls and thereby enforce security
+  properties of the system. A problem with the RC-Model and role-based
+  access control models in general is that a system administrator has
+  to specify an appropriate access control policy. The difficulty with
+  this is that \emph{``what you specify is what you get but not
+  necessarily what you want''} \cite[Page 242]{Jha08}. To overcome
+  this difficulty, a system administrator needs some kind of sanity
+  check for whether an access control policy is really securing
+  resources.  Existing works, for example \cite{sanity01,sanity02},
+  provide sanity checks for policies by specifying properties and
+  using model checking techniques to ensure a policy at hand satisfies
+  these properties. However, these checks only address the problem on
+  the level of policies---they can only check ``on the surface''
+  whether the policy reflects the intentions of the system
+  administrator---these checks are not justified by the actual
+  behaviour of the operating system. The main problem this paper addresses is to check
+  when a policy matches the intentions of a system administrator
+  \emph{and} given such a policy, the operating system actually
+  enforces this policy.
+ 
+  Our work is related to the preliminary work by Archer et al
+  \cite{Archer03} about the security model of SELinux. 
+  They also give a dynamic model of system calls on which the access
+  controls are implemented. Their dynamic model is defined in terms of
+  IO automata and mechanised in the PVS theorem prover. For specifying
+  and reasoning about automata they use the TAME tool in PVS. Their work checks
+  well-formedness properties of access policies by type-checking 
+  generated definitions in PVS. They can also ensure some ``\emph{simple
+  properties}'' (their terminology), for example whether a process
+  with a particular PID is present in every reachable state from 
+  an initial state. They also consider ``\emph{deeper properties}'', for
+  example whether only a process with root-permissions 
+  or one of its descendents ever gets permission to write to kernel
+  log files. They write that they can state such deeper 
+  properties about access policies, but about checking such properties
+  they write that ``\emph{the feasibility of doing 
+  so is currently an open question}'' \cite[Page 167]{Archer03}.
+  We improve upon their results by using our sound and complete 
+  static policy check to make this feasible. 
+
+  Our formal models and correctness proofs are mechanised in the
+  interactive theorem prover Isabelle/HOL. The mechanisation of the models is a
+  prerequisite for any correctness proof about the RC-Model, since it
+  includes a large number of interdependent concepts and very complex
+  operations that determine roles and types. In our opinion it is
+  futile to attempt to reason about them by just using ``pencil-and-paper''.
+  Following good experience in earlier mechanisation work
+  \cite{ZhangUrbanWu12}, we use Paulson's inductive method for
+  reasoning about sequences of events \cite{Paulson98}.  For example
+  we model system calls as events and reason about an inductive
+  definition of valid traces, that is lists of events.  Central to
+  this paper is a notion of a resource being \emph{tainted}, which for
+  example means it contains a virus or a back door.  We use our model
+  of system calls in order to characterise how such a tainted object
+  can ``spread'' through the system.  For a system administrator the
+  important question is whether such a tainted file, possibly
+  introduced by a user, can affect core system files and render the
+  whole system insecure, or whether it can be contained by the access
+  policy. Our results show that a corresponding check can be performed
+  statically by analysing the initial state of the system and the access policy.
+  \smallskip
+ 
+  \noindent
+  {\bf Contributions:} 
+  We give a complete formalisation of the RC-Model in the interactive
+  theorem prover Isabelle/HOL.  We also give a dynamic model of the
+  operating system by formalising all security related events that can
+  happen while the system is running.  As far as we are aware, we are
+  the first ones who formally prove that if a policy in the RC-Model
+  satisfies an access property, then there is no sequence of events
+  (system calls) that can violate this access property.  We also prove
+  the opposite: if a policy does not meet an access property, then
+  there is a sequence of events that will violate this property in our
+  model of the operating system.  With these two results in place we
+  can show that a static policy check is sufficient in order to
+  guarantee the access properties before running the system. Again as
+  far as we know, no such check that is the operating
+  system behaviour has been designed before.
+  
+
+  %Specified dynamic behaviour of the system; 
+  %we specified a static AC model; designed a tainted relation for 
+  %the system; proved that they coincide.
+  %In our paper ....
+   
+*}
+  
+section {* Preliminaries about the RC-Model *}
+
+
+text {*
+  The Role-Compatibility Model (RC-Model) is a role-based access
+  control model. It has been introduced by Ott \cite{ottrc} and is
+  used in running systems for example to secure Apache servers. It
+  provides a more fine-grained control over access permissions than
+  simple Unix-style access control models. This more fine-grained
+  control solves the problem of server processes running as root with
+  too many access permissions in order to accomplish a task at
+  hand. In the RC-Model, system administrators are able to restrict
+  what the role of server is allowed to do and in doing so reduce the
+  attack surface of a system.
+  
+  Policies in the RC-Model talk about \emph{users}, \emph{roles},
+  \emph{types} and \emph{objects}.  Objects are processes, files or
+  IPCs (interprocess communication objects---such as message queues,
+  sockets and shared memory). Objects are the resources of a system an
+  RC-policy can restrict access to.  In what follows we use the letter
+  @{term u} to stand for users, @{text r} for roles, @{term p} for
+  processes, @{term f} for files and @{term i} for IPCs. We also
+  use @{text obj} as a generic variable for objects.
+  The RC-Model has the following eight kinds of access modes to objects:
+
+  \begin{isabelle}\ \ \ \ \ %%%
+  \begin{tabular}{@ {}l}
+  @{term READ}, @{term WRITE}, @{term EXECUTE}, @{term "CHANGE_OWNER"}, 
+  @{term CREATE}, @{term SEND}, @{term RECEIVE} and @{term DELETE}
+  \end{tabular}
+  \end{isabelle}
+
+  In the RC-Model, roles group users according to tasks they need to 
+  accomplish. Users have a default role specified by the policy,
+  which is the role they start with whenever they log into the system.
+  A process contains the information about its owner (a user), its
+  role and its type, whereby a type in the RC-Model allows system
+  administrators to group resources according to a common criteria.
+  Such detailed information is needed in the RC-Model, for example, in
+  order to allow a process to change its ownership. For this the
+  RC-Model checks the role of the process and its type: if the access
+  control policy states that the role has @{term CHANGE_OWNER} access mode for
+  processes of that type, then the process is permitted to assume a
+  new owner.
+
+  Files in the RC-Model contain the information about their types.  A
+  policy then specifies whether a process with a given role can access
+  a file under a certain access mode. Files, however, also
+  include in the RC-Model information about roles. This information is
+  used when a process is permitted to execute a file. By doing so it
+  might change its role. This is often used in the context of
+  web-servers when a cgi-script is uploaded and then executed by the
+  server.  The resulting process should have much more restricted
+  access permissions.  This kind of behaviour when executing a file can
+  be specified in an RC-policy in several ways: first, the role of the
+  process does not change when executing a file; second, the process
+  takes on the role specified with the file; or third, use the role of
+  the owner, who currently owns this process. The RC-Model also makes
+  assumptions on how types can change.  For example for files and IPCs
+  the type can never change once they are created. But processes can
+  change their types according to the roles they have.
+  
+  As can be seen, the information contained in a policy in the
+  RC-Model can be rather complex: Roles and types, for example, are
+  policy-dependent, meaning each policy needs to define a set of roles and a
+  set of types. Apart from recording for each role the information
+  which type of resource it can access and under which access-mode, it
+  also needs to include a role compatibility set. This set specifies how one
+  role can change into another role. Moreover it needs to include default
+  information for cases when new processes or files are created.
+  For example, when a process clones itself, the type of the new
+  process is determined as follows: the policy might specify a default
+  type whenever a process with a certain role is cloned, or the policy
+  might specify that the cloned process inherits the type of the
+  parent process.
+  
+  Ott implemented the RC-Model on top of Linux, but only specified it
+  as a set of informal rules, partially given as logic formulas,
+  partially given as rules in ``English''. Unfortunately, some
+  presentations about the RC-Model give conflicting definitions for
+  some concepts---for example when defining the semantics of the special role
+  ``inherit parent''. In \cite{ottrc} it means inherit the initial role
+  of the parent directory, but in \cite{ottweb} it means inherit
+  the role of the parent process.  In our formalisation we mainly follow the
+  version given in \cite{ottrc}. In the next section we give a mechanised
+  model of the system calls on which the RC-Model is implemented.  
+*}
+
+
+
+section {* Dynamic Model of System Calls *}
+
+text {*
+  Central to the RC-Model are processes, since they initiate any action
+  involving resources and access control. We use natural numbers to stand for process IDs,
+  but do not model the fact that the number of processes in any practical 
+  system is limited. Similarly, IPCs and users are represented by natural 
+  numbers. The thirteen actions a process can perform are represented by
+  the following datatype of \emph{events}
+  
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{
+  \begin{tabular}{r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l@ {\hspace{3mm}}l@ 
+                     {\hspace{1.5mm}}l@ {\hspace{3mm}}l@ {\hspace{1.5mm}}l@ 
+                     {\hspace{3mm}}l@ {\hspace{1.5mm}}l}
+  event 
+  & @{text "::="} & @{term "CreateFile p f"} & @{text "|"} & @{term "ReadFile p f"} & @{text "|"} & @{term "Send p i"}   & @{text "|"} & @{term "Kill p p'"}  \\
+  & @{text "|"} & @{term "WriteFile p f"}    & @{text "|"} & @{term "Execute p f"} & @{text "|"} & @{term "Recv p i"}\\
+  & @{text "|"} & @{term "DeleteFile p f"}   & @{text "|"} & @{term "Clone p p'"} & @{text "|"} & @{term "CreateIPC p i"}  \\
+  & @{text "|"} & @{term "ChangeOwner p u"}  & @{text "|"} & @{term "ChangeRole p r"} & @{text "|"} & @{term "DeleteIPC p i"}\\
+  \end{tabular}}
+  \end{isabelle}
+
+  \noindent
+  with the idea that for example in @{term Clone} a process @{term p} is cloned
+  and the new process has the ID @{term "p'"}; with @{term Kill} the
+  intention is that the process @{term p} kills another process with
+  ID @{term p'}. We will later give the definition what the role
+  @{term r} can stand for in the constructor @{term ChangeRole}
+  (namely \emph{normal roles} only). As is custom in Unix, there is no
+  difference between a directory and a file. The files @{term f} in
+  the definition above are simply lists of strings.  For example, the
+  file @{text "/usr/bin/make"} is represented by the list @{text
+  "[make, bin, usr]"} and the @{text root}-directory is the @{text
+  Nil}-list. Following the presentation in \cite{ottrc}, our model of
+  IPCs is rather simple-minded: we only have events for creation and deletion Of IPCs,
+  as well as sending and receiving messages.
+
+  Events essentially transform one state of the system into
+  another. The system starts with an initial state determining which 
+  processes, files and IPCs are active at the start of the system. We assume the
+  users of the system are fixed in the initial state; we also assume
+  that the policy does not change while the system is running. We have 
+  three sets, namely
+  @{term init_processes},
+  @{term init_files} and
+  @{term init_ipcs}
+  specifying the processes, files and IPCs present in the initial state.
+  We will often use the abbreviation
+
+  \begin{center}
+  @{thm (lhs) Init_def} @{text "\<equiv>"} 
+  @{thm (rhs) Init_rhs(1)[where f=obj]} @{text "\<or>"}
+  @{thm (rhs) Init_rhs(2)[where p=obj]} @{text "\<or>"}
+  @{thm (rhs) Init_rhs(3)[where i=obj]}
+  \end{center}
+
+  \noindent
+  There are some assumptions we make about the files present in the initial state: we always
+  require that the @{text "root"}-directory @{term "[]"} is part of the initial state
+  and for every file in the initial state (excluding @{term "[]"}) we require that its 
+  parent is also part of the 
+  initial state.
+  After the initial state, the next states are determined
+  by a list of events, called the \emph{trace}.  We need to define
+  functions that allow us to make some observations about traces. One 
+  such function is called @{term "current_procs"} and
+  calculates the set of ``alive'' processes in a state:
+
+  %initial state:
+  %We make assumptions about the initial state, they're:
+  %1. there exists a set of processes, files, IPCs and users already in the initial state,
+  %users are not changed in system's running, we regards users adding and deleting a
+  %administration task, not the issue for our policy checker; 
+  %2. every object in the initial state have got already roles/types/owner ... information assigned;
+  %3. all the policy information are already preloaded in the initial state, including:
+  %a compatible type table, @{term compatible};
+  %a mapping function from a role to its compatible role set, @{term comproles};
+  %every role's default values is pre-set, e.g. default process create type and 
+  %and default file/directory create type.
+
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{\begin{tabular}{l@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
+  @{thm (lhs) current_procs.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) current_procs.simps(1)}\\
+  @{thm (lhs) current_procs.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) current_procs.simps(2)}\\
+  @{thm (lhs) current_procs.simps(3)} & @{text "\<equiv>"} & @{thm (rhs) current_procs.simps(3)}\\
+  @{term "current_procs (DUMMY#s)"} & @{text "\<equiv>"} & @{term "current_procs s"}
+  \end{tabular}}
+  \end{isabelle}
+    
+  \noindent
+  The first clause states that in the empty trace, that is initial
+  state, the processes are given by @{text "init_processes"}. The 
+  events for cloning a process, respectively killing a process, update this
+  set of processes appropriately. Otherwise the set of live
+  processes is unchanged. We have similar functions for alive files and 
+  IPCs, called @{term "current_files"} and @{term "current_ipcs"}. 
+
+  We can use these function in order to formally model which events are
+  \emph{admissible} by the operating system in each state. We show just three 
+  rules that give the gist of this definition. First the rule for changing 
+  an owner of a process:
+
+  \begin{center}
+  @{thm[mode=Rule] osgrant6}
+  \end{center}
+
+  \noindent
+  We require that the process @{text p} is alive in the state @{text s}
+  (first premise) and that the new owner is a user that existed in the initial state
+  (second premise).
+  Next the rule for creating a new file:
+
+  \begin{center}
+  @{thm[mode=Rule] osgrant2}
+  \end{center}
+  
+  \noindent
+  It states that
+  a file @{text f} can be created by a process @{text p} being alive in the state @{text s},
+  the new file does not exist already in this state and there exists
+  a parent file @{text "pf"} for the new file. The parent file is just
+  the tail of the list representing @{text f}. % if it exists 
+  %(@{text "Some"}-case) or @{text None} if it does not. 
+  Finally, the rule for cloning a process:
+
+  \begin{center}
+  @{thm[mode=Rule] osgrant10}
+  \end{center}
+
+  \noindent
+  Clearly the operating system should only allow to clone a process @{text p} if the 
+  process is currently alive. The cloned process will get the process
+  ID generated by the function @{term new_proc}. This process ID should
+  not already exist. Therefore we define @{term new_proc} as
+
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{\begin{tabular}{l@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
+  @{term "new_proc s"} & @{text "\<equiv>"} & @{term "Max (current_procs s) + 1"}
+  \end{tabular}}
+  \end{isabelle}
+
+  \noindent
+  namely the highest ID currently in existence increased by one. The 
+  admissibility rules for the other events impose similar conditions.
+
+  However, the admissibility check by the operating system is only one
+  ``side'' of the constraints the RC-Model imposes. We also need to
+  model the constraints of the access policy. For this we introduce
+  separate @{text granted}-rules involving the sets @{text
+  permissions} and @{text "compatible r"}: the former contains triples
+  describing access control rules; the latter specifies for each role @{text r} 
+  which roles are compatible with @{text r}. These sets are used in the
+  RC-Model when a process having a role @{text r} takes on a new role
+  @{text r'}. For example, a login-process might belong to root;
+  once the user logs in, however, the role of the process should change to
+  the user's default role. The corresponding @{text "granted"}-rule is
+  as follows
+
+  \begin{center}
+  @{thm[mode=Rule] rcgrant7}
+  \end{center}
+  
+  \noindent
+  where we check whether the process @{text p} has currently role @{text r} and
+  whether the RC-policy states that @{text r'} is in the role compatibility
+  set of @{text r}. 
+
+  The complication in the RC-Model arises from the
+  way how the current role of a process in a state @{text s} is
+  calculated---represented by the predicate @{term is_current_role} in our formalisation.  
+  For defining this predicate we need to trace the role of a process from
+  the initial state to the current state. In the
+  initial state all processes have the role given by the function
+  @{term "init_current_role"}.  If a @{term Clone} event happens then
+  the new process will inherit the role from the parent
+  process. Similarly, if a @{term ChangeRole} event happens, then 
+  as seen in the rule above we just change the role accordingly. More interesting 
+  is an @{term Execute} event in the RC-Model. For this event we have 
+  to check the role attached to the file to be executed. 
+  There are a number of cases: If the role of the file is a 
+  \emph{normal} role, then the process will just take on this role 
+  when executing the file (this is like the setuid mechanism in Unix). But 
+  there are also four \emph{special} roles in the RC-Model:
+  @{term "InheritProcessRole"}, @{term "InheritUserRole"}, 
+   @{term "InheritParentRole"} and @{term
+  InheritUpMixed}. For example, if a file to be executed has
+  @{term "InheritProcessRole"} attached to it, then the process 
+  that executes this file keeps its role regardless of the information 
+  attached to the file. In this way programs can be can quarantined;
+  @{term "InheritUserRole"} can be used for login shells
+  to make sure they run with the user's default role.
+  The purpose of the other special roles is to determine the
+  role of a process according to the directory in which the
+  files are stored.
+
+  Having the notion of current role in place, we can define the
+  granted rule for the @{term Execute}-event: Suppose a process @{term
+  p} wants to execute a file @{text f}. The RC-Model first fetches the
+  role @{text r} of this process (in the current state @{text s}) and
+  the type @{text t} of the file. It then checks if the tuple @{term
+  "(r, t, EXECUTE)"} is part of the policy, that is in our
+  formalisation being an element in the set @{term compatible}. The
+  corresponding rule is as follows
+
+  \begin{center}
+  @{thm[mode=Rule] rcgrant4}
+  \end{center}
+
+  \noindent
+  The next @{text granted}-rule concerns the @{term CreateFile} event.
+  If this event occurs, then we have two rules in our RC-Model
+  depending on how the type of the created file is derived. If the type is inherited
+  from the parent directory @{text pf}, then the @{term granted}-rule is as follows:
+
+  \begin{center}
+  @{thm[mode=Rule] rcgrant1} 
+  \end{center}
+
+  \noindent
+  We check whether @{term pf} is the parent file (directory) of @{text f} and check
+  whether the type of @{term pf} is @{term t}. We also need to fetch the 
+  the role @{text r} of the process that seeks to get permission for creating 
+  the file. If the default type of this role @{text r} states that the 
+  type of the newly created file will be inherited from the parent file
+  type, then we only need to check that the policy states that @{text r}
+  has permission to write into the directory @{text pf}.
+
+  The situation is different if the default type of role @{text r} is
+  some \emph{normal} type, like text-file or executable. In such cases we want
+  that the process creates some predetermined type of files. Therefore in the
+  rule we have to check whether the role is allowed to create a file of that
+  type, and also check whether the role is allowed to write any new
+  file into the parent file (directory). The corresponding rule is
+  as follows.
+
+  \begin{center}
+  @{thm[mode=Rule] rcgrant1'}
+  \end{center}
+  
+  \noindent
+  Interestingly, the type-information in the RC-model is also used for
+  processes, for example when they need to change their owner. For
+  this we have the rule
+
+  \begin{center}
+  @{thm[mode=Rule] rcgrant_CHO}
+  \end{center}
+
+  \noindent
+  whereby we have to obtain both the role and type of the process @{term p}, and then check
+  whether the policy allows a @{term ChangeOwner}-event for that role and type.
+
+  Overall we have 13 rules for the admissibility check by the operating system and
+  14 rules for the granted check by the RC-Model.  
+  They are used to characterise when an event @{text e} is \emph{valid} to 
+  occur in a state @{text s}. This can be inductively defined as the set of valid
+  states.
+
+  \begin{center}
+  \begin{tabular}{@ {}c@ {}}
+  \mbox{@{thm [mode=Axiom] valid.intros(1)}}\hspace{5mm}
+  \mbox{@{thm [mode=Rule] valid.intros(2)}} 
+  \end{tabular}
+  \end{center}  
+
+  The novel notion we introduce in this paper is the @{text tainted} 
+  relation. It characterises how a system can become infected when 
+  a file in the system contains, for example, a virus. We assume
+  that the initial state contains some tainted
+  objects (we call them @{term "seeds"}). Therefore in the initial state @{term "[]"}
+  an object is tainted, if it is an element in @{text "seeds"}.
+
+  \begin{center}
+  \mbox{@{thm [mode=Rule] tainted.intros(1)}}
+  \end{center}
+
+  \noindent
+  Let us first assume such a tainted object is a file @{text f}.
+  If a process reads or executes a tainted file, then this process becomes
+  tainted (in the state where the corresponding event occurs).
+
+  \begin{center}
+  \mbox{@{thm [mode=Rule] tainted.intros(3)}}\hspace{3mm}
+  \mbox{@{thm [mode=Rule] tainted.intros(6)}} 
+  \end{center}
+
+  \noindent
+  We have a similar rule for a tainted IPC, namely
+
+  \begin{center}
+  \mbox{@{thm [mode=Rule] tainted.intros(9)}}
+  \end{center}
+
+  \noindent
+  which means if we receive anything from a tainted IPC, then
+  the process becomes tainted. A process is also tainted 
+  when it is a produced by a @{text Clone}-event.
+
+  \begin{center}
+  \mbox{@{thm [mode=Rule] tainted.intros(2)}}
+  \end{center}
+
+  \noindent
+  However, the tainting relationship must also work in the 
+  ``other'' direction, meaning if a process is tainted, then
+  every file that is written or created will be tainted. 
+  This is captured by the four rules:
+
+  \begin{center}
+  \begin{tabular}{c}
+  \mbox{@{thm [mode=Rule] tainted.intros(4)}} \hspace{3mm}
+  \mbox{@{thm [mode=Rule] tainted.intros(7)}} \medskip\\
+  \mbox{@{thm [mode=Rule] tainted.intros(5)}} \hspace{3mm}
+  \mbox{@{thm [mode=Rule] tainted.intros(8)}} 
+  \end{tabular}
+  \end{center}
+  
+  \noindent
+  Finally, we have three rules that state whenever an object is tainted
+  in a state @{text s}, then it will be still tainted in the 
+  next state @{term "e#s"}, provided the object is still \emph{alive}
+  in that state. We have such a rule for each kind of objects, for
+  example for files the rule is:
+
+  \begin{center}
+  \mbox{@{thm [mode=Rule] tainted_10}} 
+  \end{center}
+
+  \noindent
+  Similarly for alive processes and IPCs (then respectively with premises 
+  @{term "p \<in> current_procs (e#s)"} and @{term "i \<in> current_ipcs (e#s)"}). 
+  When an object present in the initial state can be tainted in 
+  \emph{some} state (system run), we say it is @{text "taintable"}:
+
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{\begin{tabular}{lcl}
+  @{thm (lhs) taintable_def} & @{text "\<equiv>"} & @{term "init obj"} @{text "\<and>"} @{thm (rhs) taintable_def}
+  \end{tabular}}
+  \end{isabelle}
+
+  Before we can describe our static check deciding when a file is taintable, we
+  need to describe the notions @{term deleted} and @{term undeletable}
+  for objects. The former characterises whether there is an event that deletes  
+  these objects (files, processes or IPCs). For this we have the following 
+  four rules:
+
+  \begin{center}
+  \begin{tabular}{@ {}c@ {\hspace{10mm}}c@ {}}
+  \begin{tabular}{c}
+  @{thm [mode=Axiom] dels(1)}\\[-2mm]
+  @{thm [mode=Axiom] dels(2)}\\[-2mm]
+  @{thm [mode=Axiom] dels(3)}
+  \end{tabular} &
+  @{thm [mode=Rule] dels(4)}
+  \end{tabular}
+  \end{center}
+
+
+  \noindent
+  Note that an object cannot be deleted in the initial state @{text
+  "[]"}.  An object is then said to be @{text "undeletable"} provided
+  it did exist in the initial state and there does not exists a valid
+  state in which the object is deleted:
+  
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{\begin{tabular}{l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {}}
+  @{thm (lhs) undeletable_def} & @{text "\<equiv>"} & %%@{thm (rhs) undeletable_def}
+  @{term "init obj \<and> \<not>(\<exists> s. (valid s \<and> deleted obj s))"}
+  \end{tabular}}
+  \end{isabelle}
+
+  \noindent
+  The point of this definition is that our static taintable check will only be
+  complete for undeletable objects. But these are
+  the ones system administrators are typically interested in (for
+  example system files).  It should be clear, however, that we cannot
+  hope for a meaningful check by just trying out all possible
+  valid states in our dynamic model. The reason is that there are
+  potentially infinitely many of them and therefore the search space would be
+  infinite.  For example staring from an
+  initial state containing a process @{text p} and a file @{text pf}, 
+  we can create files @{text "f\<^isub>1"}, @{text "f\<^isub>2"}, @{text "..."} 
+  via @{text "CreateFile"}-events. This can be pictured roughly as follows:
+
+  \begin{center}
+  \begin{tabular}[t]{c@ {\hspace{-8mm}}c@ {\hspace{-8mm}}c@ {\hspace{-8mm}}c@ {\hspace{-8mm}}cc}
+  \begin{tabular}[t]{c}
+  Initial state:\\
+  @{term "{p, pf}"}
+  \end{tabular} &
+  \begin{tabular}[t]{c}
+  \\
+  @{text "\<Longrightarrow>"}\\[2mm]
+  {\small@{term "CreateFile p (f\<^isub>1#pf)"}}
+  \end{tabular}
+  &
+  \begin{tabular}[t]{c}
+  \\
+  @{term "{p, pf, f\<^isub>1#pf}"}
+  \end{tabular} 
+  &
+  \begin{tabular}[t]{c}
+  \\
+  @{text "\<Longrightarrow>"}\\[2mm]
+  {\small@{term "CreateFile p (f\<^isub>2#f\<^isub>1#pf)"}}
+  \end{tabular}
+  &
+  \begin{tabular}[t]{c}
+  \\
+  @{term "{p, pf, f\<^isub>1#pf, f\<^isub>2#f\<^isub>1#pf}"}
+  \end{tabular} &
+  \begin{tabular}[t]{c}
+  \\
+  @{text "..."}\\
+  \end{tabular}
+  \end{tabular}
+  \end{center}
+
+  \noindent
+  Instead, the idea of our static check is to use
+  the policies of the RC-model for generating an answer, since they
+  provide always a finite ``description of the system''. As we 
+  will see in the next section, this needs some care, however.
+*}
+
+section {* Our Static Check *}
+
+text {*
+  Assume there is a tainted file in the system and suppose we face the
+  problem of finding out whether this file can affect other files,
+  IPCs or processes? One idea is to work on the level of policies only, and
+  check which operations are permitted by the role and type of this
+  file. Then one builds the ``transitive closure'' of this information
+  and checks for example whether the role @{text root} has become
+  affected, in which case the whole system is compromised.  This is indeed the solution investigated
+  in~\cite{guttman2005verifying} in the context of information flow
+  and SELinux.
+
+  Unfortunately, restricting the calculations to only use policies is
+  too simplistic for obtaining a check that is sound and complete---it
+  over-approximates the dynamic tainted relation defined in the previous
+  section. To see the problem consider
+  the case where the tainted file has, say, the type @{text bin}. If
+  the RC-policy contains a role @{text r} that can both read and write
+  @{text bin}-files, we would conclude that all @{text bin}-files can potentially 
+  be tainted. That
+  is indeed the case, \emph{if} there is a process having this role @{text
+  r} running in the system. But if there is \emph{not}, then the
+  tainted file cannot ``spread''.  A similar problem arises in case there
+  are two processes having the same role @{text r},  and this role is
+  restricted to read files only.  Now if one of the processes is tainted, then
+  the simple check involving only policies would incorrectly infer
+  that all processes involving that role are tainted. But since the
+  policy for @{text r} is restricted to be read-only, there is in fact
+  no danger that both processes can become tainted.
+
+  The main idea of our sound and complete check is to find a ``middle'' ground between
+  the potentially infinite dynamic model and the too coarse
+  information contained in the RC-policies. Our solution is to
+  define a ``static'' version of the tainted relation, called @{term
+  "tainted_s"}, that records relatively precisely the information
+  about the initial state of the system (the one in which an object
+  might be a @{term seed} and therefore tainted). However,
+  we are less precise about the objects created in every subsequent
+  state. The result is that we can avoid the potential infinity of
+  the dynamic model. 
+  For the @{term tainted_s}-relation we will consider the following 
+  three kinds of \emph{items} recording the information we need about
+  processes, files and IPCs, respectively:
+
+  \begin{center}
+  \begin{tabular}{l@ {\hspace{5mm}}l}
+  & Recorded information:\smallskip\\
+  Processes: & @{term "SProc (r, dr, t, u) po"}\\
+  Files: & @{term "SFile (t, a) fo"}\\
+  IPCs: & @{term "SIPC (t) io"}
+  \end{tabular}
+  \end{center}
+  
+  \noindent
+  For a process we record its role @{text r}, its default role @{text dr} (used to determine
+  the role when executing a file or changing the owner of a process), its type @{text t} 
+  and its owner @{text u}. For a file we record
+  just the type @{text t} and its @{term "source_dir"} @{text a} (we define this
+  notion shortly). For IPCs we only record its type @{text t}. Note the superscripts
+  @{text po}, @{text fo} and @{text io} in each item. They are optional arguments and depend on
+  whether the corresponding object is present in the initial state or not.
+  If it \emph{is}, then for processes and IPCs we will record @{term "Some(id)"},
+  where @{text id} is the natural number that uniquely identifies a process or IPC;
+  for files we just record their path @{term "Some(f)"}. If the object is
+  \emph{not} present in the initial state, that is newly created, then we just have
+  @{term None} as superscript. 
+  Let us illustrate the different superscripts with the following example
+  where the initial state contains a process @{term p} and a file (directory)
+  @{term pf}. Then this
+  process creates a file via a @{term "CreateFile"}-event and after that reads
+  the created file via a @{term Read}-event:
+
+  \begin{center}
+  \begin{tabular}[t]{ccccc}
+  \begin{tabular}[t]{c}
+  Initial state:\\
+  @{term "{p, pf}"}
+  \end{tabular} &
+  \begin{tabular}[t]{c}
+  \\
+  @{text "\<Longrightarrow>"}\\
+  {\small@{term "CreateFile p (f#pf)"}}
+  \end{tabular}
+  &
+  \begin{tabular}[t]{c}
+  \\
+  @{term "{p, pf, f#pf}"}
+  \end{tabular} 
+  &
+  \begin{tabular}[t]{c}
+  \\
+  @{text "\<Longrightarrow>"}\\
+  {\small@{term "ReadFile p (f#pf)"}}
+  \end{tabular}
+  &
+  \begin{tabular}[t]{c}
+  \\
+  @{term "{p, pf, f#pf}"}
+  \end{tabular} 
+  \end{tabular}
+  \end{center}
+
+  \noindent
+  For the two objects in the initial state our static check records
+  the information @{term "SProc (r, dr, t, u) (Some(p))"} and @{term
+  "SFile (t', a) (Some(pf))"} (assuming @{text "r"}, @{text t} and so
+  on are the corresponding roles, types etc).  In both cases we have
+  the superscript @{text "Some(...)"}  since they are objects present
+  in the initial state. For the file @{term "f#pf"} created by the
+  @{term "CreateFile"}-event, we record @{term "SFile (t', a')
+  (None)"}, since it is a newly created file.  The @{text
+  "ReadFile"}-event does not change the set of objects, therefore no
+  new information needs to be recorded. The problem we are avoiding
+  with this setup of recording the precise information for the initial
+  state is where two processes have the same role and type
+  information, but only one is tainted in the initial state, but the
+  other is not. The recorded unique process IDs allows us to
+  distinguish between both processes.  For all newly created objects,
+  on the other hand, we do not care.  This is crucial, because
+  otherwise exploring all possible ``reachable'' objects can lead to
+  the potential infinity like in the dynamic model.
+
+  An @{term source_dir} for a file is the ``nearest'' directory that
+  is present in the initial state and has not been deleted in a state
+  @{text s}. Its definition is the recursive function
+  
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{\begin{tabular}{lcl}
+  @{thm (lhs) source_dir.simps(1)} & @{text "\<equiv>"} \;\; &
+  @{text "if"} @{text "\<not> deleted [] s"} @{text "then"} @{term "Some []"} @{text "else"} @{term "None"}\\
+  @{thm (lhs) source_dir.simps(2)} & @{text "\<equiv>"} & 
+  @{text "if"} @{term "(f#pf) \<in> init_files \<and> \<not>(deleted (File (f#pf)) s)"}\\
+  & & @{text "then"} @{term "Some (f#pf)"} @{text "else"} @{term "source_dir s pf"}\\
+  \end{tabular}}
+  \end{isabelle}
+  
+  \noindent
+  generating an optional value.
+  The first clause states that the anchor of the @{text
+  root}-directory is always its own anchor unless it has been
+  deleted. If a file is present in the initial state and not deleted
+  in @{text s}, then it is also its own anchor, otherwise the anchor
+  will be the anchor of the parent directory.  For example if we have
+  a directory @{text pf} in the initial state, then its anchor is @{text "Some pf"}
+  (assuming it is not deleted). If we create a new file in this directory,
+  say @{term "f#pf"}, then its anchor will also be @{text "Some pf"}.
+  The purpose of @{term source_dir} is to determine the
+  role information when a file is executed, because the role of the
+  corresponding process, according to the RC-model, is determined by the role information of the
+  anchor of the file to be executed.
+
+  There is one last problem we have to solve before we can give the
+  rules of our @{term "tainted_s"}-check. Suppose an RC-policy 
+  includes the rule @{text "(r, foo, Write) \<in> permissions"}, that is
+  a process of role @{text "r"} is allowed to write files of type @{text "foo"}.
+  If there is a tainted process with this role, we would conclude that
+  also every file of that type can potentially become tainted. However, that
+  is not the case if the initial state does not contain any file
+  with type @{text foo} and the RC-policy does not allow the
+  creation of such files, that is does not contain an access rule
+  @{text "(r, foo, Create) \<in> permissions"}. In a sense the original
+  @{text "(r, foo, Write)"} is ``useless'' and should not contribute
+  to the relation characterising the objects that are tainted.
+  To exclude such ``useless'' access rules, we define
+  a relation @{term "all_sobjs"} restricting our search space 
+  to only configurations that correspond to states in our dynamic model.
+  We first have a rule for reachable items of the form  @{text "F(t, f)\<^bsup>Some f\<^esup>"}
+  where the file @{text f} with type @{text t} is present in 
+  the initial state. 
+
+  \begin{center}
+  @{thm [mode=Rule] af_init'}
+  \end{center}
+
+  \noindent
+  We have similar reachability rules for processes and IPCs that are part of the 
+  initial state. Next is the reachability rule in case a file is created
+
+  \begin{center}
+  @{thm [mode=Rule] af_cfd[where sd=a and sf="fo" and sp="po" and fr="dr"]}
+  \end{center}
+
+  \noindent
+  where we require that we have a reachable parent directory, recorded
+  as @{text "F(t, a)\<^bsup>fo\<^esup>"}, and also a
+  process that can create the file, recorded as @{text "P(r, dr, pt,
+  u)\<^bsup>po\<^esup>"}. As can be seen, we also require that we have both @{text "(r, t,
+  Write)"} and \mbox{@{text "(r, t', Create)"}} in the @{text permissions} set 
+  for this rule to apply. If we did \emph{not} impose this requirement
+  about the RC-policy, then there would be no way to create a file
+  with @{term "NormalFileType t'"} according to our ``dynamic'' model. 
+  However in case we want to create a 
+  file of type @{term InheritPatentType}, then we only need the access-rule
+  @{text "(r, t, Write)"}: 
+  
+  \begin{center}
+  @{thm [mode=Rule] af_cfd'[where sd=a and sf="fo" and sp="po" and fr="dr"]}
+  \end{center}
+
+  \noindent
+  We also have reachability rules for processes executing files, and
+  for changing their roles and owners, for example
+
+  \begin{center}
+  @{thm [mode=Rule] ap_crole[where sp="po" and fr="dr"]}
+  \end{center}
+
+  \noindent
+  which states that when we have a process with role @{text r}, and the role
+  @{text "r'"} is in the corresponding role-compatibility set, then also
+  a process with role @{text "r'"} is reachable.
+
+  The crucial difference between between the ``dynamic'' notion of validity 
+  and the ``static'' notion of @{term "all_sobjs"}
+  is that there can be infinitely many valid states, but assuming the initial
+  state contains only finitely many objects, then also @{term "all_sobjs"} will 
+  be finite. To see the difference, consider the infinite ``chain'' of events 
+  just cloning a process @{text "p\<^isub>0"}:
+
+  \begin{center}
+  \begin{tabular}[t]{c@ {\hspace{-2mm}}c@ {\hspace{-2mm}}c@ {\hspace{-2mm}}c@ {\hspace{-2mm}}cc}
+  \begin{tabular}[t]{c}
+  Initial state:\\
+  @{term "{p\<^isub>0}"}
+  \end{tabular} &
+  \begin{tabular}[t]{c}
+  \\
+  @{text "\<Longrightarrow>"}\\[2mm]
+  {\small@{term "Clone p\<^isub>0 p\<^isub>1"}}
+  \end{tabular}
+  &
+  \begin{tabular}[t]{c}
+  \\
+  @{term "{p\<^isub>0, p\<^isub>1}"}
+  \end{tabular} 
+  &
+  \begin{tabular}[t]{c}
+  \\
+  @{text "\<Longrightarrow>"}\\[2mm]
+  {\small@{term "Clone p\<^isub>0 p\<^isub>2"}}
+  \end{tabular}
+  &
+  \begin{tabular}[t]{c}
+  \\
+  @{term "{p\<^isub>0, p\<^isub>1, p\<^isub>2}"}
+  \end{tabular} &
+  \begin{tabular}[t]{c}
+  \\
+  @{text "..."}\\
+  \end{tabular}
+  \end{tabular}
+  \end{center}
+
+  \noindent
+  The corresponding reachable objects are
+
+  \begin{center}
+  \begin{tabular}[t]{cccc}
+  \begin{tabular}[t]{c}
+  @{text "{P(r, dr, t, u)\<^bsup>Some (p\<^isub>0)\<^esup>}"}
+  \end{tabular} &
+  \begin{tabular}[t]{c}
+  @{text "\<Longrightarrow>"}
+  \end{tabular}
+  &
+  \begin{tabular}[t]{c}
+  @{text "{P(r, dr, t, u)\<^bsup>Some (p\<^isub>0)\<^esup>, P(r, dr, t, u)\<^bsup>None\<^esup>}"}
+  \end{tabular} 
+  \end{tabular}
+  \end{center}
+
+  \noindent
+  where no further progress can be made because the information
+  recorded about @{text "p\<^isub>2"}, @{text "p\<^isub>3"} and so on is just the same
+  as for @{text "p\<^isub>1"}, namely @{text "P(r, dr, t, u)\<^bsup>None\<^esup>"}.  Indeed we
+  can prove the lemma:
+
+  \begin{lemma}\label{finite}
+  If @{text "finite init"}, then @{term "finite all_sobjs"}.
+  \end{lemma}
+
+  \noindent
+  This fact of @{term all_sobjs} being finite enables us to design a
+  decidable tainted-check. For this we introduce inductive rules defining the
+  set @{term "tainted_s"}. Like in the ``dynamic'' version of tainted,
+  if an object is element of @{text seeds}, then it is @{term "tainted_s"}.
+
+  \begin{center}
+  @{thm [mode=Rule] ts_init}
+  \end{center}
+
+  \noindent
+  The function @{text "\<lbrakk>_\<rbrakk>"} extracts the static information from an object.
+  For example for a process it extracts the role, default role, type and
+  user; for a file the type and the anchor. If a process in tainted and creates
+  a file with a normal type @{text "t'"} then also the created file
+  is tainted. The corresponding rule is
+
+  \begin{center}
+  @{thm [mode=Rule] ts_cfd[where sd=a and sf="fo" and sp="po" and fr="dr"]}
+  \end{center}
+
+  \noindent
+  If a tainted process  creates a file that inherits the type of the directory,
+  then the file will also be tainted:
+
+  \begin{center}
+  @{thm [mode=Rule] ts_cfd'[where sd=a and sf="fo" and sp="po" and fr="dr"]}
+  \end{center}
+
+  \noindent
+  If a tainted process changes its role, then also with this changed role
+  it will be tainted:
+
+  \begin{center}
+  @{thm [mode=Rule] ts_crole[where pt=t and sp="po" and fr="dr"]}
+  \end{center}
+
+  \noindent 
+  Similarly when a process changes its owner. If a file is tainted, and
+  a process has read-permission to that type of files, then the 
+  process becomes tainted. The corresponding rule is
+
+  \begin{center}
+  @{thm [mode=Rule] ts_read[where sd=a and sf="fo" and sp="po" and fr="dr"]}
+  \end{center}
+  
+  \noindent
+  If a process is tainted and it has write-permission for files of type @{text t}, 
+  then these files will be tainted:
+
+  \begin{center}
+  @{thm [mode=Rule] ts_write[where sd=a and sf="fo" and sp="po" and fr="dr"]}
+  \end{center}
+
+  \noindent
+  We omit the remaining rules for executing a file, cloning a process and 
+  rules involving IPCs, which are similar. A simple consequence of our definitions 
+  is that every tainted object is also reachable:
+
+  \begin{lemma}
+  @{text "tainted\<^isup>s \<subseteq> reachable\<^isup>s"}
+  \end{lemma}
+
+  \noindent
+  which in turn means that the set of @{term "tainted_s"} items is finite by Lemma~\ref{finite}.
+  
+  Returning to our original question about whether tainted objects can spread
+  in the system. To answer this question, we take these tainted objects as
+  seeds and calculate the set of items that are @{term "tainted_s"}. We proved this
+  set is finite and can be enumerated using the rules for @{term tainted_s}.
+  However, this set is about items, not about whether objects are tainted or not. 
+  Assuming an item in @{term tainted_s} arises from an object present in the initial
+  state, we have recorded enough information to translate items back into objects 
+  via the function @{text "|_|"}:
+
+  \begin{center}
+  \begin{tabular}{lcl}
+  @{text "|P(r, dr, t, u)\<^bsup>po\<^esup>|"} & @{text "\<equiv>"} & @{text po}\\
+  @{text "|F(t, a)\<^bsup>fo\<^esup>|"} & @{text "\<equiv>"} & @{text fo}\\
+  @{text "|I(t\<^bsup>\<^esup>)\<^bsup>io\<^esup>|"} & @{text "\<equiv>"} & @{text io}
+  \end{tabular}
+  \end{center}
+
+  \noindent
+  Using this function, we can define when an object is @{term taintable_s} in terms of
+  an item being @{term tainted_s}, namely
+
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{\begin{tabular}{lcl}
+  @{thm (lhs) taintable_s_def} & @{text "\<equiv>"} & @{text "\<exists>item. item \<in> tainted\<^isup>s \<and> |item| = Some obj"}
+  \end{tabular}}
+  \end{isabelle}
+
+  \noindent
+  Note that @{term taintable_s} is only about objects that are present in
+  the initial state, because for all other items @{text "|_|"} returns @{text None}. 
+
+
+  With these definitions in place, we can state our theorem about the soundness of our 
+  static @{term taintable_s}-check for objects.
+
+  \begin{theorem}[Soundness]
+  @{thm [mode=IfThen] static_sound}
+  \end{theorem}
+
+  \noindent
+  The proof of this theorem generates for every object that is ``flagged'' as 
+  @{term taintable_s} by our check, a sequence of events which shows how the
+  object can become tainted in the dynamic model.  We can also state a completeness
+  theorem for our @{term taintable_s}-check.
+
+  \begin{theorem}[Completeness]
+  @{thm [mode=IfThen] static_complete}  
+  \end{theorem}
+
+  \noindent
+  This completeness theorem however needs to be restricted to
+  undeletebale objects. The reason is that a tainted process can be
+  killed by another process, and after that can be ``recreated'' by a
+  cloning event from an untainted process---remember we have no control
+  over which process ID a process will be assigned with.  Clearly, in
+  this case the cloned process should be considered untainted, and
+  indeed our dynamic tainted relation is defined in this way. The
+  problem is that a static test cannot know about a process being
+  killed and then recreated. Therefore the static test will not be
+  able to ``detect'' the difference.  Therefore we solve this problem
+  by considering only objects that are present in the initial state
+  and cannot be deleted. By the latter we mean that the RC-policy
+  stipulates an object cannot be deleted (for example it has been created
+  by @{term root} in single-user mode, but in the everyday running
+  of the system the RC-policy forbids to delete an object belonging to
+  @{term root}). Like @{term "taintable_s"}, we also have a static check 
+  for when a file is undeletable according to an RC-policy.
+
+  This restriction to undeletable objects might be seen as a great
+  weakness of our result, but in practice this seems to cover the
+  interesting scenarios encountered by system administrators. They
+  want to know whether a virus-infected file introduced by a user can
+  affect the core system files. Our test allows the system
+  administrator to find this out provided the RC-policy makes the core
+  system files undeletable. We assume that this provisio is already part 
+  of best practice rule for running a system.
+
+  We envisage our test to be useful in two kind of situations: First, if
+  there was a break-in into a system, then, clearly, the system
+  administrator can find out whether the existing access policy was
+  strong enough to contain the break-in, or whether core system files
+  could have been affected.  In the first case, the system
+  administrator can just plug the hole and forget about the break-in;
+  in the other case the system administrator is wise to completely
+  reinstall the system.
+  Second, the system administrator can proactively check whether an
+  RC-policy is strong enough to withstand serious break-ins. To do so
+  one has to identify the set of ``core'' system files that the policy
+  should protect and mark every possible entry point for an attacker
+  as tainted (they are the seeds of the @{term "tainted_s"} relation).  
+  Then the test will reveal
+  whether the policy is strong enough or needs to be redesigned. For
+  this redesign, the sequence of events our check generates should be
+  informative.
+*}
+
+
+
+
+
+section {*Conclusion and Related Works*}
+
+
+text {*
+  We have presented the first completely formalised dynamic model of
+  the Role-Compa\-tibility Model. This is a framework, introduced by Ott
+  \cite{ottrc}, in which role-based access control policies
+  can be formulated and is used in practice, for example, for securing Apache 
+  servers. Previously, the RC-Model was presented as a
+  collection of rules partly given in ``English'', partly given as formulas.
+  During the formalisation we uncovered an inconsistency in the
+  semantics of the special role @{term "InheritParentRole"} in
+  the existing works about the RC-Model \cite{ottrc,ottweb}.  By proving
+  the soundness and completeness of our static @{term
+  "taintable_s"}-check, we have formally related the dynamic behaviour
+  of the operating system implementing access control and the static 
+  behaviour of the access policies of the RC-Model. The
+  crucial idea in our static check is to record precisely the
+  information available about the initial state (in which some resources might be
+  tainted), but be less precise
+  about the subsequent states. The former fact essentially gives us
+  the soundness of our check, while the latter results in a finite
+  search space. 
+  
+  The two most closely related works are by Archer et al and by Guttman et al
+  \cite{Archer03,guttman2005verifying}. The first describes a
+  formalisation of the dynamic behaviour of SELinux carried out in the
+  theorem prover PVS. However, they cannot use their formalisation in
+  order to prove any ``deep'' properties about access control rules
+  \cite[Page 167]{Archer03}.  The second analyses access control
+  policies in the context of information flow. Since this work 
+  is completely on the level of policies, it does
+  not lead to a sound and complete check for files being taintable (a dynamic notion
+  defined in terms of operations performed by the operating system).
+  While our results concern the RC-Model, we expect that they 
+  equally apply to the access control model of SELinux. In fact,
+  we expect that the formalisation is simpler for SELinux, since 
+  its rules governing roles are much simpler than in the RC-Model.
+  The definition of our admissibility rules can be copied verbatim for SELinux;
+  we would need to modify our granted rules and slightly adapt our
+  static check. We leave this as future work.
+
+
+  Our formalisation is carried out in the Isabelle/HOL theorem prover.
+  It uses Paulson's inductive method for
+  reasoning about sequences of events \cite{Paulson98}.
+  We have approximately 1000 lines of code for definitions and 6000 lines of
+  code for proofs. Our formalisation is available from the
+  Mercurial repository at \url{http://www.dcs.kcl.ac.uk/staff/urbanc/cgi-bin/repos.cgi/rc/}.\\[-12mm]
+
+
+%  0. Not Policy-Analysis: cause even policy is analysed correct, there is still a gap between it and policy application to the real Access Control system. Hence here Our dynamic model is bridging this gap.  Policy-Analysis "basic" based on "Information flow", but it is not enough: the static "write" right to a certain typed file do not mean a process having this right definitely can write the file, it has to pass a "particular" "Control Flow" to achieve the state of "There are this typed file and this righted process"! 
+%  1. Both Dynamic and Statical analysis, and proved link between two \\
+%  2. Tainting Relation Formalisation  \\
+%  3. Formalisation and Verification than Model Checking  \\
+%  4. Universal Checker of Policy  \\
+%  5. source of RC rules made more precise \\
+%  6. RC example of Webserver with CGIs (key notion: Program Based Roles)  \\
+%  7. RBAC is more Policy-lever(with HUGE companies, many stablised num of roles but frequently varifying num of users); RC is more Program Base Roles, set for system with a lot of program based default value, once pre-setted, it will remains after running. which is key to policy checker.
+
+%The distinct feature of RC is to deal with program based roles, such as server behaviour. 
+%This is in contrast to usual RSBAC models where roles are modeled around a hierachy, for
+%example in a company. 
+
+
+%In a word, what the manager need is that given the
+%initial state of the system, a policy checker that make sure the under the policy
+%he made, this important file cannot: 1. be deleted; 2. be tainted. 
+%Formally speaking, this policy-checker @{text "PC"} (a function that given the 
+%initial state of system, a policy and an object, it tells whether this object 
+%will be fully protected or not) should satisfy this criteria: 
+
+%  @{text "(PC init policy obj) \<and> (exists init obj) \<longrightarrow> \<not> taintable obj"}
+%If the @{text obj} exists in the initial-state, and @{text "PC"} justify the safety
+%of this object under @{text "policy"}, this object should not be @{text taintable}.
+%We call this criteria the \emph{completeness} of @{text "PC"}. 
+%And there is the \emph{soundness} criteria of @{text "PC"} too, otherwise a "NO-to-ALL"
+%answer always satisfy the \emph{completeness}. \emph{soundness} formally is:
+%  @{text "PC init policy obj \<longrightarrow> taintable obj"}
+
+%This policy-checker should satisfy other properties:
+% 1. fully statical, that means this policy-checker should not rely on the system 
+%running information, like new created files/process, and most importantly the 
+%trace of system running. 
+% 2. decidable, that means this policy-checker should always terminate.
+
+
+%   The purpose of an operating system is to provide a common
+%  interface for accessing resources. This interface is typically 
+%  realised as a collection of system calls, for example for reading 
+%  and writing files, forking of processes, or sending and receiving 
+%  messages. Role based access control is one approach for 
+%  restricting access to such system calls: if a user has
+%  suffient rights, then a system call can be performed. 
+
+%  a user might have
+%  one or more roles and acces is granted if the role has sufficent
+%  rights
+
+%  static world...make predictions about accessing
+%  files, do they translate into actual systems behaviour.
+
+  
+*}
+
+
+(*<*)
+end
+
+end
+(*>*)
+
+(*
+
+  Central to RC-Model is the roles and types. We start with do formalisation on
+  types first.
+
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{
+  \begin{tabular}{r@ {\hspace{1mm}}l@ {\hspace{2mm}}l}
+  @{text t_client} & @{text "="} & @{text "Christian"} \\
+                   & @{text "|"} & @{text "Chunhan"}   \\
+                   & @{text "|"} & @{text " ... "}     \\
+  \end{tabular}}
+
+  \mbox{
+  \begin{tabular}{r@ {\hspace{1mm}}l@ {\hspace{2mm}}l@ {\hspace{2mm}}l}
+  @{text t_normal_file_type}  & @{text "="} & @{text "WebServerLog_file"} &             \\
+                              & @{text "|"} & @{text "WebData_file"} & @{text t_client} \\
+                              & @{text "|"} & @{text "CGI_file"}     & @{text t_client} \\
+                              & @{text "|"} & @{text "Private_file"} & @{text t_client} 
+  \end{tabular}}
+
+  \mbox{
+  \begin{tabular} {r@ {\hspace{1mm}}l@ {\hspace{5mm}}l}
+  @{text t_rc_file_type} 
+     & @{text "="} & @{term "InheritParent_file_type"}   \\
+     & @{text "|"} & @{term "NormalFile_type t_normal_file_type"}
+  \end{tabular}}
+  \end{isabelle}
+
+  @{term "type_of_file"} function calculates the current type for the files:
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{\begin{tabular}{lcl}
+  @{thm (lhs) type_of_file.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) type_of_file.simps(1)}\\
+  @{thm (lhs) type_of_file.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) type_of_file.simps(2)}\\
+  @{term "type_of_file (DUMMY#s)"} & @{text "\<equiv>"} & @{term "type_of_file s"}
+  \end{tabular}}
+  \end{isabelle}
+
+  Note that this @{term "type_of_file"} is not the function @{term "etype_of_file"} 
+  that we call in the grant check of RC-Model, @{term "rc_grant"}. The reason is
+  that file's type can be set to a special type of @{term "InheritParent_file_type"}, 
+  means that the ``efficient'' type of this file is the efficient type of its directory. 
+  \mbox{\begin{tabular}{lcl}
+  @{thm (lhs) etype_aux.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) etype_aux.simps(1)}\\
+  @{thm (lhs) etype_aux.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) etype_aux.simps(2)}\smallskip\\ 
+  @{thm (lhs) etype_of_file_def} & @{text "\<equiv>"} & @{thm (rhs) etype_of_file_def}
+  \end{tabular}}
+  Here @{term etype_aux} is an auxiliary function which do recursion
+  on the pathname of files. By the way, in our proofs, we do proved
+  that functions like @{term "etype_of_file"} will always return
+  ``normal'' values.
+  
+ 
+  We have similar observation functions calculating the current type for processes
+  and IPCs too, only diffence here is that there is no ``effcient'' type here for 
+  processes and IPCs, all types that calculated by @{term "type_of_process"} and 
+  @{term "type_of_ipc"} are alrealdy efficient types. 
+
+*}
+
+text {*
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{
+  \begin{tabular}{r@ {\hspace{1mm}}l@ {\hspace{2mm}}l@ {\hspace{2mm}}l}
+  @{text t_normal_role}  & @{text "="} & @{text "WebServer"} &                  \\
+                         & @{text "|"} & @{text "WS_client"} & @{text t_client} \\
+                         & @{text "|"} & @{text "UpLoader"}  & @{text t_client} \\
+                         & @{text "|"} & @{text "CGI   "}    & @{text t_client} 
+  \end{tabular}}
+
+  \mbox{
+  \begin{tabular} {r@ {\hspace{1mm}}l@ {\hspace{2mm}}l@ {\hspace{5mm}}l}
+  @{text t_role} 
+     & @{text "="} & @{term "InheritParentRole"}  & ``for file's initial/forced role, 
+                                                         meaning using parent directory's 
+                                                         role instead'' \\
+     & @{text "|"} & @{term "UseForcedRole"}      & ``for file's initial role'' \\
+     & @{text "|"} & @{term "InheritProcessRole"} & ``using process' current role''\\
+     & @{text "|"} & @{term "InheritUserRole"}    & ``using owner's default role''\\
+     & @{text "|"} &   ...                        & \\
+     & @{text "|"} & @{term "NormalRole t_normal_role"} & ``user-defined 
+                                                           policy roles" 
+  \end{tabular}}
+  \end{isabelle}
+
+  @{text "t_normal roles"} are normally user-defined roles in the
+  policy, where @{text "WebServer"} is the role who plays for the
+  server, while @{text "WS_client"} is the role server plays for
+  certain client, so is for @{text "UpLoader"} role. @{text "CGI"} is
+  the role that client's programme scripts play. 
+
+  @{term "currentrole"} function calculates the current role of process, here we 
+  only show 3 cases of its definition, it responses to @{term "ChangeOwner"}, 
+  @{term "ChangeRole"} events too. 
+
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{\begin{tabular}{lcl}
+  @{thm (lhs) currentrole.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) currentrole.simps(1)}\\
+  @{thm (lhs) currentrole.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) currentrole.simps(2)}\\
+  @{thm (lhs) currentrole.simps(3)} & @{text "\<equiv>"} & @{thm (rhs) currentrole.simps(3)}
+  \end{tabular}}
+  \end{isabelle}
+
+  If the event trace is @{term "[]"}, means the
+  system state currently is the initial state, then @{term "init_currentrole"} will
+  do. @{term "Execute p f"} event is one complex case, when this event happens, process
+  @{term p}'s role will be changed according to the efficient initial role of the 
+  executable file @{term f}, here ``efficient'' is like the file's type too. 
+
+  \begin{isabelle}\ \ \ \ \ %%%
+  \mbox{\begin{tabular}{lcl}
+  @{thm (lhs) initialrole.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) initialrole.simps(1)}\\
+  @{thm (lhs) initialrole.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) initialrole.simps(2)}\\
+  @{thm (lhs) initialrole.simps(3)} & @{text "\<equiv>"} & @{thm (rhs) initialrole.simps(3)}\medskip\\
+
+  @{thm (lhs) erole_functor.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) erole_functor.simps(1)}\\
+  @{thm (lhs) erole_functor.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) erole_functor.simps(2)}  
+  \end{tabular}}
+  \end{isabelle}
+
+  If this efficient initial role is normal role, then RC-Model assigns
+  this role to the process after execution finished. Otherwise if this
+  efficient initial role is using special value @{term
+  "UseForcedRole"}, then the new role for the process is then
+  determinated by the efficient forced role of the executable file
+  @{term "forcedrole"}.  When new process is created, this process'
+  role is assigned to its creator's role.
+*)
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/all_sobj_prop.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,2246 @@
+theory all_sobj_prop
+imports Main rc_theory os_rc obj2sobj_prop deleted_prop sound_defs_prop source_prop
+begin
+
+context tainting_s_complete begin
+
+lemma initf_has_effinitialrole:
+  "f \<in> init_files ==> \<exists> r. effinitialrole [] f = Some r"
+by (rule_tac f = f in file_has_effinitialrole, simp, simp add:vs_nil)
+
+lemma initf_has_effforcedrole:
+  "f \<in> init_files ==> \<exists> r. effforcedrole [] f = Some r"
+by (rule_tac f = f in file_has_effforcedrole, simp, simp add:vs_nil)
+
+lemma efffrole_sdir_some:
+  "sd \<in> init_files ==> \<exists> r. erole_functor init_file_forcedrole InheritUpMixed sd = Some r"
+apply (frule_tac s = "[]" in undel_initf_keeps_efrole, simp, simp add:vs_nil)
+by (drule initf_has_effforcedrole, simp)
+
+lemma efffrole_sdir_some':
+  "erole_functor init_file_forcedrole InheritUpMixed sd = None ==> sd \<notin> init_files"
+by (rule notI, auto dest!:efffrole_sdir_some)
+
+lemma effirole_sdir_some:
+  "sd \<in> init_files ==> \<exists> r. erole_functor init_file_initialrole UseForcedRole sd = Some r"
+apply (frule_tac s = "[]" in undel_initf_keeps_eirole, simp, simp add:vs_nil)
+by (drule initf_has_effinitialrole, simp)
+
+lemma effirole_sdir_some':
+  "erole_functor init_file_initialrole UseForcedRole sd = None ==> sd \<notin> init_files"
+by (rule notI, auto dest:effirole_sdir_some)
+
+lemma erole_func_irole_simp:
+  "erole_functor init_file_initialrole UseForcedRole sd = effinitialrole [] sd"
+by (simp add:effinitialrole_def)
+
+lemma erole_func_frole_simp:
+  "erole_functor init_file_forcedrole InheritUpMixed sd = effforcedrole [] sd"
+by (simp add:effforcedrole_def)
+
+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)"
+by (simp add:erole_func_frole_simp, erule effforcedrole_valid)
+
+lemma init_effinitialrole_valid: "erole_functor init_file_initialrole UseForcedRole sd = Some r \<Longrightarrow> r = UseForcedRole \<or> (\<exists> nr. r = NormalRole nr)" 
+by (simp add:erole_func_irole_simp, erule effinitialrole_valid)
+
+lemma exec_role_some:
+  "[|sd \<in> init_files; u \<in> init_users|] ==> \<exists> r'. exec_role_aux r sd u = Some r'"
+by (auto simp:exec_role_aux_def split:option.splits t_role.splits 
+        dest!:effirole_sdir_some' efffrole_sdir_some'
+        dest:init_effforcedrole_valid init_effinitialrole_valid
+       intro:effirole_sdir_some efffrole_sdir_some user_has_normalrole)
+
+lemma chown_role_some:
+  "u \<in> init_users ==> \<exists> r'. chown_role_aux r fr u = Some r'"
+by (auto simp:chown_role_aux_def split:option.splits t_role.splits 
+        dest!:effirole_sdir_some' efffrole_sdir_some'
+        dest:init_effforcedrole_valid init_effinitialrole_valid
+       intro:effirole_sdir_some efffrole_sdir_some user_has_normalrole)
+  
+declare obj2sobj.simps [simp del]
+
+lemma all_sobjs_I:
+  assumes ex: "exists s obj"
+  and vd: "valid s"
+  shows "obj2sobj s obj \<in> all_sobjs"
+using ex vd
+proof (induct s arbitrary:obj)
+  case Nil
+  assume ex:"exists [] obj"
+  show ?case
+  proof (cases obj)
+    case (Proc p) assume prem: "obj = Proc p" 
+    with ex have initp:"p \<in> init_processes" by simp 
+    from initp obtain r where "init_currentrole p = Some r" 
+      using init_proc_has_role by (auto simp:bidirect_in_init_def)  
+    moreover from initp obtain t where "init_process_type p = Some t" 
+      using init_proc_has_type by (auto simp:bidirect_in_init_def)  
+    moreover from initp obtain fr where "init_proc_forcedrole p = Some fr" 
+      using init_proc_has_frole by (auto simp:bidirect_in_init_def)  
+    moreover from initp obtain u where "init_owner p = Some u" 
+      using init_proc_has_owner by (auto simp:bidirect_in_init_def)  
+    ultimately have "obj2sobj [] (Proc p) \<in> all_sobjs" 
+      using initp by (auto intro!:ap_init simp:obj2sobj.simps)
+    with prem show ?thesis by simp
+  next
+    case (File f) assume prem: "obj = File f"
+    with ex have initf: "f \<in> init_files" by simp
+    from initf obtain t where "etype_aux init_file_type_aux f = Some t" 
+      using init_file_has_etype by auto
+    moreover from initf have "source_dir [] f = Some f"
+      by (simp add:source_dir_of_init')
+    ultimately have "obj2sobj [] (File f) \<in> all_sobjs"
+      using initf by (auto simp:etype_of_file_def obj2sobj.simps intro!:af_init)
+    with prem show ?thesis by simp
+  next
+    case (IPC i) assume prem: "obj = IPC i"
+    with ex have initi: "i \<in> init_ipcs" by simp
+    from initi obtain t where "init_ipc_type i = Some t" 
+      using init_ipc_has_type by (auto simp:bidirect_in_init_def)
+    hence "obj2sobj [] (IPC i) \<in> all_sobjs"
+      using initi by (auto intro!:ai_init simp:obj2sobj.simps) 
+    with prem show ?thesis by simp
+  qed
+next
+  case (Cons e s)
+  assume prem: "\<And> obj. \<lbrakk>exists s obj; valid s\<rbrakk> \<Longrightarrow> obj2sobj s obj \<in> all_sobjs"
+    and ex_cons: "exists (e # s) obj" and vs_cons: "valid (e # s)"
+  have vs: "valid s" and rc: "rc_grant s e" and os: "os_grant s e" 
+    using vs_cons by (auto intro:valid_cons valid_os valid_rc)
+  from prem and vs have prem': "\<And> obj. exists s obj \<Longrightarrow> obj2sobj s obj \<in> all_sobjs" by simp
+  show ?case
+  proof (cases e)
+    case (ChangeOwner p u)
+    assume ev: "e = ChangeOwner p u"
+    show ?thesis
+    proof (cases "obj = Proc p")
+      case True 
+      have curp: "p \<in> current_procs s" and exp: "exists s (Proc p)" using os ev by auto
+      from curp obtain r fr t u' srp where sp: "obj2sobj s (Proc p) = SProc (r,fr,t,u') (Some srp)" 
+        using vs apply (drule_tac current_proc_has_sobj, simp) by blast
+      hence sp_in: "SProc (r,fr,t,u') (Some srp) \<in> all_sobjs" using prem' exp by metis
+      have comp: "(r, Proc_type t, CHANGE_OWNER) \<in> compatible" using sp ev rc
+        by (auto simp:obj2sobj.simps split:option.splits)
+      from os ev have uinit: "u \<in> init_users" by simp
+      then obtain nr where chown: "chown_role_aux r fr u = Some nr" 
+        by (auto dest:chown_role_some)
+      hence nsp_in:"obj2sobj (e#s) (Proc p) = SProc (nr,fr,chown_type_aux r nr t, u) (Some srp)" 
+        using sp ev os
+        by (auto split:option.splits t_role.splits 
+          simp del:currentrole.simps type_of_process.simps
+          simp add:chown_role_aux_valid chown_type_aux_valid obj2sobj.simps)
+      thus ?thesis using True ev os rc sp_in sp
+        by (auto simp:chown comp intro!:ap_chown[where u'=u']) 
+    next
+      case False
+      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons
+        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps
+                              split:option.splits t_role.splits)
+      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
+    qed
+  next
+    case (Clone p p')
+    assume ev: "e = Clone p p'"
+    show ?thesis 
+    proof (cases "obj = Proc p'")
+      case True
+      from ev os have exp: "exists s (Proc p)" by (simp add:os_grant.simps)
+      from exp obtain r fr pt u sp where sproc: "cp2sproc s p = Some (r, fr, pt, u)"
+        and srp: "source_proc s p = Some sp" using vs
+        apply (simp del:cp2sproc.simps)
+        by (frule current_proc_has_sproc, simp, frule current_proc_has_srp, simp, blast)
+      hence SP: "SProc (r,fr,pt,u) (Some sp) \<in> all_sobjs" using exp prem'[where obj = "Proc p"] vs
+        by (auto split:option.splits simp add:obj2sobj.simps)
+      have "obj2sobj (e # s) (Proc p') = SProc (r,fr,clone_type_aux r pt, u) (Some sp)"
+        using ev sproc srp vs_cons 
+        by (simp add:obj2sobj.simps cp2sproc_clone del:cp2sproc.simps)
+      thus ?thesis using True SP by (simp add:ap_clone)
+    next
+      case False
+      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons
+        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps
+                              split:option.splits t_role.splits)
+      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
+    qed
+  next
+    case (Execute p f)
+    assume ev: "e = Execute p f"
+    show ?thesis
+    proof (cases "obj = Proc p")
+      case True
+      from ev os have exp: "exists s (Proc p)" and exf: "exists s (File f)" by auto
+      from exp obtain r fr pt u sp where sproc: "cp2sproc s p = Some (r, fr, pt, u)"
+        and srp: "source_proc s p = Some sp" using vs
+        apply (simp del:cp2sproc.simps)
+        by (frule current_proc_has_sproc, simp, frule current_proc_has_srp, simp, blast)
+      hence SP: "SProc (r,fr,pt,u) (Some sp) \<in> all_sobjs" using exp prem'[where obj = "Proc p"] vs
+        by (auto split:option.splits simp add:obj2sobj.simps)
+      from exf obtain sd t where srdir: "source_dir s f = Some sd" and 
+        etype: "etype_of_file s f = Some t" using vs
+        by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
+      then obtain srf where SF: "SFile (t, sd) srf \<in> all_sobjs" 
+        using exf prem'[where obj = "File f"] vs
+        by (auto split:option.splits if_splits simp:obj2sobj.simps dest:current_file_has_etype)
+      from sproc srdir have "u \<in> init_users" and "sd \<in> init_files" using vs
+        by (auto intro:source_dir_in_init owner_in_users split:option.splits)
+      then obtain nr where "exec_role_aux r sd u = Some nr" by (auto dest:exec_role_some) 
+      
+      hence "obj2sobj (e # s) (Proc p) \<in> all_sobjs" using ev vs_cons srdir sproc srp 
+        apply (auto simp:obj2sobj.simps cp2sproc_simps source_proc.simps 
+                   intro:source_dir_in_init simp del:cp2sproc.simps 
+                   split:option.splits dest!:efffrole_sdir_some')
+        apply (rule ap_exec) using SF SP rc ev etype by (auto split:option.splits)
+      with True show ?thesis by simp
+    next
+      case False
+      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons
+        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps
+                              split:option.splits t_role.splits)
+      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
+    qed
+  next
+    case (CreateFile p f)
+    assume ev: "e = CreateFile p f"
+    show ?thesis
+    proof (cases "obj = File f")
+      case True
+      from os ev obtain pf where expf: "exists s (File pf)" and parent:"parent f = Some pf" by auto
+      from expf obtain pft sd srpf where SF: "SFile (pft, sd) srpf \<in> all_sobjs"
+        and eptype: "etype_of_file s pf = Some pft" and srpf: "source_dir s pf = Some sd"
+        using prem'[where obj = "File pf"] vs
+        by (auto split:option.splits if_splits simp:obj2sobj.simps 
+                  dest:current_file_has_etype current_file_has_sd)
+      from os ev have exp: "exists s (Proc p)" by simp
+      then obtain r pt fr u srp where SP: "SProc (r, fr, pt, u) srp \<in> all_sobjs"
+        and sproc: "cp2sproc s p = Some (r, fr, pt, u)" 
+        using prem'[where obj = "Proc p"] vs
+        by (auto split:option.splits if_splits simp:obj2sobj.simps 
+                  dest:current_proc_has_sproc)
+      have "obj2sobj (e # s) (File f) \<in> all_sobjs" using ev vs_cons sproc srpf parent os
+        apply (auto simp:obj2sobj.simps source_dir_simps init_notin_curf_deleted  
+                   split:option.splits dest!:current_file_has_etype')
+        apply (case_tac "default_fd_create_type r")
+        using SF SP rc ev eptype sproc
+        apply (rule_tac sf = srpf in af_cfd', auto simp:etype_of_file_def etype_aux_prop3) [1]
+        using SF SP rc ev eptype sproc
+        apply (rule_tac sf = srpf in af_cfd)
+        apply (auto simp:etype_of_file_def etype_aux_prop4)
+        done
+      with True show ?thesis by simp
+    next
+      case False 
+      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
+        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps etype_aux_prop2
+                              split:option.splits t_role.splits )
+      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
+    qed
+  next
+    case (CreateIPC p i)
+    assume ev: "e = CreateIPC p i"
+    show ?thesis
+    proof (cases "obj = IPC i")
+      case True
+      from os ev have exp: "exists s (Proc p)" by simp
+      then obtain r pt fr u srp where SP: "SProc (r, fr, pt, u) srp \<in> all_sobjs"
+        and sproc: "cp2sproc s p = Some (r, fr, pt, u)" 
+        using prem'[where obj = "Proc p"] vs
+        by (auto split:option.splits if_splits simp:obj2sobj.simps 
+                  dest:current_proc_has_sproc)
+      have "obj2sobj (e # s) (IPC i) \<in> all_sobjs" using ev vs_cons sproc os
+        apply (auto simp:obj2sobj.simps ni_init_deled split:option.splits)
+        apply (rule ai_cipc) using SP sproc rc ev by auto
+      with True show ?thesis by simp
+    next
+      case False 
+      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
+        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
+                              split:option.splits t_role.splits )
+      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
+    qed
+  next
+    case (ChangeRole p r')
+    assume ev: "e = ChangeRole p r'"
+    show ?thesis
+    proof (cases "obj = Proc p")
+      case True
+      from os ev have exp: "exists s (Proc p)" by simp
+      then obtain r pt fr u srp where SP: "SProc (r, fr, pt, u) srp \<in> all_sobjs"
+        and sproc: "cp2sproc s p = Some (r, fr, pt, u)" and srproc: "source_proc s p = srp"
+        using prem'[where obj = "Proc p"] vs
+        by (auto split:option.splits if_splits simp:obj2sobj.simps 
+                  dest:current_proc_has_sproc)
+      have "obj2sobj (e # s) (Proc p) \<in> all_sobjs" using ev vs_cons sproc os 
+        apply (auto simp:obj2sobj.simps ni_init_deled split:option.splits)
+        apply (rule ap_crole) using SP sproc rc ev srproc by auto
+      with True show ?thesis by simp
+    next
+      case False 
+      hence "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
+        by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
+                              split:option.splits t_role.splits )
+      thus ?thesis using False prem' ex_cons ev by (case_tac obj, auto)
+    qed
+  next
+    case (ReadFile p f)
+    assume ev: "e = ReadFile p f"
+    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
+      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
+                            split:option.splits t_role.splits)
+    moreover have "exists s obj" using ev ex_cons
+      by (case_tac obj, auto)
+    ultimately show ?thesis using prem[where obj = obj] vs by simp
+  next
+    case (WriteFile p f)
+    assume ev: "e = WriteFile p f"
+    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
+      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
+                            split:option.splits t_role.splits)
+    moreover have "exists s obj" using ev ex_cons
+      by (case_tac obj, auto)
+    ultimately show ?thesis using prem[where obj = obj] vs by simp
+  next
+    case (Send p i)
+    assume ev: "e = Send p i"
+    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
+      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
+                            split:option.splits t_role.splits)
+    moreover have "exists s obj" using ev ex_cons
+      by (case_tac obj, auto)
+    ultimately show ?thesis using prem[where obj = obj] vs by simp
+  next
+    case (Recv p i)
+    assume ev: "e = Recv p i"
+    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
+      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
+                            split:option.splits t_role.splits)
+    moreover have "exists s obj" using ev ex_cons
+      by (case_tac obj, auto)
+    ultimately show ?thesis using prem[where obj = obj] vs by simp
+  next
+    case (Kill p p')
+    assume ev: "e = Kill p p'"
+    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
+      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
+                            split:option.splits t_role.splits)
+    thus ?thesis using prem[where obj = obj] vs ex_cons ev
+      by (case_tac obj, auto)
+  next
+    case (DeleteFile p f')
+    assume ev: "e = DeleteFile p f'" 
+    have "obj2sobj (e#s) obj = obj2sobj s obj"
+    proof-
+      have "\<And> f. obj = File f ==> obj2sobj (e#s) (File f) = obj2sobj s (File f)"
+        using ev vs os ex_cons vs_cons
+        by (auto simp:obj2sobj.simps etype_of_file_delete source_dir_simps
+                split:option.splits t_role.splits if_splits
+                dest!:current_file_has_etype' current_file_has_sd'
+                 dest:source_dir_prop)
+      moreover have "\<forall> f. obj \<noteq> File f ==> obj2sobj (e#s) obj = obj2sobj s obj"
+        using ev vs_cons ex_cons os vs
+        by (case_tac obj, auto simp:obj2sobj.simps split:option.splits)
+      ultimately show ?thesis by auto
+    qed
+    thus ?thesis using prem[where obj = obj] vs ex_cons ev 
+      by (case_tac obj, auto)
+  next
+    case (DeleteIPC p i)
+    assume ev: "e = DeleteIPC p i"
+    have "obj2sobj (e#s) obj = obj2sobj s obj" using ev vs_cons ex_cons os vs
+      by (case_tac obj, auto simp:obj2sobj.simps etype_of_file_def source_dir_simps 
+                            split:option.splits t_role.splits)
+    thus ?thesis using prem[where obj = obj] vs ex_cons ev
+      by (case_tac obj, auto)
+  qed
+qed
+
+declare obj2sobj.simps [simp add]
+
+lemma seeds_in_all_sobjs: 
+  assumes seed: "obj \<in> seeds" shows "init_obj2sobj obj \<in> all_sobjs"
+proof (cases obj)
+  case (Proc p)
+  assume p0: "obj = Proc p" (*?*)
+  from seed p0 have pinit: "p \<in> init_processes" by (drule_tac seeds_in_init, simp)
+  from pinit obtain r where "init_currentrole p = Some r" 
+    using init_proc_has_role by (auto simp:bidirect_in_init_def)
+  moreover from pinit obtain fr where "init_proc_forcedrole p = Some fr"
+    using init_proc_has_frole by (auto simp:bidirect_in_init_def)
+  moreover from pinit obtain pt where "init_process_type p = Some pt"
+    using init_proc_has_type by (auto simp:bidirect_in_init_def)
+  moreover from pinit obtain u where "init_owner p = Some u"
+    using init_proc_has_owner by (auto simp:bidirect_in_init_def)
+  ultimately show ?thesis using p0 by (auto intro:ap_init)
+next
+  case (File f) 
+  assume p0: "obj = File f" (*?*)
+  from seed p0 have finit: "f \<in> init_files" by (drule_tac seeds_in_init, simp)
+  then obtain t where "etype_aux init_file_type_aux f = Some t"
+    by (auto dest:init_file_has_etype)
+  with finit p0 show ?thesis by (auto intro:af_init)
+next
+  case (IPC i)
+  assume p0: "obj = IPC i" (*?*)
+  from seed p0 have iinit: "i \<in> init_ipcs" by (drule_tac seeds_in_init, simp)
+  then obtain t where "init_ipc_type i = Some t" using init_ipc_has_type
+    by (auto simp:bidirect_in_init_def)
+  with iinit p0 show ?thesis by (auto intro:ai_init)
+qed
+
+lemma tainted_s_in_all_sobjs:
+  "sobj \<in> tainted_s \<Longrightarrow> sobj \<in> all_sobjs"
+apply (erule tainted_s.induct)
+apply (erule seeds_in_all_sobjs)
+apply (auto intro:ap_crole ap_exec ap_chown ai_cipc af_cfd af_cfd' ap_clone)
+done
+
+end
+
+context tainting_s_sound begin
+
+(*** all_sobjs' equal with all_sobjs in the view of soundness ***)
+
+lemma all_sobjs'_eq1: "sobj \<in> all_sobjs \<Longrightarrow> sobj \<in> all_sobjs'"
+apply (erule all_sobjs.induct)
+apply (auto intro:af'_init af'_cfd af'_cfd' ai'_init ai'_cipc ap'_init ap'_crole ap'_exec ap'_chown)
+by (simp add:clone_type_aux_def clone_type_unchange)
+
+lemma all_sobjs'_eq2: "sobj \<in> all_sobjs' \<Longrightarrow> sobj \<in> all_sobjs"
+apply (erule all_sobjs'.induct)
+by (auto intro:af_init af_cfd af_cfd' ai_init ai_cipc ap_init ap_crole ap_exec ap_chown)
+
+lemma all_sobjs'_eq: "(sobj \<in> all_sobjs) = (sobj \<in> all_sobjs')"
+by (auto intro:iffI all_sobjs'_eq1 all_sobjs'_eq2)
+
+(************************ all_sobjs Elim Rules ********************)
+
+declare obj2sobj.simps [simp del]
+declare cp2sproc.simps [simp del]
+
+lemma all_sobjs_E0_aux[rule_format]:
+  "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))"
+proof (induct rule:all_sobjs'.induct)
+  case (af'_init f t) show ?case
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
+      and nodels': "no_del_event s'"and intacts':"initp_intact s'"
+      and exso': "exists s' obj'"
+    from nodels' af'_init(1) have exf: "f \<in> current_files s'" 
+      by (drule_tac obj = "File f" in nodel_imp_exists, simp+)    
+    have "obj2sobj s' (File f) = SFile (t, f) (Some f)"
+    proof-
+      have "obj2sobj [] (File f) = SFile (t, f) (Some f)"  using af'_init 
+        by (auto simp:etype_of_file_def source_dir_of_init' obj2sobj.simps
+                split:option.splits)
+      thus ?thesis using vss' exf nodels' af'_init(1) 
+        by (drule_tac obj2sobj_file_remains_app', auto)
+    qed
+    thus "\<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 = SFile (t, f) (Some f) \<and> exists (s @ s') obj"
+      apply (rule_tac x = "[]" in exI, rule_tac x = "File f" in exI)
+      by (simp add:vss' sobjs' nodels' intacts' exf exso')
+  qed
+next
+  case (af'_cfd t sd srf r fr pt u srp t') 
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and exobj':"exists s' obj'"
+    with af'_cfd(1,2) obtain sa pf where
+      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
+      "exists (sa@s') obj'" and "initp_intact (sa@s')" and
+      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
+      exfa: "pf \<in> current_files (sa@s')" 
+      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
+      by (frule obj2sobj_file, auto)
+    with af'_cfd(3,4) notUkn obtain sb p where
+      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
+      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
+      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
+      exsoab: "exists (sb@sa@s') obj'" and
+      intactab: "initp_intact (sb@sa@s')" and
+      nodelab: "no_del_event (sb@sa@s')"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
+      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
+      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
+    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
+      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
+    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
+      and tau: "\<tau>=sb@sa@s'" by auto
+    
+    have valid: "valid (e # \<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau af'_cfd(5,6,7) SPab SFab
+        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
+                split:if_splits option.splits t_rc_file_type.splits)
+      ultimately show ?thesis using vsab tau
+        by (rule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
+        by (case_tac obj', simp+)   moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'" 
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto simp:obj2sobj.simps cp2sproc_simps 
+             simp del:cp2sproc.simps split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
+      by (simp_all add:initp_intact_I_others)  moreover
+    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t', sd) None"
+    proof-
+      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t'"
+        using ev tau SFab SPab af'_cfd(5)
+        by (auto simp:obj2sobj.simps cp2sproc.simps etype_of_file_def
+          split:option.splits if_splits  intro!:etype_aux_prop4)
+      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
+        using ev tau SFab SPab valid ncf_parent
+        by (auto simp:source_dir_simps obj2sobj.simps 
+                split:if_splits option.splits)
+      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
+        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
+        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
+    qed
+    ultimately show "\<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 = SFile (t', sd) None \<and> exists (s @ s') obj "
+      using tau ev
+      by (rule_tac x = "e#sb@sa" in exI, rule_tac x = "File (new_childf pf \<tau>)" in exI, simp+)
+  qed
+next
+  case (af'_cfd' t sd srf r fr pt u srp)
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and exobj':"exists s' obj'"
+    with af'_cfd'(1,2) obtain sa pf where
+      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
+      "exists (sa@s') obj'" and "initp_intact (sa@s')" and
+      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
+      exfa: "pf \<in> current_files (sa@s')" 
+      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
+      by (frule obj2sobj_file, auto)
+    with af'_cfd'(3,4) notUkn obtain sb p where
+      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
+      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
+      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
+      exsoab: "exists (sb@sa@s') obj'" and
+      intactab: "initp_intact (sb@sa@s')" and
+      nodelab: "no_del_event (sb@sa@s')"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
+      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
+      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
+    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
+      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
+    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
+      and tau: "\<tau>=sb@sa@s'" by auto
+    
+    have valid: "valid (e # \<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau af'_cfd'(5,6) SPab SFab
+        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
+                split:if_splits option.splits t_rc_file_type.splits)
+      ultimately show ?thesis using vsab tau
+        by (rule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
+        by (case_tac obj', simp+)   moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'" 
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto simp:obj2sobj.simps cp2sproc_simps 
+             simp del:cp2sproc.simps split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
+      by (simp add:initp_intact_I_others)  moreover
+    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t, sd) None"
+    proof-
+      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t"
+      proof-
+        have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = etype_of_file \<tau> pf"
+          using ev tau SPab af'_cfd'(5) 
+          by (auto simp:obj2sobj.simps ncf_parent etype_of_file_def cp2sproc.simps
+            split:option.splits   intro!:etype_aux_prop3)
+        thus ?thesis using SFab tau ev
+          by (auto simp:obj2sobj.simps split:option.splits if_splits)
+      qed
+      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
+        using ev tau SFab SPab valid ncf_parent
+        by (auto simp:source_dir_simps obj2sobj.simps 
+                split:if_splits option.splits)
+      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
+        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
+        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
+    qed
+    ultimately show "\<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 = SFile (t, sd) None \<and> exists (s @ s') obj"
+      using tau ev
+      by (rule_tac x = "e#sb@sa" in exI, rule_tac x = "File (new_childf pf \<tau>)" in exI, simp+)
+  qed 
+next
+  case (ai'_init i t) 
+  hence initi: "i \<in> init_ipcs" using init_ipc_has_type
+    by (simp add:bidirect_in_init_def)
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
+      and nodels': "no_del_event s'"and intacts':"initp_intact s'"
+      and exso': "exists s' obj'"
+    from nodels' initi have exi: "i \<in> current_ipcs s'" 
+      by (drule_tac obj = "IPC i" in nodel_imp_exists, simp+)    
+    have "obj2sobj s' (IPC i) = SIPC t (Some i)"
+    proof-
+      have "obj2sobj [] (IPC i) = SIPC t (Some i)"  
+        using ai'_init initi by (auto simp:obj2sobj.simps)
+      thus ?thesis using vss' exi nodels' initi 
+        by (drule_tac obj2sobj_ipc_remains'', auto simp:obj2sobj.simps)
+    qed
+    thus "\<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 = SIPC t (Some i) \<and> exists (s @ s') obj"
+      apply (rule_tac x = "[]" in exI, rule_tac x = "IPC i" in exI)
+      by (simp add:vss' sobjs' nodels' exi exso' intacts' del:obj2sobj.simps)
+  qed
+next
+  case (ai'_cipc r fr pt u srp)
+  show ?case
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and exobj':"exists s' obj'"
+    with ai'_cipc(1,2) notUkn obtain s p where
+      SPab: "obj2sobj (s@s') (Proc p) = SProc (r,fr,pt,u) srp" and
+      expab: "exists (s@s') (Proc p)" and vsab: "valid (s@s')" and
+      soab: "obj2sobj (s@s') obj' = sobj'" and 
+      exsoab: "exists (s@s') obj'" and
+      intactab: "initp_intact (s@s')" and
+      nodelab: "no_del_event (s@s')"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    obtain e \<tau> where ev: "e = CreateIPC p (new_ipc \<tau>)" and tau: "\<tau>=s@s'" by auto
+    
+    have valid: "valid (e # \<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau expab by (simp)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ai'_cipc(3) SPab
+        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
+      ultimately show ?thesis using vsab tau
+        by (rule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
+        by (case_tac obj', simp+)   moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'" 
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto simp:obj2sobj.simps cp2sproc_simps 
+             simp del:cp2sproc.simps split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
+      by (simp add:initp_intact_I_others)  moreover
+    have "obj2sobj (e#\<tau>) (IPC (new_ipc \<tau>)) = SIPC (default_ipc_create_type r) None"
+      using ev tau SPab nodel 
+        nodel_imp_exists[where obj = "IPC (new_ipc \<tau>)" and s =\<tau>]
+      by (auto simp:obj2sobj.simps ni_notin_curi cp2sproc.simps 
+              split:option.splits  dest:no_del_event_cons_D)
+    ultimately show "\<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 = SIPC (default_ipc_create_type r) None \<and> exists (s @ s') obj"
+      using tau ev
+      by (rule_tac x = "e#s" in exI, rule_tac x = "IPC (new_ipc \<tau>)" in exI, simp+)
+  qed 
+next
+  case (ap'_init p r fr t u)
+  hence initp: "p \<in> init_processes" using init_proc_has_role
+    by (simp add:bidirect_in_init_def)
+  show ?case
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume VSs': "valid s'" and SOs': "obj2sobj s' obj' = sobj'"
+      and Nodels': "no_del_event s'"and Intacts':"initp_intact s'"
+      and exso': "exists s' obj'"
+    from Nodels' initp have exp: "p \<in> current_procs s'" 
+      apply (drule_tac obj = "Proc p" in nodel_imp_un_deleted)
+      by (drule not_deleted_imp_exists, simp+)
+    with Intacts' initp ap'_init have "obj2sobj s' (Proc p) = SProc (r, fr, t, u) (Some p)"
+      by (auto simp:initp_intact_def split:option.splits)
+    thus "\<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 = SProc (r, fr, t, u) (Some p) \<and> exists (s @ s') obj"
+      apply (rule_tac x = "[]" in exI, rule_tac x = "Proc p" in exI)
+      by (simp add:VSs' SOs' Nodels' exp exso' initp intact_imp_butp Intacts'
+               del:obj2sobj.simps)
+  qed
+next
+  case (ap'_crole r fr t u srp r')
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
+    with ap'_crole(1,2) obtain s p where
+      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
+      and nodelab: "no_del_event (s@s')"
+      and intactab: "initp_intact (s@s')"
+      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
+      and exp:"exists (s@s') (Proc p)" and exobj'ab:"exists (s@s') obj'"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    obtain e \<tau> where ev: "e = ChangeRole (new_proc (s@s')) r'" 
+      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
+    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
+
+    have valid: "valid (e#\<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau exp by (simp)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ap'_crole(3) SPab 
+        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
+      ultimately show ?thesis using vs_tau
+        by (erule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have "initp_intact (e#\<tau>)" using tau ev intactab valid
+      by (simp add:initp_intact_I_crole)   moreover
+    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
+      by (case_tac obj', simp+)  moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'"
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by (auto)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by auto
+      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
+        apply (case_tac "p' = new_proc (s @ s')") 
+        using vs_tau exobj'ab tau
+        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
+        using tau ev SOab' valid notUkn vs_tau
+        by (auto simp:obj2sobj.simps cp2sproc_simps split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (Proc (new_proc (s@s'))) = SProc (r', fr, t, u) srp"
+      using SPab tau vs_tau ev valid
+      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
+              split:option.splits)  moreover
+    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp  
+    ultimately show "\<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 = SProc (r', fr, t, u) srp \<and> exists (s @ s') obj"
+      using ev tau
+      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
+      by (rule_tac x = "Proc (new_proc (s@s'))" in exI, auto)
+  qed
+next
+  case (ap'_chown r fr t u srp u' nr)
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
+    with ap'_chown(1,2) obtain s p where
+      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
+      and nodelab: "no_del_event (s@s')" and intactab: "initp_intact (s@s')"
+      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
+      and exp:"exists (s@s') (Proc p)" and exobj'ab:"exists (s@s') obj'"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    obtain e \<tau> where ev: "e = ChangeOwner (new_proc (s@s')) u'" 
+      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
+    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
+
+    have valid: "valid (e#\<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau exp ap'_chown(3) by (simp)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ap'_chown(5) SPab 
+        by (auto simp:cp2sproc.simps obj2sobj.simps pct_def clone_type_unchange
+                split:option.splits t_rc_proc_type.splits)
+          (* here is another place of no_limit of clone event assumption *)
+      ultimately show ?thesis using vs_tau
+        by (erule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have "initp_intact (e#\<tau>)" using intactab tau ev valid 
+      by (simp add:initp_intact_I_chown)  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
+      by (case_tac obj', simp+)  moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'"
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by (auto)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by auto
+      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
+        apply (case_tac "p' = new_proc (s @ s')") 
+        using vs_tau exobj'ab tau
+        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
+        using tau ev SOab' valid notUkn vs_tau 
+        by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (Proc (new_proc (s@s'))) = 
+          SProc (nr,fr,chown_type_aux r nr t,u') srp"
+      using SPab tau vs_tau ev valid ap'_chown(4)
+      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
+              split:option.splits)  moreover
+    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp   moreover 
+    ultimately show "\<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 = SProc (nr,fr,chown_type_aux r nr t,u') srp \<and>
+      exists (s @ s') obj"
+      using ev tau
+      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
+      by (rule_tac x = "Proc (new_proc (s@s'))" in exI, auto)
+  qed
+next
+  case (ap'_exec r fr pt u sp t sd sf r' fr')
+  show ?case
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
+    with ap'_exec(3,4) obtain sa f where
+      SFa: "obj2sobj (sa @ s') (File f) = SFile (t, sd) sf" and
+      Exfa: "exists (sa @ s') (File f)" and 
+      butsa: "initp_intact (sa @ s')" and
+      othersa:"valid (sa @ s') \<and> obj2sobj (sa @ s') obj' = sobj' \<and> 
+               exists (sa @s') obj' \<and> no_del_event (sa @ s')" 
+      by (blast dest:obj2sobj_file intro:nodel_exists_remains)
+    with ap'_exec(1,2) notUkn obtain sb p where
+      VSab: "valid (sb@sa@s')" and SOab': "obj2sobj (sb@sa@s') obj' = sobj'"
+      and nodelab: "no_del_event (sb@sa@s')"
+      and intactab: "initp_intact (sb@sa@s')"
+      and SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r, fr, pt, u) sp"
+      and exp:"exists (sb@sa@s') (Proc p)" and exobj'ab:"exists (sb@sa@s') obj'"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    obtain e \<tau> where ev: "e = Execute (new_proc (sb@sa@s')) f" 
+      and tau: "\<tau> = Clone p (new_proc (sb@sa@s'))#sb@sa@s'" by auto
+    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
+    from Exfa nodelab have exf:"f \<in> current_files (sb@sa@s')" 
+      apply (drule_tac obj = "File f" in nodel_imp_un_deleted)
+      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
+    from SFa VSab Exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File f) = SFile (t,sd) sf"
+      by (rule_tac s = "sa@s'" in obj2sobj_file_remains'', simp_all)
+
+    have valid: "valid (e#\<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau exp by (simp add:exf)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ap'_exec(5) SPab SFab
+        by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
+                split:if_splits option.splits)
+      ultimately show ?thesis using vs_tau
+        by (erule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have "initp_intact (e#\<tau>)" using tau ev intactab valid
+      by (simp add:initp_intact_I_exec)  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
+      by (case_tac obj', simp+) moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'"
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
+        by (auto simp del:obj2sobj.simps)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
+        by (auto simp del:obj2sobj.simps)
+      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
+        apply (case_tac "p' = new_proc (sb @ sa @ s')") 
+        using vs_tau exobj'ab tau
+        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
+        using tau ev SOab' valid notUkn vs_tau 
+        by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (Proc (new_proc (sb @ sa @ s'))) = 
+          SProc (r',fr',exec_type_aux r pt, u) sp"
+    proof-
+      have "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) sp" using SPab tau vs_tau
+        by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
+      hence "obj2sobj \<tau> (Proc (new_proc (sb@sa@s'))) = SProc (r,fr,pt,u) sp" using tau
+      by (auto simp:obj2sobj.simps cp2sproc.simps pct_def clone_type_unchange 
+              split:option.splits)
+      moreover have "source_dir \<tau> f = Some sd" using vs_tau SFab tau 
+        by (auto simp:source_dir_simps obj2sobj.simps split:option.splits if_splits)
+      ultimately show ?thesis using valid ev ap'_exec(6,7) 
+        by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
+    qed 
+    ultimately 
+    show "\<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 = SProc (r', fr', exec_type_aux r pt, u) sp \<and>
+      exists (s @ s') obj"
+      using ev tau
+      apply (rule_tac x = "e#Clone p (new_proc (sb @ sa @ s')) #sb@sa" in exI)
+      by (rule_tac x = "Proc (new_proc (sb @ sa @ s'))" in exI, auto)
+  qed
+qed
+
+(* this is for ts2t createfile case ... *)
+lemma all_sobjs_E0:
+  "\<lbrakk>sobj \<in> all_sobjs'; valid s'; obj2sobj s' obj' = sobj'; exists s' obj'; sobj' \<noteq> Unknown; 
+    no_del_event s'; initp_intact s'\<rbrakk>
+   \<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"
+by (drule all_sobjs_E0_aux, auto)
+
+lemma all_sobjs_E1_aux[rule_format]:
+  "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))"
+proof (induct rule:all_sobjs'.induct)
+  case (af'_init f t) show ?case
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
+      and nodels': "no_del_event s'"and intacts':"initp_intact_but s' sobj'"
+      and exso': "exists s' obj'"
+    from nodels' af'_init(1) have exf: "f \<in> current_files s'" 
+      by (drule_tac obj = "File f" in nodel_imp_exists, simp+)    
+    have "obj2sobj s' (File f) = SFile (t, f) (Some f)"
+    proof-
+      have "obj2sobj [] (File f) = SFile (t, f) (Some f)"  using af'_init 
+        by (auto simp:etype_of_file_def source_dir_of_init' obj2sobj.simps
+                split:option.splits)
+      thus ?thesis using vss' exf nodels' af'_init(1) 
+        by (drule_tac obj2sobj_file_remains_app', auto)
+    qed
+    thus "\<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 = SFile (t, f) (Some f) \<and> exists (s @ s') obj"
+      apply (rule_tac x = "[]" in exI, rule_tac x = "File f" in exI)
+      by (simp add:vss' sobjs' nodels' intacts' exf exso')
+  qed
+next
+  case (af'_cfd t sd srf r fr pt u srp t') 
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" 
+      and exobj':"exists s' obj'"
+    with af'_cfd(1,2) obtain sa pf where
+      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
+      "exists (sa@s') obj'" and "initp_intact_but (sa@s') sobj'" and
+      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
+      exfa: "pf \<in> current_files (sa@s')" 
+      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
+      by (frule obj2sobj_file, auto)
+    with af'_cfd(3,4) notUkn obtain sb p where
+      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
+      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
+      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
+      exsoab: "exists (sb@sa@s') obj'" and
+      intactab: "initp_intact_but (sb@sa@s') sobj'" and
+      nodelab: "no_del_event (sb@sa@s')"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
+      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
+      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
+    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
+      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
+    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
+      and tau: "\<tau>=sb@sa@s'" by auto
+    
+    have valid: "valid (e # \<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau af'_cfd(5,6,7) SPab SFab
+        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
+                split:if_splits option.splits t_rc_file_type.splits)
+      ultimately show ?thesis using vsab tau
+        by (rule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
+        by (case_tac obj', simp+)   moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'" 
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto simp:obj2sobj.simps cp2sproc_simps' 
+             simp del:cp2sproc.simps split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t', sd) None"
+    proof-
+      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t'"
+        using ev tau SFab SPab af'_cfd(5)
+        by (auto simp:obj2sobj.simps cp2sproc.simps etype_of_file_def
+          split:option.splits if_splits  intro!:etype_aux_prop4)
+      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
+        using ev tau SFab SPab valid ncf_parent
+        by (auto simp:source_dir_simps obj2sobj.simps 
+                split:if_splits option.splits)
+      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
+        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
+        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
+    qed  moreover
+    have "initp_intact_but (e#\<tau>) sobj'" using intactab tau ev valid nodel
+        apply (case_tac sobj', case_tac option)
+        by (simp_all add:initp_intact_butp_I_others initp_intact_I_others)  
+    ultimately show "\<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 = SFile (t', sd) None \<and> exists (s @ s') obj "
+      using tau ev
+      apply (rule_tac x = "e#sb@sa" in exI)
+      by (rule_tac x = "File (new_childf pf \<tau>)" in exI, auto)
+  qed
+next
+  case (af'_cfd' t sd srf r fr pt u srp)
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" 
+      and exobj':"exists s' obj'"
+    with af'_cfd'(1,2) obtain sa pf where
+      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
+      "exists (sa@s') obj'" and "initp_intact_but (sa@s') sobj'" and
+      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
+      exfa: "pf \<in> current_files (sa@s')" 
+      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
+      by (frule obj2sobj_file, auto)
+    with af'_cfd'(3,4) notUkn obtain sb p where
+      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
+      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
+      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
+      exsoab: "exists (sb@sa@s') obj'" and
+      intactab: "initp_intact_but (sb@sa@s') sobj'" and
+      nodelab: "no_del_event (sb@sa@s')"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
+      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
+      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
+    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
+      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
+    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
+      and tau: "\<tau>=sb@sa@s'" by auto
+    
+    have valid: "valid (e # \<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau af'_cfd'(5,6) SPab SFab
+        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
+                split:if_splits option.splits t_rc_file_type.splits)
+      ultimately show ?thesis using vsab tau
+        by (rule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
+        by (case_tac obj', simp+)   moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'" 
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto simp:obj2sobj.simps cp2sproc_simps' 
+             simp del:cp2sproc.simps split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t, sd) None"
+    proof-
+      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t"
+      proof-
+        have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = etype_of_file \<tau> pf"
+          using ev tau SPab af'_cfd'(5) 
+          by (auto simp:obj2sobj.simps cp2sproc.simps ncf_parent etype_of_file_def
+            split:option.splits   intro!:etype_aux_prop3)
+        thus ?thesis using SFab tau ev
+          by (auto simp:obj2sobj.simps split:option.splits if_splits)
+      qed
+      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
+        using ev tau SFab SPab valid ncf_parent
+        by (auto simp:source_dir_simps obj2sobj.simps 
+                split:if_splits option.splits)
+      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
+        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
+        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
+    qed  moreover
+    have "initp_intact_but (e#\<tau>) sobj'" using intactab tau ev valid nodel
+        apply (case_tac sobj', case_tac option)
+        by (simp_all add:initp_intact_butp_I_others initp_intact_I_others)  
+    ultimately show "\<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 = SFile (t, sd) None \<and> exists (s @ s') obj"
+      using tau ev
+      apply (rule_tac x = "e#sb@sa" in exI)
+      by (rule_tac x = "File (new_childf pf \<tau>)" in exI, auto)
+  qed 
+next
+  case (ai'_init i t) 
+  hence initi: "i \<in> init_ipcs" using init_ipc_has_type
+    by (simp add:bidirect_in_init_def)
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
+      and nodels': "no_del_event s'"and intacts':"initp_intact_but s' sobj'"
+      and exso': "exists s' obj'"
+    from nodels' initi have exi: "i \<in> current_ipcs s'" 
+      by (drule_tac obj = "IPC i" in nodel_imp_exists, simp+)    
+    have "obj2sobj s' (IPC i) = SIPC t (Some i)"
+    proof-
+      have "obj2sobj [] (IPC i) = SIPC t (Some i)"  
+        using ai'_init initi by (auto simp:obj2sobj.simps)
+      thus ?thesis using vss' exi nodels' initi 
+        by (drule_tac obj2sobj_ipc_remains'', auto simp:obj2sobj.simps)
+    qed
+    thus "\<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 = SIPC t (Some i) \<and> exists (s @ s') obj"
+      apply (rule_tac x = "[]" in exI, rule_tac x = "IPC i" in exI)
+      by (simp add:vss' sobjs' nodels' exi exso' intacts' del:obj2sobj.simps)
+  qed
+next
+  case (ai'_cipc r fr pt u srp)
+  show ?case
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" 
+      and exobj':"exists s' obj'"
+    with ai'_cipc(1,2) notUkn obtain s p where
+      SPab: "obj2sobj (s@s') (Proc p) = SProc (r,fr,pt,u) srp" and
+      expab: "exists (s@s') (Proc p)" and vsab: "valid (s@s')" and
+      soab: "obj2sobj (s@s') obj' = sobj'" and 
+      exsoab: "exists (s@s') obj'" and
+      intactab: "initp_intact_but (s@s') sobj'" and
+      nodelab: "no_del_event (s@s')"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    obtain e \<tau> where ev: "e = CreateIPC p (new_ipc \<tau>)" and tau: "\<tau>=s@s'" by auto
+    
+    have valid: "valid (e # \<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau expab by (simp)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ai'_cipc(3) SPab
+        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
+      ultimately show ?thesis using vsab tau
+        by (rule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
+        by (case_tac obj', simp+)   moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'" 
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto simp:obj2sobj.simps cp2sproc_simps' 
+             simp del:cp2sproc.simps split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (IPC (new_ipc \<tau>)) = SIPC (default_ipc_create_type r) None"
+      using ev tau SPab nodel 
+        nodel_imp_exists[where obj = "IPC (new_ipc \<tau>)" and s =\<tau>]
+      by (auto simp:obj2sobj.simps ni_notin_curi cp2sproc.simps 
+              split:option.splits dest:no_del_event_cons_D)  moreover
+    have "initp_intact_but (e#\<tau>) sobj'" using intactab tau ev valid nodel
+        apply (case_tac sobj', case_tac option)
+        by (simp_all add:initp_intact_butp_I_others initp_intact_I_others)  
+    ultimately show "\<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 = SIPC (default_ipc_create_type r) None \<and> exists (s @ s') obj"
+      using tau ev
+      by (rule_tac x = "e#s" in exI, rule_tac x = "IPC (new_ipc \<tau>)" in exI, auto)
+  qed 
+next
+  case (ap'_init p r fr t u)  (* the big difference from other elims is in this case *)
+  hence initp: "p \<in> init_processes" using init_proc_has_role
+    by (simp add:bidirect_in_init_def)
+  show ?case
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume VSs': "valid s'" and SOs': "obj2sobj s' obj' = sobj'"
+      and Nodels': "no_del_event s'"and Intacts':"initp_intact_but s' sobj'"
+      and exso': "exists s' obj'" and notUkn: "sobj' \<noteq> Unknown"
+    from Nodels' initp have exp: "p \<in> current_procs s'" 
+      by (drule_tac obj = "Proc p" in nodel_imp_exists, simp+)
+    have "\<exists> p'. obj2sobj s' (Proc p') = SProc (r,fr,t,u) (Some p) \<and> p' \<in> current_procs s'"
+    proof (cases sobj')
+      case (SProc sp srp)
+      show ?thesis
+      proof (cases srp)
+        case None
+        with SProc Intacts' have "initp_intact s'" by simp
+        thus ?thesis using initp ap'_init
+          apply (rule_tac x = p in exI)
+          by (auto simp:initp_intact_def exp split:option.splits)
+      next
+        case (Some p')
+        show ?thesis
+        proof (cases "p' = p")
+          case True
+          with Intacts' SProc Some have "initp_alter s' p" 
+            by (simp add:initp_intact_butp_def)
+          then obtain pa where "pa \<in> current_procs s'"
+            and "obj2sobj s' (Proc pa) = init_obj2sobj (Proc p)"
+            by (auto simp only:initp_alter_def)
+          thus ?thesis using ap'_init initp 
+            by (rule_tac x = pa in exI, auto)
+        next
+          case False
+          with Intacts' SProc Some initp
+          have "obj2sobj s' (Proc p) = init_obj2sobj (Proc p)"
+            apply (simp only:initp_intact_butp_def initp_intact_but.simps)
+            by (erule conjE, erule_tac x = p in allE, simp)
+          thus ?thesis using ap'_init exp
+            by (rule_tac x = p in exI, auto split:option.splits)
+        qed
+      qed
+    next
+      case (SFile sf srf)
+      thus ?thesis using ap'_init exp Intacts' initp
+        by (rule_tac x = p in exI, auto split:option.splits simp:initp_intact_def)
+    next
+      case (SIPC si sri)
+      thus ?thesis using ap'_init exp Intacts' initp
+        by (rule_tac x = p in exI, auto split:option.splits simp:initp_intact_def)
+    next
+      case Unknown
+      thus ?thesis using notUkn by simp
+    qed
+    then obtain p' where "obj2sobj s' (Proc p') = SProc (r, fr, t, u) (Some p)"
+      and "p' \<in> current_procs s'" by blast
+    thus "\<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 = SProc (r, fr, t, u) (Some p) \<and> exists (s @ s') obj"
+      apply (rule_tac x = "[]" in exI, rule_tac x = "Proc p'" in exI)
+      by (simp add:VSs' SOs' Nodels' exp exso' Intacts') 
+  qed 
+next
+  case (ap'_crole r fr t u srp r')
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
+    with ap'_crole(1,2) obtain s p where
+      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
+      and nodelab: "no_del_event (s@s')"
+      and intactab: "initp_intact_but (s@s') sobj'"
+      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
+      and exp:"exists (s@s') (Proc p)" and exobj'ab:"exists (s@s') obj'"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    obtain e \<tau> where ev: "e = ChangeRole (new_proc (s@s')) r'" 
+      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
+    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
+    have np_not_initp: "new_proc (s@s') \<notin> init_processes" using nodelab 
+      apply (rule_tac notI, drule_tac obj = "Proc (new_proc (s@s'))" in nodel_imp_exists)
+      by (auto simp:np_notin_curp)
+
+    have valid: "valid (e#\<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau exp by (simp)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ap'_crole(3) SPab 
+        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
+      ultimately show ?thesis using vs_tau
+        by (erule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have "initp_intact_but (e#\<tau>) sobj'"
+    proof (cases sobj')
+      case (SProc sp srp)
+      show ?thesis
+      proof (cases srp)
+        case (Some p')
+        with SOab' exobj'ab VSab intactab notUkn SProc
+        have butp: "p' \<in> init_processes \<and> initp_intact_butp (s@s') p'"
+          by (case_tac obj', auto intro:source_proc_in_init simp:obj2sobj.simps
+                                split:if_splits option.splits)
+        then obtain p'' where exp': "p'' \<in> current_procs (s@s')" and 
+          SP': "obj2sobj (s@s') (Proc p'') = init_obj2sobj (Proc p')"
+          by (auto simp:initp_alter_def initp_intact_butp_def)
+        hence "initp_alter (e#\<tau>) p'" using tau ev notUkn nodel
+          apply (simp add:initp_alter_def del:init_obj2sobj.simps)
+          apply (rule_tac x = p'' in exI, rule conjI, simp)
+          apply (subgoal_tac "p'' \<noteq> new_proc (s @s')")
+          apply (auto simp:obj2sobj.simps cp2sproc.simps 
+                  simp del:init_obj2sobj.simps  split:option.splits)[1]
+          by (rule notI, simp add:np_notin_curp)
+        thus ?thesis using SProc Some intactab tau ev valid vs_tau np_not_initp
+          apply (simp add:initp_intact_butp_def del:init_obj2sobj.simps)
+          apply (rule impI|rule allI|rule conjI|erule conjE)+
+          apply (erule_tac x = pa in allE)
+          by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:init_obj2sobj.simps
+                  split:option.splits)
+      next
+        case None
+        with intactab SProc 
+        have "initp_intact (s@s')" by simp
+        hence "initp_intact (e#\<tau>)" using tau ev valid
+          by (simp add:initp_intact_I_crole)
+        thus ?thesis using SProc None by simp
+      qed
+    next
+      case (SFile sf srf)
+      with intactab SFile
+      have "initp_intact (s@s')" by simp
+      hence "initp_intact (e#\<tau>)" using tau ev valid
+        by (simp add:initp_intact_I_crole)
+      thus ?thesis using SFile by simp
+    next
+      case (SIPC si sri)
+      with intactab SIPC
+      have "initp_intact (s@s')" by simp
+      hence "initp_intact (e#\<tau>)" using tau ev valid
+        by (simp add:initp_intact_I_crole)
+      thus ?thesis using SIPC by simp
+    next
+      case Unknown 
+      with notUkn show ?thesis by simp
+    qed  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
+      by (case_tac obj', simp+)  moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'"
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by (auto)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by auto
+      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
+        apply (case_tac "p' = new_proc (s @ s')") 
+        using vs_tau exobj'ab tau
+        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
+        using tau ev SOab' valid notUkn vs_tau 
+        by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (Proc (new_proc (s@s'))) = SProc (r', fr, t, u) srp"
+      using SPab tau vs_tau ev valid
+      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
+              split:option.splits)  moreover
+    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp  
+    ultimately show "\<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 = SProc (r', fr, t, u) srp \<and> exists (s @ s') obj"
+      using ev tau
+      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
+      by (rule_tac x = "Proc (new_proc (s@s'))" in exI, auto)
+  qed 
+next
+  case (ap'_chown r fr t u srp u' nr)
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
+    with ap'_chown(1,2) obtain s p where
+      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
+      and nodelab: "no_del_event (s@s')" and intactab: "initp_intact_but (s@s') sobj'"
+      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
+      and exp:"exists (s@s') (Proc p)" and exobj'ab:"exists (s@s') obj'"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    obtain e \<tau> where ev: "e = ChangeOwner (new_proc (s@s')) u'" 
+      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
+    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
+    have np_not_initp: "new_proc (s@s') \<notin> init_processes" using nodelab 
+      apply (rule_tac notI, drule_tac obj = "Proc (new_proc (s@s'))" in nodel_imp_exists)
+      by (auto simp:np_notin_curp)
+
+    have valid: "valid (e#\<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau exp ap'_chown(3) by (simp)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ap'_chown(5) SPab 
+        by (auto simp:cp2sproc.simps obj2sobj.simps pct_def clone_type_unchange
+                split:option.splits t_rc_proc_type.splits)
+          (* here is another place of no_limit of clone event assumption *)
+      ultimately show ?thesis using vs_tau
+        by (erule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have "initp_intact_but (e#\<tau>) sobj'"
+    proof (cases sobj')
+      case (SProc sp srp)
+      show ?thesis
+      proof (cases srp)
+        case (Some p')
+        with SOab' exobj'ab VSab intactab notUkn SProc
+        have butp: "p' \<in> init_processes \<and> initp_intact_butp (s@s') p'"
+          by (case_tac obj', auto intro:source_proc_in_init simp:obj2sobj.simps
+                                split:if_splits option.splits)
+        then obtain p'' where exp': "p'' \<in> current_procs (s@s')" and 
+          SP': "obj2sobj (s@s') (Proc p'') = init_obj2sobj (Proc p')"
+          by (auto simp:initp_alter_def initp_intact_butp_def)
+        hence "initp_alter (e#\<tau>) p'" using tau ev notUkn nodel
+          apply (simp add:initp_alter_def del:init_obj2sobj.simps)
+          apply (rule_tac x = p'' in exI, rule conjI, simp)
+          apply (subgoal_tac "p'' \<noteq> new_proc (s @s')")
+          apply (auto simp:obj2sobj.simps cp2sproc.simps 
+                  simp del:init_obj2sobj.simps  split:option.splits)[1]
+          by (rule notI, simp add:np_notin_curp)
+        thus ?thesis using SProc Some intactab tau ev valid vs_tau np_not_initp
+          apply (simp add:initp_intact_butp_def del:init_obj2sobj.simps)
+          apply (rule impI|rule allI|rule conjI|erule conjE)+
+          apply (erule_tac x = pa in allE)
+          by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:init_obj2sobj.simps
+                  split:option.splits)
+      next
+        case None
+        with intactab SProc 
+        have "initp_intact (s@s')" by simp
+        hence "initp_intact (e#\<tau>)" using tau ev valid
+          by (simp add:initp_intact_I_chown)
+        thus ?thesis using SProc None by simp
+      qed
+    next
+      case (SFile sf srf)
+      with intactab SFile
+      have "initp_intact (s@s')" by simp
+      hence "initp_intact (e#\<tau>)" using tau ev valid
+        by (simp add:initp_intact_I_chown)
+      thus ?thesis using SFile by simp
+    next
+      case (SIPC si sri)
+      with intactab SIPC
+      have "initp_intact (s@s')" by simp
+      hence "initp_intact (e#\<tau>)" using tau ev valid
+        by (simp add:initp_intact_I_chown)
+      thus ?thesis using SIPC by simp
+    next
+      case Unknown 
+      with notUkn show ?thesis by simp
+    qed  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
+      by (case_tac obj', simp+)  moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'"
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by (auto)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by auto
+      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
+        apply (case_tac "p' = new_proc (s @ s')") 
+        using vs_tau exobj'ab tau
+        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
+        using tau ev SOab' valid notUkn vs_tau 
+        by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (Proc (new_proc (s@s'))) = 
+          SProc (nr,fr,chown_type_aux r nr t,u') srp"
+      using SPab tau vs_tau ev valid ap'_chown(4)
+      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
+              split:option.splits)  moreover
+    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp   moreover 
+    ultimately show "\<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 = SProc (nr,fr,chown_type_aux r nr t,u') srp \<and>
+      exists (s @ s') obj"
+      using ev tau
+      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
+      by (rule_tac x = "Proc (new_proc (s@s'))" in exI, auto)
+  qed
+next
+  case (ap'_exec r fr pt u sp t sd sf r' fr')
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and "initp_intact_but s' sobj'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
+    with ap'_exec(3,4) obtain sa f where
+      SFa: "obj2sobj (sa @ s') (File f) = SFile (t, sd) sf" and
+      Exfa: "exists (sa @ s') (File f)" and 
+      butsa: "initp_intact_but (sa @ s') sobj'" and
+      othersa:"valid (sa @ s') \<and> obj2sobj (sa @ s') obj' = sobj' \<and> 
+               exists (sa @s') obj' \<and> no_del_event (sa @ s')" 
+      by (blast dest:obj2sobj_file intro:nodel_exists_remains)
+    with ap'_exec(1,2) notUkn obtain sb p where
+      VSab: "valid (sb@sa@s')" and SOab': "obj2sobj (sb@sa@s') obj' = sobj'"
+      and nodelab: "no_del_event (sb@sa@s')"
+      and intactab: "initp_intact_but (sb@sa@s') sobj'"
+      and SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r, fr, pt, u) sp"
+      and exp:"exists (sb@sa@s') (Proc p)" and exobj'ab:"exists (sb@sa@s') obj'"
+      by (blast dest:obj2sobj_proc intro:nodel_exists_remains)
+    obtain e \<tau> where ev: "e = Execute (new_proc (sb@sa@s')) f" 
+      and tau: "\<tau> = Clone p (new_proc (sb@sa@s'))#sb@sa@s'" by auto
+    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
+    from Exfa nodelab have exf:"f \<in> current_files (sb@sa@s')" 
+      apply (drule_tac obj = "File f" in nodel_imp_un_deleted)
+      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
+    from SFa VSab Exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File f) = SFile (t,sd) sf"
+      by (rule_tac s = "sa@s'" in obj2sobj_file_remains'', simp_all)
+    have np_not_initp: "new_proc (sb@sa@s') \<notin> init_processes" using nodelab 
+      apply (rule_tac notI, drule_tac obj = "Proc (new_proc (sb@sa@s'))" in nodel_imp_exists)
+      by (auto simp:np_notin_curp)
+
+    have valid: "valid (e#\<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau exp by (simp add:exf)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ap'_exec(5) SPab SFab
+        by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
+                split:if_splits option.splits)
+      ultimately show ?thesis using vs_tau
+        by (erule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have "initp_intact_but (e#\<tau>) sobj'"
+    proof (cases sobj')
+      case (SProc sp srp)
+      show ?thesis
+      proof (cases srp)
+        case (Some p')
+        with SOab' exobj'ab VSab intactab notUkn SProc
+        have butp: "p' \<in> init_processes \<and> initp_intact_butp (sb@sa@s') p'"
+          by (case_tac obj', auto intro:source_proc_in_init simp:obj2sobj.simps
+                                split:if_splits option.splits)
+        then obtain p'' where exp': "p'' \<in> current_procs (sb@sa@s')" and 
+          SP': "obj2sobj (sb@sa@s') (Proc p'') = init_obj2sobj (Proc p')"
+          by (auto simp:initp_alter_def initp_intact_butp_def)
+        hence "initp_alter (e#\<tau>) p'" using tau ev notUkn nodel
+          apply (simp add:initp_alter_def del:init_obj2sobj.simps)
+          apply (rule_tac x = p'' in exI, rule conjI, simp)
+          apply (subgoal_tac "p'' \<noteq> new_proc (sb@sa@s')")
+          apply (auto simp:obj2sobj.simps cp2sproc.simps 
+                  simp del:init_obj2sobj.simps  split:option.splits)[1]
+          by (rule notI, simp add:np_notin_curp)
+        thus ?thesis using SProc Some intactab tau ev valid vs_tau np_not_initp
+          apply (simp add:initp_intact_butp_def del:init_obj2sobj.simps)
+          apply (rule impI|rule allI|rule conjI|erule conjE)+
+          apply (erule_tac x = pa in allE)
+          by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:init_obj2sobj.simps
+                  split:option.splits)
+      next
+        case None
+        with intactab SProc 
+        have "initp_intact (sb@sa@s')" by simp
+        hence "initp_intact (e#\<tau>)" using tau ev valid
+          by (simp add:initp_intact_I_exec)
+        thus ?thesis using SProc None by simp
+      qed
+    next
+      case (SFile sf srf)
+      with intactab SFile
+      have "initp_intact (sb@sa@s')" by simp
+      hence "initp_intact (e#\<tau>)" using tau ev valid
+        by (simp add:initp_intact_I_exec)
+      thus ?thesis using SFile by simp
+    next
+      case (SIPC si sri)
+      with intactab SIPC
+      have "initp_intact (sb@sa@s')" by simp
+      hence "initp_intact (e#\<tau>)" using tau ev valid
+        by (simp add:initp_intact_I_exec)
+      thus ?thesis using SIPC by simp
+    next
+      case Unknown 
+      with notUkn show ?thesis by simp
+    qed  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exobj'ab valid ev tau
+      by (case_tac obj', simp+) moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'"
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
+        by (auto simp del:obj2sobj.simps)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
+        by (auto simp del:obj2sobj.simps)
+      moreover have "\<And> p'. obj' = Proc p' \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"
+        apply (case_tac "p' = new_proc (sb @ sa @ s')") 
+        using vs_tau exobj'ab tau
+        apply (simp, drule_tac valid_os, simp add:np_notin_curp)
+        using tau ev SOab' valid notUkn vs_tau 
+        by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (Proc (new_proc (sb @ sa @ s'))) = 
+          SProc (r',fr',exec_type_aux r pt, u) sp"
+    proof-
+      have "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) sp" using SPab tau vs_tau
+        by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
+      hence "obj2sobj \<tau> (Proc (new_proc (sb@sa@s'))) = SProc (r,fr,pt,u) sp" using tau
+      by (auto simp:obj2sobj.simps cp2sproc.simps pct_def clone_type_unchange 
+              split:option.splits)
+      moreover have "source_dir \<tau> f = Some sd" using vs_tau SFab tau 
+        by (auto simp:source_dir_simps obj2sobj.simps split:option.splits if_splits)
+      ultimately show ?thesis using valid ev ap'_exec(6,7) 
+        by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
+    qed 
+    ultimately  show "\<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 = SProc (r', fr', exec_type_aux r pt, u) sp \<and>
+      exists (s @ s') obj"
+      using ev tau
+      apply (rule_tac x = "e#Clone p (new_proc (sb @ sa @ s')) #sb@sa" in exI)
+      by (rule_tac x = "Proc (new_proc (sb @ sa @ s'))" in exI, auto)
+  qed
+qed
+
+(* this is for all_sobjs_E2 *)
+lemma all_sobjs_E1:
+  "\<lbrakk>sobj \<in> all_sobjs'; valid s'; obj2sobj s' obj' = sobj'; exists s' obj'; sobj' \<noteq> Unknown; 
+    no_del_event s'; initp_intact_but s' sobj'\<rbrakk>
+   \<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"
+by (drule all_sobjs_E1_aux, auto)
+
+
+lemma all_sobjs_E2_aux[rule_format]:
+  "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))"
+proof (induct rule:all_sobjs'.induct)
+  case (af'_init f t) show ?case
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
+      and nodels': "no_del_event s'"and intacts':"initp_intact s'"
+      and notboth: "not_both_sproc (SFile (t, f) (Some f)) sobj'"
+      and exso': "exists s' obj'"
+    from nodels' af'_init(1) have exf: "f \<in> current_files s'" 
+      by (drule_tac obj = "File f" in nodel_imp_exists, simp+)    
+    have "obj2sobj s' (File f) = SFile (t, f) (Some f)"
+    proof-
+      have "obj2sobj [] (File f) = SFile (t, f) (Some f)"  using af'_init 
+        by (auto simp:etype_of_file_def source_dir_of_init' obj2sobj.simps
+                split:option.splits)
+      thus ?thesis using vss' exf nodels' af'_init(1) 
+        by (drule_tac obj2sobj_file_remains_app', auto)
+    qed
+    thus "\<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') (SFile (t, f) (Some f)) \<and>
+             obj2sobj (s @ s') obj = SFile (t, f) (Some f) \<and>
+             exists (s @ s') obj \<and> sobj_source_eq_obj (SFile (t, f) (Some f)) obj"
+      apply (rule_tac x = "[]" in exI, rule_tac x = "File f" in exI)
+      by (simp add:vss' sobjs' nodels' intacts' exf exso')
+  qed
+next
+  case (af'_cfd t sd srf r fr pt u srp t') 
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and Both:"not_both_sproc (SFile (t', sd) None) sobj'"
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" 
+      and exobj':"exists s' obj'"
+    with af'_cfd(1,2) obtain sa pf where
+      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
+      "exists (sa@s') obj'" and "initp_intact (sa@s')" and
+      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
+      exfa: "pf \<in> current_files (sa@s')" 
+      apply (drule_tac sf' = "(t, sd)" and srf' = srf in not_both_I_file)
+      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
+      by (frule obj2sobj_file, auto)
+    with af'_cfd(3) notUkn obtain sb p where
+      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
+      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
+      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
+      exsoab: "exists (sb@sa@s') obj'" and
+      intactab: "initp_intact (sb@sa@s')" and
+      nodelab: "no_del_event (sb@sa@s')"
+      apply (drule_tac s'= "sa@s'" and obj' = obj' in all_sobjs_E0, auto) 
+      apply (frule obj2sobj_proc, erule exE)
+      by (auto intro:nodel_exists_remains)
+    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
+      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
+      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
+    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
+      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
+    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
+      and tau: "\<tau>=sb@sa@s'" by auto
+    
+    have valid: "valid (e # \<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau af'_cfd(5,6,7) SPab SFab
+        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
+                split:if_splits option.splits t_rc_file_type.splits)
+      ultimately show ?thesis using vsab tau
+        by (rule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
+        by (case_tac obj', simp+)   moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'" 
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto simp:obj2sobj.simps cp2sproc_simps' 
+             simp del:cp2sproc.simps split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
+      by (simp add:initp_intact_I_others)  moreover
+    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t', sd) None"
+    proof-
+      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t'"
+        using ev tau SFab SPab af'_cfd(5)
+        by (auto simp:obj2sobj.simps etype_of_file_def cp2sproc.simps
+          split:option.splits if_splits  intro!:etype_aux_prop4)
+      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
+        using ev tau SFab SPab valid ncf_parent
+        by (auto simp:source_dir_simps obj2sobj.simps 
+                split:if_splits option.splits)
+      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
+        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
+        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
+    qed
+    ultimately show "\<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') (SFile (t', sd) None) \<and>
+      obj2sobj (s @ s') obj = SFile (t', sd) None \<and>
+      exists (s @ s') obj \<and> sobj_source_eq_obj (SFile (t', sd) None) obj"
+      using tau ev
+      apply (rule_tac x = "e#sb@sa" in exI)
+      by (rule_tac x = "File (new_childf pf \<tau>)" in exI, auto)
+  qed
+next
+  case (af'_cfd' t sd srf r fr pt u srp)
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and Both:"not_both_sproc (SFile (t, sd) None) sobj'"
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" 
+      and exobj':"exists s' obj'"
+    with af'_cfd'(1,2) obtain sa pf where
+      "valid (sa@s')" and "obj2sobj (sa@s') obj'=sobj' \<and> no_del_event (sa@s')" and
+      "exists (sa@s') obj'" and "initp_intact (sa@s')" and
+      SFa: "obj2sobj (sa@s') (File pf) = SFile (t, sd) srf" and
+      exfa: "pf \<in> current_files (sa@s')" 
+      apply (drule_tac sf' = "(t, sd)" and srf' = srf in not_both_I_file)
+      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE, auto)
+      by (frule obj2sobj_file, auto)
+    with af'_cfd'(3) notUkn obtain sb p where
+      SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r,fr,pt,u) srp" and
+      expab: "exists (sb@sa@s') (Proc p)" and vsab: "valid (sb@sa@s')" and
+      soab: "obj2sobj (sb@sa@s') obj' = sobj'" and 
+      exsoab: "exists (sb@sa@s') obj'" and
+      intactab: "initp_intact (sb@sa@s')" and
+      nodelab: "no_del_event (sb@sa@s')"
+      apply (drule_tac s'= "sa@s'" and obj' = obj' in all_sobjs_E0, auto) 
+      apply (frule obj2sobj_proc, erule exE)
+      by (auto intro:nodel_exists_remains)
+    from exfa nodelab have exf:"pf \<in> current_files (sb@sa@s')"
+      apply (drule_tac obj = "File pf" in nodel_imp_un_deleted)
+      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
+    from SFa vsab exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File pf) = SFile (t,sd) srf"
+      by (rule_tac s = "sa@s'" in obj2sobj_file_remains_app', simp_all)
+    obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)" 
+      and tau: "\<tau>=sb@sa@s'" by auto
+    
+    have valid: "valid (e # \<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau expab exf by (simp add:ncf_notin_curf ncf_parent)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau af'_cfd'(5,6) SPab SFab
+        by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
+                split:if_splits option.splits t_rc_file_type.splits)
+      ultimately show ?thesis using vsab tau
+        by (rule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
+        by (case_tac obj', simp+)   moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'" 
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto simp:obj2sobj.simps cp2sproc_simps' 
+             simp del:cp2sproc.simps split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
+      by (simp add:initp_intact_I_others)  moreover
+    have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t, sd) None"
+    proof-
+      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t"
+      proof-
+        have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = etype_of_file \<tau> pf"
+          using ev tau SPab af'_cfd'(5) 
+          by (auto simp:obj2sobj.simps cp2sproc.simps ncf_parent etype_of_file_def
+            split:option.splits   intro!:etype_aux_prop3)
+        thus ?thesis using SFab tau ev
+          by (auto simp:obj2sobj.simps split:option.splits if_splits)
+      qed
+      moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
+        using ev tau SFab SPab valid ncf_parent
+        by (auto simp:source_dir_simps obj2sobj.simps 
+                split:if_splits option.splits)
+      ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
+        nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
+        by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
+    qed
+    ultimately show "\<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') (SFile (t, sd) None) \<and>
+      obj2sobj (s @ s') obj = SFile (t, sd) None \<and>
+      exists (s @ s') obj \<and> sobj_source_eq_obj (SFile (t, sd) None) obj"
+      using tau ev
+      apply (rule_tac x = "e#sb@sa" in exI)
+      by (rule_tac x = "File (new_childf pf \<tau>)" in exI, auto)
+  qed 
+next
+  case (ai'_init i t) 
+  hence initi: "i \<in> init_ipcs" using init_ipc_has_type
+    by (simp add:bidirect_in_init_def)
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume vss': "valid s'" and sobjs': "obj2sobj s' obj' = sobj'"
+      and nodels': "no_del_event s'"and intacts':"initp_intact s'"
+      and notboth: "not_both_sproc (SIPC t (Some i)) sobj'"
+      and exso': "exists s' obj'"
+    from nodels' initi have exi: "i \<in> current_ipcs s'" 
+      by (drule_tac obj = "IPC i" in nodel_imp_exists, simp+)    
+    have "obj2sobj s' (IPC i) = SIPC t (Some i)"
+    proof-
+      have "obj2sobj [] (IPC i) = SIPC t (Some i)"  
+        using ai'_init initi by (auto simp:obj2sobj.simps)
+      thus ?thesis using vss' exi nodels' initi
+        by (drule_tac obj2sobj_ipc_remains'', auto simp:obj2sobj.simps)
+    qed
+    thus "\<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') (SIPC t (Some i)) \<and>
+             obj2sobj (s @ s') obj = SIPC t (Some i) \<and>
+             exists (s @ s') obj \<and> sobj_source_eq_obj (SIPC t (Some i)) obj"
+      apply (rule_tac x = "[]" in exI, rule_tac x = "IPC i" in exI)
+      by (simp add:vss' sobjs' nodels' intacts' exi exso' del:obj2sobj.simps)
+  qed
+next
+  case (ai'_cipc r fr pt u srp)
+  show ?case
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and Both:"not_both_sproc (SIPC (default_ipc_create_type r) None) sobj'"
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" 
+      and exobj':"exists s' obj'"
+    with ai'_cipc(1) notUkn obtain s p where
+      SPab: "obj2sobj (s@s') (Proc p) = SProc (r,fr,pt,u) srp" and
+      expab: "exists (s@s') (Proc p)" and vsab: "valid (s@s')" and
+      soab: "obj2sobj (s@s') obj' = sobj'" and 
+      exsoab: "exists (s@s') obj'" and
+      intactab: "initp_intact (s@s')" and
+      nodelab: "no_del_event (s@s')"
+      apply (drule_tac s'= "s'" and obj' = obj' in all_sobjs_E0, auto) 
+      apply (frule obj2sobj_proc, erule exE)
+      by (auto intro:nodel_exists_remains)
+    obtain e \<tau> where ev: "e = CreateIPC p (new_ipc \<tau>)" and tau: "\<tau>=s@s'" by auto
+    
+    have valid: "valid (e # \<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau expab by (simp)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ai'_cipc(3) SPab
+        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
+      ultimately show ?thesis using vsab tau
+        by (rule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exsoab valid ev tau
+        by (case_tac obj', simp+)   moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'" 
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto  intro!:obj2sobj_file_remains' simp:ncf_notin_curf)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto intro!:obj2sobj_ipc_remains' simp:ncf_notin_curf)
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'"  
+        using soab tau valid notUkn nodel ev exsoab
+        by (auto simp:obj2sobj.simps cp2sproc_simps' 
+             simp del:cp2sproc.simps split:option.splits)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "initp_intact (e#\<tau>)" using intactab tau ev valid nodel
+      by (simp add:initp_intact_I_others)  moreover
+    have "obj2sobj (e#\<tau>) (IPC (new_ipc \<tau>)) = SIPC (default_ipc_create_type r) None"
+      using ev tau SPab nodel 
+        nodel_imp_exists[where obj = "IPC (new_ipc \<tau>)" and s =\<tau>]
+      by (auto simp:obj2sobj.simps ni_notin_curi cp2sproc.simps 
+              split:option.splits  dest:no_del_event_cons_D)
+    ultimately show "\<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') (SIPC (default_ipc_create_type r) None) \<and>
+      obj2sobj (s @ s') obj = SIPC (default_ipc_create_type r) None \<and>
+      exists (s @ s') obj \<and> sobj_source_eq_obj (SIPC (default_ipc_create_type r) None) obj"
+      using tau ev
+      by (rule_tac x = "e#s" in exI, rule_tac x = "IPC (new_ipc \<tau>)" in exI, auto)
+  qed 
+next
+  case (ap'_init p r fr t u)
+  hence initp: "p \<in> init_processes" using init_proc_has_role
+    by (simp add:bidirect_in_init_def)
+  show ?case
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume VSs': "valid s'" and SOs': "obj2sobj s' obj' = sobj'"
+      and Nodels': "no_del_event s'"and Intacts':"initp_intact s'"
+      and notboth: "not_both_sproc (SProc (r,fr,t,u) (Some p)) sobj'"
+      and exso': "exists s' obj'"
+    from Nodels' initp have exp: "p \<in> current_procs s'" 
+      apply (drule_tac obj = "Proc p" in nodel_imp_un_deleted)
+      by (drule not_deleted_imp_exists, simp+)
+    with Intacts' initp ap'_init 
+    have "obj2sobj s' (Proc p) = SProc (r, fr, t, u) (Some p)"
+      by (auto simp:initp_intact_def split:option.splits)
+    thus "\<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') (SProc (r, fr, t, u) (Some p)) \<and>
+             obj2sobj (s @ s') obj = SProc (r, fr, t, u) (Some p) \<and>
+             exists (s @ s') obj \<and> 
+             sobj_source_eq_obj (SProc (r, fr, t, u) (Some p)) obj"
+      apply (rule_tac x = "[]" in exI, rule_tac x = "Proc p" in exI)
+      by (simp add:VSs' SOs' Nodels' Intacts' initp intact_imp_butp exp exso' 
+               del:obj2sobj.simps)
+  qed
+next
+  case (ap'_crole r fr t u srp r')
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and Both:"not_both_sproc (SProc (r', fr, t, u) srp) sobj'"
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
+    with ap'_crole(1,2) obtain s p where
+      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
+      and nodelab: "no_del_event (s@s')"
+      and butab: "initp_intact_but (s@s') (SProc (r, fr, t, u) srp)"
+      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
+      and exp:"exists (s@s') (Proc p)" and exobj':"exists (s@s') obj'"
+      and sreq: "sobj_source_eq_obj (SProc (r, fr, t, u) srp) (Proc p)"
+      by (blast dest:not_both_I obj2sobj_proc intro:nodel_exists_remains)
+    from VSab SPab sreq exp have srpeq: "srp = Some p" 
+      by (simp add:proc_source_eq_prop)
+    with exp VSab SPab have initp: "p \<in> init_processes"
+      by (auto dest:source_proc_in_init simp:obj2sobj.simps split:option.splits) 
+    obtain e \<tau> where ev: "e = ChangeRole p r'" 
+      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
+    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
+
+    have valid: "valid (e#\<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau exp by (simp)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ap'_crole(3) SPab 
+        by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
+      ultimately show ?thesis using vs_tau
+        by (erule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have "initp_intact_but (e#\<tau>) (SProc (r', fr, t, u) srp)"
+      using butab tau ev valid initp srpeq nodel
+      by (simp add:initp_intact_butp_I_crole)  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exobj' valid ev tau
+      by (case_tac obj', simp+)  moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'"
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by (auto)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by auto
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow>False"
+        using Both SOab' notUkn 
+        by (simp del:obj2sobj.simps, drule_tac obj2sobj_proc', auto)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (Proc p) = SProc (r', fr, t, u) srp"
+      using SPab tau vs_tau ev valid
+      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
+              split:option.splits)  moreover
+    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp   moreover 
+    have "sobj_source_eq_obj (SProc (r', fr, t, u) srp) (Proc p)"
+      using sreq by (case_tac srp, simp+)
+    ultimately show "\<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') (SProc (r', fr, t, u) srp) \<and> 
+      obj2sobj (s @ s') obj = SProc (r', fr, t, u) srp \<and>
+      exists (s @ s') obj \<and> sobj_source_eq_obj (SProc (r', fr, t, u) srp) obj"
+      using ev tau
+      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
+      by (rule_tac x = "Proc p" in exI, auto)
+  qed
+next
+  case (ap'_chown r fr t u srp u' nr)
+  show ?case 
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and Both:"not_both_sproc (SProc (nr,fr,chown_type_aux r nr t,u') srp) sobj'"
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
+    with ap'_chown(1,2) obtain s p where
+      VSab: "valid (s@s')" and SOab': "obj2sobj (s@s') obj' = sobj'"
+      and nodelab: "no_del_event (s@s')"
+      and butab: "initp_intact_but (s@s') (SProc (r, fr, t, u) srp)"
+      and SPab: "obj2sobj (s@s') (Proc p) = SProc (r, fr, t, u) srp"
+      and exp:"exists (s@s') (Proc p)" and exobj':"exists (s@s') obj'"
+      and sreq: "sobj_source_eq_obj (SProc (r, fr, t, u) srp) (Proc p)"
+      by (blast dest:not_both_I obj2sobj_proc intro:nodel_exists_remains)
+    from VSab SPab sreq exp have srpeq: "srp = Some p" 
+      by (simp add:proc_source_eq_prop)
+    with exp VSab SPab have initp: "p \<in> init_processes"
+      by (auto dest:source_proc_in_init simp:obj2sobj.simps split:option.splits) 
+    obtain e \<tau> where ev: "e = ChangeOwner p u'" 
+      and tau: "\<tau> = Clone p (new_proc (s@s'))#s@s'" by auto
+    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
+
+    have valid: "valid (e#\<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau exp ap'_chown(3) by (simp)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ap'_chown(5) SPab 
+        by (auto simp:cp2sproc.simps obj2sobj.simps pct_def clone_type_unchange
+                split:option.splits t_rc_proc_type.splits)
+          (* here is another place of no_limit of clone event assumption *)
+      ultimately show ?thesis using vs_tau
+        by (erule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have "initp_intact_but (e#\<tau>) (SProc (nr,fr,chown_type_aux r nr t,u') srp)"
+      using butab tau ev valid initp srpeq nodel
+      by (simp add:initp_intact_butp_I_chown)  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exobj' valid ev tau
+      by (case_tac obj', simp+)  moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'"
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by (auto)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (s @ s'))]"]
+        by auto
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow>False"
+        using Both SOab' notUkn 
+        by (simp del:obj2sobj.simps, drule_tac obj2sobj_proc', auto)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (Proc p) = SProc (nr,fr,chown_type_aux r nr t,u') srp"
+      using SPab tau vs_tau ev valid ap'_chown(4)
+      by (auto simp:obj2sobj.simps cp2sproc_simps' simp del:cp2sproc.simps
+              split:option.splits)  moreover
+    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp   moreover 
+    have "sobj_source_eq_obj (SProc (nr,fr,chown_type_aux r nr t,u') srp) (Proc p)"
+      using sreq by (case_tac srp, simp+)
+    ultimately show "\<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') (SProc (nr,fr,chown_type_aux r nr t,u') srp) \<and> 
+      obj2sobj (s @ s') obj = SProc (nr,fr,chown_type_aux r nr t,u') srp \<and>
+      exists (s @ s') obj \<and> sobj_source_eq_obj (SProc (nr,fr,chown_type_aux r nr t,u') srp) obj"
+      using ev tau
+      apply (rule_tac x = "e # Clone p (new_proc (s @ s')) # s" in exI)
+      by (rule_tac x = "Proc p" in exI, auto)
+  qed
+next
+  case (ap'_exec r fr pt u sp t sd sf r' fr')
+  show ?case
+  proof (rule allI|rule impI|erule conjE)+
+    fix s' obj' sobj'
+    assume "valid s'" and "obj2sobj s' obj' = sobj'" and "no_del_event s'" 
+      and Both:"not_both_sproc (SProc (r', fr', exec_type_aux r pt, u) sp) sobj'"
+      and "initp_intact s'" and notUkn: "sobj' \<noteq> Unknown" and "exists s' obj'"
+    with ap'_exec(3,4) obtain sa f where
+      SFa: "obj2sobj (sa @ s') (File f) = SFile (t, sd) sf" and
+      Exfa: "exists (sa @ s') (File f)" and 
+      buta: "initp_intact (sa @ s')" and
+      othersa:"valid (sa @ s') \<and> obj2sobj (sa @ s') obj' = sobj' \<and> exists (sa @s') obj' \<and>
+      no_del_event (sa @ s') \<and> sobj_source_eq_obj (SFile (t, sd) sf) (File f)" 
+      apply (simp only:not_both_sproc.simps)
+      apply (erule_tac x = s' in allE, erule_tac x = obj' in allE)
+      apply (erule_tac x = sobj' in allE, auto)
+      by (frule obj2sobj_file, auto intro:nodel_exists_remains)
+    with SFa Exfa othersa ap'_exec(1,2) Both notUkn obtain sb p where
+      VSab: "valid (sb@sa@s')" and SOab': "obj2sobj (sb@sa@s') obj' = sobj'"
+      and nodelab: "no_del_event (sb@sa@s')"
+      and butab: "initp_intact_but (sb@sa@s') (SProc (r, fr, pt, u) sp)"
+      and SPab: "obj2sobj (sb@sa@s') (Proc p) = SProc (r, fr, pt, u) sp"
+      and exp:"exists (sb@sa@s') (Proc p)" and exobj':"exists (sb@sa@s') obj'"
+      and sreq: "sobj_source_eq_obj (SProc (r, fr, pt, u) sp) (Proc p)"
+      by (blast dest:not_both_I obj2sobj_proc intro:nodel_exists_remains)
+    from VSab SPab sreq exp have srpeq: "sp = Some p" by (simp add:proc_source_eq_prop)
+    with exp VSab SPab have initp: "p \<in> init_processes"
+      by (auto dest:source_proc_in_init simp:obj2sobj.simps split:option.splits) 
+    obtain e \<tau> where ev: "e = Execute p f" 
+      and tau: "\<tau> = Clone p (new_proc (sb@sa@s'))#sb@sa@s'" by auto
+    hence vs_tau:"valid \<tau>" using exp VSab by (auto intro:clone_event_no_limit)
+    from Exfa nodelab have exf:"f \<in> current_files (sb@sa@s')" 
+      apply (drule_tac obj = "File f" in nodel_imp_un_deleted)
+      by (drule_tac s' = "sb" in not_deleted_imp_exists', auto)
+    from SFa VSab Exfa nodelab have SFab: "obj2sobj (sb@sa@s') (File f) = SFile (t,sd) sf"
+      by (rule_tac s = "sa@s'" in obj2sobj_file_remains'', simp_all)
+
+    have valid: "valid (e#\<tau>)" 
+    proof-
+      have "os_grant \<tau> e"
+        using ev tau exp by (simp add:exf)
+      moreover have "rc_grant \<tau> e" 
+        using ev tau ap'_exec(5) SPab SFab
+        by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
+                split:if_splits option.splits)
+      ultimately show ?thesis using vs_tau
+        by (erule_tac vs_step, simp+)
+    qed  moreover
+    have nodel: "no_del_event (e#\<tau>)" using nodelab tau ev by simp  moreover
+    have "initp_intact_but (e#\<tau>) (SProc (r', fr', exec_type_aux r pt, u) sp)"
+      using butab tau ev valid initp srpeq nodel
+      by (simp add:initp_intact_butp_I_exec)  moreover
+    have exobj': "exists (e#\<tau>) obj'" using exobj' valid ev tau
+      by (case_tac obj', simp+) moreover
+    have "obj2sobj (e#\<tau>) obj' = sobj'"
+    proof-
+      have "\<And> f. obj' = File f \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_file_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
+        by (auto simp del:obj2sobj.simps)
+      moreover have "\<And> i. obj' = IPC i \<Longrightarrow> obj2sobj (e#\<tau>) obj' = sobj'" 
+        using SOab' tau ev valid notUkn nodel exobj'
+        obj2sobj_ipc_remains''[where s'="[e,Clone p (new_proc (sb @ sa @ s'))]"]
+        by (auto simp del:obj2sobj.simps)
+      moreover have "\<And> p. obj' = Proc p \<Longrightarrow>False"
+        using Both SOab' notUkn 
+        by (simp del:obj2sobj.simps, drule_tac obj2sobj_proc', auto)
+      ultimately show ?thesis by (case_tac obj', auto)
+    qed  moreover 
+    have "obj2sobj (e#\<tau>) (Proc p) = SProc (r',fr',exec_type_aux r pt, u) sp"
+    proof-
+      have "obj2sobj \<tau> (Proc p) = SProc (r,fr,pt,u) sp" using SPab tau vs_tau
+        by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
+      moreover have "source_dir \<tau> f = Some sd" using vs_tau SFab tau 
+        by (auto simp:source_dir_simps obj2sobj.simps split:option.splits if_splits)
+      ultimately show ?thesis using valid ev ap'_exec(6,7) 
+        by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
+    qed  moreover
+    have "exists (e#\<tau>) (Proc p)" using exp tau ev by simp   moreover 
+    have "sobj_source_eq_obj (SProc (r',fr',exec_type_aux r pt,u) sp) (Proc p)"
+      using sreq by (case_tac sp, simp+)
+    ultimately 
+    show "\<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') (SProc (r', fr', exec_type_aux r pt, u) sp) \<and> 
+      obj2sobj (s @ s') obj = SProc (r', fr', exec_type_aux r pt, u) sp \<and>
+      exists (s @ s') obj \<and>
+      sobj_source_eq_obj (SProc (r', fr', exec_type_aux r pt, u) sp) obj"
+      using ev tau
+      apply (rule_tac x = "e#Clone p (new_proc (sb @ sa @ s')) #sb@sa" in exI)
+      by (rule_tac x = "Proc p" in exI, auto)
+  qed
+qed
+
+lemma all_sobjs_E2: 
+  "\<lbrakk>sobj \<in> all_sobjs'; valid s'; obj2sobj s' obj' = sobj'; exists s' obj'; sobj' \<noteq> Unknown; 
+    not_both_sproc sobj sobj'; no_del_event s'; initp_intact s'\<rbrakk>
+   \<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"
+by (drule all_sobjs_E2_aux, auto)
+
+end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/del_vs_del_s.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,277 @@
+theory del_vs_del_s
+imports Main rc_theory os_rc obj2sobj_prop all_sobj_prop 
+begin
+
+context tainting_s_complete begin
+
+lemma deleted_has_del_event_proc:
+  "\<lbrakk>deleted (Proc p) s; valid s\<rbrakk> \<Longrightarrow> \<exists> s' p'. Kill p' p # s' \<preceq> s \<and> \<not> deleted (Proc p) s'"
+apply (induct s, simp)
+apply (frule valid_cons, frule valid_os)
+by (case_tac a, auto simp:no_junior_def)
+
+lemma deleted_has_del_event_ipc:
+  "\<lbrakk>deleted (IPC i) s; valid s\<rbrakk> \<Longrightarrow> \<exists> s' p. DeleteIPC p i # s' \<preceq> s \<and> \<not> deleted (IPC i) s'"
+apply (induct s, simp)
+apply (frule valid_cons, frule valid_os)
+by (case_tac a, auto simp:no_junior_def)
+
+lemma deleted_has_del_event_file:
+  "\<lbrakk>deleted (File f) s; valid s\<rbrakk> \<Longrightarrow> \<exists> s' p. DeleteFile p f # s' \<preceq> s \<and> \<not> deleted (File f) s'"
+apply (induct s, simp)
+apply (frule valid_cons, frule valid_os)
+by (case_tac a, auto simp:no_junior_def)
+
+(*
+lemma noJ_Anc: "x \<prec> y = (x \<preceq> y \<and> x \<noteq> y)"
+apply (simp add: no_junior_expand)
+by (auto simp:is_ancestor_def)
+
+lemma noJ_Anc': "x \<prec> y \<Longrightarrow> (x \<preceq> y \<and> x \<noteq> y)"
+by (simp add:noJ_Anc)
+
+lemma noJ_Anc'': "\<lbrakk>x \<preceq> y; x \<noteq> y\<rbrakk> \<Longrightarrow> x \<prec> y"
+by (simp add:noJ_Anc)
+
+lemma deled_imp_no_childfs:
+  "\<lbrakk>valid (DeleteFile p f # s); f \<prec> childf\<rbrakk> \<Longrightarrow> childf \<notin> current_files s"
+apply (frule valid_cons, drule valid_os, rule notI, clarsimp dest!:noJ_Anc')
+apply (drule ancient_has_parent, simp, clarsimp)
+apply (drule_tac af = sonf and f = childf in ancient_file_in_current, simp+)
+apply (case_tac sonf, simp)
+apply (erule_tac x = a in allE, simp)
+done
+
+lemma deled_imp_childfs_deleted:
+  "\<lbrakk>valid (DeleteFile p f # s); f \<prec> childf; childf \<in> init_files\<rbrakk>
+  \<Longrightarrow> deleted (File childf) s"
+apply (drule deled_imp_no_childfs, simp+)
+apply (erule_tac P = "childf \<in> current_files s" in swap)
+by (drule not_deleted_imp_exists, simp+)
+
+lemma initf_deled_imp_childf_deled:
+  "\<lbrakk>deleted (File f) s; valid s\<rbrakk> \<Longrightarrow> \<exists> s' p. DeleteFile p f # s' \<preceq> s \<and> \<not> deleted (File f) s' \<and> (\<forall> childf \<in> init_files. f \<prec> childf \<longrightarrow> deleted (File childf) s')"
+apply (drule deleted_has_del_event_file, simp, clarify)
+apply (rule_tac x = s' in exI, rule_tac x = p in exI, simp)
+apply (rule ballI, rule impI, frule vs_history, simp)
+by (erule deled_imp_childfs_deleted, simp+)
+
+lemma initf_deled_imp_childf_deled:
+  "\<lbrakk>deleted (File f) s; valid s; f \<prec> childf; childf \<in> init_files\<rbrakk> 
+  \<Longrightarrow> \<exists> s'. s' \<preceq> s \<and> valid s' \<and> deleted (File childf) s'"
+apply (drule deleted_has_del_event_file, simp, clarify)
+apply (frule vs_history, simp, frule valid_cons)
+apply (drule deled_imp_childfs_deleted, simp, simp)
+apply (rule_tac x = s' in exI, auto elim:no_juniorE)
+done
+
+lemma deleted_has_del_event_allchildf:
+  "\<lbrakk>deleted (File f) s; valid s; f \<preceq> childf; childf \<in> init_files\<rbrakk>
+  \<Longrightarrow> \<exists> s' p. DeleteFile p childf # s' \<preceq> s \<and> \<not> deleted (File childf) s'"
+apply (case_tac "f = childf")
+apply (drule deleted_has_del_event_file, simp, simp)
+apply (drule noJ_Anc'', simp)
+apply (drule initf_deled_imp_childf_deled, simp+, clarify)
+apply (drule deleted_has_del_event_file, simp+, clarify)
+apply (rule_tac x = s'a in exI, rule conjI, rule_tac x = p in exI)
+apply (erule no_junior_trans, simp+)
+done
+*)
+
+lemma del_imp_del_s_file:
+  assumes initf: "f \<in> init_files"
+  and deled: "deleted (File f) s"
+  and vs: "valid s"
+  shows "file_deletable_s f"
+proof -
+  from deled vs obtain s' p' where
+    his: "DeleteFile p' f # s' \<preceq> s" and fstdel: "\<not> deleted (File f) s'"
+    by (drule_tac deleted_has_del_event_file, auto)
+  from his vs have "valid (DeleteFile p' f # s')" by (simp add:vs_history)
+  hence exp': "p' \<in> current_procs s'" and exf: "f \<in> current_files s'"
+    and rc: "rc_grant s' (DeleteFile p' f)" and vs': "valid s'"
+    by (auto dest:valid_os valid_rc valid_cons)
+
+  from initf obtain t where etype: "etype_of_file [] f = Some t"
+    by (drule_tac init_file_has_etype, simp add:etype_of_file_def, blast)
+  from initf have sd: "source_dir [] f = Some f"
+    by (simp add:source_dir_of_init') 
+  hence "obj2sobj [] (File f) = SFile (t, f) (Some f)"
+    using etype initf by (auto simp:obj2sobj.simps)
+  with fstdel vs' initf exf etype
+  have SF: "obj2sobj s' (File f) = SFile (t, f) (Some f)"
+    using obj2sobj_file_remains'''[where s = "[]" and s' = s']
+    by (auto simp:obj2sobj.simps)
+
+  from exp' vs' obtain r' fr' pt' u' srp' where
+    SP': "obj2sobj s' (Proc p') = SProc (r',fr',pt',u') srp'"
+    by (frule_tac current_proc_has_sobj, simp, blast)
+  with exp' vs' all_sobjs_I[where s = s' and obj = "Proc p'"]
+  have SP'_in: "SProc (r',fr',pt',u') srp' \<in> all_sobjs" by simp
+
+  show ?thesis unfolding file_deletable_s_def
+    apply (rule_tac x = t in exI, 
+           rule_tac x = "(r',fr',pt',u')" in exI,
+           rule_tac x = srp' in exI)
+    apply (simp add:SP'_in) 
+    using rc SP' SF etype
+    by (auto simp:obj2sobj.simps cp2sproc.simps split:option.splits if_splits)
+qed
+
+lemma del_imp_del_s_proc:
+  assumes initp: "p \<in> init_processes"
+  and deled: "deleted (Proc p) s"
+  and vs: "valid s"
+  shows "proc_deletable_s p"
+proof-
+  from deled vs obtain s' p' where
+    his: "Kill p' p # s' \<preceq> s" and fstdel: "\<not> deleted (Proc p) s'"
+    by (drule_tac deleted_has_del_event_proc, auto)
+  from his vs have "valid (Kill p' p # s')" by (simp add:vs_history)
+  hence exp': "p' \<in> current_procs s'" and exp: "p \<in> current_procs s'"
+    and rc: "rc_grant s' (Kill p' p)" and vs': "valid s'"
+    by (auto dest:valid_os valid_rc valid_cons)
+
+  from initp fstdel vs' have "source_proc s' p = Some p"
+    apply (induct s', simp)
+    apply (frule valid_cons, frule valid_os, frule not_deleted_cons_D, simp)
+    by (case_tac a, auto dest:not_deleted_imp_exists simp:np_notin_curp)
+  with exp initp vs' obtain r fr pt u 
+    where SP: "obj2sobj s' (Proc p) = SProc (r,fr,pt,u) (Some p)"
+    apply (frule_tac current_proc_has_sobj, simp)
+    by (simp add:obj2sobj.simps split:option.splits, blast)
+  with exp vs' all_sobjs_I[where s = s' and obj = "Proc p"]
+  have SP_in: "SProc (r,fr,pt,u) (Some p) \<in> all_sobjs" by simp
+
+  from exp' vs' obtain r' fr' pt' u' srp' where
+    SP': "obj2sobj s' (Proc p') = SProc (r',fr',pt',u') srp'"
+    by (frule_tac current_proc_has_sobj, simp, blast)
+  with exp' vs' all_sobjs_I[where s = s' and obj = "Proc p'"]
+  have SP'_in: "SProc (r',fr',pt',u') srp' \<in> all_sobjs" by simp
+
+  show ?thesis unfolding proc_deletable_s_def
+    apply (rule_tac x = r in exI, rule_tac x = fr in exI,
+           rule_tac x = pt in exI, rule_tac x = u in exI,
+           rule_tac x = "(r',fr',pt',u')" in exI,
+           rule_tac x = srp' in exI)
+    apply (simp add:SP_in SP'_in) 
+    using rc SP SP'
+    by (auto simp:obj2sobj.simps cp2sproc.simps split:option.splits)
+qed
+    
+lemma del_imp_del_s_ipc:
+  assumes initi: "i \<in> init_ipcs"
+  and deled: "deleted (IPC i) s"
+  and vs: "valid s"
+  shows "ipc_deletable_s i"
+proof-
+  from deled vs obtain s' p' where
+    his: "DeleteIPC p' i # s' \<preceq> s" and fstdel: "\<not> deleted (IPC i) s'"
+    by (drule_tac deleted_has_del_event_ipc, auto)
+  from his vs have "valid (DeleteIPC p' i # s')" by (simp add:vs_history)
+  hence exp': "p' \<in> current_procs s'" and exi: "i \<in> current_ipcs s'"
+    and rc: "rc_grant s' (DeleteIPC p' i)" and vs': "valid s'"
+    by (auto dest:valid_os valid_rc valid_cons)
+
+  from initi obtain t where type: "init_ipc_type i = Some t" 
+    using init_ipc_has_type by (auto simp:bidirect_in_init_def)
+  with fstdel vs' initi exi have SI: "obj2sobj s' (IPC i) = SIPC t (Some i)"
+    using obj2sobj_ipc_remains'''[where s = "[]" and s' = s']
+    by (auto simp:obj2sobj.simps)
+
+  from exp' vs' obtain r' fr' pt' u' srp' where
+    SP': "obj2sobj s' (Proc p') = SProc (r',fr',pt',u') srp'"
+    by (frule_tac current_proc_has_sobj, simp, blast)
+  with exp' vs' all_sobjs_I[where s = s' and obj = "Proc p'"]
+  have SP'_in: "SProc (r',fr',pt',u') srp' \<in> all_sobjs" by simp
+
+  show ?thesis unfolding ipc_deletable_s_def
+    apply (rule_tac x = t in exI, 
+           rule_tac x = "(r',fr',pt',u')" in exI,
+           rule_tac x = srp' in exI)
+    apply (simp add: SP'_in) 
+    using rc SP' SI type
+    by (auto simp:obj2sobj.simps cp2sproc.simps split:option.splits if_splits)
+qed
+
+lemma deleted_imp_deletable_s:
+  "\<lbrakk>deleted obj s; exists [] obj; valid s\<rbrakk> \<Longrightarrow> deletable_s obj"
+apply (case_tac obj)  
+apply (simp add:del_imp_del_s_proc)  
+apply (simp add:del_imp_del_s_file) 
+apply (simp add:del_imp_del_s_ipc)
+done
+
+(*
+lemma all_sobjs_E3:
+  assumes prem: "sobj \<in> all_sobjs" 
+  shows "\<exists> s obj. valid s \<and> obj2sobj s obj = sobj \<and> sobj_source_eq_obj sobj obj \<and> 
+                  initp_intact_but s sobj \<and> exists s obj \<and> no_del_event s"
+proof-
+  obtain t where "etype_of_file [] [] = Some t" 
+    using root_in_filesystem current_file_has_etype[of "[]" "[]"] vs_nil 
+    by auto
+  with root_in_filesystem
+  have "obj2sobj [] (File []) = SFile (t,[]) (Some [])"
+    by (auto simp:obj2sobj.simps source_dir_of_init' 
+            split:option.splits if_splits)
+  moreover have "initp_intact []"
+    by (auto simp:initp_intact_def cp2sproc.simps obj2sobj.simps 
+            split:option.splits)
+  ultimately show ?thesis
+    using prem vs_nil root_in_filesystem 
+    apply (drule_tac s' = "[]" and obj' = "File []" in all_sobjs_E2)
+    apply (simp+, (erule exE|erule conjE)+)
+    by (rule_tac x = s in exI, simp, rule_tac x = obj in exI, simp+)
+qed
+
+lemma del_s_imp_del_proc:
+  assumes initp: "p \<in> init_processes"
+  and del_s: "proc_deletable_s p"
+  shows "\<exists> s. valid s \<and> deleted (Proc p) s"
+proof-
+  from del_s obtain r fr pt u sp' srp' 
+    where Target: "SProc (r,fr,pt,u) (Some p) \<in> all_sobjs" 
+    and Killer: "SProc sp' srp' \<in> all_sobjs"
+    and rc: "(role_of_sproc sp', Proc_type pt, DELETE) \<in> compatible"
+    using proc_deletable_s_def by auto
+  
+  from Target obtain s where vs: "valid s" 
+    and "obj2sobj s (Proc p) = SProc (r,fr,pt,u) (Some p)" 
+    and "exists s (Proc p)"  and "no_del_event s"
+    and "initp_intact_but s (SProc (r,fr,pt,u) (Some p))"
+    apply (drule_tac all_sobjs_E3, clarsimp)
+    by (frule obj2sobj_proc, clarsimp)
+  with Killer obtain s' p' where vs': "valid (s' @ s)" and
+    SP : "obj2sobj (s' @ s) (Proc p) = SProc (r,fr,pt,u) (Some p)" and
+    exp: "exists (s' @ s) (Proc p)"  and
+    SP': "obj2sobj (s' @ s) (Proc p') = SProc sp' srp'" and
+    exp': "exists (s' @ s) (Proc p')"
+    apply (drule_tac obj' = "Proc p" and s' = s in all_sobjs_E1, auto)
+    apply (frule_tac obj = obj in obj2sobj_proc, erule exE)
+    apply (auto intro:nodel_exists_remains)
+    apply blast
+    apply (frule_tac obj = "Proc p" in nodel_exists_remains)
+
+
+lemma deletable_s_imp_deleted:
+  "deletable_s obj \<Longrightarrow> \<exists> s. valid s \<and> deleted obj s"
+apply (case_tac obj)
+apply (simp add:deletable_s.simps)
+
+
+lemma valid_kill_imp_proc_del_s:
+  "\<lbrakk>valid (Kill p' p # s); p \<in> init_processes; \<not> deleted (Proc p) s\<rbrakk> \<Longrightarrow> proc_deletable_s p"
+apply (frule valid_os, frule valid_
+
+
+lemma build_phase: "f \<in> init_files \<and> file_deletable_s f \<longrightarrow> (\<exists> s. \<forall> childf. valids s \<and> f \<preceq> childf \<and> childf \<in> init_files \<and> etype_of_file [] f = Some t \<longrightarrow> (\<exists> p. p \<in> current_procs s \<and> currentrole s p = Some r \<and> (r, File_type t, DELETE) \<in> compatible))"
+ thm all_sobjs_E0
+
+lemma del_phase: "f \<in> init_files \<and> file_deletable_s f \<and> (\<forall> childf. f \<preceq> childf \<and> childf \<notin> current_files s \<and> valid s"
+
+*)
+
+end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/deleted_prop.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,49 @@
+theory deleted_prop
+imports Main rc_theory os_rc
+begin
+
+context tainting begin
+
+lemma deleted_cons_I: "deleted obj s \<Longrightarrow> deleted obj (e # s)"
+by (case_tac e, auto)
+
+lemma not_deleted_cons_D: "\<not> deleted obj (e # s) \<Longrightarrow> \<not> deleted obj s" 
+by (auto dest:deleted_cons_I)
+
+lemma not_deleted_imp_exists:
+  "\<lbrakk>\<not> deleted obj s; exists [] obj\<rbrakk> \<Longrightarrow> exists s obj"
+apply (induct s, simp)
+apply (case_tac a, case_tac [!] obj, auto)
+done
+
+lemma cons_app_simp_aux:
+  "(a # b) @ c = a # (b @ c)" by auto
+
+lemma not_deleted_imp_exists':
+  "\<lbrakk>\<not> deleted obj (s'@s); exists s obj\<rbrakk> \<Longrightarrow> exists (s'@s) obj"
+apply (induct s', simp, simp only:cons_app_simp_aux)
+apply (frule not_deleted_cons_D)
+apply (case_tac a, case_tac [!] obj, auto)
+done
+
+lemma nodel_imp_un_deleted:
+  "no_del_event s \<Longrightarrow> \<not> deleted obj s"
+by (induct s, simp, case_tac a,auto)
+
+lemma nodel_exists_remains:
+  "\<lbrakk>no_del_event (s'@s); exists s obj\<rbrakk> \<Longrightarrow> exists (s'@s) obj"
+apply (drule_tac obj = obj in nodel_imp_un_deleted)
+by (simp add:not_deleted_imp_exists')
+
+lemma nodel_imp_exists:
+  "\<lbrakk>no_del_event s; exists [] obj\<rbrakk> \<Longrightarrow> exists s obj"
+apply (drule_tac obj = obj in nodel_imp_un_deleted)
+by (simp add:not_deleted_imp_exists)
+
+lemma no_del_event_cons_D:
+  "no_del_event (e # s) \<Longrightarrow> no_del_event s"
+by (case_tac e, auto)
+
+end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/final_theorems.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,56 @@
+theory final_theorems
+imports Main rc_theory del_vs_del_s tainted_vs_tainted_s
+begin
+
+context tainting_s_complete begin
+
+theorem static_complete: 
+  assumes undel: "undeletable obj" and tbl: "taintable obj"
+  shows "taintable_s obj"
+proof-
+  from tbl obtain s where tainted: "obj \<in> tainted s"
+    by (auto simp:taintable_def)
+  hence vs: "valid s" by (simp add:tainted_is_valid)
+  from undel vs have "\<not> deleted obj s" and "exists [] obj" 
+    by (auto simp:undeletable_def)
+  moreover from tainted have "valid s" by (rule tainted_is_valid)
+  ultimately have "source_of_sobj (obj2sobj s obj) = Some obj" 
+    using init_obj_keeps_source by auto
+  with tainted t2ts 
+  show ?thesis unfolding taintable_s_def
+    by (rule_tac x = "obj2sobj s obj" in exI, simp)
+qed
+
+theorem undeletable_s_complete:
+  "undeletable_s obj \<Longrightarrow> undeletable obj"
+apply (clarsimp simp:undeletable_s_def undeletable_def)
+apply (drule deleted_imp_deletable_s, simp+)
+done
+
+theorem final_offer:
+  "\<lbrakk>undeletable_s obj; \<not> taintable_s obj; exists [] obj\<rbrakk> \<Longrightarrow> \<not> taintable obj"
+apply (erule swap)
+by (simp add:static_complete undeletable_s_complete)
+
+end
+
+context tainting_s_sound begin
+
+theorem static_sound:
+  assumes tbl_s: "taintable_s obj"
+  shows "taintable obj"
+proof-
+  from tbl_s obtain sobj where ts: "sobj \<in> tainted_s"
+    and sreq: "source_of_sobj sobj = Some obj" 
+    by (auto simp:taintable_s_def)
+  from ts obtain obj' \<tau> where t: "obj' \<in> tainted \<tau>" 
+    and vs: "valid \<tau>" and sreq': "sobj_source_eq_obj sobj obj'"
+    by (auto dest!:tainted_s2tainted dest:tainted_is_valid)
+  from sreq' sreq have "obj = obj'" by (simp add:source_eq)
+  with vs t 
+  show ?thesis by (auto simp:taintable_def)
+qed
+
+end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/finite_static.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,270 @@
+theory finite_static
+imports Main rc_theory all_sobj_prop
+begin
+
+context tainting_s_complete begin
+
+lemma tainted_s_subset_all_sobjs:
+  "tainted_s \<subseteq> all_sobjs"
+apply (rule subsetI, erule tainted_s.induct)
+apply (auto intro:all_sobjs.intros)
+apply (drule seeds_in_init)
+apply (subgoal_tac "exists [] obj")
+apply (frule obj2sobj_nil_init)
+apply (drule all_sobjs_I)
+apply (rule vs_nil, simp) 
+apply (case_tac obj, simp+)
+done
+
+definition 
+  "init_proc_opt \<equiv> {Some p | p. p \<in> init_processes} \<union> {None}"
+
+lemma finite_init_proc_opt:
+  "finite init_proc_opt"
+unfolding init_proc_opt_def
+apply (simp add: finite_Un)
+apply (rule finite_imageI)
+by (simp add:init_finite)
+
+definition 
+  "init_file_opt \<equiv> {Some f | f. f \<in> init_files} \<union> {None}"
+
+lemma finite_init_file_opt:
+  "finite init_file_opt"
+unfolding init_file_opt_def
+apply (simp add: finite_Un)
+apply (rule finite_imageI)
+by (simp add:init_finite)
+
+definition 
+  "init_ipc_opt \<equiv> {Some i | i. i \<in> init_processes} \<union> {None}"
+
+lemma finite_init_ipc_opt:
+  "finite init_ipc_opt"
+unfolding init_ipc_opt_def
+apply (simp add: finite_Un)
+apply (rule finite_imageI)
+by (simp add:init_finite)
+
+lemma finite_t_client:
+  "finite (UNIV :: t_client set)"
+apply (subgoal_tac "UNIV = {Client1, Client2}")
+apply (metis (full_types) finite.emptyI finite_insert)
+apply auto
+apply (case_tac x, simp+)
+done
+
+lemma finite_t_normal_role:
+  "finite (UNIV :: t_normal_role set)"
+proof-
+  have p1: "finite {WebServer}" by simp
+  have p2: "finite {WS c | c. c \<in> UNIV}" using finite_t_client by auto
+  have p3: "finite {UpLoad c| c. c \<in> UNIV}" using finite_t_client by auto
+  have p4: "finite {CGI c| c. c \<in> UNIV}" using finite_t_client by auto
+  from p1 p2 p3 p4
+  have p5: "finite ({WebServer} \<union> {WS c | c. c \<in> UNIV} \<union> {UpLoad c | c. c \<in> UNIV} \<union> {CGI c | c. c \<in> UNIV})" 
+    by auto
+  have p6: "(UNIV :: t_normal_role set) = ({WebServer} \<union> {WS c | c. c \<in> UNIV} \<union> 
+     {UpLoad c | c. c \<in> UNIV} \<union> {CGI c | c. c \<in> UNIV})"
+    apply (rule set_eqI, auto split:t_normal_role.splits)
+    by (case_tac x, auto)
+  show ?thesis by (simp only:p6 p5)
+qed
+
+lemma finite_t_role: "finite (UNIV :: t_role set)"
+proof-
+  have p1: "finite {NormalRole r | r. r \<in> UNIV}" using finite_t_normal_role by auto
+  have p2: "finite {InheritParentRole, UseForcedRole, InheritUpMixed, InheritUserRole, InheritProcessRole}"
+    by auto
+  have p3: "UNIV = {InheritParentRole, UseForcedRole, InheritUpMixed, InheritUserRole, InheritProcessRole} \<union> 
+    {NormalRole r | r. r \<in> UNIV}"
+    apply (rule set_eqI, auto split:t_role.splits)
+    by (case_tac x, auto)
+  have p4: "finite ({InheritParentRole, UseForcedRole, InheritUpMixed, InheritUserRole, InheritProcessRole} \<union> 
+    {NormalRole r | r. r \<in> UNIV})" using p1 p2 by auto
+  show ?thesis by (simp only:p3 p4)
+qed
+
+lemma finite_t_normal_file_type: "finite (UNIV :: t_normal_file_type set)"
+proof-
+  have p1: "finite {WebData_file c | c. c \<in> UNIV}" using finite_t_client by auto
+  have p2: "finite {CGI_P_file c | c. c \<in> UNIV}"  using finite_t_client by auto
+  have p3: "finite {PrivateD_file c | c. c \<in> UNIV}" using finite_t_client by auto
+  have p4: "finite {Executable_file, Root_file_type, WebServerLog_file}" by auto
+  have p5: "finite ({WebData_file c | c. c \<in> UNIV} \<union> {CGI_P_file c | c. c \<in> UNIV} \<union> 
+    {PrivateD_file c | c. c \<in> UNIV} \<union> {Executable_file, Root_file_type, WebServerLog_file})"
+    using p1 p2 p3 p4 by auto
+  have p6: "UNIV = ({WebData_file c | c. c \<in> UNIV} \<union> {CGI_P_file c | c. c \<in> UNIV} \<union> 
+    {PrivateD_file c | c. c \<in> UNIV} \<union> {Executable_file, Root_file_type, WebServerLog_file})"
+    apply (rule set_eqI, auto split:t_normal_file_type.splits)
+    by (case_tac x, auto)
+  show ?thesis by (simp only:p6 p5)
+qed
+
+lemma finite_t_rc_file_type: "finite (UNIV :: t_rc_file_type set)"
+proof-
+  have p1: "finite {NormalFile_type t | t. t \<in> UNIV}" using finite_t_normal_file_type by auto
+  have p2: "finite ({InheritParent_file_type} \<union> {NormalFile_type t | t. t \<in> UNIV})"
+    using p1 by auto
+  have p3: "UNIV = ({InheritParent_file_type} \<union> {NormalFile_type t | t. t \<in> UNIV})"
+    apply (rule set_eqI, auto split:t_rc_file_type.splits)
+    by (case_tac x, auto)
+  show ?thesis by (simp only:p3 p2)
+qed
+
+lemma finite_t_normal_proc_type: "finite (UNIV :: t_normal_proc_type set)"
+proof-
+  have p1: "finite {CGI_P_proc c | c. c \<in> UNIV}" using finite_t_client by auto
+  have p2: "finite ({CGI_P_proc c | c. c \<in> UNIV} \<union> {WebServer_proc})" using p1 by auto
+  have p3: "UNIV = ({CGI_P_proc c | c. c \<in> UNIV} \<union> {WebServer_proc})"
+    apply (rule set_eqI, auto split:t_normal_proc_type.splits)
+    by (case_tac x, auto)
+  show ?thesis by (simp only:p3 p2)
+qed
+
+lemma finite_t_rc_proc_type: "finite (UNIV :: t_rc_proc_type set)"
+proof-
+  have p1: "finite {NormalProc_type t | t. t \<in> UNIV}" using finite_t_normal_proc_type by auto
+  have p2: "finite ({NormalProc_type t | t. t \<in> UNIV} \<union> {InheritParent_proc_type, UseNewRoleType})"
+    using p1 by auto
+  have p3: "UNIV = ({NormalProc_type t | t. t \<in> UNIV} \<union> {InheritParent_proc_type, UseNewRoleType})"
+    apply (rule set_eqI, auto split:t_rc_proc_type.splits)
+    by (case_tac x, auto)
+  show ?thesis by (simp only:p3 p2)
+qed
+  
+lemma finite_t_normal_ipc_type : "finite (UNIV :: t_normal_ipc_type set)"
+proof-
+  have p1: "finite {WebIPC}" by auto
+  have p2: "UNIV = {WebIPC}" apply auto by (case_tac x, auto)
+  show ?thesis by (simp only:p1 p2)
+qed
+
+definition 
+  "all_sps \<equiv> (UNIV ::t_normal_role set) \<times> (UNIV :: t_role set) \<times> (UNIV :: t_normal_proc_type set) \<times> init_users"
+
+lemma finite_all_sps: "finite all_sps"
+proof-
+  have "finite ((UNIV :: t_normal_proc_type set) \<times> init_users)"
+    using finite_t_normal_proc_type init_finite
+    by (rule_tac finite_cartesian_product, auto)
+  hence "finite ((UNIV :: t_role set) \<times> (UNIV :: t_normal_proc_type set) \<times> init_users)"
+    using finite_t_role by (rule_tac finite_cartesian_product, auto)
+  hence "finite ((UNIV::t_normal_role set) \<times> (UNIV::t_role set) \<times> (UNIV::t_normal_proc_type set) \<times> init_users)" 
+    using finite_t_normal_role by (rule_tac finite_cartesian_product, auto)
+  thus ?thesis by (simp only:all_sps_def)
+qed
+
+definition
+   "all_SPs \<equiv> {SProc sp (Some p) | sp p. sp \<in> all_sps \<and> p \<in> init_processes} \<union> {SProc sp None | sp. sp \<in> all_sps}"
+
+lemma finite_all_SPs: "finite all_SPs"
+proof-
+  have p1: "finite {SProc sp (Some p) | sp p. sp \<in> all_sps \<and> p \<in> init_processes}"
+    using finite_all_sps init_finite by auto
+  have p2: "finite {SProc sp None | sp. sp \<in> all_sps}"
+    using finite_all_sps by auto
+  have "finite ({SProc sp (Some p) | sp p. sp \<in> all_sps \<and> p \<in> init_processes} \<union> 
+    {SProc sp None | sp. sp \<in> all_sps})"
+    using p1 p2 by auto
+  thus ?thesis by (simp only:all_SPs_def)
+qed
+
+definition
+  "all_sfs \<equiv> (UNIV :: t_normal_file_type set) \<times> init_files"
+
+lemma finite_all_sfs: "finite all_sfs"
+proof-
+  have "finite ((UNIV :: t_normal_file_type set) \<times> init_files)"
+    using finite_t_normal_file_type init_finite
+    by (rule_tac finite_cartesian_product, auto)
+  thus ?thesis by (simp add:all_sfs_def)
+qed
+
+definition 
+  "all_SFs \<equiv> {SFile sf (Some f) | sf f. sf \<in> all_sfs \<and> f \<in> init_files} \<union> {SFile sf None| sf. sf \<in> all_sfs}"
+
+lemma finite_all_SFs: "finite all_SFs"
+proof-
+  have p1: "finite ({SFile sf (Some f) | sf f. sf \<in> all_sfs \<and> f \<in> init_files} \<union>
+    {SFile sf None| sf. sf \<in> all_sfs})"
+    using finite_all_sfs init_finite by auto
+  thus ?thesis by (simp only:all_SFs_def)
+qed
+
+definition 
+  "all_SIs \<equiv> {SIPC si (Some i)| si i. si \<in> UNIV \<and> i \<in> init_ipcs} \<union> {SIPC si None| si. si \<in> UNIV}"
+
+lemma finite_all_SIs: "finite all_SIs"
+proof-
+  have "finite ({SIPC si (Some i)| si i. si \<in> UNIV \<and> i \<in> init_ipcs} \<union> {SIPC si None| si. si \<in> UNIV})"
+    using finite_t_normal_ipc_type init_finite by auto
+  thus ?thesis by (simp only:all_SIs_def)
+qed
+
+lemma all_sobjs_srf_init':
+  "sobj \<in> all_sobjs \<Longrightarrow> \<forall> sf srf. sobj = SFile sf (Some srf) \<longrightarrow> srf \<in> init_files"
+by (erule all_sobjs.induct, auto) 
+
+lemma all_sobjs_srf_init:
+  "SFile sf (Some srf) \<in> all_sobjs \<Longrightarrow> srf \<in> init_files"
+by (auto dest!:all_sobjs_srf_init')
+
+lemma all_sobjs_sd_init':
+  "sobj \<in> all_sobjs \<Longrightarrow> \<forall> tf sd srf. sobj = SFile (tf, sd) srf \<longrightarrow> sd \<in> init_files"
+by (erule all_sobjs.induct, auto) 
+
+lemma all_sobjs_sd_init:
+  "SFile (tf, sd) srf \<in> all_sobjs \<Longrightarrow> sd \<in> init_files"
+by (auto dest!:all_sobjs_sd_init')
+
+lemma all_sobjs_sri_init':
+  "sobj \<in> all_sobjs \<Longrightarrow> \<forall> si sri. sobj = SIPC si (Some sri) \<longrightarrow> sri \<in> init_ipcs"
+apply (erule all_sobjs.induct, auto) using init_ipc_has_type 
+by (simp add:bidirect_in_init_def)
+
+lemma all_sobjs_sri_init:
+  "SIPC si (Some sri) \<in> all_sobjs \<Longrightarrow> sri \<in> init_ipcs"
+by (auto dest!:all_sobjs_sri_init')
+
+lemma all_sobjs_sru_init'[rule_format]:
+  "sobj \<in> all_sobjs \<Longrightarrow> \<forall> r fr pt u srp. sobj = SProc (r,fr,pt,u) srp \<longrightarrow> u \<in> init_users"
+using init_owner_valid
+by (erule_tac all_sobjs.induct, auto) 
+
+lemma all_sobjs_sru_init:
+  "SProc (r,fr,pt,u) srp \<in> all_sobjs \<Longrightarrow> u \<in> init_users"
+by (auto dest!:all_sobjs_sru_init')
+
+lemma unknown_not_in_all_sobjs':
+  "sobj \<in> all_sobjs \<Longrightarrow> sobj \<noteq> Unknown"
+by (erule_tac all_sobjs.induct, auto) 
+
+lemma unknown_not_in_all_sobjs:
+  "Unknown \<in> all_sobjs \<Longrightarrow> False"
+using unknown_not_in_all_sobjs' by auto
+
+lemma finite_all_sobjs: "finite all_sobjs"
+proof-
+  have p1: "finite (all_SPs \<union> all_SFs \<union> all_SIs)"
+    using finite_all_SPs finite_all_SFs finite_all_SIs by auto
+  have p2: "all_sobjs \<subseteq> (all_SPs \<union> all_SFs \<union> all_SIs)"
+    apply (rule subsetI)
+    using all_sobjs_sd_init all_sobjs_sri_init all_sobjs_srf_init all_sobjs_srp_init 
+      all_sobjs_sru_init unknown_not_in_all_sobjs
+    by (case_tac x, auto simp:all_SPs_def all_SFs_def all_SIs_def all_sps_def all_sfs_def)
+  show ?thesis 
+    apply (rule_tac B = "(all_SPs \<union> all_SFs \<union> all_SIs)" in finite_subset)
+    using p1 p2 by auto
+qed
+
+lemma finite_tainted_s:
+  "finite tainted_s"
+apply (rule_tac B = "all_sobjs" in finite_subset)
+apply (rule tainted_s_subset_all_sobjs)
+apply (rule finite_all_sobjs)
+done
+
+end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/my_list_prefix.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,400 @@
+(*<*)
+theory my_list_prefix
+imports "List_Prefix"
+begin
+(*>*)
+
+(* cmp:: 1:complete equal; 2:less; 3:greater; 4: len equal,but ele no equal *)
+fun cmp :: "'a list \<Rightarrow> 'a list \<Rightarrow> nat"
+where
+  "cmp [] [] = 1" |                   
+  "cmp [] (e#es) = 2" |
+  "cmp (e#es) [] = 3" |
+  "cmp (e#es) (a#as) = (let r = cmp es as in 
+                            if (e = a) then r else 4)"
+
+(* list_com:: fetch the same ele of the same left order into a new list*) 
+fun list_com :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where 
+  "list_com []  ys = []" |
+  "list_com xs [] = []" |
+  "list_com (x#xs) (y#ys) = (if x = y then x#(list_com xs ys) else [])"
+
+(* list_com_rev:: by the right order of list_com *)
+definition list_com_rev :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "\<bullet>" 50)
+where
+  "xs \<bullet> ys \<equiv> rev (list_com (rev xs) (rev ys))"
+
+(* list_diff:: list substract, once different return tailer *)
+fun list_diff :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
+  "list_diff []  xs = []" |
+  "list_diff (x#xs) [] = x#xs" |
+  "list_diff (x#xs) (y#ys) = (if x = y then list_diff xs ys else (x#xs))"
+
+(* list_diff_rev:: list substract with rev order*)
+definition list_diff_rev :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "\<setminus>" 51)
+where
+   "xs \<setminus> ys \<equiv> rev (list_diff (rev xs) (rev ys))"
+
+(* xs <= ys:: \<exists>zs. ys = xs @ zs *)
+(* no_junior:: xs is ys' tail,or equal *)
+definition no_junior  :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<preceq>" 50)
+where
+  "xs \<preceq> ys \<equiv> rev xs \<le> rev ys"
+
+(* < :: xs <= ys \<and> xs \<noteq> ys *)
+(* is_ancestor:: xs is ys' tail, but no equal  *)
+definition is_ancestor :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<prec>" 50)
+where
+  "xs \<prec> ys \<equiv> rev xs < rev ys"
+
+lemma list_com_diff [simp]: "(list_com xs  ys) @ (list_diff xs  ys) = xs" (is "?P xs ys")
+  by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_com_diff_rev [simp]: "(xs \<setminus> ys) @ (xs \<bullet> ys) = xs"
+  apply (simp only:list_com_rev_def list_diff_rev_def)
+  by (fold rev_append, simp)
+
+lemma list_com_commute: "list_com xs ys = list_com ys xs" (is "?P xs ys")
+  by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_com_ido: "xs \<le> ys \<longrightarrow> list_com xs ys = xs" (is "?P xs ys")
+  by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_com_rev_ido [simp]: "xs \<preceq> ys \<Longrightarrow> xs \<bullet> ys = xs"
+  by (cut_tac list_com_ido, auto simp: no_junior_def list_com_rev_def)
+
+lemma list_com_rev_commute [iff]: "(xs \<bullet> ys) = (ys \<bullet> xs)"
+  by (simp only:list_com_rev_def list_com_commute)
+
+lemma list_com_rev_ido1 [simp]: "xs \<preceq> ys \<Longrightarrow> ys \<bullet> xs = xs"
+  by simp
+
+lemma list_diff_le: "(list_diff xs ys = []) = (xs \<le> ys)" (is "?P xs ys")
+  by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_diff_rev_le: "(xs \<setminus> ys = []) = (xs \<preceq> ys)"
+  by (auto simp:list_diff_rev_def no_junior_def list_diff_le)
+
+lemma list_diff_lt: "(list_diff xs ys = [] \<and> list_diff ys xs \<noteq> []) = (xs < ys)" (is "?P xs ys")
+  by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_diff_rev_lt: "(xs \<setminus> ys = [] \<and> ys \<setminus> xs \<noteq> []) = (xs \<prec> ys)"
+  by (auto simp: list_diff_rev_def list_diff_lt is_ancestor_def)
+
+(* xs diff ys result not [] \<Longrightarrow> \<exists> e \<in> xs. a \<in> ys. e \<noteq> a *)
+lemma list_diff_neq: 
+  "\<forall> e es a as. list_diff xs ys = (e#es) \<and> list_diff ys xs = (a#as) \<longrightarrow> e \<noteq> a" (is "?P xs ys")
+  by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_diff_rev_neq_pre: "\<forall> e es a as. xs \<setminus> ys = rev (e#es) \<and> ys \<setminus> xs = rev (a#as) \<longrightarrow> e \<noteq> a"
+  apply (simp only:list_diff_rev_def, clarify)
+  apply (insert list_diff_neq, atomize)
+  by (erule_tac x = "rev xs" in allE, erule_tac x = "rev ys" in allE, blast)
+
+lemma list_diff_rev_neq: "\<forall> e es a as. xs \<setminus> ys = es @ [e] \<and> ys \<setminus> xs = as @ [a] \<longrightarrow> e \<noteq> a"
+  apply (rule_tac allI)+
+  apply (insert list_diff_rev_neq_pre, atomize)
+  apply (erule_tac x = "xs" in allE)
+  apply (erule_tac x = "ys" in allE)
+  apply (erule_tac x = "e" in allE)
+  apply (erule_tac x = "rev es" in allE)
+  apply (erule_tac x = "a" in allE)
+  apply (erule_tac x = "rev as" in allE)
+  by auto
+
+lemma list_com_self [simp]: "list_com zs zs = zs"
+  by (induct_tac zs, simp+)
+
+lemma list_com_rev_self [simp]: "zs \<bullet> zs = zs"
+  by (simp add:list_com_rev_def)
+
+lemma list_com_append [simp]: "(list_com (zs @ xs) (zs @ ys)) = (zs @ (list_com xs ys))"
+  by (induct_tac zs, simp+)
+
+lemma list_inter_append [simp]: "((xs @ zs) \<bullet> (ys @ zs)) = ((xs \<bullet> ys) @ zs)"
+  by (simp add:list_com_rev_def)
+
+lemma list_diff_djoin_pre: 
+  "\<forall> e es a as. list_diff xs ys = e#es \<and>  list_diff ys xs = a#as \<longrightarrow> (\<forall> zs zs'. (list_diff (xs @ zs) (ys @ zs') = [e]@es@zs))" 
+  (is "?P xs ys")
+  by (rule_tac P = ?P in cmp.induct, simp+)
+
+lemma list_diff_djoin_rev_pre:
+  "\<forall> e es a as. xs \<setminus> ys = rev (e#es) \<and>  ys \<setminus> xs = rev (a#as)  \<longrightarrow> (\<forall> zs zs'. ((zs @ xs) \<setminus> (zs' @ ys) = rev ([e]@es@rev zs)))"
+  apply (simp only: list_diff_rev_def, clarify)
+  apply (insert list_diff_djoin_pre, atomize)
+  apply (erule_tac x = "rev xs" in allE)
+  apply (erule_tac x = "rev ys" in allE)
+  apply (erule_tac x = "e" in allE)
+  apply (erule_tac x = "es" in allE)
+  apply (erule_tac x = "a" in allE)
+  apply (erule_tac x = "as" in allE)
+  by simp
+
+lemma list_diff_djoin_rev:
+  "xs \<setminus> ys = es @ [e] \<and> ys \<setminus> xs = as @ [a] \<Longrightarrow> zs @ xs \<setminus> zs' @ ys = zs @ es @ [e]"
+  apply (insert list_diff_djoin_rev_pre [rule_format, simplified])
+  apply (clarsimp, atomize)
+  apply (erule_tac x = "xs" in allE)
+  apply (erule_tac x = "ys" in allE)
+  apply (erule_tac x = "rev es" in allE)
+  apply (erule_tac x = "e" in allE)
+  apply (erule_tac x = "rev as" in allE)
+  apply (erule_tac x = "a" in allE)
+  by auto
+
+lemmas list_diff_djoin_rev_simplified = conjI [THEN list_diff_djoin_rev, simp]
+
+lemmas list_diff_djoin = conjI [THEN list_diff_djoin_pre [rule_format], simp]
+
+lemma list_diff_ext_left [simp]: "(list_diff (zs @ xs) (zs @ ys) = (list_diff xs ys))"
+  by (induct_tac zs, simp+)
+
+lemma list_diff_rev_ext_left [simp]: "((xs @ zs \<setminus> ys @ zs) = (xs \<setminus> ys))"
+  by (auto simp: list_diff_rev_def)
+
+declare no_junior_def [simp]
+
+lemma no_juniorE: "\<lbrakk>xs \<preceq> ys; \<And> zs. ys = zs @ xs \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+proof -
+  assume h: "xs \<preceq> ys"
+    and h1: "\<And> zs. ys = zs @ xs \<Longrightarrow> R"
+  show "R"
+  proof -
+    from h have "rev xs \<le> rev ys" by (simp)
+    from this obtain zs where eq_rev: "rev ys = rev xs @ zs" by (auto simp:prefix_def)
+    show R 
+    proof(rule h1 [where zs = "rev zs"])
+      from rev_rev_ident and eq_rev have "rev (rev (ys)) = rev zs @ rev (rev xs)"
+	by simp
+      thus "ys = rev zs @ xs" by simp
+    qed
+  qed
+qed
+
+lemma no_juniorI: "\<lbrakk>ys = zs @ xs\<rbrakk> \<Longrightarrow> xs \<preceq> ys"
+  by simp
+
+lemma no_junior_ident [simp]: "xs \<preceq> xs"
+  by simp
+
+lemma no_junior_expand: "xs \<preceq> ys = ((xs \<prec> ys) \<or> xs = ys)"
+  by (simp only:no_junior_def is_ancestor_def strict_prefix_def, blast)
+
+lemma no_junior_same_prefix: " e # \<tau> \<preceq> e' # \<tau>' \<Longrightarrow> \<tau> \<preceq> \<tau>'"
+apply (simp add:no_junior_def )
+apply (erule disjE, simp)
+apply (simp only:prefix_def)
+by (erule exE, rule_tac x = "[e] @ zs" in exI, auto)
+
+lemma no_junior_noteq: "\<lbrakk>\<tau> \<preceq> a # \<tau>'; \<tau> \<noteq> a # \<tau>'\<rbrakk> \<Longrightarrow> \<tau> \<preceq> \<tau>'"
+apply (erule no_juniorE)
+by (case_tac zs, simp+)
+
+lemma is_ancestor_app [simp]: "xs \<prec> ys \<Longrightarrow> xs \<prec> zs @ ys"
+  by (auto simp:is_ancestor_def strict_prefix_def)
+
+lemma is_ancestor_cons [simp]: "xs \<prec> ys \<Longrightarrow> xs \<prec> a # ys"
+  by (auto simp:is_ancestor_def strict_prefix_def)
+
+lemma no_junior_app [simp]: "xs \<preceq> ys \<Longrightarrow> xs \<preceq> zs @ ys"
+  by simp
+
+lemma is_ancestor_no_junior [simp]: "xs \<prec> ys \<Longrightarrow> xs \<preceq> ys"
+  by (simp add:is_ancestor_def)
+
+lemma is_ancestor_y [simp]: "ys \<prec> y#ys"
+  by (simp add:is_ancestor_def strict_prefix_def)
+
+lemma no_junior_cons [simp]: "xs \<preceq> ys \<Longrightarrow> xs \<prec> (y#ys)"
+  by (unfold no_junior_expand, auto)
+
+lemma no_junior_anti_sym: "\<lbrakk>xs \<preceq> ys; ys \<preceq> xs\<rbrakk> \<Longrightarrow> xs = ys"
+  by simp
+
+declare no_junior_def [simp del]
+
+(* djoin:: xs and ys is not the other's tail, not equal either *)
+definition djoin :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<asymp>" 50)
+where
+  "xs \<asymp> ys \<equiv> \<not> (xs \<preceq> ys \<or> ys \<preceq> xs)"
+
+(* dinj:: function f's returning list is not tailing when paras not equal *)
+definition dinj :: "('a \<Rightarrow> 'b list) \<Rightarrow> bool"
+where
+  "dinj f \<equiv> (\<forall> a b. a \<noteq> b \<longrightarrow> f a \<asymp> f b)"
+
+
+(* list_cmp:: list comparison: one is other's prefix or no equal at some position *)
+lemma list_cmp: "xs \<le> ys \<or> ys \<le> xs \<or>  (\<exists> zs x y a b. xs = zs @ [a] @ x  \<and> ys = zs @ [b] @ y \<and> a \<noteq> b)"
+proof(cases "list_diff xs ys")
+  assume " list_diff xs ys = []" with list_diff_le show ?thesis by blast
+next
+  fix e es
+  assume h: "list_diff xs ys = e # es"
+  show ?thesis
+  proof(cases "list_diff ys xs")
+    assume " list_diff ys xs = []" with list_diff_le show ?thesis by blast
+  next
+    fix a as assume h1: "list_diff ys xs = (a # as)"
+    have "xs = (list_com xs ys) @ [e] @ es \<and> ys = (list_com xs ys) @ [a] @ as \<and> e \<noteq> a"
+      apply (simp, fold h1, fold h)
+      apply (simp,subst list_com_commute, simp)
+      apply (rule_tac list_diff_neq[rule_format])
+      by (insert h1, insert h, blast)
+    thus ?thesis by blast
+  qed
+qed
+
+(* In fact, this is a case split *)
+lemma list_diff_ind: "\<lbrakk>list_diff xs ys = [] \<Longrightarrow> R; list_diff ys xs = [] \<Longrightarrow> R; 
+                             \<And> e es a as. \<lbrakk>list_diff xs ys = e#es; list_diff ys xs = a#as; e \<noteq> a\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+proof -
+  assume h1: "list_diff xs ys = [] \<Longrightarrow> R"
+    and h2: "list_diff ys xs = [] \<Longrightarrow> R"
+    and h3: "\<And> e es a as. \<lbrakk>list_diff xs ys = e#es; list_diff ys xs = a#as; e \<noteq> a\<rbrakk> \<Longrightarrow> R"
+  show R
+  proof(cases "list_diff xs ys")
+    assume "list_diff xs ys = []" from h1 [OF this] show R .
+  next
+    fix e es
+    assume he: "list_diff xs ys = e#es"
+    show R
+    proof(cases "list_diff ys xs")
+      assume "list_diff ys xs = []" from h2 [OF this] show R .
+    next
+      fix a as
+      assume ha: "list_diff ys xs = a#as" show R
+      proof(rule h3 [OF he ha])
+	from list_diff_neq [rule_format, OF conjI [OF he ha ]]
+	show "e \<noteq> a" .
+      qed
+    qed
+  qed
+qed
+
+lemma list_diff_rev_ind: 
+  "\<lbrakk>xs \<setminus> ys = [] \<Longrightarrow> R; ys \<setminus> xs = [] \<Longrightarrow> R; \<And> e es a as. \<lbrakk>xs \<setminus> ys = es@[e]; ys \<setminus> xs = as@[a]; e \<noteq> a\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+proof -
+  fix xs ys R
+  assume h1: "xs \<setminus> ys = [] \<Longrightarrow> R"
+    and h2: "ys \<setminus> xs = [] \<Longrightarrow> R"
+    and h3: "\<And> e es a as. \<lbrakk>xs \<setminus> ys = es@[e]; ys \<setminus> xs = as@[a]; e \<noteq> a\<rbrakk> \<Longrightarrow> R"
+  show R
+  proof (rule list_diff_ind [where xs = "rev xs" and ys = "rev ys"])
+    assume "list_diff (rev xs) (rev ys) = []" thus R by (auto intro:h1 simp:list_diff_rev_def)
+  next
+    assume "list_diff (rev ys) (rev xs) = []" thus R by (auto intro:h2 simp:list_diff_rev_def)
+  next
+    fix e es a as
+    assume "list_diff (rev xs) (rev ys) = e # es"
+      and "list_diff (rev ys) (rev xs) = a # as" 
+      and " e \<noteq> a"
+    thus R by (auto intro:h3 simp:list_diff_rev_def)
+  qed
+qed
+
+lemma djoin_diff_iff: "(xs \<asymp> ys) = (\<exists> e es a as. list_diff (rev xs) (rev ys) = e#es \<and> list_diff (rev ys) (rev xs) = a#as \<and> a \<noteq> e)"
+proof (rule list_diff_ind [where xs = "rev xs" and ys = "rev ys"])
+  assume "list_diff (rev xs) (rev ys) = []"
+  hence "xs \<preceq> ys" by (unfold no_junior_def, simp add:list_diff_le)
+  thus ?thesis 
+    apply (auto simp:djoin_def no_junior_def)
+    by (fold list_diff_le, simp)
+next
+  assume "list_diff (rev ys) (rev xs) = []"
+  hence "ys \<preceq> xs" by (unfold no_junior_def, simp add:list_diff_le)
+  thus ?thesis 
+    apply (auto simp:djoin_def no_junior_def)
+    by (fold list_diff_le, simp)
+next
+  fix e es a as
+  assume he: "list_diff (rev xs) (rev ys) = e # es"
+    and ha: "list_diff (rev ys) (rev xs) = a # as"
+    and hn: "e \<noteq> a"
+  show ?thesis
+  proof
+    from he ha hn
+    show 
+      "\<exists>e es a as. list_diff (rev xs) (rev ys) = e # es \<and> list_diff (rev ys) (rev xs) = a # as \<and> a \<noteq> e" 
+      by blast
+  next
+    from he ha hn
+    show "xs \<asymp> ys" 
+      by (auto simp:djoin_def no_junior_def, fold list_diff_le, simp+)
+  qed
+qed
+
+lemma djoin_diff_rev_iff: "(xs \<asymp> ys) = (\<exists> e es a as. xs \<setminus> ys = es@[e] \<and> ys \<setminus> xs = as@[a] \<and> a \<noteq> e)"
+  apply (auto simp:djoin_diff_iff list_diff_rev_def)
+  apply (rule_tac x = e in exI, safe)
+  apply (rule_tac x = "rev es" in exI)
+  apply (rule_tac injD[where f = rev], simp+)
+  apply (rule_tac x = "a" in exI, safe)
+  apply (rule_tac x = "rev as" in exI)
+  apply (rule_tac injD[where f = rev], simp+)
+  done
+
+lemma djoin_revE: "\<lbrakk>xs \<asymp> ys; \<And>e es a as. \<lbrakk>xs \<setminus> ys = es@[e]; ys \<setminus> xs = as@[a]; a \<noteq> e\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+  by (unfold djoin_diff_rev_iff, blast)
+
+lemma djoin_append_left[simp, intro]: "xs \<asymp> ys \<Longrightarrow> (zs' @ xs) \<asymp> (zs @ ys)"
+  by (auto simp:djoin_diff_iff intro:list_diff_djoin[simplified])
+
+lemma djoin_cons_left[simp]: "xs \<asymp> ys \<Longrightarrow> (e # xs) \<asymp> (a # ys)"
+  by (drule_tac zs' = "[e]" and zs = "[a]" in djoin_append_left, simp)
+
+lemma djoin_simp_1 [simp]: "xs \<asymp> ys \<Longrightarrow> xs \<asymp> (zs @ ys)"
+  by (drule_tac djoin_append_left [where zs' = "[]"], simp)
+
+lemma djoin_simp_2 [simp]: "xs \<asymp> ys \<Longrightarrow> (zs' @ xs) \<asymp> ys"
+  by (drule_tac djoin_append_left [where zs = "[]"], simp)
+
+lemma djoin_append_right[simp]: "xs \<asymp> ys \<Longrightarrow> (xs @ zs) \<asymp> (ys @ zs)"
+  by (simp add:djoin_diff_iff)
+
+lemma djoin_cons_append[simp]: "xs \<asymp> ys \<Longrightarrow> (x # xs) \<asymp> (zs @ ys)"
+  by (subgoal_tac "[x] @ xs \<asymp> zs @ ys", simp, blast)
+
+lemma djoin_append_cons[simp]: "xs \<asymp> ys \<Longrightarrow> (zs @ xs) \<asymp> (y # ys)"
+  by (subgoal_tac "zs @ xs \<asymp> [y] @ ys", simp, blast)
+
+lemma djoin_neq [simp]: "xs \<asymp> ys \<Longrightarrow> xs \<noteq> ys"
+  by (simp only:djoin_diff_iff, clarsimp)
+
+lemma djoin_cons [simp]: "e \<noteq> a \<Longrightarrow> e # xs \<asymp> a # xs"
+  by (unfold djoin_diff_iff, simp)
+
+lemma djoin_append_e [simp]: "e \<noteq> a \<Longrightarrow> (zs @ [e] @ xs) \<asymp> (zs' @ [a] @ xs)"
+  by (unfold djoin_diff_iff, simp)
+
+lemma djoin_mono [simp]: "\<lbrakk>xs \<asymp> ys; xs \<preceq> xs'; ys \<preceq> ys'\<rbrakk> \<Longrightarrow> xs' \<asymp> ys'"
+proof(erule_tac djoin_revE,unfold djoin_diff_rev_iff)
+  fix e es a as
+  assume hx: "xs \<preceq> xs'"
+    and hy: "ys \<preceq> ys'"
+    and hmx: "xs \<setminus> ys = es @ [e]"
+    and hmy: "ys \<setminus> xs = as @ [a]" 
+    and neq: "a \<noteq> e"
+  have "xs' \<setminus> ys' = ((xs' \<setminus> xs) @ es) @ [e] \<and> ys' \<setminus> xs' = ((ys' \<setminus> ys) @ as) @ [a] \<and> a \<noteq> e"
+  proof -
+    from hx have heqx: "(xs' \<setminus> xs) @ xs = xs'"
+      by (cut_tac list_com_diff_rev [of xs' xs], subgoal_tac "xs' \<bullet> xs = xs", simp+)
+    moreover from hy have heqy: "(ys' \<setminus> ys) @ ys = ys'" 
+      by (cut_tac list_com_diff_rev [of ys' ys], subgoal_tac "ys' \<bullet> ys = ys", simp+)
+    moreover from list_diff_djoin_rev_simplified [OF hmx hmy] 
+    have "((xs' \<setminus> xs) @ xs) \<setminus>  ((ys' \<setminus> ys) @ ys) = (xs' \<setminus> xs) @ es @ [e]" by simp
+    moreover from list_diff_djoin_rev_simplified [OF hmy hmx] 
+    have "((ys' \<setminus> ys) @ ys) \<setminus>  ((xs' \<setminus> xs) @ xs) = (ys' \<setminus> ys) @ as @ [a]" by simp
+    ultimately show ?thesis by (simp add:neq)
+  qed
+  thus "\<exists>e es a as. xs' \<setminus> ys' = es @ [e] \<and> ys' \<setminus> xs' = as @ [a] \<and> a \<noteq> e" by blast
+qed
+
+lemmas djoin_append_e_simplified [simp] = djoin_append_e [simplified]
+
+(*<*)
+end
+(*>*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/obj2sobj_prop.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,674 @@
+theory obj2sobj_prop 
+imports Main rc_theory os_rc deleted_prop
+begin
+
+context tainting_s_complete begin
+
+(** file 2 sfile   **)
+
+lemma init_son_deleted_D:
+  "\<lbrakk>deleted (File pf) s; f # pf \<in> init_files; valid s\<rbrakk> \<Longrightarrow> deleted (File (f # pf)) s"
+apply (induct s, simp)
+by (frule valid_cons, frule valid_os, case_tac a, auto dest:init_notin_curf_deleted)
+
+lemma init_parent_undeleted_I:
+  "\<lbrakk>\<not> deleted (File (f # pf)) s; f # pf \<in> init_files; valid s\<rbrakk> \<Longrightarrow> \<not> deleted (File pf) s"
+by (rule notI, simp add:init_son_deleted_D)
+
+lemma source_dir_in_init:
+  "source_dir s f = Some sd \<Longrightarrow> sd \<in> init_files"
+by (induct f, auto split:if_splits)
+
+lemma source_dir_of_init: "\<lbrakk>source_dir [] f = Some sd; f \<in> init_files\<rbrakk> \<Longrightarrow> f = sd"
+by (induct f, auto)
+
+lemma source_dir_of_init': "f \<in> init_files \<Longrightarrow> source_dir [] f = Some f"
+by (induct f, auto)
+
+lemma init_not_curf_imp_deleted:
+  "\<lbrakk>f \<in> init_files; f \<notin> current_files s; valid s\<rbrakk> \<Longrightarrow> deleted (File f) s"
+apply (induct s, simp)
+apply (frule valid_cons, frule valid_os, case_tac a, auto)
+done
+
+lemma source_dir_of_init'': 
+  "\<lbrakk>f \<in> init_files; \<not> deleted (File f) s; valid s\<rbrakk> \<Longrightarrow> source_dir s f = Some f"
+by (induct f, auto)
+
+
+lemma source_dir_createf:
+  "valid (CreateFile p (f#pf) # s) \<Longrightarrow> 
+  source_dir (CreateFile p (f#pf) # s) = (source_dir s) ((f#pf)  := source_dir s pf)"
+apply (frule valid_os, frule valid_cons)
+apply (rule ext, induct_tac x)
+apply (auto dest:init_not_curf_imp_deleted)
+done
+
+lemma source_dir_createf':
+  "valid (CreateFile p f # s) \<Longrightarrow> 
+  source_dir (CreateFile p f # s) = (source_dir s) (f := (case (parent f) of
+                                                            Some pf \<Rightarrow> source_dir s pf
+                                                          | _       \<Rightarrow> None))"
+apply (frule valid_os, case_tac f, simp+)
+apply (drule source_dir_createf, auto)
+done
+
+lemma source_dir_other:
+  "\<lbrakk>valid (e # s); \<forall> p f. e \<noteq> CreateFile p f; \<forall> p f. e \<noteq> DeleteFile p f\<rbrakk>
+   \<Longrightarrow> source_dir (e#s) = source_dir s"
+apply (rule ext, induct_tac x, simp)
+apply (auto dest:not_deleted_cons_D)
+apply (case_tac [!] e, auto)
+done
+
+lemma source_dir_deletef:
+  "valid (DeleteFile p f # s) \<Longrightarrow> source_dir (DeleteFile p f # s) f' =  
+     (if (source_dir s f') = Some f then parent f else (source_dir s f'))"
+apply (frule valid_os, frule valid_cons)
+apply (case_tac "f \<in> init_files")
+apply (induct_tac f', simp)
+apply (auto dest!:init_parent_undeleted_I intro:parent_file_in_init'
+            intro!: source_dir_of_init'')[1]
+apply (induct_tac f', auto)
+done
+
+lemma source_dir_deletef':
+  "valid (DeleteFile p f # s) \<Longrightarrow> source_dir (DeleteFile p f # s) = (\<lambda> f'.   
+     (if (source_dir s f') = Some f then parent f else (source_dir s f')) )"
+by (auto dest:source_dir_deletef)
+
+lemmas source_dir_simps = source_dir_of_init' source_dir_of_init'' source_dir_createf' 
+  source_dir_deletef' source_dir_other
+
+declare source_dir.simps [simp del]
+
+lemma source_dir_is_ancient:
+  "source_dir s f = Some sd ==> sd \<preceq> f"
+apply (induct f)
+by (auto simp:source_dir.simps no_junior_def split:if_splits)
+
+lemma no_junior_trans: "\<lbrakk>f \<preceq> f'; f' \<preceq> f''\<rbrakk> \<Longrightarrow> f \<preceq> f''"
+by (auto elim:no_juniorE)
+
+lemma ancient_has_parent:
+  "[| f \<preceq> f'; f \<noteq> f'|] ==> \<exists> sonf. parent sonf = Some f \<and> sonf \<preceq> f' "
+apply (induct f')
+apply (simp add:no_junior_def)
+apply (case_tac "f = f'")
+apply (rule_tac x = "a # f'" in exI, simp add:no_junior_def)
+apply (frule no_junior_noteq, simp)
+apply clarsimp
+apply (rule_tac x = sonf in exI, simp add:no_junior_trans)
+done
+
+lemma source_dir_prop:
+  "[|\<forall>fn. fn # f' \<notin> current_files s; source_dir s f = Some f'; f \<in> current_files s; valid s|]
+  ==> f = f'"
+  apply (drule source_dir_is_ancient)
+  apply (case_tac "f = f'", simp)
+  apply (drule ancient_has_parent, simp, clarsimp)
+  apply (drule_tac ancient_file_in_current, simp+)
+  apply (case_tac sonf, auto)
+  done
+
+lemma current_file_has_sd:
+  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> sd. source_dir s f = Some sd"
+apply (induct s arbitrary:f, simp add:source_dir_of_init')
+apply (frule valid_cons, frule valid_os, case_tac a, auto simp:source_dir_simps)
+apply (case_tac list, simp)
+apply (rule_tac f = f in cannot_del_root, simp+)
+done
+
+lemma current_file_has_sd':
+  "\<lbrakk>source_dir s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
+by (rule notI, auto dest:current_file_has_sd)
+
+lemma current_file_has_sfile:
+  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> et sd. cf2sfile s f = Some (et, sd)"
+apply (frule current_file_has_sd, simp+)
+apply (frule current_file_has_etype, auto)
+done
+
+lemma current_file_has_sfile':
+  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> sf. cf2sfile s f = Some sf"
+by (auto dest:current_file_has_sfile)
+  
+(*
+lemma not_deleted_sf_remains:
+  "\<lbrakk>f \<in> current_files s; \<not> deleted (File f) s; valid s\<rbrakk> \<Longrightarrow> "
+*)
+
+lemma current_proc_has_sproc:
+  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> r fr pt u. cp2sproc s p = Some (r, fr, pt, u)"
+apply (frule current_proc_has_role, simp+)
+apply (frule current_proc_has_type, simp)
+apply (frule current_proc_has_forcedrole, simp)
+apply (frule current_proc_has_owner, auto)
+done
+
+lemma current_proc_has_sproc':
+  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> sp. cp2sproc s p = Some sp"
+by (auto dest!:current_proc_has_sproc)
+
+lemma current_ipc_has_sipc: 
+  "\<lbrakk>i \<in> current_ipcs s; valid s\<rbrakk> \<Longrightarrow> \<exists> t. ci2sipc s i = Some t"
+by (drule current_ipc_has_type, auto)
+
+lemma init_file_has_sobj:
+  "f \<in> init_files \<Longrightarrow> \<exists> t sd. init_obj2sobj (File f) = SFile (t, sd) (Some f)"
+by (frule init_file_has_etype, clarsimp)
+
+lemma init_proc_has_sobj:
+  assumes pinit:"p \<in> init_processes"
+  shows "\<exists> r fr pt u. init_obj2sobj (Proc p) = SProc (r, fr, pt, u) (Some p)"
+proof -
+  from pinit obtain r where "init_currentrole p = Some r" 
+    using init_proc_has_role by (auto simp:bidirect_in_init_def)
+  moreover from pinit obtain fr where "init_proc_forcedrole p = Some fr"
+    using init_proc_has_frole by (auto simp:bidirect_in_init_def)
+  moreover from pinit obtain pt where "init_process_type p = Some pt"
+    using init_proc_has_type by (auto simp:bidirect_in_init_def)
+  moreover from pinit obtain u where "init_owner p = Some u"
+    using init_proc_has_owner by (auto simp:bidirect_in_init_def)
+  ultimately show ?thesis by auto 
+qed
+
+lemma init_ipc_has_sobj:
+  "i \<in> init_ipcs \<Longrightarrow> \<exists> t. init_obj2sobj (IPC i) = SIPC t (Some i)"
+using init_ipc_has_type
+by (auto simp:bidirect_in_init_def)
+
+lemma init_obj_has_sobj:
+  "exists [] obj \<Longrightarrow> init_obj2sobj obj \<noteq> Unknown"
+apply (case_tac obj)
+apply (simp_all only:exists.simps current_procs.simps current_ipcs.simps current_files.simps) 
+apply (auto dest!:init_proc_has_sobj init_file_has_sobj init_ipc_has_sobj)
+done
+
+lemma exists_obj_has_sobj:
+  "\<lbrakk>exists s obj; valid s\<rbrakk> \<Longrightarrow> obj2sobj s obj \<noteq> Unknown"
+apply (case_tac obj)
+apply (auto dest!:current_ipc_has_sipc current_proc_has_sproc' current_file_has_sfile' 
+            split:option.splits)
+done
+
+lemma current_proc_has_srp:
+  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> srp. source_proc s p = Some srp"
+apply (induct s arbitrary:p, simp)
+by (frule valid_cons, frule valid_os, case_tac a, auto)
+
+lemma current_proc_has_sobj:
+  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> r fr t u srp. obj2sobj s (Proc p) = SProc (r,fr,t,u) (Some srp)"
+apply (frule current_proc_has_sproc')
+apply (auto dest:current_proc_has_srp)
+done
+
+lemma current_file_has_sobj:
+  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> t sd srf. obj2sobj s (File f) = SFile (t, sd) srf"
+by (auto dest:current_file_has_sfile)
+
+lemma current_ipc_has_sobj:
+  "\<lbrakk>i \<in> current_ipcs s; valid s\<rbrakk> \<Longrightarrow> \<exists> t sri. obj2sobj s (IPC i) = SIPC t sri"
+by (auto dest:current_ipc_has_sipc)
+
+lemma sobj_has_proc_role:
+  "obj2sobj s (Proc p) = SProc (r, fr, t, u) srp \<Longrightarrow> currentrole s p = Some r"
+by (auto split:option.splits)
+
+lemma chown_role_aux_valid:
+  "\<lbrakk>currentrole s p = Some r; proc_forcedrole s p = Some fr\<rbrakk>
+  \<Longrightarrow> chown_role_aux r fr u = currentrole (ChangeOwner p u # s) p"
+by (auto split:t_role.splits simp:chown_role_aux_def dest:proc_forcedrole_valid)
+
+lemma chown_role_aux_valid':
+  "cp2sproc s p = Some (r, fr, t, u') \<Longrightarrow> chown_role_aux r fr u = currentrole (ChangeOwner p u # s) p"
+by (rule chown_role_aux_valid, auto split:option.splits)
+
+lemma chown_type_aux_valid:
+  "\<lbrakk>currentrole s p = Some r; currentrole (ChangeOwner p u # s) p = Some nr; type_of_process s p = Some t\<rbrakk>
+  \<Longrightarrow> type_of_process (ChangeOwner p u # s) p = Some (chown_type_aux r nr t)" 
+apply (auto split:option.splits t_rc_proc_type.splits 
+             dest:default_process_create_type_valid
+             simp:chown_type_aux_def pot_def pct_def)
+done
+
+lemma chown_type_aux_valid':
+  "\<lbrakk>cp2sproc s p = Some (r, fr, t, u'); currentrole (ChangeOwner p u # s) p = Some nr\<rbrakk> 
+   \<Longrightarrow> type_of_process (ChangeOwner p u # s) p = Some (chown_type_aux r nr t)"
+by (rule chown_type_aux_valid, auto split:option.splits)
+
+lemma exec_type_aux_valid:
+  "\<lbrakk>currentrole s p = Some r; type_of_process s p = Some t\<rbrakk>
+  \<Longrightarrow> type_of_process (Execute p f # s) p = Some (exec_type_aux r t)" 
+apply (auto split:option.splits t_rc_proc_type.splits 
+             dest:default_process_execute_type_valid
+             simp:exec_type_aux_def pet_def)
+done
+
+lemma exec_type_aux_valid':
+  "cp2sproc s p = Some (r, fr, t, u') \<Longrightarrow> type_of_process (Execute p f # s) p = Some (exec_type_aux r t)" 
+by (rule exec_type_aux_valid, auto split:option.splits)
+
+lemma non_initf_frole_inherit:
+  "\<lbrakk>f \<notin> init_files; f \<noteq> []\<rbrakk> \<Longrightarrow> forcedrole s f = Some InheritParentRole"
+apply (induct s) defer
+apply (case_tac a, auto) 
+apply (induct f, auto split:option.splits dest:init_frole_has_file)
+done
+
+lemma non_initf_irole_inherit:
+  "\<lbrakk>f \<notin> init_files; f \<noteq> []\<rbrakk> \<Longrightarrow> initialrole s f = Some InheritParentRole"
+apply (induct s) defer
+apply (case_tac a, auto) 
+apply (induct f, auto split:option.splits dest:init_irole_has_file)
+done
+
+lemma deleted_file_frole_inherit:
+  "\<lbrakk>deleted (File f) s; f \<in> current_files s\<rbrakk> \<Longrightarrow> forcedrole s f = Some InheritParentRole"
+apply (induct s, simp)
+apply (case_tac a, auto) 
+done
+
+lemma deleted_file_irole_inherit:
+  "\<lbrakk>deleted (File f) s; f \<in> current_files s\<rbrakk> \<Longrightarrow> initialrole s f = Some InheritParentRole"
+apply (induct s, simp)
+apply (case_tac a, auto) 
+done
+
+lemma sd_deter_efrole:
+  "\<lbrakk>source_dir s f = Some sd; valid s; f \<in> current_files s\<rbrakk> 
+  \<Longrightarrow> effforcedrole s f = effforcedrole s sd"
+apply (induct f)
+apply (drule source_dir_is_ancient, simp add:no_junior_def)
+apply (simp add:source_dir.simps split:if_splits)
+apply (frule parent_file_in_current', simp)
+apply (case_tac "a # f \<in> init_files", simp)
+apply (drule_tac deleted_file_frole_inherit, simp, simp add:effforcedrole_def)
+apply (drule_tac s = s in non_initf_frole_inherit, simp, simp add:effforcedrole_def)
+done
+
+lemma sd_deter_eirole:
+  "\<lbrakk>source_dir s f = Some sd; valid s; f \<in> current_files s\<rbrakk> 
+  \<Longrightarrow> effinitialrole s f = effinitialrole s sd"
+apply (induct f)
+apply (drule source_dir_is_ancient, simp add:no_junior_def)
+apply (simp add:source_dir.simps split:if_splits)
+apply (frule parent_file_in_current', simp)
+apply (case_tac "a # f \<in> init_files", simp)
+apply (drule_tac deleted_file_irole_inherit, simp, simp add:effinitialrole_def)
+apply (drule_tac s = s in non_initf_irole_inherit, simp, simp add:effinitialrole_def)
+done
+
+lemma undel_initf_keeps_frole:
+  "\<lbrakk>f \<in> init_files; \<not> deleted (File f) s; valid s\<rbrakk>
+   \<Longrightarrow> forcedrole s f = init_file_forcedrole f"
+apply (induct s, simp)
+apply (frule valid_cons, frule valid_os, case_tac a)
+apply (auto dest:init_notin_curf_deleted)
+done
+
+lemma undel_initf_keeps_efrole:
+  "\<lbrakk>f \<in> init_files; \<not> deleted (File f) s; valid s\<rbrakk>
+   \<Longrightarrow> effforcedrole s f = erole_functor init_file_forcedrole InheritUpMixed f"
+apply (induct f)
+apply (drule undel_initf_keeps_frole, simp, simp)
+apply (simp add:effforcedrole_def)
+apply (frule parent_file_in_init', frule init_parent_undeleted_I, simp+)
+apply (drule undel_initf_keeps_frole, simp, simp)
+apply (simp add:effforcedrole_def)
+done
+
+lemma undel_initf_keeps_irole:
+  "\<lbrakk>f \<in> init_files; \<not> deleted (File f) s; valid s\<rbrakk>
+   \<Longrightarrow> initialrole s f = init_file_initialrole f"
+apply (induct s, simp)
+apply (frule valid_cons, frule valid_os, case_tac a)
+apply (auto dest:init_notin_curf_deleted)
+done
+
+lemma undel_initf_keeps_eirole:
+  "\<lbrakk>f \<in> init_files; \<not> deleted (File f) s; valid s\<rbrakk>
+   \<Longrightarrow> effinitialrole s f = erole_functor init_file_initialrole UseForcedRole f"
+apply (induct f)
+apply (drule undel_initf_keeps_irole, simp, simp)
+apply (simp add:effinitialrole_def)
+apply (frule parent_file_in_init', frule init_parent_undeleted_I, simp+)
+apply (drule undel_initf_keeps_irole, simp, simp)
+apply (simp add:effinitialrole_def)
+done
+
+lemma source_dir_not_deleted:
+  "source_dir s f = Some sd \<Longrightarrow> \<not> deleted (File sd) s"
+by (induct f, auto simp:source_dir.simps split:if_splits)
+
+lemma exec_role_aux_valid:
+  "\<lbrakk>currentrole s p = Some r; source_dir s f = Some sd; owner s p = Some u; 
+  f \<in> current_files s; valid s\<rbrakk>
+  \<Longrightarrow> exec_role_aux r sd u = currentrole (Execute p f # s) p"
+apply (frule sd_deter_eirole, simp+, frule sd_deter_efrole, simp+)
+apply (frule source_dir_in_init, drule source_dir_not_deleted)
+apply (simp add:undel_initf_keeps_eirole undel_initf_keeps_efrole)
+apply (frule file_has_effinitialrole, simp, frule file_has_effforcedrole, simp)
+apply (auto split:option.splits t_role.splits simp:map_comp_def exec_role_aux_def
+             dest:effforcedrole_valid effinitialrole_valid)
+done
+
+lemma exec_role_aux_valid':
+  "\<lbrakk>cp2sproc s p = Some (r, fr, t, u); source_dir s f = Some sd; f \<in> current_files s; valid s\<rbrakk>
+  \<Longrightarrow> exec_role_aux r sd u = currentrole (Execute p f # s) p"
+by (rule exec_role_aux_valid, auto split:option.splits)
+
+lemma cp2sproc_nil_init:
+  "init_obj2sobj (Proc p) = (case (cp2sproc [] p) of 
+                               Some sp \<Rightarrow> SProc sp (Some p)
+                             | _       \<Rightarrow> Unknown)"
+by (auto split:option.splits)
+
+lemma cf2sfile_nil_init:
+  "init_obj2sobj (File f) = (case (cf2sfile [] f) of 
+                               Some sf \<Rightarrow> SFile sf (Some f)
+                             | _       \<Rightarrow> Unknown)"
+apply (auto split:option.splits simp:etype_of_file_def)
+apply (case_tac "f \<in> init_files", simp add:source_dir_of_init')
+apply (induct f, simp+)
+apply (case_tac "f \<in> init_files", simp add:source_dir_of_init')
+apply (induct f, simp+)
+done
+
+lemma ci2sipc_nil_init:
+  "init_obj2sobj (IPC i) = (case (ci2sipc [] i) of 
+                              Some si \<Rightarrow> SIPC si (Some i)
+                            | _       \<Rightarrow> Unknown)"
+by simp
+
+lemma obj2sobj_nil_init:
+  "exists [] obj \<Longrightarrow> obj2sobj [] obj = init_obj2sobj obj" 
+apply (case_tac obj) 
+apply (auto simp:cf2sfile_nil_init cp2sproc_nil_init ci2sipc_nil_init 
+                 source_dir_of_init' etype_of_file_def
+           split:if_splits option.splits)
+done
+
+(**** cp2sproc simpset ****)
+
+lemma current_proc_has_role':
+  "\<lbrakk>currentrole s p = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
+by (rule notI, auto dest:current_proc_has_role)
+
+lemma cp2sproc_chown:
+  assumes vs: "valid (ChangeOwner p u # s)"
+  shows "cp2sproc (ChangeOwner p u # s) = (cp2sproc s) 
+     (p := (case (cp2sproc s p) of
+              Some (r,fr,pt,u') \<Rightarrow> (case (chown_role_aux r fr u) of 
+                                      Some nr \<Rightarrow> Some (nr,fr,chown_type_aux r nr pt,u)
+                                    | _       \<Rightarrow> None)
+            | _                 \<Rightarrow> None)
+      )" (is "?lhs = ?rhs")
+proof-
+  have os: "os_grant s (ChangeOwner p u)" and vs': "valid s" using vs
+    by (auto dest:valid_cons valid_os)
+  have "\<And> x. x \<noteq> p \<Longrightarrow> ?lhs x = ?rhs x"
+    by (auto simp:type_of_process.simps split:option.splits t_role.splits)
+  moreover have "?lhs p = ?rhs p"
+  proof-
+    from os have p_in: "p \<in> current_procs s" by (simp+)
+    then obtain r fr t u' where csp: "cp2sproc s p = Some (r, fr, t, u')" using vs'
+      by (drule_tac current_proc_has_sproc, auto)  
+    from os have "u \<in> init_users" by simp
+    hence "defrole u \<noteq> None" using init_user_has_role by (auto simp:bidirect_in_init_def)
+    then obtain nr where nrole:"chown_role_aux r fr u = Some nr"
+      by (case_tac fr, auto simp:chown_role_aux_def)
+    have nr_eq: "currentrole (ChangeOwner p u # s) p = chown_role_aux r fr u" 
+      using csp by (auto simp:chown_role_aux_valid'[where u = u])  
+    moreover have "type_of_process (ChangeOwner p u # s) p = Some (chown_type_aux r nr t)"
+      using csp nrole nr_eq 
+      by (rule_tac fr = fr and u' = u' in chown_type_aux_valid', simp+)
+    moreover have "proc_forcedrole (ChangeOwner p u # s) p = Some fr"
+      using csp by (auto split:option.splits)
+    moreover have "owner (ChangeOwner p u # s) p = Some u" by simp
+    ultimately have "cp2sproc (ChangeOwner p u # s) p = Some (nr, fr, chown_type_aux r nr t, u)" 
+      using nrole by (simp)
+    thus ?thesis using csp nrole by simp
+  qed
+  ultimately show ?thesis by (rule_tac ext, auto)
+qed
+
+lemma cp2sproc_crole:
+  "valid (ChangeRole p r # s) \<Longrightarrow> cp2sproc (ChangeRole p r # s) = (cp2sproc s) 
+     (p := (case (cp2sproc s p) of
+              Some (r',fr,pt,u) \<Rightarrow> Some (r,fr,pt,u)
+            | _                 \<Rightarrow> None)
+      )"
+apply (frule valid_cons, frule valid_os, simp)
+apply (frule current_proc_has_sproc, simp)
+apply (rule ext, auto split:option.splits)
+done
+
+lemma cp2sproc_exec:
+  assumes vs: "valid (Execute p f # s)"
+  shows "cp2sproc (Execute p f # s) = (cp2sproc s) 
+     (p := (case (cp2sproc s p, source_dir s f) of 
+              (Some (r,fr,pt,u), Some sd) \<Rightarrow> (
+     case (exec_role_aux r sd u, erole_functor init_file_forcedrole InheritUpMixed sd) of
+       (Some r', Some fr') \<Rightarrow> Some (r', fr', exec_type_aux r pt, u)
+     | _                   \<Rightarrow> None           )
+            | _                \<Rightarrow> None))" (is "?lhs = ?rhs")
+proof-
+  have os: "os_grant s (Execute p f)" and vs': "valid s" using vs
+    by (auto dest:valid_cons valid_os)
+  have "\<And> x. x \<noteq> p \<Longrightarrow> ?lhs x = ?rhs x"
+    by (auto simp:type_of_process.simps split:option.splits t_role.splits)
+  moreover have "?lhs p = ?rhs p"
+  proof-
+    from os have p_in: "p \<in> current_procs s" by (simp+)
+    then obtain r fr t u where csp: "cp2sproc s p = Some (r, fr, t, u)" using vs'
+      by (drule_tac current_proc_has_sproc, auto)  
+    from os have f_in: "f \<in> current_files s" by simp
+    then obtain sd where sdir: "source_dir s f = Some sd" using vs'
+      by (drule_tac current_file_has_sd, auto)
+    have "currentrole (Execute p f # s) p \<noteq> None" using vs p_in
+      by (rule_tac notI, drule_tac current_proc_has_role', simp+)
+    then obtain nr where nrole: "currentrole (Execute p f # s) p = Some nr" by auto
+    have "proc_forcedrole (Execute p f # s) p \<noteq> None" using vs p_in
+      by (rule_tac notI, drule_tac current_proc_has_forcedrole', simp+)
+    then obtain nfr where nfrole: "proc_forcedrole (Execute p f # s) p = Some nfr" by auto
+    have nr_eq: "currentrole (Execute p f # s) p = exec_role_aux r sd u" 
+      using csp f_in sdir vs' by (simp only:exec_role_aux_valid')
+    moreover have "type_of_process (Execute p f # s) p = Some (exec_type_aux r t)"
+      using csp by (simp only:exec_type_aux_valid')
+    moreover have nfr_eq: "proc_forcedrole (Execute p f # s) p = 
+                           erole_functor init_file_forcedrole InheritUpMixed sd" 
+      using sdir vs' f_in
+      apply (frule_tac source_dir_in_init, drule_tac source_dir_not_deleted)
+      by (simp add:undel_initf_keeps_efrole sd_deter_efrole)
+    moreover have "owner (Execute p f # s) p = Some u" using csp
+      by (auto split:option.splits)
+    ultimately have "cp2sproc (Execute p f # s) p = Some (nr, nfr, exec_type_aux r t, u)" 
+      using nrole nfrole by (simp)
+    moreover have "exec_role_aux r sd u = Some nr" using nrole nr_eq by simp
+    moreover have "erole_functor init_file_forcedrole InheritUpMixed sd = Some nfr"
+      using nfrole nfr_eq by simp
+    ultimately show ?thesis using csp sdir by simp 
+  qed
+  ultimately show ?thesis by (rule_tac ext, auto)
+qed
+
+lemma cp2sproc_clone:
+  "valid (Clone p p' # s) \<Longrightarrow> cp2sproc (Clone p p' # s) = (cp2sproc s) (p' := 
+      (case (cp2sproc s p) of
+         Some (r, fr, pt, u) \<Rightarrow> Some (r, fr, clone_type_aux r pt, u)
+       | _                   \<Rightarrow> None))"
+apply (frule valid_cons, frule valid_os)
+apply (rule ext, auto split:option.splits t_rc_proc_type.splits 
+                       simp:pct_def clone_type_aux_def
+                       dest:current_proc_has_type default_process_create_type_valid)
+done
+
+lemma cp2sproc_other:
+  "\<lbrakk>valid (e # s); \<forall> p f. e \<noteq> Execute p f; \<forall> p p'. e \<noteq> Clone p p';
+    \<forall> p r. e \<noteq> ChangeRole p r; \<forall> p u. e \<noteq> ChangeOwner p u\<rbrakk> \<Longrightarrow> cp2sproc (e#s) = cp2sproc s"
+by (case_tac e, auto)
+
+lemmas cp2sproc_simps = cp2sproc_exec cp2sproc_chown cp2sproc_crole cp2sproc_clone cp2sproc_other
+
+lemma obj2sobj_file: "obj2sobj s obj = SFile sf fopt \<Longrightarrow> \<exists> f. obj = File f"
+by (case_tac obj, case_tac [!] s, auto split:option.splits if_splits)
+
+lemma obj2sobj_proc: "obj2sobj s obj = SProc sp popt \<Longrightarrow> \<exists> p. obj = Proc p"
+by (case_tac obj, case_tac [!] s, auto split:option.splits if_splits)
+
+lemma obj2sobj_ipc: "obj2sobj s obj = SIPC si iopt \<Longrightarrow> \<exists> i. obj = IPC i"
+by (case_tac obj, case_tac [!] s, auto split:option.splits if_splits)
+
+lemma obj2sobj_file': 
+  "\<lbrakk>obj2sobj s (File f) = sobj; sobj \<noteq> Unknown\<rbrakk> \<Longrightarrow> \<exists> sf srf. sobj = SFile sf srf"
+by (case_tac sobj, case_tac [!] s, auto split:option.splits if_splits)
+
+lemma obj2sobj_proc': 
+  "\<lbrakk>obj2sobj s (Proc p) = sobj; sobj \<noteq> Unknown\<rbrakk> \<Longrightarrow> \<exists> sp srp. sobj = SProc sp srp"
+by (case_tac sobj, case_tac [!] s, auto split:option.splits if_splits)
+
+lemma obj2sobj_ipc': 
+  "\<lbrakk>obj2sobj s (IPC i) = sobj; sobj \<noteq> Unknown\<rbrakk> \<Longrightarrow> \<exists> si sri. sobj =  SIPC si sri"
+by (case_tac sobj, case_tac [!] s, auto split:option.splits if_splits)
+
+lemma obj2sobj_file_remains_cons:
+  assumes vs: "valid (e#s)" and exf: "f \<in> current_files s"
+  and SF: "obj2sobj s (File f) = SFile sf srf" 
+  and notdeled: "\<not> deleted (File f) (e#s)"
+  shows "obj2sobj (e#s) (File f) = SFile sf srf"
+proof-
+  from vs have os:"os_grant s e" and vs': "valid s" 
+    by (auto dest:valid_cons valid_os)
+  from notdeled exf have exf': "f \<in> current_files (e#s)" by (case_tac e, auto)
+  have "etype_of_file (e # s) f = etype_of_file s f"
+    using os vs vs' exf exf' 
+    apply (case_tac e, auto simp:etype_of_file_def split:option.splits) 
+    by (auto dest:ancient_file_in_current intro!:etype_aux_prop)
+  moreover have "source_dir (e # s) f = source_dir s f"
+    using os vs vs' exf exf'
+    by (case_tac e, auto simp:source_dir_simps dest:source_dir_prop)
+  ultimately show ?thesis using vs SF notdeled 
+    by (auto split:if_splits option.splits dest:not_deleted_cons_D)
+qed
+
+lemma obj2sobj_file_remains_cons':
+  "\<lbrakk>valid (e#s); f \<in> current_files s; obj2sobj s (File f) = SFile sf srf; no_del_event (e#s)\<rbrakk>
+   \<Longrightarrow> obj2sobj (e#s) (File f) = SFile sf srf"
+by (auto intro!:obj2sobj_file_remains_cons nodel_imp_un_deleted
+       simp del:obj2sobj.simps)
+
+lemma obj2sobj_file_remains':
+  "\<lbrakk>obj2sobj s (File f) = sobj; sobj \<noteq> Unknown; valid (e#s); f \<in> current_files s;
+    no_del_event (e#s)\<rbrakk> \<Longrightarrow> obj2sobj (e#s) (File f) = sobj"
+apply (frule obj2sobj_file', simp, (erule exE)+)
+apply (simp del:obj2sobj.simps)
+apply (erule obj2sobj_file_remains_cons', simp+)
+done
+
+lemma obj2sobj_file_remains_app:
+  "\<lbrakk>obj2sobj s (File f) = SFile sf srf; valid (s' @ s); f \<in> current_files s;
+    \<not> deleted (File f) (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (File f) = SFile sf srf"
+apply (induct s', simp)
+apply (simp only:cons_app_simp_aux)
+apply (frule valid_cons, frule not_deleted_cons_D)
+apply (drule_tac s = "s'@s" in obj2sobj_file_remains_cons, auto simp del:obj2sobj.simps)
+apply (drule_tac obj = "File f" in not_deleted_imp_exists', simp+)
+done
+
+lemma obj2sobj_file_remains_app':
+  "\<lbrakk>obj2sobj s (File f) = SFile sf srf; valid (s' @ s); f \<in> current_files s;
+    no_del_event (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (File f) = SFile sf srf"
+by (auto intro!:obj2sobj_file_remains_app nodel_imp_un_deleted
+       simp del:obj2sobj.simps)
+
+lemma obj2sobj_file_remains'':
+  "\<lbrakk>obj2sobj s (File f) = sobj; sobj \<noteq> Unknown; valid (s'@s); f \<in> current_files s;
+    no_del_event (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (File f) = sobj"
+apply (frule obj2sobj_file', simp, (erule exE)+)
+apply (simp del:obj2sobj.simps)
+apply (erule obj2sobj_file_remains_app', simp+)
+done
+
+lemma obj2sobj_file_remains''':
+  "\<lbrakk>obj2sobj s (File f) = sobj; sobj \<noteq> Unknown; valid (s'@s); f \<in> current_files s;
+    \<not>deleted (File f) (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (File f) = sobj"
+apply (frule obj2sobj_file', simp, (erule exE)+)
+apply (simp del:obj2sobj.simps)
+by (erule obj2sobj_file_remains_app, simp+)
+
+lemma obj2sobj_ipc_remains_cons:
+  "\<lbrakk>valid (e#s); i \<in> current_ipcs s; obj2sobj s (IPC i) = SIPC si sri; \<not> deleted (IPC i) (e#s)\<rbrakk>
+  \<Longrightarrow> obj2sobj (e#s) (IPC i) = SIPC si sri"
+apply (frule valid_cons, frule valid_os, case_tac e)
+by (auto simp:ni_init_deled ni_notin_curi split:option.splits
+        dest!:current_proc_has_role')
+
+lemma obj2sobj_ipc_remains_cons':
+  "\<lbrakk>valid (e#s); i \<in> current_ipcs s; obj2sobj s (IPC i) = SIPC si sri; no_del_event (e#s)\<rbrakk>
+  \<Longrightarrow> obj2sobj (e#s) (IPC i) = SIPC si sri"
+by (auto intro!:obj2sobj_ipc_remains_cons nodel_imp_un_deleted
+       simp del:obj2sobj.simps)
+
+lemma obj2sobj_ipc_remains':
+  "\<lbrakk>obj2sobj s (IPC i) = sobj; sobj \<noteq> Unknown; valid (e#s); i \<in> current_ipcs s; 
+    no_del_event (e#s)\<rbrakk> \<Longrightarrow> obj2sobj (e#s) (IPC i) = sobj"
+apply (frule obj2sobj_ipc', simp, (erule exE)+)
+apply (simp del:obj2sobj.simps)
+apply (erule obj2sobj_ipc_remains_cons', simp+)
+done
+
+lemma obj2sobj_ipc_remains_app:
+  "\<lbrakk>obj2sobj s (IPC i) = SIPC si sri; valid (s'@s); i \<in> current_ipcs s; \<not> deleted (IPC i) (s'@s)\<rbrakk>
+  \<Longrightarrow> obj2sobj (s'@s) (IPC i) = SIPC si sri"
+apply (induct s', simp)
+apply (simp only:cons_app_simp_aux)
+apply (frule valid_cons, frule not_deleted_cons_D)
+apply (drule_tac s = "s'@s" in obj2sobj_ipc_remains_cons, auto simp del:obj2sobj.simps)
+apply (drule_tac obj = "IPC i" in not_deleted_imp_exists', simp+)
+done
+
+lemma obj2sobj_ipc_remains_app':
+  "\<lbrakk>obj2sobj s (IPC i) = SIPC si sri; valid (s'@s); i \<in> current_ipcs s; no_del_event (s'@s)\<rbrakk>
+  \<Longrightarrow> obj2sobj (s'@s) (IPC i) = SIPC si sri"
+by (auto intro!:obj2sobj_ipc_remains_app nodel_imp_un_deleted
+       simp del:obj2sobj.simps)
+
+lemma obj2sobj_ipc_remains'':
+  "\<lbrakk>obj2sobj s (IPC i) = sobj; sobj \<noteq> Unknown; valid (s'@s); i \<in> current_ipcs s; 
+    no_del_event (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (IPC i) = sobj"
+apply (frule obj2sobj_ipc', simp, (erule exE)+)
+apply (simp del:obj2sobj.simps)
+apply (erule obj2sobj_ipc_remains_app', simp+)
+done
+
+lemma obj2sobj_ipc_remains''':
+  "\<lbrakk>obj2sobj s (IPC i) = sobj; sobj \<noteq> Unknown; valid (s'@s); i \<in> current_ipcs s; 
+    \<not> deleted (IPC i) (s'@s)\<rbrakk> \<Longrightarrow> obj2sobj (s'@s) (IPC i) = sobj"
+apply (frule obj2sobj_ipc', simp, (erule exE)+)
+apply (simp del:obj2sobj.simps)
+apply (erule obj2sobj_ipc_remains_app, simp+)
+done
+
+end
+
+context tainting_s_sound begin
+
+lemma cp2sproc_clone':
+  "valid (Clone p p' # s) \<Longrightarrow> cp2sproc (Clone p p' # s) = (cp2sproc s) (p' := cp2sproc s p)"
+by (drule cp2sproc_clone, auto split:option.splits simp:clone_type_unchange clone_type_aux_def)
+
+lemmas cp2sproc_simps' =  cp2sproc_exec cp2sproc_chown cp2sproc_crole cp2sproc_clone' cp2sproc_other
+
+lemma clone_sobj_keeps_same:
+  "valid (Clone p p' # s) \<Longrightarrow> obj2sobj (Clone p p' # s) (Proc p') = obj2sobj s (Proc p)"
+apply (frule valid_cons, frule valid_os, clarsimp)
+apply (auto split:option.splits t_rc_proc_type.splits
+             dest:current_proc_has_role current_proc_has_forcedrole 
+                  current_proc_has_type current_proc_has_owner default_process_create_type_valid
+             simp:pct_def clone_type_unchange)
+done
+
+end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/os_rc.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,376 @@
+theory os_rc 
+imports Main rc_theory 
+begin
+
+(****** below context is for lemmas of OS and RC ******)
+context rc_basic begin
+
+inductive_cases vs_step': "valid (e # \<tau>)" 
+
+lemma valid_cons:
+  "valid (e # \<tau>) \<Longrightarrow> valid \<tau>"
+by (drule vs_step', auto)
+
+lemma valid_os: 
+  "valid (e # \<tau>) \<Longrightarrow> os_grant \<tau> e"
+by (drule vs_step', auto)
+
+lemma valid_rc:
+  "valid (e # \<tau>) \<Longrightarrow> rc_grant \<tau> e"
+by (drule vs_step', auto)
+
+lemma vs_history:
+  "\<lbrakk>s \<preceq> s'; valid s'\<rbrakk> \<Longrightarrow> valid s"
+apply (induct s', simp add:no_junior_def)
+apply (case_tac "s = a # s'", simp)
+apply (drule no_junior_noteq, simp)
+by (drule valid_cons)
+
+lemma parent_file_in_current: 
+  "\<lbrakk>parent f = Some pf; f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> pf \<in> current_files s"
+apply (induct s)
+apply (simp add:parent_file_in_init)
+apply (frule valid_cons, frule valid_rc, frule valid_os)
+apply (case_tac a, auto split:option.splits)
+apply (case_tac f, simp+)
+done
+
+lemma parent_file_in_current': 
+  "\<lbrakk>fn # pf \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> pf \<in> current_files s"
+by (auto intro!:parent_file_in_current[where pf = pf])
+
+lemma parent_file_in_init': 
+  "fn # pf \<in> init_files \<Longrightarrow> pf \<in> init_files"
+by (auto intro!:parent_file_in_init[where pf = pf])
+
+lemma ancient_file_in_current:
+  "\<lbrakk>f \<in> current_files s; valid s; af \<preceq> f\<rbrakk> \<Longrightarrow> af \<in> current_files s"
+apply (induct f)
+apply (simp add:no_junior_def)
+apply (case_tac "af = a # f", simp)
+apply (drule no_junior_noteq, simp)
+apply (drule parent_file_in_current', simp+)
+done
+
+lemma cannot_del_root:
+  "\<lbrakk>valid (DeleteFile p [] # s); f \<noteq> []; f \<in> current_files s\<rbrakk> \<Longrightarrow> False"
+apply (frule valid_cons, frule valid_os)
+apply (case_tac f rule:rev_cases, simp)
+apply (drule_tac af = "[y]" in ancient_file_in_current, simp+)
+done
+
+lemma init_file_initialrole_imp_some: "\<exists> r. init_file_initialrole f = Some r"
+by (case_tac f, auto split:option.splits)
+
+lemma file_has_initialrole: "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> (\<exists> r. initialrole s f = Some r)"
+apply (induct s arbitrary:f) 
+apply (simp, rule init_file_initialrole_imp_some)
+apply (frule valid_cons, frule valid_os, case_tac a)
+apply (auto split:if_splits option.splits)
+done
+
+lemma file_has_initialrole':
+  "\<lbrakk>initialrole s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
+by (rule notI, auto dest:file_has_initialrole)
+
+lemma file_has_effinitialrole: 
+  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> r. effinitialrole s f = Some r"
+apply (induct f)
+apply (auto simp:effinitialrole_def dest:file_has_initialrole parent_file_in_current')
+done
+
+lemma file_has_effinitialrole':
+  "\<lbrakk>effinitialrole s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
+by (rule notI, auto dest:file_has_effinitialrole)
+
+lemma init_file_forcedrole_imp_some: "\<exists> r. init_file_forcedrole f = Some r"
+by (case_tac f, auto split:option.splits)
+
+lemma file_has_forcedrole: "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> (\<exists> r. forcedrole s f = Some r)"
+apply (induct s arbitrary:f) 
+apply (simp add:init_file_forcedrole_imp_some)
+apply (frule valid_cons, frule valid_os, case_tac a, auto)
+done
+
+lemma file_has_forcedrole':
+  "\<lbrakk>forcedrole s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
+by (rule notI, auto dest:file_has_forcedrole)
+
+lemma file_has_effforcedrole: 
+  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> r. effforcedrole s f = Some r"
+apply (induct f)
+apply (auto simp:effforcedrole_def dest:file_has_forcedrole parent_file_in_current')
+done
+
+lemma file_has_effforcedrole':
+  "\<lbrakk>effforcedrole s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
+by (rule notI, auto dest:file_has_effforcedrole)
+
+lemma current_proc_has_forcedrole:
+  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> r. proc_forcedrole s p = Some r"
+apply (induct s arbitrary:p) using init_proc_has_frole
+apply (simp add:bidirect_in_init_def)
+apply (frule valid_cons, frule valid_os, case_tac a)
+apply (auto split:if_splits option.splits intro:file_has_effforcedrole)
+done
+
+lemma current_proc_has_forcedrole':
+  "\<lbrakk>proc_forcedrole s p = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
+by (rule notI, auto dest:current_proc_has_forcedrole)
+
+lemma current_proc_has_owner: "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> u. owner s p = Some u"
+apply (induct s arbitrary:p) using init_proc_has_owner
+apply (simp add:bidirect_in_init_def)
+apply (frule valid_cons, frule valid_os, case_tac a, auto)
+done
+
+lemma current_proc_has_owner':
+  "\<lbrakk>owner s p = None; valid s\<rbrakk> \<Longrightarrow> p \<notin> current_procs s"
+by (rule notI, auto dest:current_proc_has_owner)
+
+(*
+lemma effinitial_normal_intro: 
+  "\<lbrakk>f \<in> current_files \<tau>; valid \<tau>; effinitialrole \<tau> f \<noteq> Some UseForcedRole\<rbrakk> \<Longrightarrow> \<exists>nr. effinitialrole \<tau> f = Some (NormalRole nr)"
+apply (drule file_has_effinitialrole, simp)
+apply (erule exE, frule effinitialrole_valid, simp)
+done
+
+lemma effforced_normal_intro: 
+  "\<lbrakk>f \<in> current_files \<tau>; valid \<tau>; effforcedrole \<tau> f \<noteq> Some InheritUserRole; effforcedrole \<tau> f \<noteq> Some InheritProcessRole; effforcedrole \<tau> f \<noteq> Some InheritUpMixed\<rbrakk> 
+  \<Longrightarrow> \<exists>nr. effforcedrole \<tau> f = Some (NormalRole nr)"
+apply (drule file_has_effforcedrole, simp)
+apply (erule exE, frule effforcedrole_valid, simp)
+done
+*)
+
+lemma owner_in_users: "\<lbrakk>owner s p = Some u; valid s\<rbrakk> \<Longrightarrow> u \<in> init_users"
+apply (induct s arbitrary:p) defer
+apply (frule valid_cons, frule valid_os, case_tac a)
+apply (auto split:if_splits option.splits intro!:init_owner_valid)
+done
+
+lemma user_has_normalrole: 
+  "u \<in> init_users \<Longrightarrow> \<exists> nr. defrole u = Some nr" using init_user_has_role
+by (auto simp:bidirect_in_init_def)
+
+lemma user_has_normalrole':
+  "defrole u = None \<Longrightarrow> u \<notin> init_users"
+by (rule notI, auto dest:user_has_normalrole)
+
+lemma current_proc_has_role: 
+  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> nr. currentrole s p = Some nr"
+apply (induct s arbitrary:p) using init_proc_has_role
+apply (simp add:bidirect_in_init_def)
+apply (frule valid_cons, frule valid_os, case_tac a)
+apply (auto simp:map_comp_def split:if_splits option.splits t_role.splits
+           dest!:current_proc_has_owner' user_has_normalrole' current_proc_has_forcedrole'
+                 file_has_forcedrole' file_has_effforcedrole' 
+                 file_has_initialrole' file_has_effinitialrole'
+           intro:user_has_normalrole
+            dest:owner_in_users effinitialrole_valid effforcedrole_valid proc_forcedrole_valid)
+done
+
+lemma current_file_has_type: 
+  "\<lbrakk>f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> \<exists> t. type_of_file s f = Some t"
+apply (induct s)
+apply (simp split:option.splits)
+apply (frule valid_cons, frule valid_os, case_tac a)
+apply (auto split:option.splits intro:current_proc_has_role)
+done
+
+lemma init_file_has_etype: "f \<in> init_files \<Longrightarrow> \<exists> nt. etype_aux init_file_type_aux f = Some nt"
+apply (induct f) defer
+apply (frule parent_file_in_init') 
+apply (auto split:option.splits t_rc_file_type.splits)
+done
+
+lemma current_file_has_etype[rule_format]:
+  "f \<in> current_files s \<Longrightarrow> valid s \<longrightarrow> (\<exists> nt. etype_of_file s f = Some nt)"
+apply (induct f)
+apply (auto simp:etype_of_file_def dest:current_file_has_type parent_file_in_current'
+           split:option.splits t_rc_file_type.splits)
+done
+
+lemma current_file_has_etype':
+  "\<lbrakk>etype_of_file s f = None; valid s\<rbrakk> \<Longrightarrow> f \<notin> current_files s"
+by (rule notI, auto dest:current_file_has_etype)
+
+(*** etype_of_file simpset ***)
+
+lemma etype_aux_prop:
+  "\<forall> x. x \<preceq> f \<longrightarrow> func' x = func x \<Longrightarrow> etype_aux func f = etype_aux func' f"
+apply (induct f)
+by (auto split:t_rc_file_type.splits option.splits)
+
+lemma etype_aux_prop1:
+  "func' = func ((a#f) := b) \<Longrightarrow> etype_aux func f = etype_aux func' f"
+by (rule etype_aux_prop, auto simp:no_junior_def)
+
+lemma etype_aux_prop1':
+  "etype_aux func f = x \<Longrightarrow> etype_aux (func ((a#f) := b)) f = x"
+apply (subgoal_tac "etype_aux func f = etype_aux (func ((a#f) := b)) f")
+apply (simp, rule etype_aux_prop1, simp)
+done
+
+lemma etype_aux_prop2:
+  "\<lbrakk>f \<in> current_files s; f' \<notin> current_files s; valid s\<rbrakk> \<Longrightarrow>
+  etype_aux (func (f' := b)) f = etype_aux func f"
+apply (rule etype_aux_prop)
+by (auto dest:ancient_file_in_current)
+
+lemma etype_aux_prop3:
+  "parent f = Some pf 
+   \<Longrightarrow> etype_aux (func (f := Some InheritParent_file_type)) f = etype_aux func pf"
+apply (case_tac f, simp+)
+by (rule etype_aux_prop, simp add:no_junior_def)
+
+lemma etype_aux_prop4:
+  "etype_aux (func (f := Some (NormalFile_type t))) f = Some t"
+by (case_tac f, auto)
+
+lemma etype_of_file_delete:
+  "\<lbrakk>valid (DeleteFile p f # s); f' \<in> current_files s\<rbrakk>
+   \<Longrightarrow> etype_of_file (DeleteFile p f # s) f' = etype_of_file s f'"
+apply (frule valid_cons, frule valid_os)
+apply (simp add:etype_of_file_def)
+done
+
+lemma current_proc_has_type: 
+  "\<lbrakk>p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<exists> nt. type_of_process s p = Some nt"
+apply (induct s arbitrary:p) using init_proc_has_type
+apply (simp add:bidirect_in_init_def)
+
+apply (frule valid_cons, frule valid_os, case_tac a)
+
+apply (subgoal_tac "nat1 \<in> current_procs (a # s)") prefer 2 apply simp
+apply (drule_tac p = nat1 in current_proc_has_role, simp, erule exE)
+
+apply (auto simp:pct_def pot_def pet_def dest:current_proc_has_role 
+           split:option.splits t_rc_proc_type.splits 
+           dest!:default_process_create_type_valid default_process_chown_type_valid 
+                 default_process_execute_type_valid)
+done
+
+lemma current_ipc_has_type: 
+  "\<lbrakk>i \<in> current_ipcs s; valid s\<rbrakk> \<Longrightarrow> \<exists> nt. type_of_ipc s i = Some nt"
+apply (induct s) using init_ipc_has_type
+apply (simp add:bidirect_in_init_def)
+
+apply (frule valid_cons, frule valid_os, case_tac a)
+apply (auto dest:current_proc_has_role)
+done
+
+(*** finite current_*  ***)
+
+lemma finite_cf: "finite (current_files s)"
+apply (induct s) defer apply (case_tac a)
+using init_finite by auto
+
+lemma finite_cp: "finite (current_procs s)"
+apply (induct s) defer apply (case_tac a)
+using init_finite by auto
+
+lemma finite_ci: "finite (current_ipcs s)"
+apply (induct s) defer apply (case_tac a)
+using init_finite by auto
+
+(*** properties of new-proc new-ipc ... ***)
+
+lemma nn_notin_aux: "finite s \<Longrightarrow> \<forall> a \<in> s. Max s \<ge> a "
+apply (erule finite.induct, simp)
+apply (rule ballI)
+apply (case_tac "aa = a", simp+)
+done
+
+lemma nn_notin: "finite s \<Longrightarrow> next_nat s \<notin> s"
+apply (drule nn_notin_aux)
+apply (simp add:next_nat_def)
+by (auto)
+
+lemma np_notin_curp: "new_proc \<tau> \<notin> current_procs \<tau>" using finite_cp
+by (simp add:new_proc_def nn_notin)
+
+lemma np_notin_curp': "new_proc \<tau> \<in> current_procs \<tau> \<Longrightarrow> False"
+by (simp add:np_notin_curp)
+
+lemma ni_notin_curi: "new_ipc \<tau> \<notin> current_ipcs \<tau>" using finite_ci
+by (simp add:new_ipc_def nn_notin)
+
+lemma ni_notin_curi': "new_ipc \<tau> \<in> current_ipcs \<tau> \<Longrightarrow> False"
+by (simp add:ni_notin_curi)
+
+end
+
+context tainting_s_complete begin
+
+lemma init_notin_curf_deleted:
+  "\<lbrakk>f \<notin> current_files s; f \<in> init_files\<rbrakk> \<Longrightarrow> deleted (File f) s"
+by (induct s, simp, case_tac a, auto)
+
+lemma init_notin_curi_deleted:
+  "\<lbrakk>i \<notin> current_ipcs s; i \<in> init_ipcs\<rbrakk> \<Longrightarrow> deleted (IPC i) s"
+by (induct s, simp, case_tac a, auto)
+
+lemma init_notin_curp_deleted:
+  "\<lbrakk>p \<notin> current_procs s; p \<in> init_processes\<rbrakk> \<Longrightarrow> deleted (Proc p) s"
+by (induct s, simp, case_tac a, auto)
+  
+lemma ni_init_deled: "new_ipc s \<in> init_ipcs \<Longrightarrow> deleted (IPC (new_ipc s)) s"
+using ni_notin_curi[where \<tau> = s]
+by (drule_tac init_notin_curi_deleted, simp+)
+
+lemma np_init_deled: "new_proc s \<in> init_processes \<Longrightarrow> deleted (Proc (new_proc s)) s"
+using np_notin_curp[where \<tau> = s]
+by (drule_tac init_notin_curp_deleted, simp+)
+
+lemma source_dir_in_init: "source_dir s f = Some sd \<Longrightarrow> sd \<in> init_files"
+by (induct f, auto split:if_splits)
+
+lemma source_proc_in_init:
+  "\<lbrakk>source_proc s p = Some p'; p \<in> current_procs s; valid s\<rbrakk> \<Longrightarrow> p' \<in> init_processes"
+apply (induct s arbitrary:p, simp split:if_splits)
+apply (frule valid_os, frule valid_cons, case_tac a)
+by (auto simp:np_notin_curp split:if_splits)
+
+end
+
+context tainting_s_sound begin
+
+lemma len_fname_all: "length (fname_all_a len) = len"
+by (induct len, auto simp:fname_all_a.simps)
+
+lemma ncf_notin_curf: "new_childf f s \<notin> current_files s"
+apply (simp add:new_childf_def next_fname_def all_fname_under_dir_def)
+apply (rule notI)
+apply (subgoal_tac "(CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files s}))) \<in> {fn. fn # f \<in> current_files s}")
+defer apply simp
+apply (subgoal_tac "length (CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files s}))) \<in> fname_length_set {fn. fn # f \<in> current_files s}")
+defer apply (auto simp:fname_length_set_def image_def)[1]
+apply (subgoal_tac "finite (fname_length_set {fn. fn # f \<in> current_files s})")  
+defer
+apply (simp add:fname_length_set_def) 
+apply (rule finite_imageI) using finite_cf[where s = s]
+apply (drule_tac h = "\<lambda> f'. case f' of [] \<Rightarrow> '''' | fn # pf' \<Rightarrow> if (pf' = f) then fn else ''''" in finite_imageI)
+apply (rule_tac B = "(list_case [] (\<lambda>fn pf'. if pf' = f then fn else []) ` current_files s)" in finite_subset)
+unfolding image_def
+apply(auto)[1]
+apply (rule_tac x = "x # f" in bexI, simp+)
+apply (drule_tac s = "(fname_length_set {fn. fn # f \<in> current_files s})" in nn_notin_aux)
+apply (erule_tac x = "length (CHR ''a'' # fname_all_a (Max (fname_length_set {fn. fn # f \<in> current_files s})))" in ballE)
+apply (simp add:len_fname_all, simp)
+done
+
+lemma ncf_parent: "parent (new_childf f \<tau>) = Some f"
+by (simp add:new_childf_def)
+
+lemma clone_event_no_limit: 
+  "\<lbrakk>p \<in> current_procs \<tau>; valid \<tau>\<rbrakk> \<Longrightarrow> valid (Clone p (new_proc \<tau>) # \<tau>)"
+apply (rule vs_step)
+apply (auto intro:clone_no_limit split:option.splits
+            dest:current_proc_has_role current_proc_has_type)
+done 
+
+
+end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/rc_theory.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,990 @@
+theory rc_theory
+imports Main my_list_prefix
+begin
+
+
+(****************** rc policy (role, type, ... ) type definitions *****************)
+(*** normal: user-defined values for policy ***)
+(*** others: RC special values for RC internal control ***)
+ 
+datatype t_client = 
+  Client1
+| Client2
+
+datatype t_normal_role =  
+  WebServer 
+| WS     "t_client" 
+| UpLoad "t_client" 
+| CGI     "t_client"
+
+datatype t_role = 
+  InheritParentRole 
+| UseForcedRole 
+| InheritUpMixed 
+| InheritUserRole 
+| InheritProcessRole 
+| NormalRole "t_normal_role" 
+
+datatype t_normal_file_type = 
+  Root_file_type   (* special value, as 0 for root-file in the NordSec paper*)
+| WebServerLog_file
+| WebData_file      "t_client" 
+| CGI_P_file        "t_client" 
+| PrivateD_file     "t_client" 
+| Executable_file   (* basic protection of RC *)
+
+datatype t_rc_file_type = 
+  InheritParent_file_type          
+| NormalFile_type t_normal_file_type
+
+datatype t_normal_proc_type = 
+  WebServer_proc
+| CGI_P_proc "t_client"
+
+datatype t_rc_proc_type =
+  InheritParent_proc_type
+| NormalProc_type t_normal_proc_type
+| UseNewRoleType  
+
+datatype t_normal_ipc_type = 
+  WebIPC
+
+datatype t_normal_rc_type =
+  File_type t_normal_file_type
+| Proc_type t_normal_proc_type
+| IPC_type  t_normal_ipc_type
+
+
+
+(******* Operating System type definitions ********)
+
+type_synonym t_process = nat
+
+type_synonym t_user = nat
+
+type_synonym t_fname = string
+
+type_synonym t_file = "t_fname list" 
+
+type_synonym t_ipc = nat
+
+
+
+(****** Access Control related type definitions *********)
+
+datatype t_event =
+  ChangeOwner  "t_process" "t_user" 
+| Clone        "t_process" "t_process"
+| Execute      "t_process" "t_file" 
+| CreateFile   "t_process" "t_file" 
+| CreateIPC    "t_process" "t_ipc" 
+| ChangeRole   "t_process" "t_normal_role"    (* A process should change to normal role, which not specified in paper! *) 
+
+(* below are events added for tainting modelling *)
+| ReadFile     "t_process" "t_file" 
+| WriteFile    "t_process" "t_file"     
+| Send         "t_process" "t_ipc"
+| Recv         "t_process" "t_ipc"
+
+| Kill         "t_process" "t_process"
+| DeleteFile   "t_process" "t_file"
+| DeleteIPC    "t_process" "t_ipc"
+                                  
+type_synonym t_state = "t_event list"
+             
+datatype t_access_type =        (*changed by wch, original: "types access_type = nat" *)
+  READ    
+| WRITE   
+| EXECUTE 
+| CHANGE_OWNER
+| CREATE
+| SEND
+| RECEIVE
+
+| DELETE
+
+
+(****** Some global functions' definition *****)
+
+fun parent :: "'a list \<Rightarrow> ('a list) option"
+where
+   "parent []     = None"
+ | "parent (n#ns) = Some ns"
+
+definition some_in_set :: "'a set \<Rightarrow> 'a"
+where
+  "some_in_set S \<equiv> SOME x. x \<in> S"
+
+lemma nonempty_set_has_ele: "S \<noteq> {} \<Longrightarrow> \<exists> e. e \<in> S" by auto
+
+lemma some_in_set_prop: "S \<noteq> {} \<Longrightarrow> some_in_set S \<in> S"
+by (drule nonempty_set_has_ele, auto simp:some_in_set_def intro:someI)
+
+
+(*********** locale for RC+OS definitions **********)
+
+definition bidirect_in_init :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b option) \<Rightarrow> bool"
+where
+  "bidirect_in_init S f \<equiv>  (\<forall> a. a \<in> S \<longrightarrow> (\<exists> b. f a = Some b)) \<and> 
+                           (\<forall> a b. f a = Some b \<longrightarrow> a \<in> S)" 
+
+locale init =
+  fixes 
+      init_files :: "t_file set"
+  and init_file_type :: "t_file \<rightharpoonup> t_normal_file_type"
+  and init_initialrole :: "t_file \<rightharpoonup> t_role"
+  and init_forcedrole :: "t_file \<rightharpoonup> t_role"
+  
+  and init_processes :: "t_process set "
+  and init_process_type :: "t_process \<rightharpoonup> t_normal_proc_type"
+  and init_currentrole :: "t_process \<rightharpoonup> t_normal_role"
+  and init_proc_forcedrole:: "t_process \<rightharpoonup> t_role"
+
+  and init_ipcs :: "t_ipc set"
+  and init_ipc_type:: "t_ipc \<rightharpoonup> t_normal_ipc_type"
+
+  and init_users :: "t_user set"
+  and init_owner :: "t_process \<rightharpoonup> t_user"
+  and defrole :: "t_user \<rightharpoonup> t_normal_role" (* defrole should return normalrole, which is not
+specified in NordSec paper *) 
+
+  and default_fd_create_type :: "t_normal_role \<Rightarrow> t_rc_file_type" (* this func should only
+return type for normal role, for RC special role, error ! NordSec paper*)
+  and default_ipc_create_type :: "t_normal_role \<Rightarrow> t_normal_ipc_type" (* like above, NordSec paper
+does not specify the domain is normal roles *)
+  and default_process_create_type :: "t_normal_role \<Rightarrow> t_rc_proc_type"
+  and default_process_execute_type :: "t_normal_role \<Rightarrow> t_rc_proc_type"
+  and default_process_chown_type :: "t_normal_role \<Rightarrow> t_rc_proc_type"
+
+  and comproles :: "t_normal_role \<Rightarrow> t_normal_role set" (* NordSec paper do not specify all roles
+here are normal *)
+
+  and compatible :: "(t_normal_role \<times> t_normal_rc_type \<times> t_access_type) set"  (* NordSec paper do not specify all roles and all types here are normal ! *)
+
+  assumes
+      parent_file_in_init: "\<And> f pf. \<lbrakk>parent f = Some pf; f \<in> init_files\<rbrakk> \<Longrightarrow> pf \<in> init_files"
+  and root_in_filesystem:  "[] \<in> init_files"
+  and init_irole_has_file: "\<And> f r. init_initialrole f = Some r \<Longrightarrow> f \<in> init_files"
+  and init_frole_has_file: "\<And> f r. init_forcedrole f = Some r \<Longrightarrow> f \<in> init_files"
+  and init_ftype_has_file: "\<And> f t. init_file_type f = Some t \<Longrightarrow> f \<in> init_files"
+  and
+      init_proc_has_role:  "bidirect_in_init init_processes init_currentrole"
+  and init_proc_has_frole: "bidirect_in_init init_processes init_proc_forcedrole"
+  and init_proc_has_type:  "bidirect_in_init init_processes init_process_type"
+  and
+      init_ipc_has_type:   "bidirect_in_init init_ipcs init_ipc_type"
+  and
+      init_user_has_role:  "bidirect_in_init init_users defrole"
+  and init_proc_has_owner: "bidirect_in_init init_processes init_owner"
+  and init_owner_valid:    "\<And> p u. init_owner p = Some u \<Longrightarrow> u \<in> init_users"
+  and
+      init_finite: "finite init_files \<and> finite init_processes \<and> finite init_ipcs \<and> finite init_users"
+begin 
+
+(***** Operating System Listeners *****)
+
+fun current_files :: "t_state \<Rightarrow> t_file set"
+where
+  "current_files [] = init_files"
+| "current_files (CreateFile p f # s) = insert f (current_files s)"
+| "current_files (DeleteFile p f # s) = current_files s - {f}"
+| "current_files (_ # s) = current_files s"
+
+fun current_procs :: "t_state \<Rightarrow> t_process set"
+where
+  "current_procs [] = init_processes"
+| "current_procs (Clone p p' # s) = insert p' (current_procs s)"
+| "current_procs (Kill p p' # s) = current_procs s - {p'} "
+| "current_procs (_ # s) = current_procs s"
+
+fun current_ipcs :: "t_state \<Rightarrow> t_ipc set"
+where
+  "current_ipcs [] = init_ipcs" 
+| "current_ipcs (CreateIPC p i # s) = insert i (current_ipcs s)"
+| "current_ipcs (DeleteIPC p i # s) = current_ipcs s - {i}"
+| "current_ipcs (_ # s) = current_ipcs s"
+
+fun owner :: "t_state \<Rightarrow> t_process \<rightharpoonup> t_user"
+where
+  "owner [] = init_owner"
+| "owner (Clone p p' # \<tau>) = (owner \<tau>) (p' := owner \<tau> p)"
+| "owner (ChangeOwner p u # \<tau>) = (owner \<tau>) (p := Some u)"
+| "owner (_ # \<tau>) = owner \<tau>"
+
+(***** functions for rc internal *****)
+
+(*** Roles Functions ***)
+
+(* comments:
+We have fix init_initialrole already in locale, why need this function?
+ Cause, users may be lazy to specify every file in the initial state to some InheritParent as
+initialrole, they can just specify all the important files with special initial role, leaving
+all other None, it is init_file_initialrole's job to fill default value(InheritParent) for 
+other files. Accutally, this has the point of "Functional default settings" in the 2 section of
+the NordSec paper. 
+init_file_forcedrole, and init_file_type_aux are of the same meaning. 
+*)
+fun init_file_initialrole :: "t_file \<rightharpoonup> t_role"
+where
+  "init_file_initialrole [] = (case (init_initialrole []) of
+                            None   \<Rightarrow> Some UseForcedRole
+                          | Some r \<Rightarrow> Some r)"
+| "init_file_initialrole f  = (case (init_initialrole f) of
+                            None   \<Rightarrow> Some InheritParentRole
+                          | Some r \<Rightarrow> Some r)"
+
+fun initialrole :: "t_state \<Rightarrow> (t_file \<rightharpoonup> t_role)"  
+where
+  "initialrole [] = init_file_initialrole"
+| "initialrole (CreateFile p f # s) = (initialrole s) (f:= Some InheritParentRole)" 
+| "initialrole (_ # s) = initialrole s"
+
+fun erole_functor :: "(t_file \<rightharpoonup> t_role) \<Rightarrow> t_role \<Rightarrow> (t_file \<rightharpoonup> t_role)"
+where
+  "erole_functor rfunc r [] = (
+     if (rfunc [] = Some InheritParentRole) 
+     then Some r 
+     else rfunc [] )"                             
+| "erole_functor rfunc r (n#ns) = (
+     if (rfunc (n#ns) = Some InheritParentRole) 
+     then erole_functor rfunc r ns 
+     else rfunc (n#ns) )"
+
+definition effinitialrole :: "t_state \<Rightarrow> (t_file \<rightharpoonup> t_role)"
+where
+  "effinitialrole s \<equiv> erole_functor (initialrole s) UseForcedRole"
+
+fun init_file_forcedrole :: "t_file \<rightharpoonup> t_role"
+where
+  "init_file_forcedrole [] = ( case (init_forcedrole []) of
+                                 None   \<Rightarrow> Some InheritUpMixed
+                               | Some r \<Rightarrow> Some r )"
+| "init_file_forcedrole f  = ( case (init_forcedrole f) of 
+                                 None   \<Rightarrow> Some InheritParentRole
+                              |  Some r \<Rightarrow> Some r )"  
+
+fun forcedrole :: "t_state \<Rightarrow> (t_file \<rightharpoonup> t_role)"  
+where
+  "forcedrole [] = init_file_forcedrole"
+| "forcedrole (CreateFile p f # s) = (forcedrole s) (f:= Some InheritParentRole)" 
+| "forcedrole (_ # s) = forcedrole s" 
+
+definition effforcedrole :: "t_state \<Rightarrow> (t_file \<rightharpoonup> t_role)" 
+where
+  "effforcedrole s \<equiv> erole_functor (forcedrole s) InheritUpMixed"
+
+fun proc_forcedrole :: "t_state \<Rightarrow> (t_process \<rightharpoonup> t_role)" (* $6.7$ *) 
+where
+  "proc_forcedrole [] = init_proc_forcedrole"
+| "proc_forcedrole (Execute p f # s) = (proc_forcedrole s) (p := effforcedrole s f)"    
+| "proc_forcedrole (Clone p p' # s) = (proc_forcedrole s) (p' := proc_forcedrole s p)" 
+| "proc_forcedrole (e # s) = proc_forcedrole s"
+
+fun currentrole :: "t_state \<Rightarrow> (t_process \<rightharpoonup> t_normal_role)" 
+where 
+  "currentrole [] = init_currentrole"
+| "currentrole (Clone p p' # \<tau>) = (currentrole \<tau>) (p' := currentrole \<tau> p)" 
+| "currentrole (Execute p f # \<tau>) = (currentrole \<tau>) (p := 
+     case (effinitialrole \<tau> f) of 
+       Some ir \<Rightarrow> (case ir of 
+                     NormalRole r \<Rightarrow> Some r
+                   | UseForcedRole \<Rightarrow> (
+                       case (effforcedrole \<tau> f) of
+                         Some fr \<Rightarrow> (case fr of
+                                       NormalRole r \<Rightarrow> Some r
+                                     | InheritUserRole \<Rightarrow> (defrole \<circ>\<^sub>m (owner \<tau>)) p
+                                     | InheritProcessRole \<Rightarrow> currentrole \<tau> p
+                                     | InheritUpMixed \<Rightarrow> currentrole \<tau> p
+                                     | _ \<Rightarrow> None )
+                       | _ \<Rightarrow> None    )
+                   | _  \<Rightarrow> None )
+     | _ \<Rightarrow> None                                   )"       
+| "currentrole (ChangeOwner p u # \<tau>) = (currentrole \<tau>) (p := 
+     case (proc_forcedrole \<tau> p) of 
+       Some fr \<Rightarrow> (case fr of
+                     NormalRole r \<Rightarrow> Some r
+                   | InheritProcessRole \<Rightarrow> currentrole \<tau> p
+                   | InheritUserRole \<Rightarrow> defrole u
+                   | InheritUpMixed \<Rightarrow> defrole u
+                   | _ \<Rightarrow> None)
+     | _ \<Rightarrow> None                                       )"  
+| "currentrole (ChangeRole p r # \<tau>) = (currentrole \<tau>) (p := Some r)"
+| "currentrole (_ # \<tau>) = currentrole \<tau>"  
+
+(*** Types Functions ***)
+
+fun init_file_type_aux :: "t_file \<rightharpoonup> t_rc_file_type"
+where
+  "init_file_type_aux f = (if (f \<in> init_files)
+                         then (case (init_file_type f) of 
+                                 Some t \<Rightarrow> Some (NormalFile_type t)
+                               | _      \<Rightarrow> Some InheritParent_file_type)
+                         else None)"
+
+fun type_of_file :: "t_state \<Rightarrow> (t_file \<rightharpoonup> t_rc_file_type)" (* (6) *)
+where
+  "type_of_file [] = init_file_type_aux" 
+| "type_of_file (CreateFile p f # s) = (
+     case (currentrole s p) of
+       Some r \<Rightarrow> (type_of_file s) (f := Some (default_fd_create_type r))
+     | _      \<Rightarrow> (type_of_file s) (f := None))" 
+| "type_of_file (_ # s) = type_of_file s"  (* add by wch *)
+
+fun etype_aux:: "(t_file \<rightharpoonup> t_rc_file_type) \<Rightarrow> (t_file \<rightharpoonup> t_normal_file_type)"
+where
+  "etype_aux typf [] = (
+     case (typf []) of 
+       Some InheritParent_file_type \<Rightarrow> Some Root_file_type
+     | Some (NormalFile_type t) \<Rightarrow> Some t
+     | None \<Rightarrow> None )"        
+| "etype_aux typf (n#ns) = (
+     case (typf (n#ns)) of
+       Some InheritParent_file_type \<Rightarrow> etype_aux typf ns
+     | Some (NormalFile_type t) \<Rightarrow> Some t
+     | None \<Rightarrow> None )"  
+
+definition etype_of_file :: "t_state \<Rightarrow> (t_file \<rightharpoonup> t_normal_file_type)" 
+    (* etype is always normal *)
+where
+  "etype_of_file s \<equiv> etype_aux (type_of_file s)"              
+
+definition pct :: "t_state \<Rightarrow> (t_process \<rightharpoonup> t_rc_proc_type)"
+where
+  "pct s p \<equiv> (case (currentrole s p) of
+                Some r \<Rightarrow> Some (default_process_create_type r)
+              | _      \<Rightarrow> None)"
+
+definition pet :: "t_state \<Rightarrow> (t_process \<rightharpoonup> t_rc_proc_type)"
+where
+  "pet s p \<equiv> (case (currentrole s p) of
+                Some r \<Rightarrow> Some (default_process_execute_type r)
+              | _      \<Rightarrow> None)"
+
+definition pot :: "t_state \<Rightarrow> (t_process \<rightharpoonup> t_rc_proc_type)"
+where
+  "pot s p \<equiv> (case (currentrole s p) of
+                Some r \<Rightarrow> Some (default_process_chown_type r)
+              | _      \<Rightarrow> None)"
+
+fun type_of_process :: "t_state \<Rightarrow> (t_process \<rightharpoonup> t_normal_proc_type)" 
+where
+  "type_of_process [] = init_process_type"
+| "type_of_process (Clone p p' # \<tau>) = (type_of_process \<tau>) (p' := 
+     case (pct \<tau> p) of
+       Some InheritParent_proc_type \<Rightarrow> type_of_process \<tau> p
+     | Some (NormalProc_type tp)    \<Rightarrow> Some tp
+     | _                            \<Rightarrow> None               )" (*6.80*)
+| "type_of_process (Execute p f # \<tau>) = (type_of_process \<tau>) (p :=
+     case (pet \<tau> p) of
+       Some InheritParent_proc_type \<Rightarrow> type_of_process \<tau> p
+     | Some (NormalProc_type tp)    \<Rightarrow> Some tp
+     | _                            \<Rightarrow> None                )" (*6.82*)
+| "type_of_process (ChangeOwner p u # \<tau>) = (type_of_process \<tau>) (p :=
+     case (pot \<tau> p) of
+       Some InheritParent_proc_type \<Rightarrow> type_of_process \<tau> p
+     | Some UseNewRoleType          \<Rightarrow> (case (pct (ChangeOwner p u # \<tau>) p) of 
+                                          Some InheritParent_proc_type \<Rightarrow> type_of_process \<tau> p
+                                        | Some (NormalProc_type tp)    \<Rightarrow> Some tp
+                                        | _                            \<Rightarrow> None)
+     | Some (NormalProc_type tp)    \<Rightarrow> Some tp
+     | _                            \<Rightarrow> None                    )"   (* the UseNewRoleType case is refered with Nordsec paper 4 (11), and it is not right??! of just
+use "pct" *)
+| "type_of_process (_ # \<tau>) = type_of_process \<tau>" 
+
+fun type_of_ipc :: "t_state \<Rightarrow> (t_ipc \<rightharpoonup> t_normal_ipc_type)" (* (14) *)
+where
+  "type_of_ipc [] = init_ipc_type" 
+| "type_of_ipc (CreateIPC p i # s) = (type_of_ipc s) (i := 
+     case (currentrole s p) of
+       Some r \<Rightarrow> Some (default_ipc_create_type r)
+     | _      \<Rightarrow> None                                )"
+| "type_of_ipc (_ # s) = type_of_ipc s" (* add by wch *)
+
+(*** RC access control ***)
+fun rc_grant :: "t_state \<Rightarrow> t_event \<Rightarrow> bool"
+where
+  "rc_grant \<tau> (CreateFile p f) = (
+     case (parent f) of
+       Some pf \<Rightarrow> (
+         case (currentrole \<tau> p, etype_of_file \<tau> pf) of
+           (Some r, Some t) \<Rightarrow> (
+              case (default_fd_create_type r) of
+                InheritParent_file_type \<Rightarrow> (r, File_type t, WRITE) \<in> compatible 
+              | NormalFile_type t' \<Rightarrow> (r, File_type t, WRITE) \<in> compatible \<and> (r, File_type t', CREATE) \<in> compatible
+                               )
+         | _               \<Rightarrow> False )
+     | _       \<Rightarrow> False                         )"  
+| "rc_grant \<tau> (ReadFile p f) = (
+     case (currentrole \<tau> p, etype_of_file \<tau> f) of
+        (Some r, Some t) \<Rightarrow> (r, File_type t, READ) \<in> compatible 
+     | _                \<Rightarrow> False                )"
+| "rc_grant \<tau> (WriteFile p f) = (
+     case (currentrole \<tau> p, etype_of_file \<tau> f) of
+       (Some r, Some t) \<Rightarrow> (r, File_type t, WRITE) \<in> compatible 
+     | _      \<Rightarrow> False                          )"
+| "rc_grant \<tau> (Execute p f) = (
+     case (currentrole \<tau> p, etype_of_file \<tau> f) of
+       (Some r, Some t) \<Rightarrow> (r, File_type t, EXECUTE) \<in> compatible
+     | _                \<Rightarrow> False               )"     
+| "rc_grant \<tau> (ChangeOwner p u) = (
+     case (currentrole \<tau> p, type_of_process \<tau> p) of
+       (Some r, Some t) \<Rightarrow> (r, Proc_type t, CHANGE_OWNER) \<in> compatible
+     | _      \<Rightarrow> False                         )" 
+| "rc_grant \<tau> (Clone p newproc) = (
+     case (currentrole \<tau> p, type_of_process \<tau> p) of
+       (Some r, Some t) \<Rightarrow> (r, Proc_type t, CREATE) \<in> compatible
+     | _      \<Rightarrow> False                         )" (* premiss of no limit to clone is removed to locale assumptions *)
+| "rc_grant \<tau> (ChangeRole p r) = (
+     case (currentrole \<tau> p) of
+       Some cr \<Rightarrow> r \<in> comproles cr
+     | _       \<Rightarrow> False                        )"
+| "rc_grant \<tau> (Send p i) = (
+     case (currentrole \<tau> p, type_of_ipc \<tau> i) of 
+       (Some r, Some t) \<Rightarrow> (r, IPC_type t, SEND) \<in> compatible
+     | _                \<Rightarrow> False               )"
+| "rc_grant \<tau> (Recv p i) = (
+     case (currentrole \<tau> p, type_of_ipc \<tau> i) of 
+       (Some r, Some t) \<Rightarrow> (r, IPC_type t, RECEIVE) \<in> compatible
+     | _                \<Rightarrow> False               )"
+| "rc_grant \<tau> (CreateIPC p i) = (
+     case (currentrole \<tau> p) of
+       Some r \<Rightarrow> (r, IPC_type (default_ipc_create_type r), CREATE) \<in> compatible
+     | _      \<Rightarrow> False                         )"
+| "rc_grant \<tau> (DeleteFile p f) = (
+     case (currentrole \<tau> p, etype_of_file \<tau> f) of 
+       (Some r, Some t) \<Rightarrow> (r, File_type t, DELETE) \<in> compatible
+     | _                \<Rightarrow> False               )"
+| "rc_grant \<tau> (DeleteIPC p i) = (
+     case (currentrole \<tau> p, type_of_ipc \<tau> i) of 
+       (Some r, Some t) \<Rightarrow> (r, IPC_type t, DELETE) \<in> compatible
+     | _                \<Rightarrow> False               )"
+| "rc_grant \<tau> (Kill p p') = (
+     case (currentrole \<tau> p, type_of_process \<tau> p') of 
+       (Some r, Some t) \<Rightarrow> (r, Proc_type t, DELETE) \<in> compatible
+     | _                \<Rightarrow> False               )"
+
+
+(**** OS' job: checking resources existence & grant new resource ****)
+
+definition next_nat :: "nat set \<Rightarrow> nat"
+where
+  "next_nat ps = (Max ps) + 1"
+
+definition new_proc :: "t_state \<Rightarrow> t_process"
+where 
+  "new_proc \<tau> = next_nat (current_procs \<tau>)"
+
+definition new_ipc :: "t_state \<Rightarrow> t_ipc"
+where
+  "new_ipc \<tau> = next_nat (current_ipcs \<tau>)"
+
+(* new file pathname is user-defined, so os just check if its unexistence *) 
+fun os_grant :: "t_state \<Rightarrow> t_event \<Rightarrow> bool"
+where
+  "os_grant \<tau> (Execute p f)    = (p \<in> current_procs \<tau> \<and> f \<in> current_files \<tau>)"
+| "os_grant \<tau> (CreateFile p f) = (p \<in> current_procs \<tau> \<and> f \<notin> current_files \<tau> \<and> (\<exists> pf. (parent f = Some pf) \<and> pf \<in> current_files \<tau>))"  (*cannot create disk, ?? or f = [] ??*)
+| "os_grant \<tau> (ChangeRole p r) = (p \<in> current_procs \<tau>)"
+| "os_grant \<tau> (ReadFile p f)   = (p \<in> current_procs \<tau> \<and> f \<in> current_files \<tau>)"
+| "os_grant \<tau> (WriteFile p f)  = (p \<in> current_procs \<tau> \<and> f \<in> current_files \<tau>)"
+| "os_grant \<tau> (ChangeOwner p u)= (p \<in> current_procs \<tau> \<and> u \<in> init_users)"
+| "os_grant \<tau> (CreateIPC p i)  = (p \<in> current_procs \<tau> \<and> i = new_ipc \<tau>)"
+| "os_grant \<tau> (Send p i)       = (p \<in> current_procs \<tau> \<and> i \<in> current_ipcs \<tau>)" 
+| "os_grant \<tau> (Recv p i)       = (p \<in> current_procs \<tau> \<and> i \<in> current_ipcs \<tau>)"
+| "os_grant \<tau> (Clone p p')     = (p \<in> current_procs \<tau> \<and> p' = new_proc \<tau>)"
+| "os_grant \<tau> (Kill p p')      = (p \<in> current_procs \<tau> \<and> p' \<in> current_procs \<tau>)" 
+| "os_grant \<tau> (DeleteFile p f) = (p \<in> current_procs \<tau> \<and> f \<in> current_files \<tau> \<and> \<not> (\<exists>fn. fn # f \<in> current_files \<tau>))"
+| "os_grant \<tau> (DeleteIPC p i)  = (p \<in> current_procs \<tau> \<and> i \<in> current_ipcs \<tau>)" 
+
+(**** system valid state ****)
+
+inductive valid :: "t_state \<Rightarrow> bool"
+where
+  vs_nil:  "valid []"
+| vs_step: "\<lbrakk>valid s; os_grant s e; rc_grant s e\<rbrakk> \<Longrightarrow> valid (e # s)"
+
+end
+
+(*** more RC constrains of type system in formalisation ***)
+
+locale rc_basic = init +
+  assumes 
+      init_initialrole_valid: "\<And> f r. init_initialrole f = Some r \<Longrightarrow> r = InheritParentRole \<or> r = UseForcedRole \<or> (\<exists> nr. r = NormalRole nr)"   (* 6.10 *)
+  and init_forcedrole_valid:  "\<And> f r. init_forcedrole f = Some r \<Longrightarrow> r = InheritUserRole \<or> r = InheritProcessRole \<or> r = InheritParentRole \<or> r = InheritUpMixed \<or> (\<exists> nr. r = NormalRole nr)"  (* 6.11 *)
+  and init_proc_forcedrole_valid: "\<And> p r. init_proc_forcedrole p = Some r \<Longrightarrow> r = InheritUserRole \<or> r = InheritProcessRole \<or> r = InheritUpMixed \<or> (\<exists> nr. r = NormalRole nr)"  (* 6.7 *)
+  and default_fd_create_type_valid: "\<And> nr t. default_fd_create_type nr = t \<Longrightarrow> t = InheritParent_file_type \<or> (\<exists> nt. t = NormalFile_type nt)" (*6.16*)
+  and default_process_create_type_valid: "\<And> nr t. default_process_create_type nr = t \<Longrightarrow> t = InheritParent_proc_type \<or> (\<exists> nt. t = NormalProc_type nt)" (*6.18*)
+  and default_process_chown_type_valid: "\<And> nr t. default_process_chown_type nr = t \<Longrightarrow> t = InheritParent_proc_type \<or> t = UseNewRoleType \<or> (\<exists> nt. t = NormalProc_type nt)" (*6.19*)
+  and default_process_execute_type_valid: "\<And> nr t. default_process_execute_type nr = t \<Longrightarrow> t = InheritParent_proc_type \<or> (\<exists> nt. t = NormalProc_type nt)" (*6.20*)
+
+begin
+
+lemma init_file_initialrole_valid: "init_file_initialrole f = Some r \<Longrightarrow> r = InheritParentRole \<or> r = UseForcedRole \<or> (\<exists> nr. r = NormalRole nr)"
+apply (induct f)
+by (auto simp:init_file_initialrole.simps dest:init_initialrole_valid split:option.splits)
+
+lemma initialrole_valid: "initialrole \<tau> f = Some r \<Longrightarrow> r = InheritParentRole \<or> r = UseForcedRole \<or> (\<exists> nr. r = NormalRole nr)"   (* 6.10 *)
+apply (induct \<tau> arbitrary:f)  defer
+apply (case_tac a)
+apply (auto simp:initialrole.simps dest:init_file_initialrole_valid 
+            split:option.splits if_splits)
+done
+
+lemma effinitialrole_valid: "effinitialrole \<tau> f = Some r \<Longrightarrow> r = UseForcedRole \<or> (\<exists> nr. r = NormalRole nr)" 
+apply (induct f)
+apply (auto simp:effinitialrole_def dest:initialrole_valid split:option.splits if_splits)
+done
+
+lemma init_file_forcedrole_valid: 
+  "init_file_forcedrole f = Some r \<Longrightarrow> r = InheritUserRole \<or> r = InheritProcessRole \<or> r = InheritParentRole \<or> r = InheritUpMixed \<or> (\<exists> nr. r = NormalRole nr)" 
+apply (induct f)
+by (auto simp:init_file_forcedrole.simps dest:init_forcedrole_valid split:option.splits)
+
+lemma forcedrole_valid: "forcedrole \<tau> f = Some r \<Longrightarrow> r = InheritUserRole \<or> r = InheritProcessRole \<or> r = InheritParentRole \<or> r = InheritUpMixed \<or> (\<exists> nr. r = NormalRole nr)"  (* 6.10 *)
+apply (induct \<tau> arbitrary:f) defer 
+apply (case_tac a)
+apply (auto simp:forcedrole.simps dest:init_file_forcedrole_valid 
+           split:option.splits if_splits)
+done
+
+lemma effforcedrole_valid: "effforcedrole \<tau> f = Some r \<Longrightarrow> r = InheritUserRole \<or> r = InheritProcessRole  \<or> r = InheritUpMixed \<or> (\<exists> nr. r = NormalRole nr)"
+apply (induct f)
+apply (auto simp:effforcedrole_def dest:forcedrole_valid split:option.splits if_splits)
+done
+
+lemma proc_forcedrole_valid: "proc_forcedrole \<tau> p = Some r \<Longrightarrow> r = InheritUserRole \<or> r = InheritProcessRole \<or> r = InheritUpMixed \<or> (\<exists> nr. r = NormalRole nr)"  (* 6.7 *)
+apply (induct \<tau> arbitrary:p) defer
+apply (case_tac a)
+apply (auto simp:proc_forcedrole.simps dest:init_proc_forcedrole_valid effforcedrole_valid 
+           split:option.splits if_splits)
+done
+
+lemma pct_valid: "pct \<tau> p = Some t \<Longrightarrow> t = InheritParent_proc_type \<or> (\<exists> nt. t = NormalProc_type nt)"
+by (auto simp:pct_def default_process_create_type_valid split:option.splits)
+  
+lemma pot_valid: "pot \<tau> p = Some t \<Longrightarrow> t = InheritParent_proc_type \<or> t = UseNewRoleType \<or> (\<exists> nt. t = NormalProc_type nt)" 
+by (auto simp:pot_def default_process_chown_type_valid split:option.splits)
+
+lemma pet_valid: "pet \<tau> p = Some t \<Longrightarrow> t = InheritParent_proc_type \<or> (\<exists> nt. t = NormalProc_type nt)" 
+by (simp add:pet_def default_process_execute_type_valid map_comp_def split:option.splits)
+
+end
+
+(******* Reachable/tainting related type definitions  ********)
+ 
+datatype t_object = 
+  Proc "t_process" 
+| File "t_file"
+| IPC  "t_ipc"
+
+fun object_in_init :: "t_object \<Rightarrow> t_file set \<Rightarrow> t_process set \<Rightarrow> t_ipc set \<Rightarrow> bool"
+where
+  "object_in_init (Proc p) init_files init_procs init_ipcs = (p \<in> init_procs)"
+| "object_in_init (File f) init_files init_procs init_ipcs = (f \<in> init_files)"
+| "object_in_init (IPC i) init_files init_procs init_ipcs = (i \<in> init_ipcs)"
+
+(*** locale for dynamic tainting ***)
+
+locale tainting = rc_basic + 
+
+fixes seeds :: "t_object set"
+
+assumes
+  seeds_in_init: "\<And> obj. obj \<in> seeds \<Longrightarrow> object_in_init obj init_files init_processes init_ipcs"
+begin
+
+lemma finite_seeds: "finite seeds"
+proof-
+  have "finite {obj. object_in_init obj init_files init_processes init_ipcs}"
+  proof-
+    have "finite {File f| f. f \<in> init_files}"
+      using init_finite by auto
+    moreover have "finite {Proc p| p. p \<in> init_processes}"
+      using init_finite by auto
+    moreover have "finite {IPC i| i. i \<in> init_ipcs}"
+      using init_finite by auto
+    ultimately have "finite ({File f| f. f \<in> init_files} \<union> {Proc p| p. p \<in> init_processes} \<union> {IPC i| i. i \<in> init_ipcs})"
+      by auto
+    thus ?thesis
+      apply (rule_tac B = "({File f| f. f \<in> init_files} \<union> {Proc p| p. p \<in> init_processes} \<union> {IPC i| i. i \<in> init_ipcs})" in finite_subset)
+      by (auto, case_tac x, simp+)
+  qed
+  with seeds_in_init
+  show ?thesis
+    by (rule_tac finite_subset, auto)
+qed
+
+fun exists :: "t_state \<Rightarrow> t_object \<Rightarrow> bool"
+where
+  "exists s (File f) = (f \<in> current_files s)"
+| "exists s (Proc p) = (p \<in> current_procs s)"
+| "exists s (IPC  i) = (i \<in> current_ipcs  s)"
+
+inductive tainted :: "t_object \<Rightarrow> t_state \<Rightarrow> bool" ("_ \<in> tainted _" [100, 100] 100)
+where
+  t_init:  "obj \<in> seeds \<Longrightarrow> obj \<in> tainted []"
+| t_clone: "\<lbrakk>Proc p \<in> tainted s; valid (Clone p p' # s)\<rbrakk> \<Longrightarrow> Proc p' \<in> tainted (Clone p p' # s)"
+| t_exec:  "\<lbrakk>File f \<in> tainted s; valid (Execute p f # s)\<rbrakk> \<Longrightarrow> Proc p \<in> tainted (Execute p f # s)"
+| t_cfile: "\<lbrakk>Proc p \<in> tainted s; valid (CreateFile p f # s)\<rbrakk> \<Longrightarrow> File f \<in> tainted (CreateFile p f # s)"
+| t_cipc:  "\<lbrakk>Proc p \<in> tainted s; valid (CreateIPC p i # s)\<rbrakk> \<Longrightarrow> IPC i \<in> tainted (CreateIPC p i # s)"
+| t_read:  "\<lbrakk>File f \<in> tainted s; valid (ReadFile p f # s)\<rbrakk> \<Longrightarrow> Proc p \<in> tainted (ReadFile p f # s)"
+| t_write: "\<lbrakk>Proc p \<in> tainted s; valid (WriteFile p f # s)\<rbrakk> \<Longrightarrow> File f \<in> tainted (WriteFile p f # s)"
+| t_send:  "\<lbrakk>Proc p \<in> tainted s; valid (Send p i # s)\<rbrakk> \<Longrightarrow> IPC i \<in> tainted (Send p i # s)"
+| t_recv:  "\<lbrakk>IPC i \<in> tainted s; valid (Recv p i # s)\<rbrakk> \<Longrightarrow> Proc p \<in> tainted (Recv p i # s)"
+| t_remain:"\<lbrakk>obj \<in> tainted s; valid (e # s); exists (e # s) obj\<rbrakk> \<Longrightarrow> obj \<in> tainted (e # s)"
+
+definition taintable:: "t_object \<Rightarrow> bool"
+where
+  "taintable obj \<equiv> \<exists> s. obj \<in> tainted s"
+
+fun deleted :: "t_object \<Rightarrow> t_event list \<Rightarrow> bool"
+where
+  "deleted obj [] = False"
+| "deleted obj (Kill p p' # \<tau>) = ((obj = Proc p') \<or> deleted obj \<tau>)"
+| "deleted obj (DeleteFile p f # \<tau>) = ((obj = File f) \<or> deleted obj \<tau>)"
+| "deleted obj (DeleteIPC p i # \<tau>) = ((obj = IPC i) \<or> deleted obj \<tau>)"
+| "deleted obj (_ # \<tau>) = deleted obj \<tau>"
+
+definition undeletable :: "t_object \<Rightarrow> bool"
+where
+  "undeletable obj \<equiv> exists [] obj \<and> \<not> (\<exists> s. valid s \<and> deleted obj s)"
+
+fun no_del_event:: "t_event list \<Rightarrow> bool"
+where
+  "no_del_event [] = True"
+| "no_del_event (Kill p p' # \<tau>) = False"
+| "no_del_event (DeleteFile p f # \<tau>) = False"
+| "no_del_event (DeleteIPC p i # \<tau>) = False"
+| "no_del_event (_ # \<tau>) = no_del_event \<tau>"
+
+end
+
+
+(***** locale for statical world *****)
+
+type_synonym t_sprocess = "t_normal_role \<times> t_role \<times> t_normal_proc_type \<times> t_user"
+
+type_synonym t_sfile = "t_normal_file_type \<times> t_file"
+
+type_synonym t_sipc = "t_normal_ipc_type"
+
+datatype t_sobject =
+  SProc "t_sprocess" "t_process option"
+| SFile "t_sfile"    "t_file option"
+| SIPC  "t_sipc"     "t_ipc option"
+| Unknown
+
+locale tainting_s_complete = tainting
+
+begin 
+
+definition chown_role_aux:: "t_normal_role \<Rightarrow> t_role \<Rightarrow> t_user \<Rightarrow> t_normal_role option"
+where
+  "chown_role_aux cr fr u \<equiv> (
+     case fr of
+       NormalRole r \<Rightarrow> Some r
+     | InheritProcessRole \<Rightarrow> Some cr
+     | _ \<Rightarrow> defrole u       )"
+
+definition chown_type_aux:: "t_normal_role \<Rightarrow> t_normal_role \<Rightarrow> t_normal_proc_type \<Rightarrow> t_normal_proc_type"
+where
+  "chown_type_aux cr nr t \<equiv> (
+     case (default_process_chown_type cr) of
+       InheritParent_proc_type   \<Rightarrow> t
+     | UseNewRoleType            \<Rightarrow> (case (default_process_create_type nr) of 
+                                      InheritParent_proc_type \<Rightarrow> t
+                                     | NormalProc_type tp      \<Rightarrow> tp)
+     | NormalProc_type tp        \<Rightarrow> tp )"
+
+definition exec_type_aux :: "t_normal_role \<Rightarrow> t_normal_proc_type \<Rightarrow> t_normal_proc_type"
+where
+  "exec_type_aux cr t \<equiv> (case (default_process_execute_type cr) of 
+                           InheritParent_proc_type \<Rightarrow> t
+                         | NormalProc_type tp      \<Rightarrow>tp)"
+
+definition exec_role_aux :: "t_normal_role \<Rightarrow> t_file \<Rightarrow> t_user \<Rightarrow> t_normal_role option"
+where
+ "exec_role_aux cr sd u \<equiv> (
+    case (erole_functor init_file_initialrole UseForcedRole sd) of
+      Some ir \<Rightarrow> (case ir of 
+                     NormalRole r \<Rightarrow> Some r
+                   | UseForcedRole \<Rightarrow> (
+                       case (erole_functor init_file_forcedrole InheritUpMixed sd) of
+                         Some fr \<Rightarrow> (case fr of 
+                                       NormalRole r \<Rightarrow> Some r
+                                     | InheritUserRole \<Rightarrow> defrole u 
+                                     | _ \<Rightarrow> Some cr )
+                       | None    \<Rightarrow> None )     
+                   | _  \<Rightarrow> None )
+     | _ \<Rightarrow> None            )"       
+
+definition clone_type_aux :: "t_normal_role \<Rightarrow> t_normal_proc_type \<Rightarrow> t_normal_proc_type"
+where
+  "clone_type_aux r t \<equiv> (case (default_process_create_type r) of
+                           InheritParent_proc_type \<Rightarrow> t
+                         | NormalProc_type tp      \<Rightarrow> tp)"
+
+(* all the static objects that dynamic objects referring to *)
+(* they should all be in a finite set ? *)
+(* besides, they're no clone case for all_sobjs, we have no "SProc ... None" case, Clone p p', p' statically viewed as p too, so the many2many mapping becomes to many2one, which makes all succeeded*)
+inductive_set all_sobjs :: "t_sobject set"
+where
+  af_init:  "\<lbrakk>f \<in> init_files; etype_aux init_file_type_aux f = Some t\<rbrakk> \<Longrightarrow> SFile (t, f) (Some f) \<in> all_sobjs"
+| af_cfd:   "\<lbrakk>SFile (t, sd) sf \<in> all_sobjs; SProc (r, fr, pt, u) sp \<in> all_sobjs; default_fd_create_type r = NormalFile_type t'; (r, File_type t, WRITE) \<in> compatible; (r, File_type t', CREATE) \<in> compatible\<rbrakk> \<Longrightarrow> SFile (t', sd) None \<in> all_sobjs"
+| af_cfd':  "\<lbrakk>SFile (t, sd) sf \<in> all_sobjs; SProc (r, fr, pt, u) sp \<in> all_sobjs; default_fd_create_type r = InheritParent_file_type; (r, File_type t, WRITE) \<in> compatible\<rbrakk> \<Longrightarrow> SFile (t, sd) None \<in> all_sobjs"
+
+| ai_init:  "init_ipc_type i = Some t \<Longrightarrow> SIPC t (Some i) \<in> all_sobjs"
+| ai_cipc:  "\<lbrakk>SProc (r, fr, pt, u) sp \<in> all_sobjs; (r, IPC_type (default_ipc_create_type r), CREATE) \<in> compatible\<rbrakk> \<Longrightarrow> SIPC (default_ipc_create_type r) None \<in> all_sobjs"
+
+| ap_init:  "\<lbrakk>init_currentrole p = Some r; init_proc_forcedrole p = Some fr; init_process_type p = Some t; init_owner p = Some u\<rbrakk> \<Longrightarrow> SProc (r, fr, t, u) (Some p) \<in> all_sobjs"
+| ap_crole: "\<lbrakk>SProc (r, fr, t, u) sp \<in> all_sobjs; r' \<in> comproles r\<rbrakk> \<Longrightarrow> SProc (r', fr, t, u) sp \<in> all_sobjs"
+| ap_chown: "\<lbrakk>SProc (r, fr, t, u') sp \<in> all_sobjs; u \<in> init_users; chown_role_aux r fr u = Some nr; (r, Proc_type t, CHANGE_OWNER) \<in> compatible\<rbrakk> \<Longrightarrow> SProc (nr, fr, chown_type_aux r nr t, u) sp \<in> all_sobjs"
+| ap_exec:  "\<lbrakk>SProc (r, fr, pt, u) sp \<in> all_sobjs; SFile (t, sd) sf \<in> all_sobjs; (r, File_type t, EXECUTE) \<in> compatible; exec_role_aux r sd u = Some r'; erole_functor init_file_forcedrole InheritUpMixed sd = Some fr'\<rbrakk> \<Longrightarrow> SProc (r', fr', exec_type_aux r pt, u) sp \<in> all_sobjs" 
+| ap_clone: "SProc (r, fr, pt, u) sp \<in> all_sobjs \<Longrightarrow> SProc (r, fr, clone_type_aux r pt, u) sp \<in> all_sobjs"
+
+fun init_obj2sobj :: "t_object \<Rightarrow> t_sobject"
+where
+  "init_obj2sobj (File f) = (
+     case (etype_aux init_file_type_aux f) of
+       Some t \<Rightarrow> SFile (t, f) (Some f)
+     | _      \<Rightarrow> Unknown      )"
+| "init_obj2sobj (Proc p) = (
+     case (init_currentrole p, init_proc_forcedrole p, init_process_type p, init_owner p) of
+       (Some r, Some fr, Some t, Some u) \<Rightarrow> SProc (r, fr, t, u) (Some p)
+     | _  \<Rightarrow> Unknown          )"
+| "init_obj2sobj (IPC i) = (
+     case (init_ipc_type i) of 
+       Some t \<Rightarrow> SIPC t (Some i)
+     | _      \<Rightarrow> Unknown      )"
+
+inductive_set tainted_s :: "t_sobject set"
+where
+  ts_init: "obj \<in> seeds \<Longrightarrow> init_obj2sobj obj \<in> tainted_s"
+| ts_exec1: "\<lbrakk>SFile (t, sd) sf \<in> tainted_s; SProc (r, fr, pt, u) sp \<in> all_sobjs; (r, File_type t, EXECUTE) \<in> compatible; exec_role_aux r sd u = Some r'; erole_functor init_file_forcedrole InheritUpMixed sd = Some fr'\<rbrakk> \<Longrightarrow> SProc (r', fr', exec_type_aux r pt, u) sp \<in> tainted_s"
+| ts_exec2: "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s; SFile (t, sd) sf \<in> all_sobjs; (r, File_type t, EXECUTE) \<in> compatible; exec_role_aux r sd u = Some r'; erole_functor init_file_forcedrole InheritUpMixed sd = Some fr'\<rbrakk> \<Longrightarrow> SProc (r', fr', exec_type_aux r pt, u) sp \<in> tainted_s"
+| ts_cfd:  "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s; SFile (t, sd) sf \<in> all_sobjs; default_fd_create_type r = NormalFile_type t'; (r, File_type t, WRITE) \<in> compatible; (r, File_type t', CREATE) \<in> compatible\<rbrakk> \<Longrightarrow> SFile (t', sd) None \<in> tainted_s"
+| ts_cfd': "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s; SFile (t, sd) sf \<in> all_sobjs; default_fd_create_type r = InheritParent_file_type; (r, File_type t, WRITE) \<in> compatible\<rbrakk> \<Longrightarrow> SFile (t, sd) None \<in> tainted_s"
+| ts_cipc: "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s; (r, IPC_type (default_ipc_create_type r), CREATE) \<in> compatible\<rbrakk> \<Longrightarrow> SIPC (default_ipc_create_type r) None \<in> tainted_s"
+| ts_read: "\<lbrakk>SFile (t, sd) sf \<in> tainted_s; SProc (r, fr, pt, u) sp \<in> all_sobjs; (r, File_type t, READ) \<in> compatible\<rbrakk> \<Longrightarrow> SProc (r, fr, pt, u) sp \<in> tainted_s"
+| ts_write: "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s; SFile (t, sd) sf \<in> all_sobjs; (r, File_type t, WRITE) \<in> compatible\<rbrakk> \<Longrightarrow> SFile (t, sd) sf \<in> tainted_s"
+| ts_recv: "\<lbrakk>SIPC t si \<in> tainted_s; SProc (r, fr, pt, u) sp \<in> all_sobjs; (r, IPC_type t, RECEIVE) \<in> compatible\<rbrakk> \<Longrightarrow> SProc (r, fr, pt, u) sp \<in> tainted_s"
+| ts_send: "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s; SIPC t si \<in> all_sobjs; (r, IPC_type t, SEND) \<in> compatible\<rbrakk> \<Longrightarrow> SIPC t si \<in> tainted_s"
+| ts_crole: "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s; r' \<in> comproles r\<rbrakk> \<Longrightarrow> SProc (r', fr, pt, u) sp \<in> tainted_s"
+| ts_chown: "\<lbrakk>SProc (r, fr, t, u') sp \<in> tainted_s; u \<in> init_users; chown_role_aux r fr u = Some nr; (r, Proc_type t, CHANGE_OWNER) \<in> compatible\<rbrakk> \<Longrightarrow> SProc (nr, fr, chown_type_aux r nr t, u) sp \<in> tainted_s"
+| ts_clone: "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s; (r, Proc_type pt, CREATE) \<in> compatible\<rbrakk> \<Longrightarrow> SProc (r, fr, clone_type_aux r pt, u) sp \<in> tainted_s"
+
+(*** mapping function from dynamic 2 statical ****)
+
+fun source_dir:: "t_state \<Rightarrow> t_file \<Rightarrow> t_file option"
+where
+   "source_dir s []     = (if ([] \<in> init_files \<and> \<not> (deleted (File []) s)) 
+                           then Some [] 
+                           else None
+                          )"                                 
+ | "source_dir s (f#pf) = (if ((f#pf) \<in> init_files \<and> \<not> (deleted (File (f#pf)) s)) 
+                           then Some (f#pf)
+                           else source_dir s pf
+                          )"
+
+(* cf2sfile's properities should all be under condition: f is not deleted in \<tau>*)
+fun cf2sfile:: "t_state \<Rightarrow> t_file \<Rightarrow> t_sfile option"
+where 
+  "cf2sfile s f = (case (etype_of_file s f, source_dir s f) of
+                     (Some t, Some sd) \<Rightarrow> Some (t, sd)
+                   | _                 \<Rightarrow> None)"
+
+fun cp2sproc:: "t_state \<Rightarrow> t_process \<Rightarrow> t_sprocess option"
+where
+  "cp2sproc s p = (case (currentrole s p, proc_forcedrole s p, type_of_process s p, owner s p) of
+                     (Some r, Some  fr, Some t, Some u) \<Rightarrow> Some (r, fr, t, u)
+                   | _                                  \<Rightarrow> None)"
+
+fun ci2sipc:: "t_state \<Rightarrow> t_ipc \<Rightarrow> t_sipc option"
+where
+  "ci2sipc s i = (type_of_ipc s i)"
+
+(*** in statical view, clone event is process to try different runs(different sprocess:role/type...),
+so statically these sprocesses should have the same source: the process in the initial state *********)
+fun source_proc :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process option"
+where
+  "source_proc [] p = (if (p \<in> init_processes) then Some p else None)"
+| "source_proc (Clone p p' # s) p'' = (if (p'' = p') then source_proc s p else source_proc s p'')"
+| "source_proc (e # s) p = source_proc s p"
+
+fun obj2sobj :: "t_state \<Rightarrow> t_object \<Rightarrow> t_sobject"
+where
+  "obj2sobj s (File f) = (
+     case (cf2sfile s f) of
+       Some sf \<Rightarrow> (if (f \<in> init_files \<and> (\<not> deleted (File f) s)) 
+                   then SFile sf (Some f)  
+                   else SFile sf None)
+     | _       \<Rightarrow> Unknown )"
+| "obj2sobj s (Proc p) = (
+     case (cp2sproc s p) of 
+       Some sp \<Rightarrow> SProc sp (source_proc s p)
+     | _       \<Rightarrow> Unknown )"
+| "obj2sobj s (IPC i) = (
+     case (ci2sipc s i) of 
+       Some si \<Rightarrow> (if (i \<in> init_ipcs \<and> (\<not> deleted (IPC i) s)) 
+                   then SIPC si (Some i) 
+                   else SIPC si None)
+     | _       \<Rightarrow> Unknown )"
+
+fun role_of_sproc :: "t_sprocess \<Rightarrow> t_normal_role"
+where
+  "role_of_sproc (r, fr, pt, u) = r"
+
+fun source_of_sobj :: "t_sobject \<Rightarrow> t_object option"
+where
+  "source_of_sobj (SFile sf tag) = (case tag of
+                                        Some f \<Rightarrow> Some (File f)
+                                      | _      \<Rightarrow> None)"
+| "source_of_sobj (SProc sp tag) = (case tag of
+                                        Some p \<Rightarrow> Some (Proc p)
+                                      | _      \<Rightarrow> None)"
+| "source_of_sobj (SIPC  si tag) = (case tag of
+                                        Some i \<Rightarrow> Some (IPC i)
+                                      | _      \<Rightarrow> None)"
+| "source_of_sobj Unknown        = None"
+
+definition taintable_s :: "t_object \<Rightarrow> bool"
+where
+  "taintable_s obj \<equiv> \<exists>sobj. sobj \<in> tainted_s \<and> source_of_sobj sobj = Some obj"
+
+(*
+definition file_deletable_s:: "t_file \<Rightarrow> bool"
+where 
+  "file_deletable_s f \<equiv> \<forall> t sd srf. (SFile (t,sd) srf \<in> all_sobjs \<and> f \<preceq> sd \<longrightarrow> (\<exists> sp srp. SProc sp srp \<in> all_sobjs \<and> (role_of_sproc sp, File_type t, DELETE) \<in> compatible))"
+*)
+(*
+definition file_deletable_s:: "t_file \<Rightarrow> bool"
+where 
+  "file_deletable_s sd \<equiv> \<forall> f. (f \<in> init_files \<and> sd \<preceq> f \<longrightarrow> (\<exists> t sp srp. SProc sp srp \<in> all_sobjs \<and> etype_of_file [] f = Some t \<and> (role_of_sproc sp, File_type t, DELETE) \<in> compatible))"*)
+
+definition file_deletable_s:: "t_file \<Rightarrow> bool"
+where 
+  "file_deletable_s f \<equiv> \<exists> t sp srp. SProc sp srp \<in> all_sobjs \<and> etype_of_file [] f = Some t \<and> (role_of_sproc sp, File_type t, DELETE) \<in> compatible"
+
+definition proc_deletable_s:: "t_process \<Rightarrow> bool"
+where
+  "proc_deletable_s p \<equiv> \<exists> r fr pt u sp' srp'. SProc (r,fr,pt,u) (Some p) \<in> all_sobjs \<and> SProc sp' srp' \<in> all_sobjs \<and> (role_of_sproc sp', Proc_type pt, DELETE) \<in> compatible"
+
+definition ipc_deletable_s:: "t_ipc \<Rightarrow> bool"
+where
+  "ipc_deletable_s i \<equiv> \<exists> t sp srp. SProc sp srp \<in> all_sobjs \<and> type_of_ipc [] i = Some t \<and> (role_of_sproc sp, IPC_type t, DELETE) \<in> compatible"
+
+fun deletable_s :: "t_object \<Rightarrow> bool"
+where
+  "deletable_s (Proc p) = (p \<in> init_processes \<and> proc_deletable_s p)"
+| "deletable_s (File f) = (f \<in> init_files     \<and> file_deletable_s f)"
+| "deletable_s (IPC  i) = (i \<in> init_ipcs      \<and> ipc_deletable_s  i)"
+
+definition undeletable_s:: "t_object \<Rightarrow> bool"
+where
+  "undeletable_s obj \<equiv> exists [] obj \<and> \<not> deletable_s obj"
+
+end
+
+locale tainting_s_sound = tainting_s_complete +
+
+assumes
+  clone_type_unchange: "\<And> r. default_process_create_type r = InheritParent_proc_type"
+(* this is for statically clone do not change anything ! ! so that, there're no rule for 
+clone in all_sobjs and tainted_s *)
+and clone_no_limit: "\<And> r t. (r, Proc_type t, CREATE) \<in> compatible"
+
+begin
+
+(* the all_sobjs': the soundness view of all_sobjs, just remove the clone case, cause
+in this locale, clone doesn't change any information of such a sprocess*)
+inductive_set all_sobjs' :: "t_sobject set"
+where
+  af'_init:  "\<lbrakk>f \<in> init_files; etype_aux init_file_type_aux f = Some t\<rbrakk> \<Longrightarrow> SFile (t, f) (Some f) \<in> all_sobjs'"
+| af'_cfd:   "\<lbrakk>SFile (t, sd) sf \<in> all_sobjs'; SProc (r, fr, pt, u) sp \<in> all_sobjs'; default_fd_create_type r = NormalFile_type t'; (r, File_type t, WRITE) \<in> compatible; (r, File_type t', CREATE) \<in> compatible\<rbrakk> \<Longrightarrow> SFile (t', sd) None \<in> all_sobjs'"
+| af'_cfd':  "\<lbrakk>SFile (t, sd) sf \<in> all_sobjs'; SProc (r, fr, pt, u) sp \<in> all_sobjs'; default_fd_create_type r = InheritParent_file_type; (r, File_type t, WRITE) \<in> compatible\<rbrakk> \<Longrightarrow> SFile (t, sd) None \<in> all_sobjs'"
+
+| ai'_init:  "init_ipc_type i = Some t \<Longrightarrow> SIPC t (Some i) \<in> all_sobjs'"
+| ai'_cipc:  "\<lbrakk>SProc (r, fr, pt, u) sp \<in> all_sobjs'; (r, IPC_type (default_ipc_create_type r), CREATE) \<in> compatible\<rbrakk> \<Longrightarrow> SIPC (default_ipc_create_type r) None \<in> all_sobjs'"
+
+| ap'_init:  "\<lbrakk>init_currentrole p = Some r; init_proc_forcedrole p = Some fr; init_process_type p = Some t; init_owner p = Some u\<rbrakk> \<Longrightarrow> SProc (r, fr, t, u) (Some p) \<in> all_sobjs'"
+| ap'_crole: "\<lbrakk>SProc (r, fr, t, u) sp \<in> all_sobjs'; r' \<in> comproles r\<rbrakk> \<Longrightarrow> SProc (r', fr, t, u) sp \<in> all_sobjs'"
+| ap'_chown: "\<lbrakk>SProc (r, fr, t, u') sp \<in> all_sobjs'; u \<in> init_users; chown_role_aux r fr u = Some nr; (r, Proc_type t, CHANGE_OWNER) \<in> compatible\<rbrakk> \<Longrightarrow> SProc (nr, fr, chown_type_aux r nr t, u) sp \<in> all_sobjs'"
+| ap'_exec:  "\<lbrakk>SProc (r, fr, pt, u) sp \<in> all_sobjs'; SFile (t, sd) sf \<in> all_sobjs'; (r, File_type t, EXECUTE) \<in> compatible; exec_role_aux r sd u = Some r'; erole_functor init_file_forcedrole InheritUpMixed sd = Some fr'\<rbrakk> \<Longrightarrow> SProc (r', fr', exec_type_aux r pt, u) sp \<in> all_sobjs'" 
+
+(* the tainted_s': the soundness view of all_sobjs, just remove the clone case, cause
+in this locale, clone doesn't change any information of such a sprocess*)
+inductive_set tainted_s' :: "t_sobject set"
+where
+  ts'_init: "obj \<in> seeds \<Longrightarrow> init_obj2sobj obj \<in> tainted_s'"
+| ts'_exec1: "\<lbrakk>SFile (t, sd) sf \<in> tainted_s'; SProc (r, fr, pt, u) sp \<in> all_sobjs'; (r, File_type t, EXECUTE) \<in> compatible; exec_role_aux r sd u = Some r'; erole_functor init_file_forcedrole InheritUpMixed sd = Some fr'\<rbrakk> \<Longrightarrow> SProc (r', fr', exec_type_aux r pt, u) sp \<in> tainted_s'"
+| ts'_exec2: "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s'; SFile (t, sd) sf \<in> all_sobjs'; (r, File_type t, EXECUTE) \<in> compatible; exec_role_aux r sd u = Some r'; erole_functor init_file_forcedrole InheritUpMixed sd = Some fr'\<rbrakk> \<Longrightarrow> SProc (r', fr', exec_type_aux r pt, u) sp \<in> tainted_s'"
+| ts'_cfd:  "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s'; SFile (t, sd) sf \<in> all_sobjs'; default_fd_create_type r = NormalFile_type t'; (r, File_type t, WRITE) \<in> compatible; (r, File_type t', CREATE) \<in> compatible\<rbrakk> \<Longrightarrow> SFile (t', sd) None \<in> tainted_s'"
+| ts'_cfd': "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s'; SFile (t, sd) sf \<in> all_sobjs'; default_fd_create_type r = InheritParent_file_type; (r, File_type t, WRITE) \<in> compatible\<rbrakk> \<Longrightarrow> SFile (t, sd) None \<in> tainted_s'"
+| ts'_cipc: "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s'; (r, IPC_type (default_ipc_create_type r), CREATE) \<in> compatible\<rbrakk> \<Longrightarrow> SIPC (default_ipc_create_type r) None \<in> tainted_s'"
+| ts'_read: "\<lbrakk>SFile (t, sd) sf \<in> tainted_s'; SProc (r, fr, pt, u) sp \<in> all_sobjs'; (r, File_type t, READ) \<in> compatible\<rbrakk> \<Longrightarrow> SProc (r, fr, pt, u) sp \<in> tainted_s'"
+| ts'_write: "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s'; SFile (t, sd) sf \<in> all_sobjs'; (r, File_type t, WRITE) \<in> compatible\<rbrakk> \<Longrightarrow> SFile (t, sd) sf \<in> tainted_s'"
+| ts'_recv: "\<lbrakk>SIPC t si \<in> tainted_s'; SProc (r, fr, pt, u) sp \<in> all_sobjs'; (r, IPC_type t, RECEIVE) \<in> compatible\<rbrakk> \<Longrightarrow> SProc (r, fr, pt, u) sp \<in> tainted_s'"
+| ts'_send: "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s'; SIPC t si \<in> all_sobjs'; (r, IPC_type t, SEND) \<in> compatible\<rbrakk> \<Longrightarrow> SIPC t si \<in> tainted_s'"
+| ts'_crole: "\<lbrakk>SProc (r, fr, pt, u) sp \<in> tainted_s'; r' \<in> comproles r\<rbrakk> \<Longrightarrow> SProc (r', fr, pt, u) sp \<in> tainted_s'"
+| ts'_chown: "\<lbrakk>SProc (r, fr, t, u') sp \<in> tainted_s'; u \<in> init_users; chown_role_aux r fr u = Some nr; (r, Proc_type t, CHANGE_OWNER) \<in> compatible\<rbrakk> \<Longrightarrow> SProc (nr, fr, chown_type_aux r nr t, u) sp \<in> tainted_s'"
+
+fun sobj_source_eq_obj :: "t_sobject \<Rightarrow> t_object \<Rightarrow> bool"
+where
+  "sobj_source_eq_obj (SFile sf None)      (File f) = True"
+| "sobj_source_eq_obj (SFile sf (Some f')) (File f) = (f' = f)"
+| "sobj_source_eq_obj (SProc sp None)      (Proc p) = True"
+| "sobj_source_eq_obj (SProc sp (Some p')) (Proc p) = (p' = p)"
+| "sobj_source_eq_obj (SIPC  si None)      (IPC  i) = True"
+| "sobj_source_eq_obj (SIPC  si (Some i')) (IPC  i) = (i' = i)"
+| "sobj_source_eq_obj _                    _        = False"
+
+fun not_both_sproc:: "t_sobject \<Rightarrow> t_sobject \<Rightarrow> bool"
+where
+  "not_both_sproc (SProc sp srp) (SProc sp' srp') = False"
+| "not_both_sproc _              _                = True"
+
+(*** definitions for init processes statical informations reservation ***)
+
+definition initp_intact :: "t_state => bool"
+where
+  "initp_intact s \<equiv> \<forall> p. p \<in> init_processes --> p \<in> current_procs s \<and> obj2sobj s (Proc p) = init_obj2sobj (Proc p)"
+
+definition initp_alter :: "t_state \<Rightarrow> t_process \<Rightarrow> bool"
+where
+ "initp_alter s p \<equiv> \<exists> p'. p' \<in> current_procs s \<and> obj2sobj s (Proc p') = init_obj2sobj (Proc p)"
+
+definition initp_intact_butp :: "t_state \<Rightarrow> t_process \<Rightarrow> bool"
+where
+  "initp_intact_butp s proc \<equiv>  (\<forall> p. p \<in> init_processes \<and> p \<noteq> proc \<longrightarrow> p \<in> current_procs s \<and> obj2sobj s (Proc p) = init_obj2sobj (Proc p)) \<and> initp_alter s proc"
+
+fun initp_intact_but :: "t_state \<Rightarrow> t_sobject \<Rightarrow> bool"
+where
+  "initp_intact_but s (SProc sp (Some p)) = initp_intact_butp s p"
+| "initp_intact_but s _                   = initp_intact s"
+
+(*** how to generating new valid pathname for examples of CreateFile ***)
+
+definition all_fname_under_dir:: "t_file \<Rightarrow> t_state \<Rightarrow> t_fname set"
+where
+  "all_fname_under_dir d \<tau> = {fn. \<exists> f. fn # d = f \<and> f \<in> current_files \<tau>}"
+
+fun fname_all_a:: "nat \<Rightarrow> t_fname"
+where
+  "fname_all_a 0 = []" |
+  "fname_all_a (Suc n) = ''a''@(fname_all_a n)"
+
+definition fname_length_set :: "t_fname set \<Rightarrow> nat set"
+where
+  "fname_length_set fns = length`fns"
+
+definition next_fname:: "t_file \<Rightarrow> t_state \<Rightarrow> t_fname"
+where
+  "next_fname pf \<tau> = fname_all_a ((Max (fname_length_set (all_fname_under_dir pf \<tau>))) + 1)"
+
+definition new_childf:: "t_file \<Rightarrow> t_state \<Rightarrow> t_file"
+where
+  "new_childf pf \<tau> = next_fname pf \<tau> # pf"
+
+end 
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/sound_defs_prop.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,198 @@
+theory sound_defs_prop
+imports rc_theory Main os_rc obj2sobj_prop
+begin
+
+context tainting_s_sound begin
+
+lemma not_both_I:
+  "not_both_sproc (SProc sp srp) sobj \<Longrightarrow> not_both_sproc sobj' sobj"
+by (case_tac sobj, auto)
+
+lemma not_both_I_file:
+  "not_both_sproc (SFile sf srf) sobj \<Longrightarrow> not_both_sproc (SFile sf' srf') sobj"
+by (case_tac sobj, auto)
+
+lemma not_both_I_ipc:
+  "not_both_sproc (SIPC si sri) sobj \<Longrightarrow> not_both_sproc (SIPC si' sri') sobj"
+by (case_tac sobj, auto)
+
+lemma intact_imp_butp:
+  "\<lbrakk>p \<in> init_processes; initp_intact s\<rbrakk> \<Longrightarrow> initp_intact_butp s p "
+by (auto simp:initp_intact_def initp_intact_butp_def initp_alter_def)
+
+lemma no_sproc_imp_intact:
+  "\<lbrakk>not_both_sproc (SProc sp srp) sobj; initp_intact_but s sobj\<rbrakk> \<Longrightarrow> initp_intact s"
+by (case_tac sobj, simp_all)
+
+lemma initp_intact_but_nil:"initp_intact_but [] (init_obj2sobj obj)"  
+apply (case_tac obj)
+apply (auto simp:initp_intact_def initp_intact_butp_def initp_alter_def split:option.splits)
+apply (rule_tac x = nat in exI) using init_proc_has_role 
+by (auto simp:bidirect_in_init_def)
+
+lemma init_alterp_exec:
+  "\<lbrakk>initp_alter s p; p \<in> init_processes; valid (Execute p f # Clone p (new_proc s) # s)\<rbrakk>
+  \<Longrightarrow> initp_alter (Execute p f # Clone p (new_proc s) # s) p"
+apply (frule valid_cons, frule_tac \<tau> = s in valid_cons, frule_tac \<tau> = s in valid_os)
+apply (clarsimp simp add:initp_alter_def simp del:obj2sobj.simps init_obj2sobj.simps)
+apply (subgoal_tac "p' \<noteq> new_proc s") defer apply (rule notI,simp add:np_notin_curp)
+apply (case_tac "p' = p")
+apply (rule_tac x = "new_proc s" in exI) defer
+apply (rule_tac x = p' in exI)
+apply (auto simp:cp2sproc_simps' simp del:cp2sproc.simps 
+           split:option.splits dest!:current_proc_has_sproc')
+done
+
+lemma init_alterp_chown:
+  "\<lbrakk>initp_alter s p; p \<in> init_processes; valid (ChangeOwner p u # Clone p (new_proc s) # s)\<rbrakk>
+  \<Longrightarrow> initp_alter (ChangeOwner p u # Clone p (new_proc s) # s) p"
+apply (frule valid_cons, frule_tac \<tau> = s in valid_cons, frule_tac \<tau> = s in valid_os)
+apply (clarsimp simp add:initp_alter_def simp del:obj2sobj.simps init_obj2sobj.simps)
+apply (subgoal_tac "p' \<noteq> new_proc s") defer apply (rule notI,simp add:np_notin_curp)
+apply (case_tac "p' = p")
+apply (rule_tac x = "new_proc s" in exI) defer
+apply (rule_tac x = p' in exI)
+apply (auto simp:cp2sproc_simps' simp del:cp2sproc.simps 
+           split:option.splits dest!:current_proc_has_sproc')
+done
+
+lemma init_alterp_crole:
+  "\<lbrakk>initp_alter s p; p \<in> init_processes; valid (ChangeRole p r # Clone p (new_proc s) # s)\<rbrakk>
+  \<Longrightarrow> initp_alter (ChangeRole p r # Clone p (new_proc s) # s) p"
+apply (frule valid_cons, frule_tac \<tau> = s in valid_cons, frule_tac \<tau> = s in valid_os)
+apply (clarsimp simp add:initp_alter_def simp del:obj2sobj.simps init_obj2sobj.simps)
+apply (subgoal_tac "p' \<noteq> new_proc s") defer apply (rule notI,simp add:np_notin_curp)
+apply (case_tac "p' = p")
+apply (rule_tac x = "new_proc s" in exI) defer
+apply (rule_tac x = p' in exI)
+apply (auto simp:cp2sproc_simps' simp del:cp2sproc.simps 
+           split:option.splits dest!:current_proc_has_sproc')
+done
+
+lemma init_alterp_other:
+  "\<lbrakk>initp_alter s p; valid (e # s); \<forall> p f. e \<noteq> Execute p f;
+    \<forall> p u. e \<noteq> ChangeOwner p u; \<forall> p r. e \<noteq> ChangeRole p r; no_del_event (e # s)\<rbrakk>
+  \<Longrightarrow> initp_alter (e # s) p"
+apply (frule valid_cons, frule valid_os)
+apply (clarsimp simp add:initp_alter_def simp del:obj2sobj.simps init_obj2sobj.simps)
+apply (subgoal_tac "p' \<noteq> new_proc s") defer apply (rule notI,simp add:np_notin_curp)
+apply (rule_tac x = p' in exI, case_tac e)
+apply (auto simp add: cp2sproc_simps' simp del:cp2sproc.simps split:option.splits)
+done
+
+lemma initp_intact_butp_I_exec:
+  "\<lbrakk>initp_intact_butp s p; p \<in> init_processes; valid (Execute p f # Clone p (new_proc s) # s); 
+    no_del_event s\<rbrakk>
+  \<Longrightarrow> initp_intact_butp (Execute p f # Clone p (new_proc s) # s) p"
+apply (frule valid_cons, frule_tac \<tau> = s in valid_cons)
+apply (auto simp add:initp_intact_def initp_intact_butp_def 
+            simp del:obj2sobj.simps init_obj2sobj.simps 
+               intro:init_alterp_exec)
+apply (erule_tac x = pa in allE)
+apply (subgoal_tac "pa \<noteq> (new_proc s)")
+apply (auto simp:cp2sproc_simps' simp del:cp2sproc.simps 
+           split:option.splits dest!:current_proc_has_sproc')[1]
+apply (rule notI, simp add:np_notin_curp)
+done
+
+lemma initp_intact_butp_I_chown:
+  "\<lbrakk>initp_intact_butp s p; p \<in> init_processes; no_del_event s;
+   valid (ChangeOwner p u # Clone p (new_proc s) # s)\<rbrakk>
+  \<Longrightarrow> initp_intact_butp (ChangeOwner p u # Clone p (new_proc s) # s) p"
+apply (frule valid_cons, frule_tac \<tau> = s in valid_cons)
+apply (auto simp add:initp_intact_def initp_intact_butp_def 
+            simp del:obj2sobj.simps init_obj2sobj.simps 
+               intro:init_alterp_chown)
+apply (erule_tac x = pa in allE)
+apply (subgoal_tac "pa \<noteq> (new_proc s)")
+apply (auto simp:cp2sproc_simps simp del:cp2sproc.simps 
+           split:option.splits dest!:current_proc_has_sproc')[1]
+apply (rule notI, simp add:np_notin_curp)
+done
+
+lemma initp_intact_butp_I_crole:
+  "\<lbrakk>initp_intact_butp s p; p \<in> init_processes; 
+    valid (ChangeRole p r # Clone p (new_proc s) # s); no_del_event s\<rbrakk>
+  \<Longrightarrow> initp_intact_butp (ChangeRole p r # Clone p (new_proc s) # s) p"
+apply (frule valid_cons, frule_tac \<tau> = s in valid_cons)
+apply (auto simp add:initp_intact_def initp_intact_butp_def 
+            simp del:obj2sobj.simps init_obj2sobj.simps 
+               intro:init_alterp_crole)
+apply (erule_tac x = pa in allE)
+apply (subgoal_tac "pa \<noteq> (new_proc s)")
+apply (auto simp:cp2sproc_simps' simp del:cp2sproc.simps 
+           split:option.splits dest!:current_proc_has_sproc')[1]
+apply (rule notI, simp add:np_notin_curp)
+done
+
+lemma initp_intact_butp_I_others:
+  "\<lbrakk>initp_intact_butp s p; valid (e # s); \<forall> p f. e \<noteq> Execute p f;
+    \<forall> p u. e \<noteq> ChangeOwner p u; \<forall> p r. e \<noteq> ChangeRole p r; no_del_event (e # s)\<rbrakk>
+  \<Longrightarrow> initp_intact_butp (e # s) p"
+apply (frule valid_os, frule valid_cons)
+apply (simp add:initp_intact_butp_def init_alterp_other 
+            del:obj2sobj.simps init_obj2sobj.simps)
+apply (rule impI|rule allI|erule conjE|rule conjI)+
+apply (drule_tac obj = "Proc pa" in nodel_imp_exists, simp, simp)
+apply (frule no_del_event_cons_D, drule_tac obj = "Proc pa" and s = s in nodel_imp_exists, simp)
+apply (subgoal_tac "pa \<noteq> new_proc s") defer apply (rule notI,simp add:np_notin_curp)
+apply (rotate_tac 6, erule_tac x = pa in allE, case_tac e)
+apply (auto simp:cp2sproc_simps' simp del:cp2sproc.simps 
+           split:option.splits dest!:current_proc_has_sproc')
+done
+
+lemma initp_intact_I_exec:
+  "\<lbrakk>initp_intact s; valid (Execute (new_proc s) f # Clone p (new_proc s) # s)\<rbrakk>
+  \<Longrightarrow> initp_intact (Execute (new_proc s) f # Clone p (new_proc s) # s)"
+apply (frule valid_cons, frule_tac \<tau> = s in valid_cons)
+apply (auto simp add:initp_intact_def 
+            simp del:obj2sobj.simps init_obj2sobj.simps)
+apply (erule_tac x = pa in allE)
+apply (subgoal_tac "pa \<noteq> (new_proc s)")
+apply (auto simp:cp2sproc_simps' cp2sproc.simps 
+           split:option.splits dest!:current_proc_has_sproc')[1]
+apply (rule notI, simp add:np_notin_curp)
+done
+
+lemma initp_intact_I_chown:
+  "\<lbrakk>initp_intact s; valid (ChangeOwner (new_proc s) u # Clone p (new_proc s) # s)\<rbrakk>
+  \<Longrightarrow> initp_intact (ChangeOwner (new_proc s) u # Clone p (new_proc s) # s)"
+apply (frule valid_cons, frule_tac \<tau> = s in valid_cons)
+apply (auto simp add:initp_intact_def 
+            simp del:obj2sobj.simps init_obj2sobj.simps)
+apply (erule_tac x = pa in allE)
+apply (subgoal_tac "pa \<noteq> (new_proc s)")
+apply (auto simp:cp2sproc_simps cp2sproc.simps 
+           split:option.splits dest!:current_proc_has_sproc')[1]
+apply (rule notI, simp add:np_notin_curp)
+done
+
+lemma initp_intact_I_crole:
+  "\<lbrakk>initp_intact s; valid (ChangeRole (new_proc s) r # Clone p (new_proc s) # s)\<rbrakk>
+  \<Longrightarrow> initp_intact (ChangeRole (new_proc s) r # Clone p (new_proc s) # s)"
+apply (frule valid_cons, frule_tac \<tau> = s in valid_cons)
+apply (auto simp add:initp_intact_def initp_intact_butp_def 
+            simp del:obj2sobj.simps init_obj2sobj.simps)
+apply (erule_tac x = pa in allE)
+apply (subgoal_tac "pa \<noteq> (new_proc s)")
+apply (auto simp:cp2sproc_simps' cp2sproc.simps 
+           split:option.splits dest!:current_proc_has_sproc')[1]
+apply (rule notI, simp add:np_notin_curp)
+done
+
+lemma initp_intact_I_others:
+  "\<lbrakk>initp_intact s; valid (e # s); \<forall> p f. e \<noteq> Execute p f;
+    \<forall> p u. e \<noteq> ChangeOwner p u; \<forall> p r. e \<noteq> ChangeRole p r; no_del_event (e # s)\<rbrakk>
+  \<Longrightarrow> initp_intact (e # s)"
+apply (frule valid_os, frule valid_cons)
+apply (clarsimp simp add:initp_intact_def simp del:obj2sobj.simps init_obj2sobj.simps)
+apply (frule no_del_event_cons_D, drule_tac obj = "Proc p" and s = s in nodel_imp_exists, simp)
+apply (subgoal_tac "p \<noteq> new_proc s") defer apply (rule notI,simp add:np_notin_curp)
+apply (erule_tac x = p in allE, case_tac e)
+apply (auto simp:cp2sproc_simps' simp del:cp2sproc.simps 
+           split:option.splits dest!:current_proc_has_sproc')
+done
+
+end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/source_prop.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,75 @@
+theory source_prop
+imports Main rc_theory os_rc deleted_prop obj2sobj_prop
+begin
+
+context tainting_s_complete begin
+
+lemma all_sobjs_srp_init'[rule_format]:
+  "sobj \<in> all_sobjs \<Longrightarrow> \<forall> sp srp. sobj = SProc sp (Some srp) \<longrightarrow> srp \<in> init_processes"
+apply (erule all_sobjs.induct, auto) 
+using init_proc_has_type 
+by (simp add:bidirect_in_init_def)
+
+lemma all_sobjs_srp_init:
+  "SProc sp (Some srp) \<in> all_sobjs \<Longrightarrow> srp \<in> init_processes"
+by (auto dest!:all_sobjs_srp_init')
+
+(*
+lemma tainted_sproc_srp_init:
+  "SProc sp (Some srp) \<in> tainted_s \<Longrightarrow> srp \<in> init_processes"
+by (auto dest:tainted_s_in_all_sobjs intro:all_sobjs_srp_init) *)
+
+lemma source_proc_of_init_remains:
+  "\<lbrakk>p \<in> init_processes; \<not> deleted (Proc p) s; valid s\<rbrakk> \<Longrightarrow> source_proc s p = Some p"
+apply (induct s, simp)
+apply (frule valid_cons, frule valid_os, frule not_deleted_cons_D, simp)
+apply (case_tac a, auto simp:np_notin_curp dest:not_deleted_imp_exists)
+done
+
+lemma init_proc_keeps_source:
+  "\<lbrakk>p \<in> init_processes; \<not> deleted (Proc p) s; valid s\<rbrakk> 
+   \<Longrightarrow> source_of_sobj (obj2sobj s (Proc p)) = Some (Proc p)"
+apply (frule not_deleted_imp_exists, simp, simp only:exists.simps)
+apply (frule current_proc_has_sproc, simp)
+apply (clarsimp split:option.splits simp:source_proc_of_init_remains)
+done
+
+lemma init_file_keeps_source:
+  "\<lbrakk>f \<in> init_files; \<not> deleted (File f) s; valid s\<rbrakk> 
+   \<Longrightarrow> source_of_sobj (obj2sobj s (File f)) = Some (File f)"
+apply (frule not_deleted_imp_exists, simp, simp only:exists.simps)
+apply (frule current_file_has_sfile, auto)
+done
+
+lemma init_ipc_keeps_source:
+  "\<lbrakk>i \<in> init_ipcs; \<not> deleted (IPC i) s; valid s\<rbrakk> 
+   \<Longrightarrow> source_of_sobj (obj2sobj s (IPC i)) = Some (IPC i)"
+apply (frule not_deleted_imp_exists, simp, simp only:exists.simps)
+apply (frule current_ipc_has_sipc, auto)
+done
+ 
+lemma init_obj_keeps_source: 
+  "\<lbrakk>exists [] obj; \<not> deleted obj s; valid s\<rbrakk> \<Longrightarrow> source_of_sobj (obj2sobj s obj) = Some obj"
+apply (case_tac obj)
+by (auto intro:init_ipc_keeps_source init_proc_keeps_source init_file_keeps_source 
+      simp del:obj2sobj.simps)
+
+end
+
+context tainting_s_sound begin
+
+lemma source_eq:
+  "\<lbrakk>source_of_sobj sobj = Some obj; sobj_source_eq_obj sobj obj'\<rbrakk> \<Longrightarrow> obj = obj'"
+apply (case_tac obj, case_tac [!] sobj, case_tac [!] obj')
+by (auto split:option.splits)
+
+lemma proc_source_eq_prop:
+  "\<lbrakk>obj2sobj s (Proc p) = SProc sp srp; p \<in> current_procs s;
+    sobj_source_eq_obj (SProc sp srp) (Proc p); valid s\<rbrakk>
+  \<Longrightarrow> srp = Some p"
+apply (frule current_proc_has_srp,auto split:option.splits)
+by (case_tac "source_proc s p", simp+)
+
+end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/tainted.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,53 @@
+theory tainted 
+imports Main rc_theory os_rc deleted_prop obj2sobj_prop
+begin
+
+context tainting begin
+
+lemma tainted_is_valid:
+  "obj \<in> tainted s \<Longrightarrow> valid s"
+by (erule tainted.induct, auto simp:vs_nil)
+
+lemma tainted_is_current:
+  "obj \<in> tainted s \<Longrightarrow> exists s obj"
+apply (erule tainted.induct)
+apply (auto dest:valid_os)
+apply (drule seeds_in_init, case_tac obj, auto)
+done
+
+lemma nodel_tainted_exists:
+  "\<lbrakk>no_del_event (s'@s); obj \<in> tainted s\<rbrakk> \<Longrightarrow> exists (s'@s) obj"
+apply (drule_tac obj = obj in nodel_imp_un_deleted)
+by (drule tainted_is_current, simp add:not_deleted_imp_exists')
+
+lemma t_remain_app:
+  "\<lbrakk>obj \<in> tainted s'; no_del_event (s @ s'); valid (s @ s')\<rbrakk> 
+  \<Longrightarrow> obj \<in> tainted (s @ s')"
+apply (induct s, simp)
+apply (simp only:cons_app_simp_aux)
+apply (frule valid_cons, frule no_del_event_cons_D, simp)
+apply (rule t_remain, simp+)
+apply (drule tainted_is_current)
+apply (case_tac a, case_tac [!] obj, auto)
+done
+
+end
+
+context tainting_s_complete begin
+
+lemma unknown_notin_tainted_s':
+  "sobj \<in> tainted_s \<Longrightarrow> sobj \<noteq> Unknown" 
+apply (erule tainted_s.induct, auto) 
+apply (drule seeds_in_init) 
+apply (subgoal_tac "exists [] obj")
+apply (drule init_obj_has_sobj, simp+)
+apply (case_tac obj, simp+)
+done
+
+lemma unknown_notin_tainted_s:
+  "Unknown \<in> tainted_s \<Longrightarrow> False"
+by (auto dest:unknown_notin_tainted_s')
+
+end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/tainted_vs_tainted_s.thy	Fri Apr 12 10:43:11 2013 +0100
@@ -0,0 +1,1466 @@
+theory tainted_vs_tainted_s
+imports Main rc_theory os_rc deleted_prop tainted obj2sobj_prop source_prop all_sobj_prop 
+begin
+
+context tainting_s_complete begin
+
+lemma t2ts[rule_format]: 
+  "obj \<in> tainted s \<Longrightarrow> obj2sobj s obj \<in> tainted_s "  
+proof (induct rule:tainted.induct)
+  case (t_init obj) 
+  assume seed: "obj \<in> seeds" 
+  hence"exists [] obj" by (drule_tac seeds_in_init, case_tac obj, auto)
+  thus ?case using seed by (simp add:ts_init obj2sobj_nil_init)
+next
+  case (t_clone p s p') 
+  assume p1: "Proc p \<in> tainted s" 
+  and p2: "obj2sobj s (Proc p) \<in> tainted_s"
+  and p3: "valid (Clone p p' # s)"
+  from p3 have os: "os_grant s (Clone p p')" and rc: "rc_grant s (Clone p p')"
+    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
+  from os have exp: "exists s (Proc p)" by (simp add:os_grant.simps)
+  from exp obtain r fr pt u sp where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
+    and srp: "source_proc s p = Some sp" using vs
+    apply (simp del:cp2sproc.simps)
+    by (frule current_proc_has_sproc, simp, frule current_proc_has_srp, simp, blast)
+  with exp all_sobjs_I[where obj = "Proc p" and s = s] vs
+
+  have "obj2sobj (Clone p p' # s) (Proc p') = SProc (r,fr,clone_type_aux r pt, u) (Some sp)"
+    using sproc srp p3
+    by (simp add:obj2sobj.simps cp2sproc_clone del:cp2sproc.simps)
+  moreover have "(r, Proc_type pt, CREATE) \<in> compatible" using rc sproc
+    by (auto split:option.splits)
+  moreover have "SProc (r, fr, pt, u) (Some sp) \<in> tainted_s" using p2 sproc srp
+    by simp
+  ultimately show ?case by (auto intro:ts_clone) 
+next
+  case (t_exec f s p)
+  assume p1: "File f \<in> tainted s"
+  and p2: "obj2sobj s (File f) \<in> tainted_s"
+  and p3: "valid (Execute p f # s)"
+  from p3 have os: "os_grant s (Execute p f)" and rc: "rc_grant s (Execute p f)"
+    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
+  from os have exp: "exists s (Proc p)" and exf: "exists s (File f)" by auto 
+  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
+    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
+  with exp all_sobjs_I[where obj = "Proc p" and s = s] vs
+  have SP: "SProc (r,fr,pt,u) (source_proc s p) \<in> all_sobjs"
+    by (auto simp:obj2sobj.simps)
+  from exf obtain ft sd where etype: "etype_of_file s f = Some ft"
+    and srdir: "source_dir s f = Some sd" using vs
+    by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
+  with p2 obtain srf where TF: "SFile (ft, sd) srf \<in> tainted_s" 
+    by (auto simp:obj2sobj.simps split:if_splits)
+  from sproc srdir have "u \<in> init_users" and "sd \<in> init_files" using vs
+    by (auto simp:obj2sobj.simps cp2sproc.simps 
+            intro:source_dir_in_init owner_in_users 
+            split:option.splits)
+  then obtain nr nfr where nrole: "exec_role_aux r sd u = Some nr" 
+    and nfrole: "erole_functor init_file_forcedrole InheritUpMixed sd = Some nfr"
+    by (frule_tac r = r in exec_role_some, simp, frule_tac efffrole_sdir_some, auto)
+  hence "cp2sproc (Execute p f # s) p = Some (nr, nfr, exec_type_aux r pt, u)" 
+    using p3 srdir sproc by (simp add:cp2sproc_exec)
+  with nrole nfrole TF SP rc sproc etype
+  show ?case 
+    by (auto simp:obj2sobj.simps cp2sproc.simps intro!:ts_exec1 split:option.splits)
+next
+  case (t_cfile p s f)
+  assume p1: "Proc p \<in> tainted s" and p2: "obj2sobj s (Proc p) \<in> tainted_s"
+    and p3: "valid (CreateFile p f # s)"
+  from p3 have os: "os_grant s (CreateFile p f)" and rc: "rc_grant s (CreateFile p f)"
+    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
+  from os obtain pf where parent: "parent f = Some pf" and exp: "exists s (Proc p)"
+    and expf: "exists s (File pf)" by auto
+  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
+    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
+  with p2 have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
+    by (auto simp:obj2sobj.simps cp2sproc.simps)
+  from expf obtain pft sd where etype: "etype_of_file s pf = Some pft"
+    and srdir: "source_dir s pf = Some sd" using vs
+    by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
+  with expf all_sobjs_I[where obj = "File pf" and s = s] vs
+  obtain srpf where SF: "SFile (pft, sd) srpf \<in> all_sobjs"
+    by (auto simp:obj2sobj.simps cp2sproc.simps split:if_splits)
+  show ?case using etype srdir p3 parent os
+    apply (auto simp:source_dir_simps init_notin_curf_deleted obj2sobj.simps split:option.splits
+               dest!:current_file_has_etype')
+    apply (case_tac "default_fd_create_type r")
+    using SF TP rc sproc
+    apply (rule_tac sf = srpf in ts_cfd',
+           auto simp:etype_of_file_def etype_aux_prop3 obj2sobj.simps cp2sproc.simps
+               split:option.splits) [1]
+    using SF TP rc sproc
+    apply (rule_tac sf = srpf in ts_cfd)
+    apply (auto simp:etype_of_file_def etype_aux_prop4 cp2sproc.simps split:option.splits)
+    done
+next
+  case (t_cipc p s i)
+  assume p1: "Proc p \<in> tainted s" and p2: "obj2sobj s (Proc p) \<in> tainted_s"
+    and p3: "valid (CreateIPC p i # s)"
+  from p3 have os: "os_grant s (CreateIPC p i)" and rc: "rc_grant s (CreateIPC p i)"
+    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
+  from os have exp: "exists s (Proc p)" by simp
+  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
+    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
+  with p2 have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
+    by (auto simp:obj2sobj.simps cp2sproc.simps)
+  show ?case using p3 sproc os rc TP 
+    by (auto simp:ni_init_deled obj2sobj.simps cp2sproc.simps 
+            split:option.splits intro!:ts_cipc)
+next
+  case (t_read f s p)
+  assume p1: "File f \<in> tainted s" and p2: "obj2sobj s (File f) \<in> tainted_s"
+    and p3: "valid (ReadFile p f # s)"
+  from p3 have os: "os_grant s (ReadFile p f)" and rc: "rc_grant s (ReadFile p f)"
+    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
+  from os have exp: "exists s (Proc p)" and exf: "exists s (File f)" by auto 
+  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
+    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
+  with exp all_sobjs_I[where obj = "Proc p" and s = s] vs
+  have SP: "SProc (r,fr,pt,u) (source_proc s p) \<in> all_sobjs"
+    by (auto simp:obj2sobj.simps)
+  from exf obtain ft sd where etype: "etype_of_file s f = Some ft"
+    and srdir: "source_dir s f = Some sd" using vs
+    by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
+  with p2 obtain srf where TF: "SFile (ft, sd) srf \<in> tainted_s"
+    by (auto simp:obj2sobj.simps cp2sproc.simps split:if_splits)
+  show ?case using sproc SP TF rc etype 
+    by (auto simp:obj2sobj.simps cp2sproc.simps 
+            split:option.splits intro!:ts_read) 
+next
+  case (t_write p s f)
+  assume p1: "Proc p \<in> tainted s" and p2: "obj2sobj s (Proc p) \<in> tainted_s"
+    and p3: "valid (WriteFile p f # s)"
+  from p3 have os: "os_grant s (WriteFile p f)" and rc: "rc_grant s (WriteFile p f)"
+    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
+  from os have exp: "exists s (Proc p)" and exf: "exists s (File f)" by auto 
+  from exf obtain ft sd where etype: "etype_of_file s f = Some ft"
+    and srdir: "source_dir s f = Some sd" using vs
+    by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
+  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
+    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
+  with exp p2 have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
+    by (auto simp:obj2sobj.simps)
+  from etype p3 have etype':"etype_of_file (WriteFile p f # s) f = Some ft"
+    by (case_tac f, auto simp:etype_of_file_def)
+  have SF: "obj2sobj s (File f) \<in> all_sobjs" using exf vs 
+    by (rule_tac all_sobjs_I, simp+) 
+  show ?case using sproc TP rc etype p3 srdir etype' SF
+    by (auto simp:source_dir_simps obj2sobj.simps cp2sproc.simps
+            split:option.splits intro!:ts_write) 
+next
+  case (t_send p s i)
+  assume p1: "Proc p \<in> tainted s" and p2: "obj2sobj s (Proc p) \<in> tainted_s"
+    and p3: "valid (Send p i # s)"
+  from p3 have os: "os_grant s (Send p i)" and rc: "rc_grant s (Send p i)"
+    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
+  from os have exp: "exists s (Proc p)" and exi: "exists s (IPC i)" by auto 
+  from exi obtain t where etype: "type_of_ipc s i = Some t" using vs
+    by (simp, drule_tac current_ipc_has_type, auto)
+  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
+    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
+  with exp p2 have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
+    by (auto simp:obj2sobj.simps cp2sproc.simps)
+  have SI: "obj2sobj s (IPC i) \<in> all_sobjs" using exi vs
+    by (simp add:all_sobjs_I del:obj2sobj.simps)
+  show ?case using sproc TP rc etype p3 vs SI
+    by (auto simp:obj2sobj.simps cp2sproc.simps 
+            split:option.splits intro!:ts_send)
+next
+  case (t_recv i s p)
+  assume p1: "IPC i \<in> tainted s" and p2: "obj2sobj s (IPC i) \<in> tainted_s"
+    and p3: "valid (Recv p i # s)"
+  from p3 have os: "os_grant s (Recv p i)" and rc: "rc_grant s (Recv p i)"
+    and vs: "valid s" by (auto dest:valid_cons valid_os valid_rc)
+  from os have exp: "exists s (Proc p)" and exi: "exists s (IPC i)" by auto 
+  from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
+    using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
+  with exp all_sobjs_I[where obj = "Proc p" and s = s] vs
+  have SP: "SProc (r,fr,pt,u) (source_proc s p) \<in> all_sobjs"
+    by (auto simp:obj2sobj.simps)
+  from exi obtain t where etype: "type_of_ipc s i = Some t"using vs
+    by (simp, drule_tac current_ipc_has_type, auto)
+  with p2 obtain sri where TI: "SIPC t sri \<in> tainted_s"
+    by (auto simp:obj2sobj.simps split:if_splits)
+  show ?case using sproc SP TI rc etype 
+    by (auto simp:obj2sobj.simps cp2sproc.simps 
+            split:option.splits intro!:ts_recv) 
+next
+  case (t_remain obj s e)
+  from t_remain(1) have p5: "exists s obj" by (rule tainted_is_current)
+  from t_remain(3) have os: "os_grant s e" and vs: "valid s" and rc: "rc_grant s e"
+    by (auto dest:valid_os valid_cons valid_rc)
+  show ?case
+  proof (cases obj)
+    case (File f)
+    have "obj2sobj (e # s) (File f) = obj2sobj s (File f)"
+    proof-
+      have "etype_of_file (e # s) f = etype_of_file s f"
+        using p5 os vs File t_remain(3,4)
+        apply (case_tac e, auto simp:etype_of_file_def split:option.splits)
+        by (auto dest:ancient_file_in_current intro!:etype_aux_prop)
+      moreover have "source_dir (e # s) f = source_dir s f"
+        using p5 os vs File t_remain(3,4)
+        by (case_tac e, auto simp:source_dir_simps dest:source_dir_prop)
+      ultimately show ?thesis using vs t_remain(4) File
+        apply (auto simp:obj2sobj.simps cp2sproc.simps 
+                   split:if_splits option.splits dest:not_deleted_cons_D)
+        by (case_tac e, auto)
+    qed
+    with File t_remain(2) show ?thesis by simp
+  next
+    case (IPC i) 
+    have "obj2sobj (e # s) (IPC i) = obj2sobj s (IPC i)" using p5 t_remain(3,4) os IPC
+      by (case_tac e, auto simp:ni_init_deled ni_notin_curi obj2sobj.simps 
+                          split:option.splits
+                          dest!:current_proc_has_role')
+    with IPC t_remain(2) show ?thesis by simp
+  next
+    case (Proc p) 
+    show ?thesis 
+    proof-
+      have "\<And> f. e = Execute p f \<Longrightarrow> obj2sobj (e # s) (Proc p) \<in> tainted_s"
+      proof-
+        fix f
+        assume ev: "e = Execute p f"
+        show "obj2sobj (e # s) (Proc p) \<in> tainted_s"
+        proof-
+          from os ev have exp: "exists s (Proc p)" and exf: "exists s (File f)" by auto 
+          from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
+            using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
+          with Proc t_remain(2) 
+          have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
+            by (auto simp:obj2sobj.simps cp2sproc.simps)
+          from exf obtain ft sd where etype: "etype_of_file s f = Some ft"
+            and srdir: "source_dir s f = Some sd" using vs
+            by (simp, drule_tac current_file_has_sd, auto dest:current_file_has_etype)
+          with exf all_sobjs_I[where obj = "File f" and s = s] vs
+          obtain srf where SF: "SFile (ft, sd) srf \<in> all_sobjs"
+            by (auto simp:obj2sobj.simps split:if_splits) 
+          from sproc srdir have "u \<in> init_users" and "sd \<in> init_files" using vs
+            by (auto simp:obj2sobj.simps cp2sproc.simps
+                    intro:source_dir_in_init owner_in_users split:option.splits)
+          then obtain nr nfr where nrole: "exec_role_aux r sd u = Some nr" 
+            and nfrole: "erole_functor init_file_forcedrole InheritUpMixed sd = Some nfr"
+            by (frule_tac r = r in exec_role_some, simp, frule_tac efffrole_sdir_some, auto)
+          hence "cp2sproc (Execute p f # s) p = Some (nr, nfr, exec_type_aux r pt, u)" 
+            using t_remain(3) srdir sproc ev by (simp add:cp2sproc_exec)
+          with nrole nfrole SF TP rc sproc etype ev
+          show ?thesis
+            by (auto simp:obj2sobj.simps cp2sproc.simps 
+                   intro!:ts_exec2 split:option.splits)
+        qed
+      qed  moreover 
+      have "\<And> r'. e = ChangeRole p r' \<Longrightarrow> obj2sobj (e # s) (Proc p) \<in> tainted_s"
+      proof-
+        fix r' 
+        assume ev: "e = ChangeRole p r'"
+        show "obj2sobj (e # s) (Proc p) \<in> tainted_s"
+        proof-
+          from os ev have exp: "exists s (Proc p)" by auto 
+          from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
+            using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
+          with Proc t_remain(2)
+          have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
+            by (auto simp:obj2sobj.simps cp2sproc.simps)
+          with rc sproc ev show ?thesis 
+            apply (simp add:obj2sobj.simps cp2sproc.simps split:option.splits)
+            by (rule_tac ts_crole, simp+)
+        qed
+      qed  moreover 
+      have "\<And> u'. e = ChangeOwner p u' \<Longrightarrow> obj2sobj (e # s) (Proc p) \<in> tainted_s"
+      proof-
+        fix u'
+        assume ev: "e = ChangeOwner p u'"
+        show "obj2sobj (e # s) (Proc p) \<in> tainted_s"
+        proof-        
+          from os ev have exp: "exists s (Proc p)" by auto 
+          from exp obtain r fr pt u where sproc: "cp2sproc s p = Some (r,fr,pt,u)"
+            using vs by (auto simp del:cp2sproc.simps dest!:current_proc_has_sproc)
+          with Proc t_remain(2)
+          have TP: "SProc (r,fr,pt,u) (source_proc s p) \<in> tainted_s"
+            by (auto simp:obj2sobj.simps cp2sproc.simps)
+          from os ev have uinit: "u' \<in> init_users" by simp
+          then obtain nr where chown: "chown_role_aux r fr u' = Some nr" 
+            by (auto dest:chown_role_some)
+          hence nsproc:"cp2sproc (e#s) p = Some (nr,fr,chown_type_aux r nr pt, u')" 
+            using sproc ev os
+            by (auto split:option.splits t_role.splits 
+                  simp del:currentrole.simps type_of_process.simps
+                  simp add:obj2sobj.simps cp2sproc.simps 
+                           chown_role_aux_valid chown_type_aux_valid)
+          with rc sproc ev TP uinit chown
+          show ?thesis
+            by (auto simp:obj2sobj.simps cp2sproc.simps 
+                   intro!:ts_chown split:option.splits)
+        qed
+      qed  moreover 
+      have "\<lbrakk>\<forall> f. e \<noteq> Execute p f; \<forall> r. e \<noteq> ChangeRole p r; \<forall> u. e \<noteq> ChangeOwner p u\<rbrakk>
+        \<Longrightarrow> obj2sobj (e # s) (Proc p) \<in> tainted_s"
+        using t_remain(2,3,4) os p5 Proc
+        by (case_tac e, auto simp add:obj2sobj.simps cp2sproc_simps np_notin_curp 
+                             simp del:cp2sproc.simps  split:option.splits) 
+      ultimately show ?thesis using Proc
+        by (case_tac e, auto simp del:obj2sobj.simps) 
+    qed
+  qed
+qed
+
+end
+
+context tainting_s_sound begin
+
+lemma tainted_s'_eq1: "sobj \<in> tainted_s \<Longrightarrow> sobj \<in> tainted_s'"
+apply (erule tainted_s.induct)
+apply (auto elim:ts'_init ts'_exec1 ts'_exec2 ts'_cfd ts'_cfd' ts'_cipc ts'_read ts'_write ts'_recv ts'_send ts'_crole ts'_chown simp:all_sobjs'_eq)
+by (simp add:clone_type_aux_def clone_type_unchange)
+
+lemma tainted_s'_eq2: "sobj \<in> tainted_s' \<Longrightarrow> sobj \<in> tainted_s"
+apply (erule tainted_s'.induct)
+by (auto intro:ts_init ts_exec1 ts_exec2 ts_cfd ts_cfd' ts_cipc ts_read ts_write ts_recv ts_send ts_crole ts_chown ts_clone simp:all_sobjs'_eq)
+
+lemma tainted_s'_eq: "(sobj \<in> tainted_s) = (sobj \<in> tainted_s')"
+by (auto intro:iffI tainted_s'_eq1 tainted_s'_eq2)
+
+(* cause sobj_source_eq_sobj may conflict with initp_intact, so remove it too. *) 
+lemma ts2t_intact:
+  "sobj \<in> tainted_s' \<Longrightarrow> \<exists> obj s. obj2sobj s obj = sobj \<and> obj \<in> tainted s \<and> 
+                                 no_del_event s \<and> initp_intact s"
+
+proof (induct rule:tainted_s'.induct)
+  case (ts'_init obj)
+  hence ex: "exists [] obj"
+    apply (drule_tac seeds_in_init) 
+    by (case_tac obj, simp+)
+  have "obj2sobj [] obj = init_obj2sobj obj" using ex 
+    by (simp add:obj2sobj_nil_init)
+  moreover have "obj \<in> tainted []" using ts'_init by (simp add:t_init)
+  moreover have "initp_intact []"
+    by (auto simp:initp_intact_def obj2sobj.simps cp2sproc.simps split:option.splits)
+  ultimately show ?case 
+    by (rule_tac x = obj in exI, rule_tac x = "[]" in exI, simp+)
+next
+  case (ts'_exec1 t sd srf r fr pt u srp r' fr')
+  then obtain f s  where TF: "(File f) \<in> tainted s" and vds: "valid s" 
+    and "f \<in> current_files s \<and> obj2sobj s (File f) = SFile (t, sd) srf \<and> 
+         no_del_event s \<and> initp_intact s" 
+    apply (clarsimp, frule_tac obj2sobj_file)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_exec1(3) obtain p s' where vds's: "valid (s' @ s)"
+    and etype: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)" 
+    and ISP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "f \<in> current_files (s' @ s)"
+    apply (drule_tac s' =s and obj' = "File f" in all_sobjs_E0, auto)
+    apply (frule_tac obj = "File f" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_proc, clarsimp)
+  obtain ev \<tau> where ev: "ev = Execute (new_proc (s'@s)) f"
+    and tau: "\<tau> = Clone p (new_proc (s'@s)) # (s' @ s)" by auto
+  hence vs_tau:"valid \<tau>" using exp vds's by (auto intro:clone_event_no_limit)
+
+  have valid: "valid (ev # \<tau>)"
+  proof-
+    have "os_grant \<tau> ev"
+      using ev tau by (simp add:exf exp)
+    moreover have "rc_grant \<tau> ev" 
+      using ev tau ts'_exec1 ISP etype
+      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis using vs_tau
+      by (erule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (ev # \<tau>) (Proc (new_proc (s'@s))) = 
+        SProc (r', fr', exec_type_aux r pt, u) srp"
+  proof-
+    have "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp" using ISP tau vs_tau
+      by (auto split:option.splits simp:cp2sproc_simps obj2sobj.simps)
+    hence "obj2sobj \<tau> (Proc (new_proc (s'@s))) = SProc (r,fr,pt,u) srp" using tau
+      by (auto simp:obj2sobj.simps cp2sproc.simps pct_def clone_type_unchange 
+              split:option.splits)
+    moreover have "source_dir \<tau> f = Some sd" using vs_tau etype tau 
+      by (auto simp:obj2sobj.simps source_dir_simps split:option.splits if_splits)
+    ultimately show ?thesis using valid ts'_exec1(5) ts'_exec1(6) ev
+      by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
+  qed  moreover
+  have "Proc (new_proc (s'@s)) \<in> tainted (ev # \<tau>)"
+  proof-
+    have "(File f) \<in> tainted \<tau>" using TF nodels' tau vs_tau
+      by (drule_tac s = "Clone p (new_proc (s' @ s)) # s'" in t_remain_app,auto)
+    thus ?thesis using ev valid
+      by (drule_tac t_exec, simp+)
+  qed  moreover
+  have "no_del_event (ev # \<tau>)" using ev tau nodels' by simp 
+  moreover have "initp_intact (ev#\<tau>)" using ev tau intacts' valid
+    by (simp add:initp_intact_I_exec)
+  ultimately show ?case
+    apply (rule_tac x = "Proc (new_proc (s'@s))" in exI)
+    by (rule_tac x = "ev#\<tau>" in exI, auto)
+next
+  case (ts'_exec2 r fr pt u srp t sd srf r' fr')
+  then obtain p s where TP: "(Proc p) \<in> tainted s" and vds: "valid s" 
+    and "p \<in> current_procs s \<and> obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp \<and> 
+         no_del_event s \<and> initp_intact s" 
+    apply (clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_exec2(3) obtain f s' where vds's: "valid (s' @ s)"
+    and etype: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)" 
+    and ISP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "f \<in> current_files (s' @ s)"
+    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
+    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_file, clarsimp)
+  obtain ev \<tau> where ev: "ev = Execute (new_proc (s'@s)) f"
+    and tau: "\<tau> = Clone p (new_proc (s'@s)) # (s' @ s)" by auto
+  hence vs_tau:"valid \<tau>" using exp vds's by (auto intro:clone_event_no_limit)
+
+  have valid: "valid (ev # \<tau>)"
+  proof-
+    have "os_grant \<tau> ev"
+      using ev tau by (simp add:exf exp)
+    moreover have "rc_grant \<tau> ev" 
+      using ev tau ts'_exec2(4) ISP etype
+      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps 
+              split:if_splits option.splits)
+    ultimately show ?thesis using vs_tau
+      by (erule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (ev # \<tau>) (Proc (new_proc (s'@s))) = 
+        SProc (r', fr', exec_type_aux r pt, u) srp"
+  proof-
+    have "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp" using ISP tau vs_tau
+      by (auto split:option.splits simp:cp2sproc_simps obj2sobj.simps)
+    hence "obj2sobj \<tau> (Proc (new_proc (s'@s))) = SProc (r,fr,pt,u) srp" using tau
+      by (auto simp:obj2sobj.simps cp2sproc.simps pct_def clone_type_unchange 
+              split:option.splits)
+    moreover have "source_dir \<tau> f = Some sd" using vs_tau etype tau 
+      by (auto simp:source_dir_simps obj2sobj.simps split:option.splits if_splits)
+    ultimately show ?thesis using valid ts'_exec2(5) ts'_exec2(6) ev
+      by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
+  qed  moreover
+  have "Proc (new_proc (s'@s)) \<in> tainted (ev # \<tau>)"
+  proof-
+    have "Proc p \<in> tainted (s' @ s)" using TP vds's nodels'
+      by (drule_tac s = s' in t_remain_app, auto)
+    hence "Proc (new_proc (s'@s)) \<in> tainted \<tau>" using TP tau vs_tau
+      by (auto intro:t_clone)
+    thus ?thesis using ev valid
+      by (drule_tac t_remain, auto dest:valid_os)
+  qed  moreover
+  have "no_del_event (ev # \<tau>)" using ev tau nodels' by simp 
+  moreover have "initp_intact (ev#\<tau>)" using ev tau intacts' valid
+    by (simp add:initp_intact_I_exec)
+  ultimately show ?case
+    by (rule_tac x = "Proc (new_proc (s'@s))" in exI, rule_tac x = "ev#\<tau>" in exI, auto)
+next
+  case (ts'_cfd r fr pt u srp t sd srf t')
+  then obtain p s where TP: "(Proc p) \<in> tainted s" 
+    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
+         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
+    apply (clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_cfd(3) obtain pf s' where vds': "valid (s' @ s)"
+    and SF: "obj2sobj (s'@s) (File pf) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "pf \<in> current_files (s' @ s)"
+    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
+    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_file, clarsimp)
+  obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)"
+    and tau: "\<tau> = (s' @ s)" by auto
+
+  have valid: "valid (e# \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp exf by (simp add:ncf_notin_curf ncf_parent)
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_cfd(4,5,6) SP SF
+      by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
+        split:if_splits option.splits t_rc_file_type.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
+  have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t', sd) None"
+  proof-
+    have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t'"
+      using ev tau SF SP ts'_cfd(4)
+      by (auto simp:obj2sobj.simps etype_of_file_def cp2sproc.simps
+              split:option.splits if_splits  intro!:etype_aux_prop4)
+    moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
+      using ev tau SF SP valid ncf_parent
+      by (auto simp:source_dir_simps obj2sobj.simps cp2sproc.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
+      nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
+      by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
+  qed  moreover
+  have "initp_intact (e#\<tau>)" using ev tau intacts' valid nodels'
+    by (auto intro!:initp_intact_I_others)  moreover
+  have "File (new_childf pf \<tau>) \<in> tainted (e#\<tau>)" 
+  proof- 
+    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
+      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
+    thus ?thesis using ev tau valid by (auto intro:t_cfile)
+  qed
+  ultimately show ?case using tau ev
+    apply (rule_tac x = "File (new_childf pf \<tau>)" in exI)
+    by (rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_cfd' r fr pt u srp t sd srf)
+  then obtain p s where TP: "(Proc p) \<in> tainted s" 
+    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
+         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
+    apply (clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_cfd'(3) obtain pf s' where vds': "valid (s' @ s)"
+    and SF: "obj2sobj (s'@s) (File pf) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "pf \<in> current_files (s' @ s)"
+    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
+    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_file, clarsimp)
+  obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)"
+    and tau: "\<tau> = (s' @ s)" by auto
+
+  have valid: "valid (e# \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp exf by (simp add:ncf_notin_curf ncf_parent)
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_cfd'(4,5) SP SF
+      by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
+        split:if_splits option.splits t_rc_file_type.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
+  have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t, sd) None"
+  proof-
+      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t"
+      proof-
+        have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = etype_of_file \<tau> pf"
+          using ev tau SP ts'_cfd'(4)
+          by (auto simp:obj2sobj.simps ncf_parent etype_of_file_def cp2sproc.simps
+            split:option.splits   intro!:etype_aux_prop3)
+        thus ?thesis using SF tau ev
+          by (auto simp:obj2sobj.simps split:option.splits if_splits)
+      qed
+    moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
+      using ev tau SF SP valid ncf_parent
+      by (auto simp:source_dir_simps obj2sobj.simps cp2sproc.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
+      nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
+      by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
+  qed  moreover
+  have "initp_intact (e#\<tau>)" using ev tau intacts' valid nodels'
+    by (auto intro!:initp_intact_I_others)  moreover
+  have "File (new_childf pf \<tau>) \<in> tainted (e#\<tau>)" 
+  proof- 
+    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
+      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
+    thus ?thesis using ev tau valid by (auto intro:t_cfile)
+  qed
+  ultimately show ?case using tau ev
+    apply (rule_tac x = "File (new_childf pf \<tau>)" in exI)
+    by (rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_cipc r fr pt u srp)
+  then obtain p s where TP: "(Proc p) \<in> tainted s" 
+    and vds: "valid s" and exp: "p \<in> current_procs s" 
+    and nodels: "no_del_event s" and intacts: "initp_intact s"
+    and SP: "obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
+    apply (clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  obtain e where ev: "e = CreateIPC p (new_ipc s)" by simp
+      
+  have valid: "valid (e # s)" 
+  proof-
+    have "os_grant s e"
+      using ev exp by (simp)
+    moreover have "rc_grant s e" 
+      using ev ts'_cipc(3) SP
+      by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
+    ultimately show ?thesis using vds
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have nodel: "no_del_event (e#s)" using nodels ev by simp  moreover
+  have "initp_intact (e#s)" using ev intacts valid nodels
+    by (auto intro!:initp_intact_I_others)  moreover
+  have "IPC (new_ipc s) \<in> tainted (e#s)" using TP ev valid
+    by (auto intro:t_cipc)  moreover
+  have "obj2sobj (e#s) (IPC (new_ipc s)) = SIPC (default_ipc_create_type r) None"
+    using ev SP nodel nodel_imp_exists[where obj = "IPC (new_ipc s)" and s=s]
+    by (auto simp:obj2sobj.simps ni_notin_curi cp2sproc.simps
+            split:option.splits dest:no_del_event_cons_D)
+  ultimately show ?case using ev
+    apply (rule_tac x = "IPC (new_ipc s)" in exI)
+    by (rule_tac x = "e # s" in exI, auto)
+next
+  case (ts'_read t sd srf r fr pt u srp)
+  then obtain f s  where TF: "(File f) \<in> tainted s" and vds: "valid s" 
+    and "f \<in> current_files s \<and> obj2sobj s (File f) = SFile (t, sd) srf \<and> 
+         no_del_event s \<and> initp_intact s" 
+    apply (clarsimp, frule_tac obj2sobj_file)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_read(3) obtain p s' where vds': "valid (s' @ s)"
+    and SF: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "f \<in> current_files (s' @ s)"
+    apply (drule_tac s' =s and obj' = "File f" in all_sobjs_E0, auto)
+    apply (frule_tac obj = "File f" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_proc, clarsimp)
+  obtain e \<tau> where ev: "e = ReadFile p f" and tau: "\<tau> = s' @ s" by auto
+
+  have valid: "valid (e # \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau by (simp add:exf exp)
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_read(4) SP SF
+      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (e # \<tau>) (Proc p) = SProc (r, fr, pt, u) srp" using valid tau ev SP
+    by (auto simp:obj2sobj.simps cp2sproc_simps split:option.splits)
+  moreover  have "Proc p \<in> tainted (e # \<tau>)"
+  proof-
+    have "(File f) \<in> tainted \<tau>" using TF nodels' tau vds'
+      by (drule_tac s = "s'" in t_remain_app,auto)
+    thus ?thesis using ev valid
+      by (drule_tac t_read, simp+)
+  qed  moreover
+  have "no_del_event (e # \<tau>)" using ev tau nodels' by simp  moreover
+  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
+    by (auto intro!:initp_intact_I_others) 
+  ultimately show ?case
+    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_write r fr pt u srp t sd srf)
+  then obtain p s where TP: "(Proc p) \<in> tainted s" 
+    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
+         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
+    apply (clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_write(3) obtain f s' where vds': "valid (s' @ s)"
+    and SF: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "f \<in> current_files (s' @ s)"
+    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
+    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_file, clarsimp)
+  obtain e \<tau> where ev: "e = WriteFile p f" and tau: "\<tau> = (s' @ s)" by auto
+
+  have valid: "valid (e# \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp exf by simp
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_write(4) SP SF
+      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
+  have "obj2sobj (e#\<tau>) (File f) = SFile (t, sd) srf"
+    using tau ev SF valid
+    by (auto simp:obj2sobj.simps source_dir_simps etype_of_file_def
+            split:option.splits if_splits)
+  moreover  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
+    by (auto intro!:initp_intact_I_others)   moreover
+  have "File f \<in> tainted (e#\<tau>)" 
+  proof- 
+    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
+      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
+    thus ?thesis using ev tau valid by (auto intro:t_write)
+  qed
+  ultimately show ?case using tau ev
+    by (rule_tac x = "File f" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+next 
+  case (ts'_recv t sri r fr pt u srp)
+  then obtain i s  where TI: "(IPC i) \<in> tainted s" and vds: "valid s" 
+    and "i \<in> current_ipcs s \<and> obj2sobj s (IPC i) = SIPC t sri \<and> 
+         no_del_event s \<and> initp_intact s" 
+    apply (clarsimp, frule_tac obj2sobj_ipc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_recv(3) obtain p s' where vds': "valid (s' @ s)"
+    and SI: "obj2sobj (s'@s) (IPC i) = SIPC t sri"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exi: "i \<in> current_ipcs (s' @ s)"
+    apply (drule_tac s' =s and obj' = "IPC i" in all_sobjs_E0, auto)
+    apply (frule_tac obj = "IPC i" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_proc, clarsimp)
+  obtain e \<tau> where ev: "e = Recv p i" and tau: "\<tau> = s' @ s" by auto
+
+  have valid: "valid (e # \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau by (simp add:exi exp)
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_recv(4) SP SI
+      by (auto simp:cp2sproc.simps obj2sobj.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (e # \<tau>) (Proc p) = SProc (r, fr, pt, u) srp" using valid tau ev SP
+    by (auto simp:obj2sobj.simps cp2sproc_simps split:option.splits)
+  moreover  have "Proc p \<in> tainted (e # \<tau>)"
+  proof-
+    have "(IPC i) \<in> tainted \<tau>" using TI nodels' tau vds'
+      by (drule_tac s = "s'" in t_remain_app,auto)
+    thus ?thesis using ev valid
+      by (drule_tac t_recv, simp+)
+  qed  moreover
+  have "no_del_event (e # \<tau>)" using ev tau nodels' by simp 
+  moreover have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
+    by (auto intro!:initp_intact_I_others)   
+  ultimately show ?case
+    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_send r fr pt u srp t sri)
+  then obtain p s where TP: "(Proc p) \<in> tainted s" 
+    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
+         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
+    apply (clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_send(3) obtain i s' where vds': "valid (s' @ s)"
+    and SI: "obj2sobj (s'@s) (IPC i) = SIPC t sri"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exi: "i \<in> current_ipcs (s' @ s)"
+    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
+    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_ipc, clarsimp)
+  obtain e \<tau> where ev: "e = Send p i" and tau: "\<tau> = (s' @ s)" by auto
+
+  have valid: "valid (e# \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp exi by simp
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_send(4) SP SI
+      by (auto simp:cp2sproc.simps obj2sobj.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
+  have "obj2sobj (e#\<tau>) (IPC i) = SIPC t sri"
+    using tau ev SI valid
+    by (auto simp:obj2sobj.simps split:option.splits if_splits)
+  moreover have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
+    by (auto intro!:initp_intact_I_others)  moreover
+  have "IPC i \<in> tainted (e#\<tau>)" 
+  proof- 
+    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
+      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
+    thus ?thesis using ev tau valid by (auto intro:t_send)
+  qed
+  ultimately show ?case using tau ev 
+    by (rule_tac x = "IPC i" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_crole r fr pt u srp r')
+  then obtain p s where exp: "p \<in> current_procs s"
+    and TP: "(Proc p) \<in> tainted s" and vds: "valid s" and nodels: "no_del_event s"
+    and SP: "obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp"
+    and intacts: "initp_intact s" 
+    apply (clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  obtain e \<tau> where ev: "e = ChangeRole (new_proc s) r'"
+    and tau: "\<tau> = Clone p (new_proc s) # s" by auto
+  hence vs_tau:"valid \<tau>" using exp vds by (auto intro:clone_event_no_limit)
+
+  have valid: "valid (e # \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp by simp
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_crole(3) SP
+      by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
+    ultimately show ?thesis using vs_tau
+      by (erule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (e # \<tau>) (Proc (new_proc s)) = SProc (r',fr,pt,u) srp"
+  proof-
+    have "obj2sobj \<tau> (Proc (new_proc s)) = SProc (r,fr,pt,u) srp" using SP tau vs_tau
+      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
+    thus ?thesis using valid ev
+      by (auto simp:cp2sproc_crole obj2sobj.simps split:option.splits)
+  qed  moreover
+  have "Proc (new_proc s) \<in> tainted (e # \<tau>)"
+  proof-
+    have "(Proc (new_proc s)) \<in> tainted \<tau>" using TP tau vs_tau
+      by (auto intro!:t_clone)
+    thus ?thesis using ev valid
+      by (drule_tac t_remain, auto dest:valid_os)
+  qed  moreover
+  have "no_del_event (e # \<tau>)" using ev tau nodels by simp 
+  moreover have "initp_intact (e#\<tau>)" using ev intacts valid nodels tau
+    by (simp add:initp_intact_I_crole)  
+  ultimately show ?case
+    by (rule_tac x = "Proc (new_proc s)" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_chown r fr pt u srp u' nr)
+  then obtain p s where exp: "p \<in> current_procs s"
+    and TP: "(Proc p) \<in> tainted s" and vds: "valid s" and nodels: "no_del_event s"
+    and SP: "obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp"
+    and intacts: "initp_intact s" 
+    apply (clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  obtain e \<tau> where ev: "e = ChangeOwner (new_proc s) u'"
+    and tau: "\<tau> = Clone p (new_proc s) # s" by auto
+  hence vs_tau:"valid \<tau>" using exp vds by (auto intro:clone_event_no_limit)
+
+  have valid: "valid (e # \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp ts'_chown(3) by simp
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_chown(5) SP
+      by (auto simp:cp2sproc.simps obj2sobj.simps pct_def clone_type_unchange
+              split:option.splits t_rc_proc_type.splits)
+        (* here is another place of no_limit of clone event assumption *)
+    ultimately show ?thesis using vs_tau
+      by (erule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (e # \<tau>) (Proc (new_proc s)) = SProc (nr, fr, chown_type_aux r nr pt, u') srp"
+  proof-
+    have "obj2sobj \<tau> (Proc (new_proc s)) = SProc (r,fr,pt,u) srp" using SP tau vs_tau
+      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
+    thus ?thesis using valid ev ts'_chown(4)
+      by (auto simp:cp2sproc_chown obj2sobj.simps split:option.splits)
+  qed  moreover
+  have "Proc (new_proc s) \<in> tainted (e # \<tau>)"
+  proof-
+    have "Proc (new_proc s) \<in> tainted \<tau>" using TP tau vs_tau exp
+      by (auto intro!:t_clone)
+    thus ?thesis using ev valid
+      by (drule_tac t_remain, auto dest:valid_os)
+  qed  moreover
+  have "no_del_event (e # \<tau>)" using ev tau nodels by simp 
+  moreover have "initp_intact (e#\<tau>)"  using ev intacts valid nodels tau
+    by (simp add:initp_intact_I_chown)
+  ultimately show ?case
+    by (rule_tac x = "Proc (new_proc s)" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+qed
+
+lemma ts2t:
+  "sobj \<in> tainted_s' \<Longrightarrow> \<exists> obj s. obj2sobj s obj = sobj \<and> obj \<in> tainted s \<and> 
+                                 sobj_source_eq_obj sobj obj \<and> no_del_event s \<and> 
+                                 initp_intact_but s sobj"
+proof (induct rule:tainted_s'.induct)
+  case (ts'_init obj)
+  hence ex: "exists [] obj"
+    apply (drule_tac seeds_in_init) 
+    by (case_tac obj, simp+)
+  have "obj2sobj [] obj = init_obj2sobj obj" using ex 
+    by (simp add:obj2sobj_nil_init)
+  moreover have "obj \<in> tainted []" using ts'_init by (simp add:t_init)
+  moreover have "sobj_source_eq_obj (init_obj2sobj obj) obj" using ex
+    apply (frule_tac init_obj_has_sobj)
+    apply (case_tac "init_obj2sobj obj", case_tac[!] obj)
+    by (auto split:option.splits)
+  ultimately show ?case 
+    apply (rule_tac x = obj in exI, rule_tac x = "[]" in exI)
+    by (auto simp:initp_intact_but_nil)
+next
+  case (ts'_exec1 t sd srf r fr pt u srp r' fr')
+  then obtain f s  where "sobj_source_eq_obj (SFile (t, sd) srf) (File f)"
+    and TF: "(File f) \<in> tainted s" and vds: "valid s" and "f \<in> current_files s"
+    and "obj2sobj s (File f) = SFile (t, sd) srf \<and> no_del_event s \<and> initp_intact s" 
+    apply (clarsimp, frule_tac obj2sobj_file)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_exec1(3) obtain p s' where vds's: "valid (s' @ s)"
+    and etype: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact_but (s'@s) (SProc (r,fr,pt,u) srp)" 
+    and ISP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "f \<in> current_files (s' @ s)"
+    and sreq: "sobj_source_eq_obj (SProc (r,fr,pt,u) srp) (Proc p)"
+    apply (drule_tac s' =s and obj' = "File f" in all_sobjs_E2, auto)
+    apply (frule_tac obj = "File f" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_proc, clarsimp)
+  from vds's ISP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
+  with exp vds's ISP have initp: "p \<in> init_processes" 
+    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
+  obtain ev \<tau> where ev: "ev = Execute p f"
+    and tau: "\<tau> = Clone p (new_proc (s'@s)) # (s' @ s)" by auto
+  hence vs_tau:"valid \<tau>" using exp vds's by (auto intro:clone_event_no_limit)
+
+  have valid: "valid (ev # \<tau>)"
+  proof-
+    have "os_grant \<tau> ev"
+      using ev tau by (simp add:exf exp)
+    moreover have "rc_grant \<tau> ev" 
+      using ev tau ts'_exec1 ISP etype
+      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis using vs_tau
+      by (erule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (ev # \<tau>) (Proc p) = SProc (r', fr', exec_type_aux r pt, u) srp"
+  proof-
+    have "obj2sobj \<tau> (Proc p) = SProc (r,fr,pt,u) srp" using ISP tau vs_tau
+      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
+    moreover have "source_dir \<tau> f = Some sd" using vs_tau etype tau 
+      by (auto simp:obj2sobj.simps source_dir_simps split:option.splits if_splits)
+    ultimately show ?thesis using valid ts'_exec1(5) ts'_exec1(6) ev
+      by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
+  qed  moreover
+  have "Proc p \<in> tainted (ev # \<tau>)"
+  proof-
+    have "(File f) \<in> tainted \<tau>" using TF nodels' tau vs_tau
+      by (drule_tac s = "Clone p (new_proc (s' @ s)) # s'" in t_remain_app,auto)
+    thus ?thesis using ev valid
+      by (drule_tac t_exec, simp+)
+  qed  moreover
+  have "no_del_event (ev # \<tau>)" using ev tau nodels' by simp 
+  moreover have "initp_intact_but (ev#\<tau>) (SProc (r', fr', exec_type_aux r pt, u) srp)"
+    using ev tau nodels' intacts' srpeq valid initp
+    by (simp, rule_tac initp_intact_butp_I_exec, simp_all)
+  moreover have "sobj_source_eq_obj (SProc (r', fr', exec_type_aux r pt, u) srp) (Proc p)"
+    using sreq by (case_tac srp, simp+)
+  ultimately show ?case
+    by (rule_tac x = "Proc p" in exI, rule_tac x = "ev#\<tau>" in exI, auto)
+next
+  case (ts'_exec2 r fr pt u srp t sd srf r' fr')
+  then obtain p s  where sreq:"sobj_source_eq_obj (SProc (r, fr, pt, u) srp) (Proc p)"
+    and TP: "(Proc p) \<in> tainted s" and vds: "valid s" and "p \<in> current_procs s"
+    and "obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp \<and> no_del_event s \<and> 
+         initp_intact_but s (SProc (r, fr, pt, u) srp)" 
+    apply (clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_exec2(3) obtain f s' where vds's: "valid (s' @ s)"
+    and etype: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact_but (s'@s) (SProc (r,fr,pt,u) srp)" 
+    and ISP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "f \<in> current_files (s' @ s)"
+    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E1, auto)
+    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_file, clarsimp)
+  from vds's ISP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
+  with exp vds's ISP have initp: "p \<in> init_processes" 
+    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
+  obtain ev \<tau> where ev: "ev = Execute p f"
+    and tau: "\<tau> = Clone p (new_proc (s'@s)) # (s' @ s)" by auto
+  hence vs_tau:"valid \<tau>" using exp vds's by (auto intro:clone_event_no_limit)
+
+  have valid: "valid (ev # \<tau>)"
+  proof-
+    have "os_grant \<tau> ev"
+      using ev tau by (simp add:exf exp)
+    moreover have "rc_grant \<tau> ev" 
+      using ev tau ts'_exec2 ISP etype
+      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps 
+              split:if_splits option.splits)
+    ultimately show ?thesis using vs_tau
+      by (erule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (ev # \<tau>) (Proc p) = SProc (r', fr', exec_type_aux r pt, u) srp"
+  proof-
+    have "obj2sobj \<tau> (Proc p) = SProc (r,fr,pt,u) srp" using ISP tau vs_tau
+      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
+    moreover have "source_dir \<tau> f = Some sd" using vs_tau etype tau 
+      by (auto simp:source_dir_simps obj2sobj.simps split:option.splits if_splits)
+    ultimately show ?thesis using valid ts'_exec2(5) ts'_exec2(6) ev
+      by (auto simp:cp2sproc_exec obj2sobj.simps split:option.splits)
+  qed  moreover
+  have "Proc p \<in> tainted (ev # \<tau>)"
+  proof-
+    have "(Proc p) \<in> tainted \<tau>" using TP nodels' tau vs_tau
+      by (drule_tac s = "Clone p (new_proc (s' @ s)) # s'" in t_remain_app,auto)
+    thus ?thesis using ev valid
+      by (drule_tac t_remain, auto dest:valid_os)
+  qed  moreover
+  have "no_del_event (ev # \<tau>)" using ev tau nodels' by simp 
+  moreover have "initp_intact_but (ev#\<tau>) (SProc (r', fr', exec_type_aux r pt, u) srp)"
+    using ev tau nodels' intacts' srpeq valid initp
+    by (simp, rule_tac initp_intact_butp_I_exec, simp_all)
+  moreover have "sobj_source_eq_obj (SProc (r', fr', exec_type_aux r pt, u) srp) (Proc p)"
+    using sreq by (case_tac srp, simp+)
+  ultimately show ?case
+    by (rule_tac x = "Proc p" in exI, rule_tac x = "ev#\<tau>" in exI, auto)
+next
+  case (ts'_cfd r fr pt u srp t sd srf t')
+  from ts'_cfd(1) obtain p s where TP: "(Proc p) \<in> tainted s" 
+    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
+         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
+    apply (drule_tac ts2t_intact, clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_cfd(3) obtain pf s' where vds': "valid (s' @ s)"
+    and SF: "obj2sobj (s'@s) (File pf) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "pf \<in> current_files (s' @ s)"
+    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
+    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_file, clarsimp)
+  obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)"
+    and tau: "\<tau> = (s' @ s)" by auto
+
+  have valid: "valid (e# \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp exf by (simp add:ncf_notin_curf ncf_parent)
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_cfd(4,5,6) SP SF
+      by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
+        split:if_splits option.splits t_rc_file_type.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
+  have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t', sd) None"
+  proof-
+    have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t'"
+      using ev tau SF SP ts'_cfd(4)
+      by (auto simp:obj2sobj.simps etype_of_file_def cp2sproc.simps
+              split:option.splits if_splits  intro!:etype_aux_prop4)
+    moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
+      using ev tau SF SP valid ncf_parent
+      by (auto simp:source_dir_simps obj2sobj.simps cp2sproc.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
+      nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
+      by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
+  qed  moreover
+  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
+    by (auto intro!:initp_intact_I_others)  moreover
+  have "File (new_childf pf \<tau>) \<in> tainted (e#\<tau>)" 
+  proof- 
+    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
+      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
+    thus ?thesis using ev tau valid by (auto intro:t_cfile)
+  qed
+  ultimately show ?case using tau ev
+    apply (rule_tac x = "File (new_childf pf \<tau>)" in exI)
+    by (rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_cfd' r fr pt u srp t sd srf)
+  from ts'_cfd'(1) obtain p s where TP: "(Proc p) \<in> tainted s" 
+    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
+         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
+    apply (drule_tac ts2t_intact, clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_cfd'(3) obtain pf s' where vds': "valid (s' @ s)"
+    and SF: "obj2sobj (s'@s) (File pf) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "pf \<in> current_files (s' @ s)"
+    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E0, auto)
+    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_file, clarsimp)
+  obtain e \<tau> where ev: "e = CreateFile p (new_childf pf \<tau>)"
+    and tau: "\<tau> = (s' @ s)" by auto
+
+  have valid: "valid (e# \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp exf by (simp add:ncf_notin_curf ncf_parent)
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_cfd'(4,5) SP SF
+      by (auto simp:etype_of_file_def cp2sproc.simps ncf_parent obj2sobj.simps
+        split:if_splits option.splits t_rc_file_type.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
+  have "obj2sobj (e#\<tau>) (File (new_childf pf \<tau>)) = SFile (t, sd) None"
+  proof-
+      have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = Some t"
+      proof-
+        have "etype_of_file (e#\<tau>) (new_childf pf \<tau>) = etype_of_file \<tau> pf"
+          using ev tau SP ts'_cfd'(4)
+          by (auto simp:obj2sobj.simps ncf_parent etype_of_file_def cp2sproc.simps
+            split:option.splits   intro!:etype_aux_prop3)
+        thus ?thesis using SF tau ev
+          by (auto simp:obj2sobj.simps split:option.splits if_splits)
+      qed
+    moreover have "source_dir (e#\<tau>) (new_childf pf \<tau>) = Some sd"
+      using ev tau SF SP valid ncf_parent
+      by (auto simp:source_dir_simps obj2sobj.simps cp2sproc.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis  using nodel ncf_notin_curf[where s = \<tau>]
+      nodel_imp_exists[where obj = "File (new_childf pf \<tau>)" and s =\<tau>]
+      by (auto simp:obj2sobj.simps dest:no_del_event_cons_D)
+  qed  moreover
+  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
+    by (auto intro!:initp_intact_I_others)  moreover
+  have "File (new_childf pf \<tau>) \<in> tainted (e#\<tau>)" 
+  proof- 
+    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
+      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
+    thus ?thesis using ev tau valid by (auto intro:t_cfile)
+  qed
+  ultimately show ?case using tau ev
+    apply (rule_tac x = "File (new_childf pf \<tau>)" in exI)
+    by (rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_cipc r fr pt u srp)
+  from ts'_cipc(1) obtain p s where TP: "(Proc p) \<in> tainted s" 
+    and vds: "valid s" and exp: "p \<in> current_procs s" 
+    and nodels: "no_del_event s" and intacts: "initp_intact s"
+    and SP: "obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
+    apply (drule_tac ts2t_intact, clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  obtain e where ev: "e = CreateIPC p (new_ipc s)" by simp
+      
+  have valid: "valid (e # s)" 
+  proof-
+    have "os_grant s e"
+      using ev exp by (simp)
+    moreover have "rc_grant s e" 
+      using ev ts'_cipc(3) SP
+      by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
+    ultimately show ?thesis using vds
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have nodel: "no_del_event (e#s)" using nodels ev by simp  moreover
+  have "initp_intact (e#s)" using ev intacts valid nodels
+    by (auto intro!:initp_intact_I_others)  moreover
+  have "IPC (new_ipc s) \<in> tainted (e#s)" using TP ev valid
+    by (auto intro:t_cipc)  moreover
+  have "obj2sobj (e#s) (IPC (new_ipc s)) = SIPC (default_ipc_create_type r) None"
+    using ev SP nodel nodel_imp_exists[where obj = "IPC (new_ipc s)" and s=s]
+    by (auto simp:obj2sobj.simps ni_notin_curi cp2sproc.simps
+            split:option.splits dest:no_del_event_cons_D)
+  ultimately show ?case using ev
+    apply (rule_tac x = "IPC (new_ipc s)" in exI)
+    by (rule_tac x = "e # s" in exI, auto)
+next
+  case (ts'_read t sd srf r fr pt u srp)
+  then obtain f s  where "sobj_source_eq_obj (SFile (t, sd) srf) (File f)"
+    and TF: "(File f) \<in> tainted s" and vds: "valid s" and "f \<in> current_files s"
+    and "obj2sobj s (File f) = SFile (t, sd) srf \<and> no_del_event s \<and> initp_intact s" 
+    apply (clarsimp, frule_tac obj2sobj_file)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_read(3) obtain p s' where vds': "valid (s' @ s)"
+    and SF: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact_but (s'@s) (SProc (r,fr,pt,u) srp)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "f \<in> current_files (s' @ s)"
+    and sreq: "sobj_source_eq_obj (SProc (r,fr,pt,u) srp) (Proc p)"
+    apply (drule_tac s' =s and obj' = "File f" in all_sobjs_E2, auto)
+    apply (frule_tac obj = "File f" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_proc, clarsimp)
+  from vds' SP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
+  with exp vds' SP have initp: "p \<in> init_processes" 
+    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
+  obtain e \<tau> where ev: "e = ReadFile p f" and tau: "\<tau> = s' @ s" by auto
+
+  have valid: "valid (e # \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau by (simp add:exf exp)
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_read(4) SP SF
+      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (e # \<tau>) (Proc p) = SProc (r, fr, pt, u) srp" using valid tau ev SP
+    by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
+  moreover  have "Proc p \<in> tainted (e # \<tau>)"
+  proof-
+    have "(File f) \<in> tainted \<tau>" using TF nodels' tau vds'
+      by (drule_tac s = "s'" in t_remain_app,auto)
+    thus ?thesis using ev valid
+      by (drule_tac t_read, simp+)
+  qed  moreover
+  have "no_del_event (e # \<tau>)" using ev tau nodels' by simp 
+  moreover  have "initp_intact_but (e#\<tau>) (SProc (r, fr, pt, u) srp)"
+    using ev tau nodels' intacts' srpeq valid initp
+    by (simp, rule_tac initp_intact_butp_I_others, simp_all)
+  moreover have "sobj_source_eq_obj (SProc (r,fr,pt,u) srp) (Proc p)"
+    using sreq by (case_tac srp, simp+)
+  ultimately show ?case
+    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_write r fr pt u srp t sd srf)
+  from ts'_write(1) obtain p s where TP: "(Proc p) \<in> tainted s" 
+    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
+         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
+    apply (drule_tac ts2t_intact, clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_write(3) obtain f s' where vds': "valid (s' @ s)"
+    and SF: "obj2sobj (s'@s) (File f) = SFile (t, sd) srf"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exf: "f \<in> current_files (s' @ s)"
+    and sreq: "sobj_source_eq_obj (SFile (t, sd) srf) (File f)"
+    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E2, auto)
+    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_file, clarsimp)
+  obtain e \<tau> where ev: "e = WriteFile p f" and tau: "\<tau> = (s' @ s)" by auto
+
+  have valid: "valid (e# \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp exf by simp
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_write(4) SP SF
+      by (auto simp:etype_of_file_def cp2sproc.simps obj2sobj.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
+  have "obj2sobj (e#\<tau>) (File f) = SFile (t, sd) srf"
+    using tau ev SF valid
+    by (auto simp:obj2sobj.simps source_dir_simps etype_of_file_def
+            split:option.splits if_splits)
+  moreover
+  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
+    by (auto intro!:initp_intact_I_others)  moreover
+  have "File f \<in> tainted (e#\<tau>)" 
+  proof- 
+    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
+      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
+    thus ?thesis using ev tau valid by (auto intro:t_write)
+  qed
+  ultimately show ?case using tau ev sreq
+    by (rule_tac x = "File f" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+next 
+  case (ts'_recv t sri r fr pt u srp)
+  then obtain i s  where "sobj_source_eq_obj (SIPC t sri) (IPC i)"
+    and TI: "(IPC i) \<in> tainted s" and vds: "valid s" and "i \<in> current_ipcs s"
+    and "obj2sobj s (IPC i) = SIPC t sri \<and> no_del_event s \<and> initp_intact s" 
+    apply (clarsimp, frule_tac obj2sobj_ipc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_recv(3) obtain p s' where vds': "valid (s' @ s)"
+    and SI: "obj2sobj (s'@s) (IPC i) = SIPC t sri"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact_but (s'@s) (SProc (r,fr,pt,u) srp)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exi: "i \<in> current_ipcs (s' @ s)"
+    and sreq: "sobj_source_eq_obj (SProc (r,fr,pt,u) srp) (Proc p)"
+    apply (drule_tac s' =s and obj' = "IPC i" in all_sobjs_E2, auto)
+    apply (frule_tac obj = "IPC i" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_proc, clarsimp)
+  from vds' SP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
+  with exp vds' SP have initp: "p \<in> init_processes" 
+    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
+  obtain e \<tau> where ev: "e = Recv p i" and tau: "\<tau> = s' @ s" by auto
+
+  have valid: "valid (e # \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau by (simp add:exi exp)
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_recv(4) SP SI
+      by (auto simp:cp2sproc.simps obj2sobj.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (e # \<tau>) (Proc p) = SProc (r, fr, pt, u) srp" using valid tau ev SP
+    by (auto simp:obj2sobj.simps cp2sproc_simps' split:option.splits)
+  moreover  have "Proc p \<in> tainted (e # \<tau>)"
+  proof-
+    have "(IPC i) \<in> tainted \<tau>" using TI nodels' tau vds'
+      by (drule_tac s = "s'" in t_remain_app,auto)
+    thus ?thesis using ev valid
+      by (drule_tac t_recv, simp+)
+  qed  moreover
+  have "no_del_event (e # \<tau>)" using ev tau nodels' by simp 
+  moreover  have "initp_intact_but (e#\<tau>) (SProc (r, fr, pt, u) srp)"
+    using ev tau nodels' intacts' srpeq valid initp
+    by (simp, rule_tac initp_intact_butp_I_others, simp_all)
+  moreover have "sobj_source_eq_obj (SProc (r,fr,pt,u) srp) (Proc p)"
+    using sreq by (case_tac srp, simp+)
+  ultimately show ?case
+    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_send r fr pt u srp t sri)
+  from ts'_send(1) obtain p s where TP: "(Proc p) \<in> tainted s" 
+    and "valid s \<and> p \<in> current_procs s \<and> no_del_event s \<and> initp_intact s \<and>
+         obj2sobj s (Proc p) = SProc (r,fr,pt,u) srp"
+    apply (drule_tac ts2t_intact, clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  with ts'_send(3) obtain i s' where vds': "valid (s' @ s)"
+    and SI: "obj2sobj (s'@s) (IPC i) = SIPC t sri"
+    and nodels': "no_del_event (s'@s)"
+    and intacts': "initp_intact (s'@s)"
+    and SP: "obj2sobj (s'@s) (Proc p) = SProc (r,fr,pt,u) srp"
+    and exp: "p \<in> current_procs (s' @ s)"
+    and exi: "i \<in> current_ipcs (s' @ s)"
+    and sreq: "sobj_source_eq_obj (SIPC t sri) (IPC i)"
+    apply (drule_tac s' =s and obj' = "Proc p" in all_sobjs_E2, auto)
+    apply (frule_tac obj = "Proc p" in nodel_tainted_exists, simp)
+    by (frule obj2sobj_ipc, clarsimp)
+  obtain e \<tau> where ev: "e = Send p i" and tau: "\<tau> = (s' @ s)" by auto
+
+  have valid: "valid (e# \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp exi by simp
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_send(4) SP SI
+      by (auto simp:cp2sproc.simps obj2sobj.simps
+              split:if_splits option.splits)
+    ultimately show ?thesis using vds' tau
+      by (rule_tac vs_step, simp+)
+  qed  moreover
+  have nodel: "no_del_event (e#\<tau>)" using nodels' tau ev by simp  moreover
+  have "obj2sobj (e#\<tau>) (IPC i) = SIPC t sri"
+    using tau ev SI valid
+    by (auto simp:obj2sobj.simps split:option.splits if_splits)
+  moreover  have "initp_intact (e#\<tau>)" using ev intacts' valid nodels' tau
+    by (auto intro!:initp_intact_I_others)  
+  moreover  have "IPC i \<in> tainted (e#\<tau>)" 
+  proof- 
+    have "Proc p \<in> tainted \<tau>" using nodel vds' TP ev tau
+      by (drule_tac s = "s'" and s' = s in t_remain_app, auto)
+    thus ?thesis using ev tau valid by (auto intro:t_send)
+  qed
+  ultimately show ?case using tau ev sreq
+    by (rule_tac x = "IPC i" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_crole r fr pt u srp r')
+  then obtain p s where exp: "p \<in> current_procs s"
+    and sreq:"sobj_source_eq_obj (SProc (r, fr, pt, u) srp) (Proc p)"
+    and TP: "(Proc p) \<in> tainted s" and vds: "valid s" and nodels: "no_del_event s"
+    and SP: "obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp"
+    and intacts: "initp_intact_but s (SProc (r, fr, pt, u) srp)" 
+    apply (clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  from vds SP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
+  with exp vds SP have initp: "p \<in> init_processes" 
+    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
+  obtain e \<tau> where ev: "e = ChangeRole p r'"
+    and tau: "\<tau> = Clone p (new_proc s) # s" by auto
+  hence vs_tau:"valid \<tau>" using exp vds by (auto intro:clone_event_no_limit)
+
+  have valid: "valid (e # \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp by simp
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_crole(3) SP
+      by (auto simp:cp2sproc.simps obj2sobj.simps split:option.splits)
+    ultimately show ?thesis using vs_tau
+      by (erule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (e # \<tau>) (Proc p) = SProc (r', fr, pt, u) srp"
+  proof-
+    have "obj2sobj \<tau> (Proc p) = SProc (r,fr,pt,u) srp" using SP tau vs_tau
+      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
+    thus ?thesis using valid ev
+      by (auto simp:cp2sproc_crole obj2sobj.simps split:option.splits)
+  qed  moreover
+  have "Proc p \<in> tainted (e # \<tau>)"
+  proof-
+    have "(Proc p) \<in> tainted \<tau>" using TP tau vs_tau exp
+      by (auto intro!:t_remain)
+    thus ?thesis using ev valid
+      by (drule_tac t_remain, auto dest:valid_os)
+  qed  moreover
+  have "no_del_event (e # \<tau>)" using ev tau nodels by simp 
+  moreover have "initp_intact_but (e#\<tau>) (SProc (r', fr, pt, u) srp)"
+    using ev tau nodels intacts srpeq valid initp
+    by (simp, rule_tac initp_intact_butp_I_crole, simp_all)
+  moreover have "sobj_source_eq_obj (SProc (r',fr,pt,u) srp) (Proc p)"
+    using sreq by (case_tac srp, simp+)
+  ultimately show ?case
+    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+next
+  case (ts'_chown r fr pt u srp u' nr)
+  then obtain p s where exp: "p \<in> current_procs s"
+    and sreq:"sobj_source_eq_obj (SProc (r, fr, pt, u) srp) (Proc p)"
+    and TP: "(Proc p) \<in> tainted s" and vds: "valid s" and nodels: "no_del_event s"
+    and SP: "obj2sobj s (Proc p) = SProc (r, fr, pt, u) srp"
+    and intacts: "initp_intact_but s (SProc (r, fr, pt, u) srp)" 
+    apply (clarsimp, frule_tac obj2sobj_proc)
+    by (frule tainted_is_valid, frule tainted_is_current, auto)
+  from vds SP sreq exp have srpeq: "srp = Some p" by (simp add:proc_source_eq_prop)
+  with exp vds SP have initp: "p \<in> init_processes" 
+    by (auto simp:obj2sobj.simps dest:source_proc_in_init split:option.splits)
+  obtain e \<tau> where ev: "e = ChangeOwner p u'"
+    and tau: "\<tau> = Clone p (new_proc s) # s" by auto
+  hence vs_tau:"valid \<tau>" using exp vds by (auto intro:clone_event_no_limit)
+
+  have valid: "valid (e # \<tau>)"
+  proof-
+    have "os_grant \<tau> e"
+      using ev tau exp ts'_chown(3) by simp
+    moreover have "rc_grant \<tau> e" 
+      using ev tau ts'_chown(5) SP
+      by (auto simp:cp2sproc.simps obj2sobj.simps pct_def clone_type_unchange
+              split:option.splits t_rc_proc_type.splits)
+        (* here is another place of no_limit of clone event assumption *)
+    ultimately show ?thesis using vs_tau
+      by (erule_tac vs_step, simp+)
+  qed  moreover
+  have "obj2sobj (e # \<tau>) (Proc p) = SProc (nr, fr, chown_type_aux r nr pt, u') srp"
+  proof-
+    have "obj2sobj \<tau> (Proc p) = SProc (r,fr,pt,u) srp" using SP tau vs_tau
+      by (auto split:option.splits simp:cp2sproc_simps' obj2sobj.simps)
+    thus ?thesis using valid ev ts'_chown(4)
+      by (auto simp:cp2sproc_chown obj2sobj.simps split:option.splits)
+  qed  moreover
+  have "Proc p \<in> tainted (e # \<tau>)"
+  proof-
+    have "(Proc p) \<in> tainted \<tau>" using TP tau vs_tau exp
+      by (auto intro!:t_remain)
+    thus ?thesis using ev valid
+      by (drule_tac t_remain, auto dest:valid_os)
+  qed  moreover
+  have "no_del_event (e # \<tau>)" using ev tau nodels by simp 
+  moreover have "initp_intact_but (e#\<tau>) (SProc (nr, fr, chown_type_aux r nr pt, u') srp)"
+    using ev tau nodels intacts srpeq valid initp
+    by (simp, rule_tac initp_intact_butp_I_chown, simp_all)
+  moreover have "sobj_source_eq_obj (SProc (nr, fr, chown_type_aux r nr pt, u') srp) (Proc p)"
+    using sreq by (case_tac srp, simp+)
+  ultimately show ?case
+    by (rule_tac x = "Proc p" in exI, rule_tac x = "e#\<tau>" in exI, auto)
+qed
+
+lemma tainted_s2tainted:
+  "sobj \<in> tainted_s \<Longrightarrow> \<exists> obj s. obj2sobj s obj = sobj \<and> obj \<in> tainted s \<and> 
+                                 sobj_source_eq_obj sobj obj \<and> no_del_event s \<and> 
+                                 initp_intact_but s sobj"
+apply (simp add:tainted_s'_eq)
+by (erule ts2t)
+
+end
+
+end
\ No newline at end of file