CpsG.thy retrofiting almost completed. An important mile stone.
authorzhangx
Wed, 27 Jan 2016 19:26:56 +0800
changeset 80 17305a85493d
parent 79 8067efcb43da
child 81 c495eb16beb6
CpsG.thy retrofiting almost completed. An important mile stone.
CpsG.thy~
PIPBasics.thy
PIPBasics.thy~
PIPDefs.thy
PIPDefs.thy~
RTree.thy
RTree.thy~
--- a/CpsG.thy~	Sun Jan 17 22:18:35 2016 +0800
+++ b/CpsG.thy~	Wed Jan 27 19:26:56 2016 +0800
@@ -1,7 +1,38 @@
 theory CpsG
-imports PIPDefs 
+imports PIPDefs
 begin
 
+lemma Max_fg_mono:
+  assumes "finite A"
+  and "\<forall> a \<in> A. f a \<le> g a"
+  shows "Max (f ` A) \<le> Max (g ` A)"
+proof(cases "A = {}")
+  case True
+  thus ?thesis by auto
+next
+  case False
+  show ?thesis
+  proof(rule Max.boundedI)
+    from assms show "finite (f ` A)" by auto
+  next
+    from False show "f ` A \<noteq> {}" by auto
+  next
+    fix fa
+    assume "fa \<in> f ` A"
+    then obtain a where h_fa: "a \<in> A" "fa = f a" by auto
+    show "fa \<le> Max (g ` A)"
+    proof(rule Max_ge_iff[THEN iffD2])
+      from assms show "finite (g ` A)" by auto
+    next
+      from False show "g ` A \<noteq> {}" by auto
+    next
+      from h_fa have "g a \<in> g ` A" by auto
+      moreover have "fa \<le> g a" using h_fa assms(2) by auto
+      ultimately show "\<exists>a\<in>g ` A. fa \<le> a" by auto
+    qed
+  qed
+qed 
+
 lemma Max_f_mono:
   assumes seq: "A \<subseteq> B"
   and np: "A \<noteq> {}"
@@ -15,6 +46,52 @@
   from fnt and seq show "finite (f ` B)" by auto
 qed
 
+lemma eq_RAG: 
+  "RAG (wq s) = RAG s"
+  by (unfold cs_RAG_def s_RAG_def, auto)
+
+lemma waiting_holding:
+  assumes "waiting (s::state) th cs"
+  obtains th' where "holding s th' cs"
+proof -
+  from assms[unfolded s_waiting_def, folded wq_def]
+  obtain th' where "th' \<in> set (wq s cs)" "th' = hd (wq s cs)"
+    by (metis empty_iff hd_in_set list.set(1))
+  hence "holding s th' cs" 
+    by (unfold s_holding_def, fold wq_def, auto)
+  from that[OF this] show ?thesis .
+qed
+
+lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
+unfolding cp_def wq_def
+apply(induct s rule: schs.induct)
+apply(simp add: Let_def cpreced_initial)
+apply(simp add: Let_def)
+apply(simp add: Let_def)
+apply(simp add: Let_def)
+apply(subst (2) schs.simps)
+apply(simp add: Let_def)
+apply(subst (2) schs.simps)
+apply(simp add: Let_def)
+done
+
+lemma cp_alt_def:
+  "cp s th =  
+           Max ((the_preced s) ` {th'. Th th' \<in> (subtree (RAG s) (Th th))})"
+proof -
+  have "Max (the_preced s ` ({th} \<union> dependants (wq s) th)) =
+        Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})" 
+          (is "Max (_ ` ?L) = Max (_ ` ?R)")
+  proof -
+    have "?L = ?R" 
+    by (auto dest:rtranclD simp:cs_dependants_def cs_RAG_def s_RAG_def subtree_def)
+    thus ?thesis by simp
+  qed
+  thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp)
+qed
+
+(* ccc *)
+
 
 locale valid_trace = 
   fixes s
@@ -114,19 +191,6 @@
 
 end
 
-lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
-unfolding cp_def wq_def
-apply(induct s rule: schs.induct)
-apply(simp add: Let_def cpreced_initial)
-apply(simp add: Let_def)
-apply(simp add: Let_def)
-apply(simp add: Let_def)
-apply(subst (2) schs.simps)
-apply(simp add: Let_def)
-apply(subst (2) schs.simps)
-apply(simp add: Let_def)
-done
-
 lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
   by (unfold s_RAG_def, auto)
 
@@ -232,6 +296,17 @@
   show ?thesis by (cases, simp)
 qed
 
+lemma ready_th_s: "th \<in> readys s"
+  using runing_th_s
+  by (unfold runing_def, auto)
+
+lemma live_th_s: "th \<in> threads s"
+  using readys_threads ready_th_s by auto
+
+lemma live_th_es: "th \<in> threads (e#s)"
+  using live_th_s 
+  by (unfold is_p, simp)
+
 lemma th_not_waiting: 
   "\<not> waiting s th c"
 proof -
@@ -735,14 +810,25 @@
 apply (unfold s_RAG_def s_waiting_def wq_def)
 by (simp add:Let_def)
 
+lemma (in valid_trace_set)
+   RAG_unchanged: "(RAG (e # s)) = RAG s"
+   by (unfold is_set RAG_set_unchanged, simp)
+
 lemma RAG_create_unchanged: "(RAG (Create th prio # s)) = RAG s"
 apply (unfold s_RAG_def s_waiting_def wq_def)
 by (simp add:Let_def)
 
+lemma (in valid_trace_create)
+   RAG_unchanged: "(RAG (e # s)) = RAG s"
+   by (unfold is_create RAG_create_unchanged, simp)
+
 lemma RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
 apply (unfold s_RAG_def s_waiting_def wq_def)
 by (simp add:Let_def)
 
+lemma (in valid_trace_exit)
+   RAG_unchanged: "(RAG (e # s)) = RAG s"
+   by (unfold is_exit RAG_exit_unchanged, simp)
 
 context valid_trace_v
 begin
@@ -766,12 +852,14 @@
 lemma distinct_wq': "distinct wq'"
   by (metis (mono_tags, lifting) distinct_rest  some_eq_ex wq'_def)
   
+lemma set_wq': "set wq' = set rest"
+  by (metis (mono_tags, lifting) distinct_rest rest_def 
+      some_eq_ex wq'_def)
+    
 lemma th'_in_inv:
   assumes "th' \<in> set wq'"
   shows "th' \<in> set rest"
-  using assms
-  by (metis (mono_tags, lifting) distinct.simps(2) 
-        rest_def some_eq_ex wq'_def wq_distinct wq_s_cs) 
+  using assms set_wq' by simp
 
 lemma neq_t_th: 
   assumes "waiting (e#s) t c"
@@ -1380,81 +1468,18 @@
   from wq_threads [OF this] show ?thesis .
 qed
 
-lemma  cp_le:
-  assumes th_in: "th \<in> threads s"
-  shows "cp s th \<le> Max ((\<lambda> th. (preced th s)) ` threads s)"
-proof(unfold cp_eq_cpreced cpreced_def cs_dependants_def)
-  show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}))
-         \<le> Max ((\<lambda>th. preced th s) ` threads s)"
-    (is "Max (?f ` ?A) \<le> Max (?f ` ?B)")
-  proof(rule Max_f_mono)
-    show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}" by simp
-  next
-    from finite_threads
-    show "finite (threads s)" .
-  next
-    from th_in
-    show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> threads s"
-      apply (auto simp:Domain_def)
-      apply (rule_tac dm_RAG_threads)
-      apply (unfold trancl_domain [of "RAG s", symmetric])
-      by (unfold cs_RAG_def s_RAG_def, auto simp:Domain_def)
-  qed
-qed
-
-lemma max_cp_eq: 
-  shows "Max ((cp s) ` threads s) = Max ((\<lambda> th. (preced th s)) ` threads s)"
-  (is "?l = ?r")
-proof(cases "threads s = {}")
-  case True
-  thus ?thesis by auto
-next
-  case False
-  have "?l \<in> ((cp s) ` threads s)"
-  proof(rule Max_in)
-    from finite_threads
-    show "finite (cp s ` threads s)" by auto
-  next
-    from False show "cp s ` threads s \<noteq> {}" by auto
-  qed
-  then obtain th 
-    where th_in: "th \<in> threads s" and eq_l: "?l = cp s th" by auto
-  have "\<dots> \<le> ?r" by (rule cp_le[OF th_in])
-  moreover have "?r \<le> cp s th" (is "Max (?f ` ?A) \<le> cp s th")
-  proof -
-    have "?r \<in> (?f ` ?A)"
-    proof(rule Max_in)
-      from finite_threads
-      show " finite ((\<lambda>th. preced th s) ` threads s)" by auto
-    next
-      from False show " (\<lambda>th. preced th s) ` threads s \<noteq> {}" by auto
-    qed
-    then obtain th' where 
-      th_in': "th' \<in> ?A " and eq_r: "?r = ?f th'" by auto
-    from le_cp [of th']  eq_r
-    have "?r \<le> cp s th'" by auto
-    moreover have "\<dots> \<le> cp s th"
-    proof(fold eq_l)
-      show " cp s th' \<le> Max (cp s ` threads s)"
-      proof(rule Max_ge)
-        from th_in' show "cp s th' \<in> cp s ` threads s"
-          by auto
-      next
-        from finite_threads
-        show "finite (cp s ` threads s)" by auto
-      qed
-    qed
-    ultimately show ?thesis by auto
-  qed
-  ultimately show ?thesis using eq_l by auto
-qed
-
-lemma max_cp_eq_the_preced:
-  shows "Max ((cp s) ` threads s) = Max (the_preced s ` threads s)"
-  using max_cp_eq using the_preced_def by presburger 
+lemma rg_RAG_threads: 
+  assumes "(Th th) \<in> Range (RAG s)"
+  shows "th \<in> threads s"
+  using assms
+  by (unfold s_RAG_def cs_waiting_def cs_holding_def, 
+       auto intro:wq_threads)
 
 end
 
+
+
+
 lemma preced_v [simp]: "preced th' (V th cs#s) = preced th' s"
   by (unfold preced_def, simp)
 
@@ -1476,7 +1501,7 @@
 context valid_trace_p
 begin
 
-lemma not_holding_es_th_cs: "\<not> holding s th cs"
+lemma not_holding_s_th_cs: "\<not> holding s th cs"
 proof
   assume otherwise: "holding s th cs"
   from pip_e[unfolded is_p]
@@ -1490,13 +1515,39 @@
   qed
 qed
 
+lemma waiting_kept:
+  assumes "waiting s th' cs'"
+  shows "waiting (e#s) th' cs'"
+  using assms
+  by (metis cs_waiting_def hd_append2 list.sel(1) list.set_intros(2) 
+      rotate1.simps(2) self_append_conv2 set_rotate1 
+        th_not_in_wq waiting_eq wq_es_cs wq_neq_simp)
+  
+lemma holding_kept:
+  assumes "holding s th' cs'"
+  shows "holding (e#s) th' cs'"
+proof(cases "cs' = cs")
+  case False
+  hence "wq (e#s) cs' = wq s cs'" by simp
+  with assms show ?thesis using cs_holding_def holding_eq by auto 
+next
+  case True
+  from assms[unfolded s_holding_def, folded wq_def]
+  obtain rest where eq_wq: "wq s cs' = th'#rest"
+    by (metis empty_iff list.collapse list.set(1)) 
+  hence "wq (e#s) cs' = th'#(rest@[th])"
+    by (simp add: True wq_es_cs) 
+  thus ?thesis
+    by (simp add: cs_holding_def holding_eq) 
+qed
+
 end
 
 locale valid_trace_p_h = valid_trace_p +
   assumes we: "wq s cs = []"
 
 locale valid_trace_p_w = valid_trace_p +
-  assumes we: "wq s cs \<noteq> []"
+  assumes wne: "wq s cs \<noteq> []"
 begin
 
 definition "holder = hd (wq s cs)"
@@ -1504,7 +1555,7 @@
 definition "waiters' = waiters @ [th]"
 
 lemma wq_s_cs: "wq s cs = holder#waiters"
-    by (simp add: holder_def waiters_def we)
+    by (simp add: holder_def waiters_def wne)
     
 lemma wq_es_cs': "wq (e#s) cs = holder#waiters@[th]"
   by (simp add: wq_es_cs wq_s_cs)
@@ -1515,6 +1566,95 @@
 lemma RAG_edge: "(Th th, Cs cs) \<in> RAG (e#s)"
    by (unfold s_RAG_def, fold waiting_eq, insert waiting_es_th_cs, auto)
 
+lemma holding_esE:
+  assumes "holding (e#s) th' cs'"
+  obtains "holding s th' cs'"
+  using assms 
+proof(cases "cs' = cs")
+  case False
+  hence "wq (e#s) cs' = wq s cs'" by simp
+  with assms show ?thesis
+    using cs_holding_def holding_eq that by auto 
+next
+  case True
+  with assms show ?thesis
+  by (metis cs_holding_def holding_eq list.sel(1) list.set_intros(1) that 
+        wq_es_cs' wq_s_cs) 
+qed
+
+lemma waiting_esE:
+  assumes "waiting (e#s) th' cs'"
+  obtains "th' \<noteq> th" "waiting s th' cs'"
+     |  "th' = th" "cs' = cs"
+proof(cases "waiting s th' cs'")
+  case True
+  have "th' \<noteq> th"
+  proof
+    assume otherwise: "th' = th"
+    from True[unfolded this]
+    show False by (simp add: th_not_waiting) 
+  qed
+  from that(1)[OF this True] show ?thesis .
+next
+  case False
+  hence "th' = th \<and> cs' = cs"
+      by (metis assms cs_waiting_def holder_def list.sel(1) rotate1.simps(2) 
+        set_ConsD set_rotate1 waiting_eq wq_es_cs wq_es_cs' wq_neq_simp)
+  with that(2) show ?thesis by metis
+qed
+
+lemma RAG_es: "RAG (e # s) =  RAG s \<union> {(Th th, Cs cs)}" (is "?L = ?R")
+proof(rule rel_eqI)
+  fix n1 n2
+  assume "(n1, n2) \<in> ?L"
+  thus "(n1, n2) \<in> ?R" 
+  proof(cases rule:in_RAG_E)
+    case (waiting th' cs')
+    from this(3)
+    show ?thesis
+    proof(cases rule:waiting_esE)
+      case 1
+      thus ?thesis using waiting(1,2)
+        by (unfold s_RAG_def, fold waiting_eq, auto)
+    next
+      case 2
+      thus ?thesis using waiting(1,2) by auto
+    qed
+  next
+    case (holding th' cs')
+    from this(3)
+    show ?thesis
+    proof(cases rule:holding_esE)
+      case 1
+      with holding(1,2)
+      show ?thesis by (unfold s_RAG_def, fold holding_eq, auto)
+    qed
+  qed
+next
+  fix n1 n2
+  assume "(n1, n2) \<in> ?R"
+  hence "(n1, n2) \<in> RAG s \<or> (n1 = Th th \<and> n2 = Cs cs)" by auto
+  thus "(n1, n2) \<in> ?L"
+  proof
+    assume "(n1, n2) \<in> RAG s"
+    thus ?thesis
+    proof(cases rule:in_RAG_E)
+      case (waiting th' cs')
+      from waiting_kept[OF this(3)]
+      show ?thesis using waiting(1,2)
+         by (unfold s_RAG_def, fold waiting_eq, auto)
+    next
+      case (holding th' cs')
+      from holding_kept[OF this(3)]
+      show ?thesis using holding(1,2)
+         by (unfold s_RAG_def, fold holding_eq, auto)
+    qed
+  next
+    assume "n1 = Th th \<and> n2 = Cs cs"
+    thus ?thesis using RAG_edge by auto
+  qed
+qed
+
 end
 
 context valid_trace_p_h
@@ -1557,14 +1697,6 @@
   from that(1)[OF False this] show ?thesis .
 qed
 
-lemma waiting_kept:
-  assumes "waiting s th' cs'"
-  shows "waiting (e#s) th' cs'"
-  using assms
-  by (metis cs_waiting_def list.sel(1) list.set_intros(2) 
-        th_not_in_wq waiting_eq we wq_es_cs' wq_neq_simp)
-    
-
 lemma RAG_es: "RAG (e # s) =  RAG s \<union> {(Cs cs, Th th)}" (is "?L = ?R")
 proof(rule rel_eqI)
   fix n1 n2
@@ -1602,45 +1734,1938 @@
     thus ?thesis
     proof(cases rule:in_RAG_E)
       case (waiting th' cs')
-      find_theorems waiting e s
+      from waiting_kept[OF this(3)]
+      show ?thesis using waiting(1,2)
+         by (unfold s_RAG_def, fold waiting_eq, auto)
+    next
+      case (holding th' cs')
+      from holding_kept[OF this(3)]
+      show ?thesis using holding(1,2)
+         by (unfold s_RAG_def, fold holding_eq, auto)
+    qed
+  next
+    assume "n1 = Cs cs \<and> n2 = Th th"
+    with holding_es_th_cs
+    show ?thesis 
+      by (unfold s_RAG_def, fold holding_eq, auto)
+  qed
+qed
+
+end
+
+context valid_trace_p
+begin
+
+lemma RAG_es': "RAG (e # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
+                                                  else RAG s \<union> {(Th th, Cs cs)})"
+proof(cases "wq s cs = []")
+  case True
+  interpret vt_p: valid_trace_p_h using True
+    by (unfold_locales, simp)
+  show ?thesis by (simp add: vt_p.RAG_es vt_p.we) 
+next
+  case False
+  interpret vt_p: valid_trace_p_w using False
+    by (unfold_locales, simp)
+  show ?thesis by (simp add: vt_p.RAG_es vt_p.wne) 
+qed
+
+end
+
+lemma (in valid_trace_v_n) finite_waiting_set:
+  "finite {(Th th', Cs cs) |th'. next_th s th cs th'}"
+    by (simp add: waiting_set_eq)
+
+lemma (in valid_trace_v_n) finite_holding_set:
+  "finite {(Cs cs, Th th') |th'. next_th s th cs th'}"
+    by (simp add: holding_set_eq)
+
+lemma (in valid_trace_v_e) finite_waiting_set:
+  "finite {(Th th', Cs cs) |th'. next_th s th cs th'}"
+    by (simp add: waiting_set_eq)
+
+lemma (in valid_trace_v_e) finite_holding_set:
+  "finite {(Cs cs, Th th') |th'. next_th s th cs th'}"
+    by (simp add: holding_set_eq)
+
+context valid_trace_v
+begin
+
+lemma 
+  finite_RAG_kept:
+  assumes "finite (RAG s)"
+  shows "finite (RAG (e#s))"
+proof(cases "rest = []")
+  case True
+  interpret vt: valid_trace_v_e using True
+    by (unfold_locales, simp)
+  show ?thesis using assms
+    by  (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp)
+next
+  case False
+  interpret vt: valid_trace_v_n using False
+    by (unfold_locales, simp)
+  show ?thesis using assms
+    by  (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp)
+qed
+
+end
+
+context valid_trace_v_e
+begin 
+
+lemma 
+  acylic_RAG_kept:
+  assumes "acyclic (RAG s)"
+  shows "acyclic (RAG (e#s))"
+proof(rule acyclic_subset[OF assms])
+  show "RAG (e # s) \<subseteq> RAG s"
+      by (unfold RAG_es waiting_set_eq holding_set_eq, auto)
+qed
+
+end
+
+context valid_trace_v_n
+begin 
+
+lemma waiting_taker: "waiting s taker cs"
+  apply (unfold s_waiting_def, fold wq_def, unfold wq_s_cs taker_def)
+  using eq_wq' th'_in_inv wq'_def by fastforce
+
+lemma 
+  acylic_RAG_kept:
+  assumes "acyclic (RAG s)"
+  shows "acyclic (RAG (e#s))"
+proof -
+  have "acyclic ((RAG s - {(Cs cs, Th th)} - {(Th taker, Cs cs)}) \<union> 
+                 {(Cs cs, Th taker)})" (is "acyclic (?A \<union> _)")
+  proof -
+    from assms
+    have "acyclic ?A"
+       by (rule acyclic_subset, auto)
+    moreover have "(Th taker, Cs cs) \<notin> ?A^*"
+    proof
+      assume otherwise: "(Th taker, Cs cs) \<in> ?A^*"
+      hence "(Th taker, Cs cs) \<in> ?A^+"
+        by (unfold rtrancl_eq_or_trancl, auto)
+      from tranclD[OF this]
+      obtain cs' where h: "(Th taker, Cs cs') \<in> ?A" 
+                          "(Th taker, Cs cs') \<in> RAG s"
+        by (unfold s_RAG_def, auto)
+      from this(2) have "waiting s taker cs'" 
+        by (unfold s_RAG_def, fold waiting_eq, auto)
+      from waiting_unique[OF this waiting_taker]
+      have "cs' = cs" .
+      from h(1)[unfolded this] show False by auto
+    qed
+    ultimately show ?thesis by auto
+  qed
+  thus ?thesis 
+    by (unfold RAG_es waiting_set_eq holding_set_eq, simp)
+qed
+
+end
+
+context valid_trace_p_h
+begin
+
+lemma 
+  acylic_RAG_kept:
+  assumes "acyclic (RAG s)"
+  shows "acyclic (RAG (e#s))"
+proof -
+  have "acyclic (RAG s \<union> {(Cs cs, Th th)})" (is "acyclic (?A \<union> _)") 
+  proof -
+    from assms
+    have "acyclic ?A"
+       by (rule acyclic_subset, auto)
+    moreover have "(Th th, Cs cs) \<notin> ?A^*"
+    proof
+      assume otherwise: "(Th th, Cs cs) \<in> ?A^*"
+      hence "(Th th, Cs cs) \<in> ?A^+"
+        by (unfold rtrancl_eq_or_trancl, auto)
+      from tranclD[OF this]
+      obtain cs' where h: "(Th th, Cs cs') \<in> RAG s"
+        by (unfold s_RAG_def, auto)
+      hence "waiting s th cs'" 
+        by (unfold s_RAG_def, fold waiting_eq, auto)
+      with th_not_waiting show False by auto
+    qed
+    ultimately show ?thesis by auto
+  qed
+  thus ?thesis by (unfold RAG_es, simp)
+qed
+
+end
+
+context valid_trace_p_w
+begin
+
+lemma 
+  acylic_RAG_kept:
+  assumes "acyclic (RAG s)"
+  shows "acyclic (RAG (e#s))"
+proof -
+  have "acyclic (RAG s \<union> {(Th th, Cs cs)})" (is "acyclic (?A \<union> _)") 
+  proof -
+    from assms
+    have "acyclic ?A"
+       by (rule acyclic_subset, auto)
+    moreover have "(Cs cs, Th th) \<notin> ?A^*"
+    proof
+      assume otherwise: "(Cs cs, Th th) \<in> ?A^*"
+      from pip_e[unfolded is_p]
+      show False
+      proof(cases)
+        case (thread_P)
+        moreover from otherwise have "(Cs cs, Th th) \<in> ?A^+"
+            by (unfold rtrancl_eq_or_trancl, auto)
+        ultimately show ?thesis by auto
+      qed
+    qed
+    ultimately show ?thesis by auto
+  qed
+  thus ?thesis by (unfold RAG_es, simp)
+qed
+
+end
+
+context valid_trace
+begin
+
+lemma finite_RAG:
+  shows "finite (RAG s)"
+proof(induct rule:ind)
+  case Nil
+  show ?case 
+  by (auto simp: s_RAG_def cs_waiting_def 
+                   cs_holding_def wq_def acyclic_def)
+next
+  case (Cons s e)
+  interpret vt_e: valid_trace_e s e using Cons by simp
+  show ?case
+  proof(cases e)
+    case (Create th prio)
+    interpret vt: valid_trace_create s e th prio using Create
+      by (unfold_locales, simp)
+    show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
+  next
+    case (Exit th)
+    interpret vt: valid_trace_exit s e th using Exit
+      by (unfold_locales, simp)
+    show ?thesis using Cons by (simp add: vt.RAG_unchanged)
+  next
+    case (P th cs)
+    interpret vt: valid_trace_p s e th cs using P
+      by (unfold_locales, simp)
+    show ?thesis using Cons using vt.RAG_es' by auto 
+  next
+    case (V th cs)
+    interpret vt: valid_trace_v s e th cs using V
+      by (unfold_locales, simp)
+    show ?thesis using Cons by (simp add: vt.finite_RAG_kept) 
+  next
+    case (Set th prio)
+    interpret vt: valid_trace_set s e th prio using Set
+      by (unfold_locales, simp)
+    show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
+  qed
+qed
+
+lemma acyclic_RAG:
+  shows "acyclic (RAG s)"
+proof(induct rule:ind)
+  case Nil
+  show ?case 
+  by (auto simp: s_RAG_def cs_waiting_def 
+                   cs_holding_def wq_def acyclic_def)
+next
+  case (Cons s e)
+  interpret vt_e: valid_trace_e s e using Cons by simp
+  show ?case
+  proof(cases e)
+    case (Create th prio)
+    interpret vt: valid_trace_create s e th prio using Create
+      by (unfold_locales, simp)
+    show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
+  next
+    case (Exit th)
+    interpret vt: valid_trace_exit s e th using Exit
+      by (unfold_locales, simp)
+    show ?thesis using Cons by (simp add: vt.RAG_unchanged)
+  next
+    case (P th cs)
+    interpret vt: valid_trace_p s e th cs using P
+      by (unfold_locales, simp)
+    show ?thesis
+    proof(cases "wq s cs = []")
+      case True
+      then interpret vt_h: valid_trace_p_h s e th cs
+        by (unfold_locales, simp)
+      show ?thesis using Cons by (simp add: vt_h.acylic_RAG_kept) 
+    next
+      case False
+      then interpret vt_w: valid_trace_p_w s e th cs
+        by (unfold_locales, simp)
+      show ?thesis using Cons by (simp add: vt_w.acylic_RAG_kept) 
+    qed
+  next
+    case (V th cs)
+    interpret vt: valid_trace_v s e th cs using V
+      by (unfold_locales, simp)
+    show ?thesis
+    proof(cases "vt.rest = []")
+      case True
+      then interpret vt_e: valid_trace_v_e s e th cs
+        by (unfold_locales, simp)
+      show ?thesis by (simp add: Cons.hyps(2) vt_e.acylic_RAG_kept) 
+    next
+      case False
+      then interpret vt_n: valid_trace_v_n s e th cs
+        by (unfold_locales, simp)
+      show ?thesis by (simp add: Cons.hyps(2) vt_n.acylic_RAG_kept) 
+    qed
+  next
+    case (Set th prio)
+    interpret vt: valid_trace_set s e th prio using Set
+      by (unfold_locales, simp)
+    show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
+  qed
+qed
+
+lemma wf_RAG: "wf (RAG s)"
+proof(rule finite_acyclic_wf)
+  from finite_RAG show "finite (RAG s)" .
+next
+  from acyclic_RAG show "acyclic (RAG s)" .
+qed
+
+lemma sgv_wRAG: "single_valued (wRAG s)"
+  using waiting_unique
+  by (unfold single_valued_def wRAG_def, auto)
+
+lemma sgv_hRAG: "single_valued (hRAG s)"
+  using held_unique 
+  by (unfold single_valued_def hRAG_def, auto)
+
+lemma sgv_tRAG: "single_valued (tRAG s)"
+  by (unfold tRAG_def, rule single_valued_relcomp, 
+              insert sgv_wRAG sgv_hRAG, auto)
+
+lemma acyclic_tRAG: "acyclic (tRAG s)"
+proof(unfold tRAG_def, rule acyclic_compose)
+  show "acyclic (RAG s)" using acyclic_RAG .
+next
+  show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto
+next
+  show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto
+qed
+
+lemma unique_RAG: "\<lbrakk>(n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2"
+  apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq)
+  by(auto elim:waiting_unique held_unique)
+
+lemma sgv_RAG: "single_valued (RAG s)"
+  using unique_RAG by (auto simp:single_valued_def)
+
+lemma rtree_RAG: "rtree (RAG s)"
+  using sgv_RAG acyclic_RAG
+  by (unfold rtree_def rtree_axioms_def sgv_def, auto)
+
+end
+
+sublocale valid_trace < rtree_RAG: rtree "RAG s"
+proof
+  show "single_valued (RAG s)"
+  apply (intro_locales)
+    by (unfold single_valued_def, 
+        auto intro:unique_RAG)
+
+  show "acyclic (RAG s)"
+     by (rule acyclic_RAG)
+qed
+
+sublocale valid_trace < rtree_s: rtree "tRAG s"
+proof(unfold_locales)
+  from sgv_tRAG show "single_valued (tRAG s)" .
+next
+  from acyclic_tRAG show "acyclic (tRAG s)" .
+qed
+
+sublocale valid_trace < fsbtRAGs : fsubtree "RAG s"
+proof -
+  show "fsubtree (RAG s)"
+  proof(intro_locales)
+    show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] .
+  next
+    show "fsubtree_axioms (RAG s)"
+    proof(unfold fsubtree_axioms_def)
+      from wf_RAG show "wf (RAG s)" .
     qed
   qed
 qed
 
+context valid_trace
+begin
+
+lemma finite_subtree_threads:
+    "finite {th'. Th th' \<in> subtree (RAG s) (Th th)}" (is "finite ?A")
+proof -
+  have "?A = the_thread ` {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}"
+        by (auto, insert image_iff, fastforce)
+  moreover have "finite {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}"
+        (is "finite ?B")
+  proof -
+     have "?B = (subtree (RAG s) (Th th)) \<inter> {Th th' | th'. True}"
+      by auto
+     moreover have "... \<subseteq> (subtree (RAG s) (Th th))" by auto
+     moreover have "finite ..." by (simp add: finite_subtree) 
+     ultimately show ?thesis by auto
+  qed
+  ultimately show ?thesis by auto
+qed
+
+lemma le_cp:
+  shows "preced th s \<le> cp s th"
+  proof(unfold cp_alt_def, rule Max_ge)
+    show "finite (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})"
+      by (simp add: finite_subtree_threads)
+  next
+    show "preced th s \<in> the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)}"
+      by (simp add: subtree_def the_preced_def)   
+  qed
+
+lemma cp_le:
+  assumes th_in: "th \<in> threads s"
+  shows "cp s th \<le> Max (the_preced s ` threads s)"
+proof(unfold cp_alt_def, rule Max_f_mono)
+  show "finite (threads s)" by (simp add: finite_threads) 
+next
+  show " {th'. Th th' \<in> subtree (RAG s) (Th th)} \<noteq> {}"
+    using subtree_def by fastforce
+next
+  show "{th'. Th th' \<in> subtree (RAG s) (Th th)} \<subseteq> threads s"
+    using assms
+    by (smt Domain.DomainI dm_RAG_threads mem_Collect_eq 
+        node.inject(1) rtranclD subsetI subtree_def trancl_domain) 
+qed
+
+lemma max_cp_eq: 
+  shows "Max ((cp s) ` threads s) = Max (the_preced s ` threads s)"
+  (is "?L = ?R")
+proof -
+  have "?L \<le> ?R" 
+  proof(cases "threads s = {}")
+    case False
+    show ?thesis 
+      by (rule Max.boundedI, 
+          insert cp_le, 
+          auto simp:finite_threads False)
+  qed auto
+  moreover have "?R \<le> ?L"
+    by (rule Max_fg_mono, 
+        simp add: finite_threads,
+        simp add: le_cp the_preced_def)
+  ultimately show ?thesis by auto
+qed
+
+lemma max_cp_eq_the_preced:
+  shows "Max ((cp s) ` threads s) = Max (the_preced s ` threads s)"
+  using max_cp_eq using the_preced_def by presburger 
+
+lemma wf_RAG_converse: 
+  shows "wf ((RAG s)^-1)"
+proof(rule finite_acyclic_wf_converse)
+  from finite_RAG 
+  show "finite (RAG s)" .
+next
+  from acyclic_RAG
+  show "acyclic (RAG s)" .
+qed
+
+lemma chain_building:
+  assumes "node \<in> Domain (RAG s)"
+  obtains th' where "th' \<in> readys s" "(node, Th th') \<in> (RAG s)^+"
+proof -
+  from assms have "node \<in> Range ((RAG s)^-1)" by auto
+  from wf_base[OF wf_RAG_converse this]
+  obtain b where h_b: "(b, node) \<in> ((RAG s)\<inverse>)\<^sup>+" "\<forall>c. (c, b) \<notin> (RAG s)\<inverse>" by auto
+  obtain th' where eq_b: "b = Th th'"
+  proof(cases b)
+    case (Cs cs)
+    from h_b(1)[unfolded trancl_converse] 
+    have "(node, b) \<in> ((RAG s)\<^sup>+)" by auto
+    from tranclE[OF this]
+    obtain n where "(n, b) \<in> RAG s" by auto
+    from this[unfolded Cs]
+    obtain th1 where "waiting s th1 cs"
+      by (unfold s_RAG_def, fold waiting_eq, auto)
+    from waiting_holding[OF this]
+    obtain th2 where "holding s th2 cs" .
+    hence "(Cs cs, Th th2) \<in> RAG s"
+      by (unfold s_RAG_def, fold holding_eq, auto)
+    with h_b(2)[unfolded Cs, rule_format]
+    have False by auto
+    thus ?thesis by auto
+  qed auto
+  have "th' \<in> readys s" 
+  proof -
+    from h_b(2)[unfolded eq_b]
+    have "\<forall>cs. \<not> waiting s th' cs"
+      by (unfold s_RAG_def, fold waiting_eq, auto)
+    moreover have "th' \<in> threads s"
+    proof(rule rg_RAG_threads)
+      from tranclD[OF h_b(1), unfolded eq_b]
+      obtain z where "(z, Th th') \<in> (RAG s)" by auto
+      thus "Th th' \<in> Range (RAG s)" by auto
+    qed
+    ultimately show ?thesis by (auto simp:readys_def)
+  qed
+  moreover have "(node, Th th') \<in> (RAG s)^+" 
+    using h_b(1)[unfolded trancl_converse] eq_b by auto
+  ultimately show ?thesis using that by metis
+qed
+
+text {* \noindent
+  The following is just an instance of @{text "chain_building"}.
+*}
+lemma th_chain_to_ready:
+  assumes th_in: "th \<in> threads s"
+  shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (RAG s)^+)"
+proof(cases "th \<in> readys s")
+  case True
+  thus ?thesis by auto
+next
+  case False
+  from False and th_in have "Th th \<in> Domain (RAG s)" 
+    by (auto simp:readys_def s_waiting_def s_RAG_def wq_def cs_waiting_def Domain_def)
+  from chain_building [rule_format, OF this]
+  show ?thesis by auto
+qed
+
+end
+
+lemma count_rec1 [simp]: 
+  assumes "Q e"
+  shows "count Q (e#es) = Suc (count Q es)"
+  using assms
+  by (unfold count_def, auto)
+
+lemma count_rec2 [simp]: 
+  assumes "\<not>Q e"
+  shows "count Q (e#es) = (count Q es)"
+  using assms
+  by (unfold count_def, auto)
+
+lemma count_rec3 [simp]: 
+  shows "count Q [] =  0"
+  by (unfold count_def, auto)
+
+lemma cntP_simp1[simp]:
+  "cntP (P th cs'#s) th = cntP s th + 1"
+  by (unfold cntP_def, simp)
+
+lemma cntP_simp2[simp]:
+  assumes "th' \<noteq> th"
+  shows "cntP (P th cs'#s) th' = cntP s th'"
+  using assms
+  by (unfold cntP_def, simp)
+
+lemma cntP_simp3[simp]:
+  assumes "\<not> isP e"
+  shows "cntP (e#s) th' = cntP s th'"
+  using assms
+  by (unfold cntP_def, cases e, simp+)
+
+lemma cntV_simp1[simp]:
+  "cntV (V th cs'#s) th = cntV s th + 1"
+  by (unfold cntV_def, simp)
+
+lemma cntV_simp2[simp]:
+  assumes "th' \<noteq> th"
+  shows "cntV (V th cs'#s) th' = cntV s th'"
+  using assms
+  by (unfold cntV_def, simp)
+
+lemma cntV_simp3[simp]:
+  assumes "\<not> isV e"
+  shows "cntV (e#s) th' = cntV s th'"
+  using assms
+  by (unfold cntV_def, cases e, simp+)
+
+lemma cntP_diff_inv:
+  assumes "cntP (e#s) th \<noteq> cntP s th"
+  shows "isP e \<and> actor e = th"
+proof(cases e)
+  case (P th' pty)
+  show ?thesis
+  by (cases "(\<lambda>e. \<exists>cs. e = P th cs) (P th' pty)", 
+        insert assms P, auto simp:cntP_def)
+qed (insert assms, auto simp:cntP_def)
+  
+lemma cntV_diff_inv:
+  assumes "cntV (e#s) th \<noteq> cntV s th"
+  shows "isV e \<and> actor e = th"
+proof(cases e)
+  case (V th' pty)
+  show ?thesis
+  by (cases "(\<lambda>e. \<exists>cs. e = V th cs) (V th' pty)", 
+        insert assms V, auto simp:cntV_def)
+qed (insert assms, auto simp:cntV_def)
+
+lemma children_RAG_alt_def:
+  "children (RAG (s::state)) (Th th) = Cs ` {cs. holding s th cs}"
+  by (unfold s_RAG_def, auto simp:children_def holding_eq)
+
+fun the_cs :: "node \<Rightarrow> cs" where
+  "the_cs (Cs cs) = cs"
+
+lemma holdents_alt_def:
+  "holdents s th = the_cs ` (children (RAG (s::state)) (Th th))"
+  by (unfold children_RAG_alt_def holdents_def, simp add: image_image)
+
+lemma cntCS_alt_def:
+  "cntCS s th = card (children (RAG s) (Th th))"
+  apply (unfold children_RAG_alt_def cntCS_def holdents_def)
+  by (rule card_image[symmetric], auto simp:inj_on_def)
+
+context valid_trace
+begin
+
+lemma finite_holdents: "finite (holdents s th)"
+  by (unfold holdents_alt_def, insert finite_children, auto)
+  
+end
+
+context valid_trace_p_w
+begin
+
+lemma holding_s_holder: "holding s holder cs"
+  by (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto)
+
+lemma holding_es_holder: "holding (e#s) holder cs"
+  by (unfold s_holding_def, fold wq_def, unfold wq_es_cs wq_s_cs, auto)
+
+lemma holdents_es:
+  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") 
+proof -
+  { fix cs'
+    assume "cs' \<in> ?L"
+    hence h: "holding (e#s) th' cs'" by (auto simp:holdents_def)
+    have "holding s th' cs'"
+    proof(cases "cs' = cs")
+      case True
+      from held_unique[OF h[unfolded True] holding_es_holder]
+      have "th' = holder" .
+      thus ?thesis 
+        by (unfold True holdents_def, insert holding_s_holder, simp)
+    next
+      case False
+      hence "wq (e#s) cs' = wq s cs'" by simp
+      from h[unfolded s_holding_def, folded wq_def, unfolded this]
+      show ?thesis
+       by (unfold s_holding_def, fold wq_def, auto)
+    qed 
+    hence "cs' \<in> ?R" by (auto simp:holdents_def)
+  } moreover {
+    fix cs'
+    assume "cs' \<in> ?R"
+    hence h: "holding s th' cs'" by (auto simp:holdents_def)
+    have "holding (e#s) th' cs'"
+    proof(cases "cs' = cs")
+      case True
+      from held_unique[OF h[unfolded True] holding_s_holder]
+      have "th' = holder" .
+      thus ?thesis 
+        by (unfold True holdents_def, insert holding_es_holder, simp)
+    next
+      case False
+      hence "wq s cs' = wq (e#s) cs'" by simp
+      from h[unfolded s_holding_def, folded wq_def, unfolded this]
+      show ?thesis
+       by (unfold s_holding_def, fold wq_def, auto)
+    qed 
+    hence "cs' \<in> ?L" by (auto simp:holdents_def)
+  } ultimately show ?thesis by auto
+qed
+
+lemma cntCS_es_th[simp]: "cntCS (e#s) th' = cntCS s th'"
+ by (unfold cntCS_def holdents_es, simp)
+
+lemma th_not_ready_es: 
+  shows "th \<notin> readys (e#s)"
+  using waiting_es_th_cs 
+  by (unfold readys_def, auto)
+
+end
+  
+context valid_trace_p_h
+begin
+
+lemma th_not_waiting':
+  "\<not> waiting (e#s) th cs'"
+proof(cases "cs' = cs")
+  case True
+  show ?thesis
+    by (unfold True s_waiting_def, fold wq_def, unfold wq_es_cs', auto)
+next
+  case False
+  from th_not_waiting[of cs', unfolded s_waiting_def, folded wq_def]
+  show ?thesis
+    by (unfold s_waiting_def, fold wq_def, insert False, simp)
+qed
+
+lemma ready_th_es: 
+  shows "th \<in> readys (e#s)"
+  using th_not_waiting'
+  by (unfold readys_def, insert live_th_es, auto)
+
+lemma holdents_es_th:
+  "holdents (e#s) th = (holdents s th) \<union> {cs}" (is "?L = ?R")
+proof -
+  { fix cs'
+    assume "cs' \<in> ?L" 
+    hence "holding (e#s) th cs'"
+      by (unfold holdents_def, auto)
+    hence "cs' \<in> ?R"
+     by (cases rule:holding_esE, auto simp:holdents_def)
+  } moreover {
+    fix cs'
+    assume "cs' \<in> ?R"
+    hence "holding s th cs' \<or> cs' = cs" 
+      by (auto simp:holdents_def)
+    hence "cs' \<in> ?L"
+    proof
+      assume "holding s th cs'"
+      from holding_kept[OF this]
+      show ?thesis by (auto simp:holdents_def)
+    next
+      assume "cs' = cs"
+      thus ?thesis using holding_es_th_cs
+        by (unfold holdents_def, auto)
+    qed
+  } ultimately show ?thesis by auto
+qed
+
+lemma cntCS_es_th: "cntCS (e#s) th = cntCS s th + 1"
+proof -
+  have "card (holdents s th \<union> {cs}) = card (holdents s th) + 1"
+  proof(subst card_Un_disjoint)
+    show "holdents s th \<inter> {cs} = {}"
+      using not_holding_s_th_cs by (auto simp:holdents_def)
+  qed (auto simp:finite_holdents)
+  thus ?thesis
+   by (unfold cntCS_def holdents_es_th, simp)
+qed
+
+lemma no_holder: 
+  "\<not> holding s th' cs"
+proof
+  assume otherwise: "holding s th' cs"
+  from this[unfolded s_holding_def, folded wq_def, unfolded we]
+  show False by auto
+qed
+
+lemma holdents_es_th':
+  assumes "th' \<noteq> th"
+  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
+proof -
+  { fix cs'
+    assume "cs' \<in> ?L"
+    hence h_e: "holding (e#s) th' cs'" by (auto simp:holdents_def)
+    have "cs' \<noteq> cs"
+    proof
+      assume "cs' = cs"
+      from held_unique[OF h_e[unfolded this] holding_es_th_cs]
+      have "th' = th" .
+      with assms show False by simp
+    qed
+    from h_e[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp[OF this]]
+    have "th' \<in> set (wq s cs') \<and> th' = hd (wq s cs')" .
+    hence "cs' \<in> ?R" 
+      by (unfold holdents_def s_holding_def, fold wq_def, auto)
+  } moreover {
+    fix cs'
+    assume "cs' \<in> ?R"
+    hence "holding s th' cs'" by (auto simp:holdents_def)
+    from holding_kept[OF this]
+    have "holding (e # s) th' cs'" .
+    hence "cs' \<in> ?L"
+      by (unfold holdents_def, auto)
+  } ultimately show ?thesis by auto
+qed
+
+lemma cntCS_es_th'[simp]: 
+  assumes "th' \<noteq> th"
+  shows "cntCS (e#s) th' = cntCS s th'"
+  by (unfold cntCS_def holdents_es_th'[OF assms], simp)
+
+end
+
+context valid_trace_p
+begin
+
+lemma readys_kept1: 
+  assumes "th' \<noteq> th"
+  and "th' \<in> readys (e#s)"
+  shows "th' \<in> readys s"
+proof -
+  { fix cs'
+    assume wait: "waiting s th' cs'"
+    have n_wait: "\<not> waiting (e#s) th' cs'" 
+        using assms(2)[unfolded readys_def] by auto
+    have False
+    proof(cases "cs' = cs")
+      case False
+      with n_wait wait
+      show ?thesis 
+        by (unfold s_waiting_def, fold wq_def, auto)
+    next
+      case True
+      show ?thesis
+      proof(cases "wq s cs = []")
+        case True
+        then interpret vt: valid_trace_p_h
+          by (unfold_locales, simp)
+        show ?thesis using n_wait wait waiting_kept by auto 
+      next
+        case False
+        then interpret vt: valid_trace_p_w by (unfold_locales, simp)
+        show ?thesis using n_wait wait waiting_kept by blast 
+      qed
+    qed
+  } with assms(2) show ?thesis  
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_kept2: 
+  assumes "th' \<noteq> th"
+  and "th' \<in> readys s"
+  shows "th' \<in> readys (e#s)"
+proof -
+  { fix cs'
+    assume wait: "waiting (e#s) th' cs'"
+    have n_wait: "\<not> waiting s th' cs'" 
+        using assms(2)[unfolded readys_def] by auto
+    have False
+    proof(cases "cs' = cs")
+      case False
+      with n_wait wait
+      show ?thesis 
+        by (unfold s_waiting_def, fold wq_def, auto)
+    next
+      case True
+      show ?thesis
+      proof(cases "wq s cs = []")
+        case True
+        then interpret vt: valid_trace_p_h
+          by (unfold_locales, simp)
+        show ?thesis using n_wait vt.waiting_esE wait by blast 
+      next
+        case False
+        then interpret vt: valid_trace_p_w by (unfold_locales, simp)
+        show ?thesis using assms(1) n_wait vt.waiting_esE wait by auto 
+      qed
+    qed
+  } with assms(2) show ?thesis  
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_simp [simp]:
+  assumes "th' \<noteq> th"
+  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
+  using readys_kept1[OF assms] readys_kept2[OF assms]
+  by metis
+
+lemma cnp_cnv_cncs_kept:
+  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
+  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
+proof(cases "th' = th")
+  case True
+  note eq_th' = this
+  show ?thesis
+  proof(cases "wq s cs = []")
+    case True
+    then interpret vt: valid_trace_p_h by (unfold_locales, simp)
+    show ?thesis
+      using assms eq_th' is_p ready_th_s vt.cntCS_es_th vt.ready_th_es pvD_def by auto 
+  next
+    case False
+    then interpret vt: valid_trace_p_w by (unfold_locales, simp)
+    show ?thesis
+      using add.commute add.left_commute assms eq_th' is_p live_th_s 
+            ready_th_s vt.th_not_ready_es pvD_def
+      apply (auto)
+      by (fold is_p, simp)
+  qed
+next
+  case False
+  note h_False = False
+  thus ?thesis
+  proof(cases "wq s cs = []")
+    case True
+    then interpret vt: valid_trace_p_h by (unfold_locales, simp)
+    show ?thesis using assms
+      by (insert True h_False pvD_def, auto split:if_splits,unfold is_p, auto)
+  next
+    case False
+    then interpret vt: valid_trace_p_w by (unfold_locales, simp)
+    show ?thesis using assms
+      by (insert False h_False pvD_def, auto split:if_splits,unfold is_p, auto)
+  qed
+qed
+
 end
 
 
-
-lemma "RAG (e # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
-                                         else RAG s \<union> {(Th th, Cs cs)})"
-  proof(cases "wq s cs = []")
-    case True
-    from wq_es_cs[unfolded this]
-    have "th \<in> set (wq (e#s) cs)" "th = hd (wq (e#s) cs)" by auto
-    hence "holding (e#s) th cs"
-      using cs_holding_def holding_eq by blast 
-    thus 
+context valid_trace_v (* ccc *)
+begin
+
+lemma holding_th_cs_s: 
+  "holding s th cs" 
+ by  (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto)
+
+lemma th_ready_s [simp]: "th \<in> readys s"
+  using runing_th_s
+  by (unfold runing_def readys_def, auto)
+
+lemma th_live_s [simp]: "th \<in> threads s"
+  using th_ready_s by (unfold readys_def, auto)
+
+lemma th_ready_es [simp]: "th \<in> readys (e#s)"
+  using runing_th_s neq_t_th
+  by (unfold is_v runing_def readys_def, auto)
+
+lemma th_live_es [simp]: "th \<in> threads (e#s)"
+  using th_ready_es by (unfold readys_def, auto)
+
+lemma pvD_th_s[simp]: "pvD s th = 0"
+  by (unfold pvD_def, simp)
+
+lemma pvD_th_es[simp]: "pvD (e#s) th = 0"
+  by (unfold pvD_def, simp)
+
+lemma cntCS_s_th [simp]: "cntCS s th > 0"
+proof -
+  have "cs \<in> holdents s th" using holding_th_cs_s
+    by (unfold holdents_def, simp)
+  moreover have "finite (holdents s th)" using finite_holdents
+    by simp
+  ultimately show ?thesis
+    by (unfold cntCS_def, 
+        auto intro!:card_gt_0_iff[symmetric, THEN iffD1])
+qed
+
+end
+
+context valid_trace_v_n
+begin
+
+lemma not_ready_taker_s[simp]: 
+  "taker \<notin> readys s"
+  using waiting_taker
+  by (unfold readys_def, auto)
+
+lemma taker_live_s [simp]: "taker \<in> threads s"
+proof -
+  have "taker \<in> set wq'" by (simp add: eq_wq') 
+  from th'_in_inv[OF this]
+  have "taker \<in> set rest" .
+  hence "taker \<in> set (wq s cs)" by (simp add: wq_s_cs) 
+  thus ?thesis using wq_threads by auto 
+qed
+
+lemma taker_live_es [simp]: "taker \<in> threads (e#s)"
+  using taker_live_s threads_es by blast
+
+lemma taker_ready_es [simp]:
+  shows "taker \<in> readys (e#s)"
+proof -
+  { fix cs'
+    assume "waiting (e#s) taker cs'"
+    hence False
+    proof(cases rule:waiting_esE)
+      case 1
+      thus ?thesis using waiting_taker waiting_unique by auto 
+    qed simp
+  } thus ?thesis by (unfold readys_def, auto)
+qed
+
+lemma neq_taker_th: "taker \<noteq> th"
+  using th_not_waiting waiting_taker by blast
+
+lemma not_holding_taker_s_cs:
+  shows "\<not> holding s taker cs"
+  using holding_cs_eq_th neq_taker_th by auto
+
+lemma holdents_es_taker:
+  "holdents (e#s) taker = holdents s taker \<union> {cs}" (is "?L = ?R")
+proof -
+  { fix cs'
+    assume "cs' \<in> ?L"
+    hence "holding (e#s) taker cs'" by (auto simp:holdents_def)
+    hence "cs' \<in> ?R"
+    proof(cases rule:holding_esE)
+      case 2
+      thus ?thesis by (auto simp:holdents_def)
+    qed auto
+  } moreover {
+    fix cs'
+    assume "cs' \<in> ?R"
+    hence "holding s taker cs' \<or> cs' = cs" by (auto simp:holdents_def)
+    hence "cs' \<in> ?L" 
+    proof
+      assume "holding s taker cs'"
+      hence "holding (e#s) taker cs'" 
+          using holding_esI2 holding_taker by fastforce 
+      thus ?thesis by (auto simp:holdents_def)
+    next
+      assume "cs' = cs"
+      with holding_taker
+      show ?thesis by (auto simp:holdents_def)
+    qed
+  } ultimately show ?thesis by auto
+qed
+
+lemma cntCS_es_taker [simp]: "cntCS (e#s) taker = cntCS s taker + 1"
+proof -
+  have "card (holdents s taker \<union> {cs}) = card (holdents s taker) + 1"
+  proof(subst card_Un_disjoint)
+    show "holdents s taker \<inter> {cs} = {}"
+      using not_holding_taker_s_cs by (auto simp:holdents_def)
+  qed (auto simp:finite_holdents)
+  thus ?thesis 
+    by (unfold cntCS_def, insert holdents_es_taker, simp)
+qed
+
+lemma pvD_taker_s[simp]: "pvD s taker = 1"
+  by (unfold pvD_def, simp)
+
+lemma pvD_taker_es[simp]: "pvD (e#s) taker = 0"
+  by (unfold pvD_def, simp)  
+
+lemma pvD_th_s[simp]: "pvD s th = 0"
+  by (unfold pvD_def, simp)
+
+lemma pvD_th_es[simp]: "pvD (e#s) th = 0"
+  by (unfold pvD_def, simp)
+
+lemma holdents_es_th:
+  "holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R")
+proof -
+  { fix cs'
+    assume "cs' \<in> ?L"
+    hence "holding (e#s) th cs'" by (auto simp:holdents_def)
+    hence "cs' \<in> ?R"
+    proof(cases rule:holding_esE)
+      case 2
+      thus ?thesis by (auto simp:holdents_def)
+    qed (insert neq_taker_th, auto)
+  } moreover {
+    fix cs'
+    assume "cs' \<in> ?R"
+    hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def)
+    from holding_esI2[OF this]
+    have "cs' \<in> ?L" by (auto simp:holdents_def)
+  } ultimately show ?thesis by auto
+qed
+
+lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1"
+proof -
+  have "card (holdents s th - {cs}) = card (holdents s th) - 1"
+  proof -
+    have "cs \<in> holdents s th" using holding_th_cs_s
+      by (auto simp:holdents_def)
+    moreover have "finite (holdents s th)"
+        by (simp add: finite_holdents) 
+    ultimately show ?thesis by auto
   qed
+  thus ?thesis by (unfold cntCS_def holdents_es_th)
+qed
+
+lemma holdents_kept:
+  assumes "th' \<noteq> taker"
+  and "th' \<noteq> th"
+  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
+proof -
+  { fix cs'
+    assume h: "cs' \<in> ?L"
+    have "cs' \<in> ?R"
+    proof(cases "cs' = cs")
+      case False
+      hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
+      from h have "holding (e#s) th' cs'" by (auto simp:holdents_def)
+      from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
+      show ?thesis
+        by (unfold holdents_def s_holding_def, fold wq_def, auto)
+    next
+      case True
+      from h[unfolded this]
+      have "holding (e#s) th' cs" by (auto simp:holdents_def)
+      from held_unique[OF this holding_taker]
+      have "th' = taker" .
+      with assms show ?thesis by auto
+    qed
+  } moreover {
+    fix cs'
+    assume h: "cs' \<in> ?R"
+    have "cs' \<in> ?L"
+    proof(cases "cs' = cs")
+      case False
+      hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
+      from h have "holding s th' cs'" by (auto simp:holdents_def)
+      from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
+      show ?thesis
+        by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp)
+    next
+      case True
+      from h[unfolded this]
+      have "holding s th' cs" by (auto simp:holdents_def)
+      from held_unique[OF this holding_th_cs_s]
+      have "th' = th" .
+      with assms show ?thesis by auto
+    qed
+  } ultimately show ?thesis by auto
+qed
+
+lemma cntCS_kept [simp]:
+  assumes "th' \<noteq> taker"
+  and "th' \<noteq> th"
+  shows "cntCS (e#s) th' = cntCS s th'"
+  by (unfold cntCS_def holdents_kept[OF assms], simp)
+
+lemma readys_kept1: 
+  assumes "th' \<noteq> taker"
+  and "th' \<in> readys (e#s)"
+  shows "th' \<in> readys s"
+proof -
+  { fix cs'
+    assume wait: "waiting s th' cs'"
+    have n_wait: "\<not> waiting (e#s) th' cs'" 
+        using assms(2)[unfolded readys_def] by auto
+    have False
+    proof(cases "cs' = cs")
+      case False
+      with n_wait wait
+      show ?thesis 
+        by (unfold s_waiting_def, fold wq_def, auto)
+    next
+      case True
+      have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" 
+        using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] .
+      moreover have "\<not> (th' \<in> set rest \<and> th' \<noteq> hd (taker # rest'))" 
+        using n_wait[unfolded True s_waiting_def, folded wq_def, 
+                    unfolded wq_es_cs set_wq', unfolded eq_wq'] .
+      ultimately have "th' = taker" by auto
+      with assms(1)
+      show ?thesis by simp
+    qed
+  } with assms(2) show ?thesis  
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_kept2: 
+  assumes "th' \<noteq> taker"
+  and "th' \<in> readys s"
+  shows "th' \<in> readys (e#s)"
+proof -
+  { fix cs'
+    assume wait: "waiting (e#s) th' cs'"
+    have n_wait: "\<not> waiting s th' cs'" 
+        using assms(2)[unfolded readys_def] by auto
+    have False
+    proof(cases "cs' = cs")
+      case False
+      with n_wait wait
+      show ?thesis 
+        by (unfold s_waiting_def, fold wq_def, auto)
+    next
+      case True
+      have "th' \<in> set rest \<and> th' \<noteq> hd (taker # rest')"
+          using  wait [unfolded True s_waiting_def, folded wq_def, 
+                    unfolded wq_es_cs set_wq', unfolded eq_wq']  .
+      moreover have "\<not> (th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest))"
+          using n_wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] .
+      ultimately have "th' = taker" by auto
+      with assms(1)
+      show ?thesis by simp
+    qed
+  } with assms(2) show ?thesis  
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_simp [simp]:
+  assumes "th' \<noteq> taker"
+  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
+  using readys_kept1[OF assms] readys_kept2[OF assms]
+  by metis
+
+lemma cnp_cnv_cncs_kept:
+  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
+  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
+proof -
+  { assume eq_th': "th' = taker"
+    have ?thesis
+      apply (unfold eq_th' pvD_taker_es cntCS_es_taker)
+      by (insert neq_taker_th assms[unfolded eq_th'], unfold is_v, simp)
+  } moreover {
+    assume eq_th': "th' = th"
+    have ?thesis 
+      apply (unfold eq_th' pvD_th_es cntCS_es_th)
+      by (insert assms[unfolded eq_th'], unfold is_v, simp)
+  } moreover {
+    assume h: "th' \<noteq> taker" "th' \<noteq> th"
+    have ?thesis using assms
+      apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp)
+      by (fold is_v, unfold pvD_def, simp)
+  } ultimately show ?thesis by metis
+qed
+
+end
+
+context valid_trace_v_e
+begin
+
+lemma holdents_es_th:
+  "holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R")
+proof -
+  { fix cs'
+    assume "cs' \<in> ?L"
+    hence "holding (e#s) th cs'" by (auto simp:holdents_def)
+    hence "cs' \<in> ?R"
+    proof(cases rule:holding_esE)
+      case 1
+      thus ?thesis by (auto simp:holdents_def)
+    qed 
+  } moreover {
+    fix cs'
+    assume "cs' \<in> ?R"
+    hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def)
+    from holding_esI2[OF this]
+    have "cs' \<in> ?L" by (auto simp:holdents_def)
+  } ultimately show ?thesis by auto
+qed
+
+lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1"
+proof -
+  have "card (holdents s th - {cs}) = card (holdents s th) - 1"
+  proof -
+    have "cs \<in> holdents s th" using holding_th_cs_s
+      by (auto simp:holdents_def)
+    moreover have "finite (holdents s th)"
+        by (simp add: finite_holdents) 
+    ultimately show ?thesis by auto
+  qed
+  thus ?thesis by (unfold cntCS_def holdents_es_th)
+qed
+
+lemma holdents_kept:
+  assumes "th' \<noteq> th"
+  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
+proof -
+  { fix cs'
+    assume h: "cs' \<in> ?L"
+    have "cs' \<in> ?R"
+    proof(cases "cs' = cs")
+      case False
+      hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
+      from h have "holding (e#s) th' cs'" by (auto simp:holdents_def)
+      from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
+      show ?thesis
+        by (unfold holdents_def s_holding_def, fold wq_def, auto)
+    next
+      case True
+      from h[unfolded this]
+      have "holding (e#s) th' cs" by (auto simp:holdents_def)
+      from this[unfolded s_holding_def, folded wq_def, 
+            unfolded wq_es_cs nil_wq']
+      show ?thesis by auto
+    qed
+  } moreover {
+    fix cs'
+    assume h: "cs' \<in> ?R"
+    have "cs' \<in> ?L"
+    proof(cases "cs' = cs")
+      case False
+      hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
+      from h have "holding s th' cs'" by (auto simp:holdents_def)
+      from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
+      show ?thesis
+        by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp)
+    next
+      case True
+      from h[unfolded this]
+      have "holding s th' cs" by (auto simp:holdents_def)
+      from held_unique[OF this holding_th_cs_s]
+      have "th' = th" .
+      with assms show ?thesis by auto
+    qed
+  } ultimately show ?thesis by auto
+qed
+
+lemma cntCS_kept [simp]:
+  assumes "th' \<noteq> th"
+  shows "cntCS (e#s) th' = cntCS s th'"
+  by (unfold cntCS_def holdents_kept[OF assms], simp)
+
+lemma readys_kept1: 
+  assumes "th' \<in> readys (e#s)"
+  shows "th' \<in> readys s"
+proof -
+  { fix cs'
+    assume wait: "waiting s th' cs'"
+    have n_wait: "\<not> waiting (e#s) th' cs'" 
+        using assms(1)[unfolded readys_def] by auto
+    have False
+    proof(cases "cs' = cs")
+      case False
+      with n_wait wait
+      show ?thesis 
+        by (unfold s_waiting_def, fold wq_def, auto)
+    next
+      case True
+      have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" 
+        using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] . 
+      hence "th' \<in> set rest" by auto
+      with set_wq' have "th' \<in> set wq'" by metis
+      with nil_wq' show ?thesis by simp
+    qed
+  } thus ?thesis using assms
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_kept2: 
+  assumes "th' \<in> readys s"
+  shows "th' \<in> readys (e#s)"
+proof -
+  { fix cs'
+    assume wait: "waiting (e#s) th' cs'"
+    have n_wait: "\<not> waiting s th' cs'" 
+        using assms[unfolded readys_def] by auto
+    have False
+    proof(cases "cs' = cs")
+      case False
+      with n_wait wait
+      show ?thesis 
+        by (unfold s_waiting_def, fold wq_def, auto)
+    next
+      case True
+      have "th' \<in> set [] \<and> th' \<noteq> hd []"
+        using wait[unfolded True s_waiting_def, folded wq_def, 
+              unfolded wq_es_cs nil_wq'] .
+      thus ?thesis by simp
+    qed
+  } with assms show ?thesis  
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_simp [simp]:
+  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
+  using readys_kept1[OF assms] readys_kept2[OF assms]
+  by metis
+
+lemma cnp_cnv_cncs_kept:
+  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
+  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
+proof -
+  {
+    assume eq_th': "th' = th"
+    have ?thesis 
+      apply (unfold eq_th' pvD_th_es cntCS_es_th)
+      by (insert assms[unfolded eq_th'], unfold is_v, simp)
+  } moreover {
+    assume h: "th' \<noteq> th"
+    have ?thesis using assms
+      apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp)
+      by (fold is_v, unfold pvD_def, simp)
+  } ultimately show ?thesis by metis
+qed
+
+end
+
+context valid_trace_v
+begin
+
+lemma cnp_cnv_cncs_kept:
+  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
+  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
+proof(cases "rest = []")
+  case True
+  then interpret vt: valid_trace_v_e by (unfold_locales, simp)
+  show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast 
+next
+  case False
+  then interpret vt: valid_trace_v_n by (unfold_locales, simp)
+  show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast 
+qed
+
 end
 
-text {* 
-  The following @{text "step_RAG_p"} lemma charaterizes how @{text "RAG"} is changed
-  with the happening of @{text "P"}-events:
-*}
-lemma step_RAG_p:
-  "vt (P th cs#s) \<Longrightarrow>
-  RAG (P th cs # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
-                                             else RAG s \<union> {(Th th, Cs cs)})"
-  apply(simp only: s_RAG_def wq_def)
-  apply (auto split:list.splits prod.splits simp:Let_def wq_def cs_waiting_def cs_holding_def)
-  apply(case_tac "csa = cs", auto)
-  apply(fold wq_def)
-  apply(drule_tac step_back_step)
-  apply(ind_cases " step s (P (hd (wq s cs)) cs)")
-  apply(simp add:s_RAG_def wq_def cs_holding_def)
-  apply(auto)
-  done
-
-
-
-end
\ No newline at end of file
+context valid_trace_create
+begin
+
+lemma th_not_live_s [simp]: "th \<notin> threads s"
+proof -
+  from pip_e[unfolded is_create]
+  show ?thesis by (cases, simp)
+qed
+
+lemma th_not_ready_s [simp]: "th \<notin> readys s"
+  using th_not_live_s by (unfold readys_def, simp)
+
+lemma th_live_es [simp]: "th \<in> threads (e#s)"
+  by (unfold is_create, simp)
+
+lemma not_waiting_th_s [simp]: "\<not> waiting s th cs'"
+proof
+  assume "waiting s th cs'"
+  from this[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
+  have "th \<in> set (wq s cs')" by auto
+  from wq_threads[OF this] have "th \<in> threads s" .
+  with th_not_live_s show False by simp
+qed
+
+lemma not_holding_th_s [simp]: "\<not> holding s th cs'"
+proof
+  assume "holding s th cs'"
+  from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
+  have "th \<in> set (wq s cs')" by auto
+  from wq_threads[OF this] have "th \<in> threads s" .
+  with th_not_live_s show False by simp
+qed
+
+lemma not_waiting_th_es [simp]: "\<not> waiting (e#s) th cs'"
+proof
+  assume "waiting (e # s) th cs'"
+  from this[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
+  have "th \<in> set (wq s cs')" by auto
+  from wq_threads[OF this] have "th \<in> threads s" .
+  with th_not_live_s show False by simp
+qed
+
+lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'"
+proof
+  assume "holding (e # s) th cs'"
+  from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
+  have "th \<in> set (wq s cs')" by auto
+  from wq_threads[OF this] have "th \<in> threads s" .
+  with th_not_live_s show False by simp
+qed
+
+lemma ready_th_es [simp]: "th \<in> readys (e#s)"
+  by (simp add:readys_def)
+
+lemma holdents_th_s: "holdents s th = {}"
+  by (unfold holdents_def, auto)
+
+lemma holdents_th_es: "holdents (e#s) th = {}"
+  by (unfold holdents_def, auto)
+
+lemma cntCS_th_s [simp]: "cntCS s th = 0"
+  by (unfold cntCS_def, simp add:holdents_th_s)
+
+lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0"
+  by (unfold cntCS_def, simp add:holdents_th_es)
+
+lemma pvD_th_s [simp]: "pvD s th = 0"
+  by (unfold pvD_def, simp)
+
+lemma pvD_th_es [simp]: "pvD (e#s) th = 0"
+  by (unfold pvD_def, simp)
+
+lemma holdents_kept:
+  assumes "th' \<noteq> th"
+  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
+proof -
+  { fix cs'
+    assume h: "cs' \<in> ?L"
+    hence "cs' \<in> ?R"
+      by (unfold holdents_def s_holding_def, fold wq_def, 
+             unfold wq_neq_simp, auto)
+  } moreover {
+    fix cs'
+    assume h: "cs' \<in> ?R"
+    hence "cs' \<in> ?L"
+      by (unfold holdents_def s_holding_def, fold wq_def, 
+             unfold wq_neq_simp, auto)
+  } ultimately show ?thesis by auto
+qed
+
+lemma cntCS_kept [simp]:
+  assumes "th' \<noteq> th"
+  shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
+  using holdents_kept[OF assms]
+  by (unfold cntCS_def, simp)
+
+lemma readys_kept1: 
+  assumes "th' \<noteq> th"
+  and "th' \<in> readys (e#s)"
+  shows "th' \<in> readys s"
+proof -
+  { fix cs'
+    assume wait: "waiting s th' cs'"
+    have n_wait: "\<not> waiting (e#s) th' cs'" 
+      using assms by (auto simp:readys_def)
+    from wait[unfolded s_waiting_def, folded wq_def]
+         n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
+    have False by auto
+  } thus ?thesis using assms
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_kept2: 
+  assumes "th' \<noteq> th"
+  and "th' \<in> readys s"
+  shows "th' \<in> readys (e#s)"
+proof -
+  { fix cs'
+    assume wait: "waiting (e#s) th' cs'"
+    have n_wait: "\<not> waiting s th' cs'"
+      using assms(2) by (auto simp:readys_def)
+    from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
+         n_wait[unfolded s_waiting_def, folded wq_def]
+    have False by auto
+  } with assms show ?thesis  
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_simp [simp]:
+  assumes "th' \<noteq> th"
+  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
+  using readys_kept1[OF assms] readys_kept2[OF assms]
+  by metis
+
+lemma pvD_kept [simp]:
+  assumes "th' \<noteq> th"
+  shows "pvD (e#s) th' = pvD s th'"
+  using assms
+  by (unfold pvD_def, simp)
+
+lemma cnp_cnv_cncs_kept:
+  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
+  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
+proof -
+  {
+    assume eq_th': "th' = th"
+    have ?thesis using assms
+      by (unfold eq_th', simp, unfold is_create, simp)
+  } moreover {
+    assume h: "th' \<noteq> th"
+    hence ?thesis using assms
+      by (simp, simp add:is_create)
+  } ultimately show ?thesis by metis
+qed
+
+end
+
+context valid_trace_exit
+begin
+
+lemma th_live_s [simp]: "th \<in> threads s"
+proof -
+  from pip_e[unfolded is_exit]
+  show ?thesis
+  by (cases, unfold runing_def readys_def, simp)
+qed
+
+lemma th_ready_s [simp]: "th \<in> readys s"
+proof -
+  from pip_e[unfolded is_exit]
+  show ?thesis
+  by (cases, unfold runing_def, simp)
+qed
+
+lemma th_not_live_es [simp]: "th \<notin> threads (e#s)"
+  by (unfold is_exit, simp)
+
+lemma not_holding_th_s [simp]: "\<not> holding s th cs'"
+proof -
+  from pip_e[unfolded is_exit]
+  show ?thesis 
+   by (cases, unfold holdents_def, auto)
+qed
+
+lemma cntCS_th_s [simp]: "cntCS s th = 0"
+proof -
+  from pip_e[unfolded is_exit]
+  show ?thesis 
+   by (cases, unfold cntCS_def, simp)
+qed
+
+lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'"
+proof
+  assume "holding (e # s) th cs'"
+  from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
+  have "holding s th cs'" 
+    by (unfold s_holding_def, fold wq_def, auto)
+  with not_holding_th_s 
+  show False by simp
+qed
+
+lemma ready_th_es [simp]: "th \<notin> readys (e#s)"
+  by (simp add:readys_def)
+
+lemma holdents_th_s: "holdents s th = {}"
+  by (unfold holdents_def, auto)
+
+lemma holdents_th_es: "holdents (e#s) th = {}"
+  by (unfold holdents_def, auto)
+
+lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0"
+  by (unfold cntCS_def, simp add:holdents_th_es)
+
+lemma pvD_th_s [simp]: "pvD s th = 0"
+  by (unfold pvD_def, simp)
+
+lemma pvD_th_es [simp]: "pvD (e#s) th = 0"
+  by (unfold pvD_def, simp)
+
+lemma holdents_kept:
+  assumes "th' \<noteq> th"
+  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
+proof -
+  { fix cs'
+    assume h: "cs' \<in> ?L"
+    hence "cs' \<in> ?R"
+      by (unfold holdents_def s_holding_def, fold wq_def, 
+             unfold wq_neq_simp, auto)
+  } moreover {
+    fix cs'
+    assume h: "cs' \<in> ?R"
+    hence "cs' \<in> ?L"
+      by (unfold holdents_def s_holding_def, fold wq_def, 
+             unfold wq_neq_simp, auto)
+  } ultimately show ?thesis by auto
+qed
+
+lemma cntCS_kept [simp]:
+  assumes "th' \<noteq> th"
+  shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
+  using holdents_kept[OF assms]
+  by (unfold cntCS_def, simp)
+
+lemma readys_kept1: 
+  assumes "th' \<noteq> th"
+  and "th' \<in> readys (e#s)"
+  shows "th' \<in> readys s"
+proof -
+  { fix cs'
+    assume wait: "waiting s th' cs'"
+    have n_wait: "\<not> waiting (e#s) th' cs'" 
+      using assms by (auto simp:readys_def)
+    from wait[unfolded s_waiting_def, folded wq_def]
+         n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
+    have False by auto
+  } thus ?thesis using assms
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_kept2: 
+  assumes "th' \<noteq> th"
+  and "th' \<in> readys s"
+  shows "th' \<in> readys (e#s)"
+proof -
+  { fix cs'
+    assume wait: "waiting (e#s) th' cs'"
+    have n_wait: "\<not> waiting s th' cs'"
+      using assms(2) by (auto simp:readys_def)
+    from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
+         n_wait[unfolded s_waiting_def, folded wq_def]
+    have False by auto
+  } with assms show ?thesis  
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_simp [simp]:
+  assumes "th' \<noteq> th"
+  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
+  using readys_kept1[OF assms] readys_kept2[OF assms]
+  by metis
+
+lemma pvD_kept [simp]:
+  assumes "th' \<noteq> th"
+  shows "pvD (e#s) th' = pvD s th'"
+  using assms
+  by (unfold pvD_def, simp)
+
+lemma cnp_cnv_cncs_kept:
+  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
+  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
+proof -
+  {
+    assume eq_th': "th' = th"
+    have ?thesis using assms
+      by (unfold eq_th', simp, unfold is_exit, simp)
+  } moreover {
+    assume h: "th' \<noteq> th"
+    hence ?thesis using assms
+      by (simp, simp add:is_exit)
+  } ultimately show ?thesis by metis
+qed
+
+end
+
+context valid_trace_set
+begin
+
+lemma th_live_s [simp]: "th \<in> threads s"
+proof -
+  from pip_e[unfolded is_set]
+  show ?thesis
+  by (cases, unfold runing_def readys_def, simp)
+qed
+
+lemma th_ready_s [simp]: "th \<in> readys s"
+proof -
+  from pip_e[unfolded is_set]
+  show ?thesis
+  by (cases, unfold runing_def, simp)
+qed
+
+lemma th_not_live_es [simp]: "th \<in> threads (e#s)"
+  by (unfold is_set, simp)
+
+
+lemma holdents_kept:
+  shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
+proof -
+  { fix cs'
+    assume h: "cs' \<in> ?L"
+    hence "cs' \<in> ?R"
+      by (unfold holdents_def s_holding_def, fold wq_def, 
+             unfold wq_neq_simp, auto)
+  } moreover {
+    fix cs'
+    assume h: "cs' \<in> ?R"
+    hence "cs' \<in> ?L"
+      by (unfold holdents_def s_holding_def, fold wq_def, 
+             unfold wq_neq_simp, auto)
+  } ultimately show ?thesis by auto
+qed
+
+lemma cntCS_kept [simp]:
+  shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
+  using holdents_kept
+  by (unfold cntCS_def, simp)
+
+lemma threads_kept[simp]:
+  "threads (e#s) = threads s"
+  by (unfold is_set, simp)
+
+lemma readys_kept1: 
+  assumes "th' \<in> readys (e#s)"
+  shows "th' \<in> readys s"
+proof -
+  { fix cs'
+    assume wait: "waiting s th' cs'"
+    have n_wait: "\<not> waiting (e#s) th' cs'" 
+      using assms by (auto simp:readys_def)
+    from wait[unfolded s_waiting_def, folded wq_def]
+         n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
+    have False by auto
+  } moreover have "th' \<in> threads s" 
+    using assms[unfolded readys_def] by auto
+  ultimately show ?thesis 
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_kept2: 
+  assumes "th' \<in> readys s"
+  shows "th' \<in> readys (e#s)"
+proof -
+  { fix cs'
+    assume wait: "waiting (e#s) th' cs'"
+    have n_wait: "\<not> waiting s th' cs'"
+      using assms by (auto simp:readys_def)
+    from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
+         n_wait[unfolded s_waiting_def, folded wq_def]
+    have False by auto
+  } with assms show ?thesis  
+    by (unfold readys_def, auto)
+qed
+
+lemma readys_simp [simp]:
+  shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
+  using readys_kept1 readys_kept2
+  by metis
+
+lemma pvD_kept [simp]:
+  shows "pvD (e#s) th' = pvD s th'"
+  by (unfold pvD_def, simp)
+
+lemma cnp_cnv_cncs_kept:
+  assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
+  shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
+  using assms
+  by (unfold is_set, simp, fold is_set, simp)
+
+end
+
+context valid_trace
+begin
+
+lemma cnp_cnv_cncs: "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
+proof(induct rule:ind)
+  case Nil
+  thus ?case 
+    by (unfold cntP_def cntV_def pvD_def cntCS_def holdents_def 
+              s_holding_def, simp)
+next
+  case (Cons s e)
+  interpret vt_e: valid_trace_e s e using Cons by simp
+  show ?case
+  proof(cases e)
+    case (Create th prio)
+    interpret vt_create: valid_trace_create s e th prio 
+      using Create by (unfold_locales, simp)
+    show ?thesis using Cons by (simp add: vt_create.cnp_cnv_cncs_kept) 
+  next
+    case (Exit th)
+    interpret vt_exit: valid_trace_exit s e th  
+        using Exit by (unfold_locales, simp)
+   show ?thesis using Cons by (simp add: vt_exit.cnp_cnv_cncs_kept) 
+  next
+    case (P th cs)
+    interpret vt_p: valid_trace_p s e th cs using P by (unfold_locales, simp)
+    show ?thesis using Cons by (simp add: vt_p.cnp_cnv_cncs_kept) 
+  next
+    case (V th cs)
+    interpret vt_v: valid_trace_v s e th cs using V by (unfold_locales, simp)
+    show ?thesis using Cons by (simp add: vt_v.cnp_cnv_cncs_kept) 
+  next
+    case (Set th prio)
+    interpret vt_set: valid_trace_set s e th prio
+        using Set by (unfold_locales, simp)
+    show ?thesis using Cons by (simp add: vt_set.cnp_cnv_cncs_kept) 
+  qed
+qed
+
+lemma not_thread_holdents:
+  assumes not_in: "th \<notin> threads s" 
+  shows "holdents s th = {}"
+proof -
+  { fix cs
+    assume "cs \<in> holdents s th"
+    hence "holding s th cs" by (auto simp:holdents_def)
+    from this[unfolded s_holding_def, folded wq_def]
+    have "th \<in> set (wq s cs)" by auto
+    with wq_threads have "th \<in> threads s" by auto
+    with assms
+    have False by simp
+  } thus ?thesis by auto
+qed
+
+lemma not_thread_cncs:
+  assumes not_in: "th \<notin> threads s" 
+  shows "cntCS s th = 0"
+  using not_thread_holdents[OF assms]
+  by (simp add:cntCS_def)
+
+lemma cnp_cnv_eq:
+  assumes "th \<notin> threads s"
+  shows "cntP s th = cntV s th"
+  using assms cnp_cnv_cncs not_thread_cncs pvD_def
+  by (auto)
+
+lemma runing_unique:
+  assumes runing_1: "th1 \<in> runing s"
+  and runing_2: "th2 \<in> runing s"
+  shows "th1 = th2"
+proof -
+  from runing_1 and runing_2 have "cp s th1 = cp s th2"
+    unfolding runing_def by auto
+  from this[unfolded cp_alt_def]
+  have eq_max: 
+    "Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th1)}) =
+     Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th2)})" 
+        (is "Max ?L = Max ?R") .
+  have "Max ?L \<in> ?L"
+  proof(rule Max_in)
+    show "finite ?L" by (simp add: finite_subtree_threads)
+  next
+    show "?L \<noteq> {}" using subtree_def by fastforce 
+  qed
+  then obtain th1' where 
+    h_1: "Th th1' \<in> subtree (RAG s) (Th th1)" "the_preced s th1' = Max ?L"
+    by auto
+  have "Max ?R \<in> ?R"
+  proof(rule Max_in)
+    show "finite ?R" by (simp add: finite_subtree_threads)
+  next
+    show "?R \<noteq> {}" using subtree_def by fastforce 
+  qed
+  then obtain th2' where 
+    h_2: "Th th2' \<in> subtree (RAG s) (Th th2)" "the_preced s th2' = Max ?R"
+    by auto
+  have "th1' = th2'"
+  proof(rule preced_unique)
+    from h_1(1)
+    show "th1' \<in> threads s"
+    proof(cases rule:subtreeE)
+      case 1
+      hence "th1' = th1" by simp
+      with runing_1 show ?thesis by (auto simp:runing_def readys_def)
+    next
+      case 2
+      from this(2)
+      have "(Th th1', Th th1) \<in> (RAG s)^+" by (auto simp:ancestors_def)
+      from tranclD[OF this]
+      have "(Th th1') \<in> Domain (RAG s)" by auto
+      from dm_RAG_threads[OF this] show ?thesis .
+    qed
+  next
+    from h_2(1)
+    show "th2' \<in> threads s"
+    proof(cases rule:subtreeE)
+      case 1
+      hence "th2' = th2" by simp
+      with runing_2 show ?thesis by (auto simp:runing_def readys_def)
+    next
+      case 2
+      from this(2)
+      have "(Th th2', Th th2) \<in> (RAG s)^+" by (auto simp:ancestors_def)
+      from tranclD[OF this]
+      have "(Th th2') \<in> Domain (RAG s)" by auto
+      from dm_RAG_threads[OF this] show ?thesis .
+    qed
+  next
+    have "the_preced s th1' = the_preced s th2'" 
+     using eq_max h_1(2) h_2(2) by metis
+    thus "preced th1' s = preced th2' s" by (simp add:the_preced_def)
+  qed
+  from h_1(1)[unfolded this]
+  have star1: "(Th th2', Th th1) \<in> (RAG s)^*" by (auto simp:subtree_def)
+  from h_2(1)[unfolded this]
+  have star2: "(Th th2', Th th2) \<in> (RAG s)^*" by (auto simp:subtree_def)
+  from star_rpath[OF star1] obtain xs1 
+    where rp1: "rpath (RAG s) (Th th2') xs1 (Th th1)"
+    by auto
+  from star_rpath[OF star2] obtain xs2 
+    where rp2: "rpath (RAG s) (Th th2') xs2 (Th th2)"
+    by auto
+  show ?thesis
+  proof(cases rule:rtree_RAG.rpath_overlap[OF rp1 rp2])
+    case (1 xs')
+    moreover have "xs' = []"
+    proof(rule ccontr)
+      assume otherwise: "xs' \<noteq> []"
+      find_theorems rpath "_@_"
+    qed
+    ultimately have "xs2 = xs1" by simp
+    from rpath_dest_eq[OF rp1 rp2[unfolded this]]
+    show ?thesis by simp
+  next
+    case (2 xs')
+    moreover have "xs' = []" sorry
+    ultimately have "xs2 = xs1" by simp
+    from rpath_dest_eq[OF rp1 rp2[unfolded this]]
+    show ?thesis by simp
+  qed
+qed
+
+end
+
+
+
+end
+
--- a/PIPBasics.thy	Sun Jan 17 22:18:35 2016 +0800
+++ b/PIPBasics.thy	Wed Jan 27 19:26:56 2016 +0800
@@ -3775,6 +3775,13 @@
 definition cps:: "state \<Rightarrow> (thread \<times> precedence) set"
 where "cps s = {(th, cp s th) | th . th \<in> threads s}"
 
-find_theorems holding wq
+
+find_theorems "waiting" holding
+context valid_trace
+begin
+
+find_theorems "waiting" holding
 
 end
+
+end
--- a/PIPBasics.thy~	Sun Jan 17 22:18:35 2016 +0800
+++ b/PIPBasics.thy~	Wed Jan 27 19:26:56 2016 +0800
@@ -3775,6 +3775,13 @@
 definition cps:: "state \<Rightarrow> (thread \<times> precedence) set"
 where "cps s = {(th, cp s th) | th . th \<in> threads s}"
 
-find_theorems readys runing
+
+find_theorems "waiting" holding
+context valid_trace
+begin
+
+find_theorems "waiting" holding
 
 end
+
+end
--- a/PIPDefs.thy	Sun Jan 17 22:18:35 2016 +0800
+++ b/PIPDefs.thy	Wed Jan 27 19:26:56 2016 +0800
@@ -626,14 +626,18 @@
 definition cntV :: "state \<Rightarrow> thread \<Rightarrow> nat"
   where "cntV s th = count (\<lambda> e. \<exists> cs. e = V th cs) s"
 
+definition "pvD s th = (if (th \<in> readys s \<or> th \<notin> threads s) then 0 else (1::nat))"
+
 text {* @{text "the_preced"} is also the same as @{text "preced"}, the only
        difference is the order of arguemts. *}
 definition "the_preced s th = preced th s"
 
+
 text {* @{term "the_thread"} extracts thread out of RAG node. *}
 fun the_thread :: "node \<Rightarrow> thread" where
    "the_thread (Th th) = th"
 
+
 text {* The following @{text "wRAG"} is the waiting sub-graph of @{text "RAG"}. *}
 definition "wRAG (s::state) = {(Th th, Cs cs) | th cs. waiting s th cs}"
 
--- a/PIPDefs.thy~	Sun Jan 17 22:18:35 2016 +0800
+++ b/PIPDefs.thy~	Wed Jan 27 19:26:56 2016 +0800
@@ -626,6 +626,8 @@
 definition cntV :: "state \<Rightarrow> thread \<Rightarrow> nat"
   where "cntV s th = count (\<lambda> e. \<exists> cs. e = V th cs) s"
 
+definition "pvD s th = (if (th \<in> readys s \<or> th \<notin> threads s) then 0 else (1::nat))"
+
 text {* @{text "the_preced"} is also the same as @{text "preced"}, the only
        difference is the order of arguemts. *}
 definition "the_preced s th = preced th s"
--- a/RTree.thy	Sun Jan 17 22:18:35 2016 +0800
+++ b/RTree.thy	Wed Jan 27 19:26:56 2016 +0800
@@ -980,6 +980,37 @@
   from that[OF this] show ?thesis .
 qed
 
+lemma rpath_overlap_oneside':
+  assumes "rpath r x xs1 x1" 
+  and "rpath r x xs2 x2"
+  and "length xs1 \<le> length xs2"
+  obtains xs3 where 
+    "xs2 = xs1 @ xs3" "rpath r x xs1 x1" "rpath r x1 xs3 x2"
+proof -
+  from rpath_overlap_oneside[OF assms]
+  obtain xs3 where eq_xs: "xs2 = xs1 @ xs3" by auto
+  show ?thesis
+  proof(cases "xs1 = []")
+    case True
+    from rpath_nilE[OF assms(1)[unfolded this]]
+    have eq_x1: "x1 = x" .
+    have "xs2 = xs3" using True eq_xs by simp
+    from that[OF eq_xs assms(1) assms(2)[folded eq_x1, unfolded this]]
+    show ?thesis .
+  next
+    case False  
+    from rpath_nnl_lastE[OF assms(1) False]
+    obtain xs' where eq_xs1: "xs1 = xs'@[x1]" by auto
+    from assms(2)[unfolded eq_xs this]
+    have "rpath r x (xs' @ [x1] @ xs3) x2" by simp
+    from rpath_appendE[OF this]
+    have "rpath r x (xs' @ [x1]) x1" "rpath r x1 xs3 x2" by auto
+    from that [OF eq_xs this(1)[folded eq_xs1] this(2)]
+    show ?thesis .
+  qed
+qed
+
+
 lemma rpath_overlap [consumes 2, cases pred:rpath]:
   assumes "rpath r x xs1 x1"
   and "rpath r x xs2 x2"
@@ -990,6 +1021,16 @@
   with assms rpath_overlap_oneside that show ?thesis by metis
 qed
 
+lemma rpath_overlap' [consumes 2, cases pred:rpath]:
+  assumes "rpath r x xs1 x1"
+  and "rpath r x xs2 x2"
+  obtains (less_1) xs3 where "xs2 = xs1 @ xs3" "rpath r x xs1 x1" "rpath r x1 xs3 x2"
+     |    (less_2) xs3 where "xs1 = xs2 @ xs3" "rpath r x xs2 x2" "rpath r x2 xs3 x1"
+proof -
+  have "length xs1 \<le> length xs2 \<or> length xs2 \<le> length xs1" by auto
+  with assms rpath_overlap_oneside' that show ?thesis by metis
+qed
+
 text {*
   As a corollary of @{thm "rpath_overlap_oneside"}, 
   the following two lemmas gives one important property of relation tree, 
@@ -1399,8 +1440,27 @@
   qed
 qed
 
+end (* of rtree *)
 
-end (* of rtree *)
+lemma subtree_trancl:
+  "subtree r x = {x} \<union> {y. (y, x) \<in> r^+}" (is "?L = ?R")
+proof -
+  { fix z
+    assume "z \<in> ?L"
+    hence "z \<in> ?R"
+    proof(cases rule:subtreeE)
+      case 2
+      thus ?thesis  
+        by (unfold ancestors_def, auto)
+    qed auto
+  } moreover
+  { fix z
+    assume "z \<in> ?R"
+    hence "z \<in> ?L" 
+      by (unfold subtree_def, auto)
+  } ultimately show ?thesis by auto
+qed
+
 
 lemma subtree_children:
   "subtree r x = {x} \<union> (\<Union> (subtree r ` (children r x)))" (is "?L = ?R")
@@ -1742,4 +1802,45 @@
   using assms
   by (auto simp:children_def)
 
+lemma wf_rbase:
+  assumes "wf r"
+  obtains b where "(b, a) \<in> r^*" "\<forall> c. (c, b) \<notin> r"
+proof -
+  from assms
+  have "\<exists> b. (b, a) \<in> r^* \<and> (\<forall> c. (c, b) \<notin> r)"
+  proof(induct) 
+    case (less x)
+    thus ?case
+    proof(cases "\<exists> z. (z, x) \<in> r")
+      case False
+      moreover have "(x, x) \<in> r^*" by auto
+      ultimately show ?thesis by metis
+    next
+      case True
+      then obtain z where h_z: "(z, x) \<in> r" by auto
+      from less[OF this]
+      obtain b where "(b, z) \<in> r^*" "(\<forall>c. (c, b) \<notin> r)"
+        by auto
+      moreover from this(1) h_z have "(b, x) \<in> r^*" by auto
+      ultimately show ?thesis by metis
+    qed
+  qed
+  with that show ?thesis by metis
+qed
+
+lemma wf_base:
+  assumes "wf r"
+  and "a \<in> Range r"
+  obtains b where "(b, a) \<in> r^+" "\<forall> c. (c, b) \<notin> r"
+proof -
+  from assms(2) obtain a' where h_a: "(a', a) \<in> r" by auto
+  from wf_rbase[OF assms(1), of a]
+  obtain b where h_b: "(b, a) \<in> r\<^sup>*" "\<forall>c. (c, b) \<notin> r" by auto
+  from rtranclD[OF this(1)]
+  have "b = a \<or>  b \<noteq> a \<and> (b, a) \<in> r\<^sup>+" by auto
+  moreover have "b \<noteq> a" using h_a h_b(2) by auto
+  ultimately have "(b, a) \<in> r\<^sup>+" by auto
+  with h_b(2) and that show ?thesis by metis
+qed
+
 end
\ No newline at end of file
--- a/RTree.thy~	Sun Jan 17 22:18:35 2016 +0800
+++ b/RTree.thy~	Wed Jan 27 19:26:56 2016 +0800
@@ -154,6 +154,7 @@
 lemma index_minimize:
   assumes "P (i::nat)"
   obtains j where "P j" and "\<forall> k < j. \<not> P k" 
+using assms
 proof -
   have "\<exists> j. P j \<and> (\<forall> k < j. \<not> P k)"
   using assms
@@ -597,10 +598,6 @@
   with that[unfolded ancestors_def] show ?thesis by auto
 qed
 
-lemma subtree_Field:
-  assumes "a \<in> Field r"
-  shows "subtree r a \<subseteq> Field r"
-by (metis Field_def UnI1 ancestors_Field assms subsetI subtreeE)
 
 lemma subtree_Field:
   "subtree r x \<subseteq> Field r \<union> {x}"
@@ -891,6 +888,100 @@
 qed
 
 lemma rpath_overlap_oneside: (* ddd *)
+  assumes "rpath r x xs1 x1" (* ccc *)
+  and "rpath r x xs2 x2"
+  and "length xs1 \<le> length xs2"
+  obtains xs3 where 
+    "xs2 = xs1 @ xs3" "rpath r x xs1 x1" "rpath r x1 xs3 x2"
+proof(cases "xs1 = []")
+  case True
+  with that show ?thesis by auto
+next
+  case False
+  have "\<forall> i \<le> length xs1. take i xs1 = take i xs2"
+  proof -
+     { assume "\<not> (\<forall> i \<le> length xs1. take i xs1 = take i xs2)"
+       then obtain i where "i \<le> length xs1 \<and> take i xs1 \<noteq> take i xs2" by auto
+       from this(1) have "False"
+       proof(rule index_minimize)
+          fix j
+          assume h1: "j \<le> length xs1 \<and> take j xs1 \<noteq> take j xs2"
+          and h2: " \<forall>k<j. \<not> (k \<le> length xs1 \<and> take k xs1 \<noteq> take k xs2)"
+          -- {* @{text "j - 1"} is the branch point between @{text "xs1"} and @{text "xs2"} *}
+          let ?idx = "j - 1"
+          -- {* A number of inequalities concerning @{text "j - 1"} are derived first *}
+          have lt_i: "?idx < length xs1" using False h1 
+            by (metis Suc_diff_1 le_neq_implies_less length_greater_0_conv lessI less_imp_diff_less)
+          have lt_i': "?idx < length xs2" using lt_i and assms(3) by auto
+          have lt_j: "?idx < j" using h1 by (cases j, auto)
+          -- {* From thesis inequalities, a number of equations concerning @{text "xs1"}
+                 and @{text "xs2"} are derived *}
+          have eq_take: "take ?idx xs1 = take ?idx xs2"
+            using h2[rule_format, OF lt_j] and h1 by auto
+          have eq_xs1: " xs1 = take ?idx xs1 @ xs1 ! (?idx) # drop (Suc (?idx)) xs1" 
+            using id_take_nth_drop[OF lt_i] .
+          have eq_xs2: "xs2 = take ?idx xs2 @ xs2 ! (?idx) # drop (Suc (?idx)) xs2" 
+              using id_take_nth_drop[OF lt_i'] .
+          -- {* The branch point along the path is finally pinpointed *}
+          have neq_idx: "xs1!?idx \<noteq> xs2!?idx" 
+          proof -
+            have "take j xs1 = take ?idx xs1 @ [xs1 ! ?idx]"
+                using eq_xs1 Suc_diff_1 lt_i lt_j take_Suc_conv_app_nth by fastforce 
+            moreover have eq_tk2: "take j xs2 = take ?idx xs2 @ [xs2 ! ?idx]"
+                using Suc_diff_1 lt_i' lt_j take_Suc_conv_app_nth by fastforce 
+            ultimately show ?thesis using eq_take h1 by auto
+          qed
+          show ?thesis
+          proof(cases " take (j - 1) xs1 = []")
+            case True
+            have "(x, xs1!?idx) \<in> r"
+            proof -
+                from eq_xs1[unfolded True, simplified, symmetric] assms(1) 
+                have "rpath r x ( xs1 ! ?idx # drop (Suc ?idx) xs1) x1" by simp
+                from this[unfolded rpath_def]
+                show ?thesis by (auto simp:pred_of_def)
+            qed
+            moreover have "(x, xs2!?idx) \<in> r"
+            proof -
+              from eq_xs2[folded eq_take, unfolded True, simplified, symmetric] assms(2)
+              have "rpath r x ( xs2 ! ?idx # drop (Suc ?idx) xs2) x2" by simp
+              from this[unfolded rpath_def]
+              show ?thesis by (auto simp:pred_of_def)
+            qed
+            ultimately show ?thesis using neq_idx sgv[unfolded single_valued_def] by metis
+        next
+           case False
+           then obtain e es where eq_es: "take ?idx xs1 = es@[e]" 
+            using rev_exhaust by blast 
+           have "(e, xs1!?idx) \<in> r"
+           proof -
+            from eq_xs1[unfolded eq_es] 
+            have "xs1 = es@[e, xs1!?idx]@drop (Suc ?idx) xs1" by simp
+            hence "(e, xs1!?idx) \<in> edges_on xs1" by (simp add:edges_on_def, metis)
+            with rpath_edges_on[OF assms(1)] edges_on_Cons_mono[of xs1 x]
+            show ?thesis by auto
+           qed moreover have "(e, xs2!?idx) \<in> r"
+           proof -
+            from eq_xs2[folded eq_take, unfolded eq_es]
+            have "xs2 = es@[e, xs2!?idx]@drop (Suc ?idx) xs2" by simp
+            hence "(e, xs2!?idx) \<in> edges_on xs2" by (simp add:edges_on_def, metis)
+            with rpath_edges_on[OF assms(2)] edges_on_Cons_mono[of xs2 x]
+            show ?thesis by auto
+           qed
+           ultimately show ?thesis 
+              using sgv[unfolded single_valued_def] neq_idx by metis
+        qed
+       qed
+     } thus ?thesis by auto
+  qed
+  from this[rule_format, of "length xs1"]
+  have "take (length xs1) xs1 = take (length xs1) xs2" by simp
+  moreover have "xs2 = take (length xs1) xs2 @ drop (length xs1) xs2" by simp
+  ultimately have "xs2 = xs1 @ drop (length xs1) xs2" by auto
+  from that[OF this] show ?thesis .
+qed
+
+lemma rpath_overlap_oneside: (* ddd *)
   assumes "rpath r x xs1 x1"
   and "rpath r x xs2 x2"
   and "length xs1 \<le> length xs2"
@@ -1745,4 +1836,45 @@
   using assms
   by (auto simp:children_def)
 
+lemma wf_rbase:
+  assumes "wf r"
+  obtains b where "(b, a) \<in> r^*" "\<forall> c. (c, b) \<notin> r"
+proof -
+  from assms
+  have "\<exists> b. (b, a) \<in> r^* \<and> (\<forall> c. (c, b) \<notin> r)"
+  proof(induct) 
+    case (less x)
+    thus ?case
+    proof(cases "\<exists> z. (z, x) \<in> r")
+      case False
+      moreover have "(x, x) \<in> r^*" by auto
+      ultimately show ?thesis by metis
+    next
+      case True
+      then obtain z where h_z: "(z, x) \<in> r" by auto
+      from less[OF this]
+      obtain b where "(b, z) \<in> r^*" "(\<forall>c. (c, b) \<notin> r)"
+        by auto
+      moreover from this(1) h_z have "(b, x) \<in> r^*" by auto
+      ultimately show ?thesis by metis
+    qed
+  qed
+  with that show ?thesis by metis
+qed
+
+lemma wf_base:
+  assumes "wf r"
+  and "a \<in> Range r"
+  obtains b where "(b, a) \<in> r^+" "\<forall> c. (c, b) \<notin> r"
+proof -
+  from assms(2) obtain a' where h_a: "(a', a) \<in> r" by auto
+  from wf_rbase[OF assms(1), of a]
+  obtain b where h_b: "(b, a) \<in> r\<^sup>*" "\<forall>c. (c, b) \<notin> r" by auto
+  from rtranclD[OF this(1)]
+  have "b = a \<or>  b \<noteq> a \<and> (b, a) \<in> r\<^sup>+" by auto
+  moreover have "b \<noteq> a" using h_a h_b(2) by auto
+  ultimately have "(b, a) \<in> r\<^sup>+" by auto
+  with h_b(2) and that show ?thesis by metis
+qed
+
 end
\ No newline at end of file