Removed *.*~, #***#, log, etc.
authorzhangx
Fri, 29 Jan 2016 09:46:07 +0800
changeset 91 0525670d8e6a
parent 90 ed938e2246b9
child 92 4763aa246dbd
Removed *.*~, #***#, log, etc.
#PIPBasics.thy#
.hgignore
Correctness.thy
CpsG - 副本.thy~
CpsG.thy~
ExtGG.thy~
Implementation.thy~
Moment.thy~
PIPBasics.thy~
PIPDefs.thy~
Precedence_ord.thy~
PrioG.thy~
PrioGDef.thy~
RTree.thy~
log
--- a/#PIPBasics.thy#	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3787 +0,0 @@
-theory PIPBasics
-imports PIPDefs 
-begin
-
-locale valid_trace = 
-  fixes s
-  assumes vt : "vt s"
-
-locale valid_trace_e = valid_trace +
-  fixes e
-  assumes vt_e: "vt (e#s)"
-begin
-
-lemma pip_e: "PIP s e"
-  using vt_e by (cases, simp)  
-
-end
-
-lemma runing_ready: 
-  shows "runing s \<subseteq> readys s"
-  unfolding runing_def readys_def
-  by auto 
-
-lemma readys_threads:
-  shows "readys s \<subseteq> threads s"
-  unfolding readys_def
-  by auto
-
-lemma wq_v_neq:
-   "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
-  by (auto simp:wq_def Let_def cp_def split:list.splits)
-
-lemma runing_head:
-  assumes "th \<in> runing s"
-  and "th \<in> set (wq_fun (schs s) cs)"
-  shows "th = hd (wq_fun (schs s) cs)"
-  using assms
-  by (simp add:runing_def readys_def s_waiting_def wq_def)
-
-context valid_trace
-begin
-
-lemma actor_inv: 
-  assumes "PIP s e"
-  and "\<not> isCreate e"
-  shows "actor e \<in> runing s"
-  using assms
-  by (induct, auto)
-
-lemma ind [consumes 0, case_names Nil Cons, induct type]:
-  assumes "PP []"
-     and "(\<And>s e. valid_trace s \<Longrightarrow> valid_trace (e#s) \<Longrightarrow>
-                   PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))"
-     shows "PP s"
-proof(rule vt.induct[OF vt])
-  from assms(1) show "PP []" .
-next
-  fix s e
-  assume h: "vt s" "PP s" "PIP s e"
-  show "PP (e # s)"
-  proof(cases rule:assms(2))
-    from h(1) show v1: "valid_trace s" by (unfold_locales, simp)
-  next
-    from h(1,3) have "vt (e#s)" by auto
-    thus "valid_trace (e # s)" by (unfold_locales, simp)
-  qed (insert h, auto)
-qed
-
-lemma wq_distinct: "distinct (wq s cs)"
-proof(induct rule:ind)
-  case (Cons s e)
-  from Cons(4,3)
-  show ?case 
-  proof(induct)
-    case (thread_P th s cs1)
-    show ?case 
-    proof(cases "cs = cs1")
-      case True
-      thus ?thesis (is "distinct ?L")
-      proof - 
-        have "?L = wq_fun (schs s) cs1 @ [th]" using True
-          by (simp add:wq_def wf_def Let_def split:list.splits)
-        moreover have "distinct ..."
-        proof -
-          have "th \<notin> set (wq_fun (schs s) cs1)"
-          proof
-            assume otherwise: "th \<in> set (wq_fun (schs s) cs1)"
-            from runing_head[OF thread_P(1) this]
-            have "th = hd (wq_fun (schs s) cs1)" .
-            hence "(Cs cs1, Th th) \<in> (RAG s)" using otherwise
-              by (simp add:s_RAG_def s_holding_def wq_def cs_holding_def)
-            with thread_P(2) show False by auto
-          qed
-          moreover have "distinct (wq_fun (schs s) cs1)"
-              using True thread_P wq_def by auto 
-          ultimately show ?thesis by auto
-        qed
-        ultimately show ?thesis by simp
-      qed
-    next
-      case False
-      with thread_P(3)
-      show ?thesis
-        by (auto simp:wq_def wf_def Let_def split:list.splits)
-    qed
-  next
-    case (thread_V th s cs1)
-    thus ?case
-    proof(cases "cs = cs1")
-      case True
-      show ?thesis (is "distinct ?L")
-      proof(cases "(wq s cs)")
-        case Nil
-        thus ?thesis
-          by (auto simp:wq_def wf_def Let_def split:list.splits)
-      next
-        case (Cons w_hd w_tl)
-        moreover have "distinct (SOME q. distinct q \<and> set q = set w_tl)"
-        proof(rule someI2)
-          from thread_V(3)[unfolded Cons]
-          show  "distinct w_tl \<and> set w_tl = set w_tl" by auto
-        qed auto
-        ultimately show ?thesis
-          by (auto simp:wq_def wf_def Let_def True split:list.splits)
-      qed 
-    next
-      case False
-      with thread_V(3)
-      show ?thesis
-        by (auto simp:wq_def wf_def Let_def split:list.splits)
-    qed
-  qed (insert Cons, auto simp: wq_def Let_def split:list.splits)
-qed (unfold wq_def Let_def, simp)
-
-end
-
-
-context valid_trace_e
-begin
-
-text {*
-  The following lemma shows that only the @{text "P"}
-  operation can add new thread into waiting queues. 
-  Such kind of lemmas are very obvious, but need to be checked formally.
-  This is a kind of confirmation that our modelling is correct.
-*}
-
-lemma block_pre: 
-  assumes s_ni: "thread \<notin> set (wq s cs)"
-  and s_i: "thread \<in> set (wq (e#s) cs)"
-  shows "e = P thread cs"
-proof(cases e)
-  -- {* This is the only non-trivial case: *}
-  case (V th cs1)
-  have False
-  proof(cases "cs1 = cs")
-    case True
-    show ?thesis
-    proof(cases "(wq s cs1)")
-      case (Cons w_hd w_tl)
-      have "set (wq (e#s) cs) \<subseteq> set (wq s cs)"
-      proof -
-        have "(wq (e#s) cs) = (SOME q. distinct q \<and> set q = set w_tl)"
-          using  Cons V by (auto simp:wq_def Let_def True split:if_splits)
-        moreover have "set ... \<subseteq> set (wq s cs)"
-        proof(rule someI2)
-          show "distinct w_tl \<and> set w_tl = set w_tl"
-            by (metis distinct.simps(2) local.Cons wq_distinct)
-        qed (insert Cons True, auto)
-        ultimately show ?thesis by simp
-      qed
-      with assms show ?thesis by auto
-    qed (insert assms V True, auto simp:wq_def Let_def split:if_splits)
-  qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
-  thus ?thesis by auto
-qed (insert assms, auto simp:wq_def Let_def split:if_splits)
-
-end
-
-text {*
-  The following lemmas is also obvious and shallow. It says
-  that only running thread can request for a critical resource 
-  and that the requested resource must be one which is
-  not current held by the thread.
-*}
-
-lemma p_pre: "\<lbrakk>vt ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
-  thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (RAG s)^+"
-apply (ind_cases "vt ((P thread cs)#s)")
-apply (ind_cases "step s (P thread cs)")
-by auto
-
-lemma abs1:
-  assumes ein: "e \<in> set es"
-  and neq: "hd es \<noteq> hd (es @ [x])"
-  shows "False"
-proof -
-  from ein have "es \<noteq> []" by auto
-  then obtain e ess where "es = e # ess" by (cases es, auto)
-  with neq show ?thesis by auto
-qed
-
-lemma q_head: "Q (hd es) \<Longrightarrow> hd es = hd [th\<leftarrow>es . Q th]"
-  by (cases es, auto)
-
-inductive_cases evt_cons: "vt (a#s)"
-
-context valid_trace_e
-begin
-
-lemma abs2:
-  assumes inq: "thread \<in> set (wq s cs)"
-  and nh: "thread = hd (wq s cs)"
-  and qt: "thread \<noteq> hd (wq (e#s) cs)"
-  and inq': "thread \<in> set (wq (e#s) cs)"
-  shows "False"
-proof -
-  from vt_e assms show "False"
-    apply (cases e)
-    apply ((simp split:if_splits add:Let_def wq_def)[1])+
-    apply (insert abs1, fast)[1]
-    apply (auto simp:wq_def simp:Let_def split:if_splits list.splits)
-  proof -
-    fix th qs
-    assume vt: "vt (V th cs # s)"
-      and th_in: "thread \<in> set (SOME q. distinct q \<and> set q = set qs)"
-      and eq_wq: "wq_fun (schs s) cs = thread # qs"
-    show "False"
-    proof -
-      from wq_distinct[of cs]
-        and eq_wq[folded wq_def] have "distinct (thread#qs)" by simp
-      moreover have "thread \<in> set qs"
-      proof -
-        have "set (SOME q. distinct q \<and> set q = set qs) = set qs"
-        proof(rule someI2)
-          from wq_distinct [of cs]
-          and eq_wq [folded wq_def]
-          show "distinct qs \<and> set qs = set qs" by auto
-        next
-          fix x assume "distinct x \<and> set x = set qs"
-          thus "set x = set qs" by auto
-        qed
-        with th_in show ?thesis by auto
-      qed
-      ultimately show ?thesis by auto
-    qed
-  qed
-qed
-
-end
-
-
-context valid_trace
-begin
-lemma  vt_moment: "\<And> t. vt (moment t s)"
-proof(induct rule:ind)
-  case Nil
-  thus ?case by (simp add:vt_nil)
-next
-  case (Cons s e t)
-  show ?case
-  proof(cases "t \<ge> length (e#s)")
-    case True
-    from True have "moment t (e#s) = e#s" by simp
-    thus ?thesis using Cons
-      by (simp add:valid_trace_def)
-  next
-    case False
-    from Cons have "vt (moment t s)" by simp
-    moreover have "moment t (e#s) = moment t s"
-    proof -
-      from False have "t \<le> length s" by simp
-      from moment_app [OF this, of "[e]"] 
-      show ?thesis by simp
-    qed
-    ultimately show ?thesis by simp
-  qed
-qed
-end
-
-locale valid_moment = valid_trace + 
-  fixes i :: nat
-
-sublocale valid_moment < vat_moment: valid_trace "(moment i s)"
-  by (unfold_locales, insert vt_moment, auto)
-
-context valid_trace
-begin
-
-
-text {* (* ddd *)
-  The nature of the work is like this: since it starts from a very simple and basic 
-  model, even intuitively very `basic` and `obvious` properties need to derived from scratch.
-  For instance, the fact 
-  that one thread can not be blocked by two critical resources at the same time
-  is obvious, because only running threads can make new requests, if one is waiting for 
-  a critical resource and get blocked, it can not make another resource request and get 
-  blocked the second time (because it is not running). 
-
-  To derive this fact, one needs to prove by contraction and 
-  reason about time (or @{text "moement"}). The reasoning is based on a generic theorem
-  named @{text "p_split"}, which is about status changing along the time axis. It says if 
-  a condition @{text "Q"} is @{text "True"} at a state @{text "s"},
-  but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} 
-  in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history 
-  of events leading to it), such that @{text "Q"} switched 
-  from being @{text "False"} to @{text "True"} and kept being @{text "True"}
-  till the last moment of @{text "s"}.
-
-  Suppose a thread @{text "th"} is blocked
-  on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, 
-  since no thread is blocked at the very beginning, by applying 
-  @{text "p_split"} to these two blocking facts, there exist 
-  two moments @{text "t1"} and @{text "t2"}  in @{text "s"}, such that 
-  @{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} 
-  and kept on blocked on them respectively ever since.
- 
-  Without lost of generality, we assume @{text "t1"} is earlier than @{text "t2"}.
-  However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
-  in blocked state at moment @{text "t2"} and could not
-  make any request and get blocked the second time: Contradiction.
-*}
-
-lemma waiting_unique_pre: (* ccc *)
-  assumes h11: "thread \<in> set (wq s cs1)"
-  and h12: "thread \<noteq> hd (wq s cs1)"
-  assumes h21: "thread \<in> set (wq s cs2)"
-  and h22: "thread \<noteq> hd (wq s cs2)"
-  and neq12: "cs1 \<noteq> cs2"
-  shows "False"
-proof -
-  let "?Q cs s" = "thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
-  from h11 and h12 have q1: "?Q cs1 s" by simp
-  from h21 and h22 have q2: "?Q cs2 s" by simp
-  have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
-  have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
-  from p_split [of "?Q cs1", OF q1 nq1]
-  obtain t1 where lt1: "t1 < length s"
-    and np1: "\<not>(thread \<in> set (wq (moment t1 s) cs1) \<and>
-        thread \<noteq> hd (wq (moment t1 s) cs1))"
-    and nn1: "(\<forall>i'>t1. thread \<in> set (wq (moment i' s) cs1) \<and>
-             thread \<noteq> hd (wq (moment i' s) cs1))" by auto
-  from p_split [of "?Q cs2", OF q2 nq2]
-  obtain t2 where lt2: "t2 < length s"
-    and np2: "\<not>(thread \<in> set (wq (moment t2 s) cs2) \<and>
-        thread \<noteq> hd (wq (moment t2 s) cs2))"
-    and nn2: "(\<forall>i'>t2. thread \<in> set (wq (moment i' s) cs2) \<and>
-             thread \<noteq> hd (wq (moment i' s) cs2))" by auto
-  show ?thesis
-  proof -
-    { 
-      assume lt12: "t1 < t2"
-      let ?t3 = "Suc t2"
-      from lt2 have le_t3: "?t3 \<le> length s" by auto
-      from moment_plus [OF this] 
-      obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
-      have "t2 < ?t3" by simp
-      from nn2 [rule_format, OF this] and eq_m
-      have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
-        h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
-      have "vt (e#moment t2 s)"
-      proof -
-        from vt_moment 
-        have "vt (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      then interpret vt_e: valid_trace_e "moment t2 s" "e"
-        by (unfold_locales, auto, cases, simp)
-      have ?thesis
-      proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
-        case True
-        from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
-          by auto 
-        from vt_e.abs2 [OF True eq_th h2 h1]
-        show ?thesis by auto
-      next
-        case False
-        from vt_e.block_pre[OF False h1]
-        have "e = P thread cs2" .
-        with vt_e.vt_e have "vt ((P thread cs2)# moment t2 s)" by simp
-        from p_pre [OF this] have "thread \<in> runing (moment t2 s)" by simp
-        with runing_ready have "thread \<in> readys (moment t2 s)" by auto
-        with nn1 [rule_format, OF lt12]
-        show ?thesis  by (simp add:readys_def wq_def s_waiting_def, auto)
-      qed
-    } moreover {
-      assume lt12: "t2 < t1"
-      let ?t3 = "Suc t1"
-      from lt1 have le_t3: "?t3 \<le> length s" by auto
-      from moment_plus [OF this] 
-      obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
-      have lt_t3: "t1 < ?t3" by simp
-      from nn1 [rule_format, OF this] and eq_m
-      have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
-        h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
-      have "vt  (e#moment t1 s)"
-      proof -
-        from vt_moment
-        have "vt (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      then interpret vt_e: valid_trace_e "moment t1 s" e
-        by (unfold_locales, auto, cases, auto)
-      have ?thesis
-      proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
-        case True
-        from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
-          by auto
-        from vt_e.abs2 True eq_th h2 h1
-        show ?thesis by auto
-      next
-        case False
-        from vt_e.block_pre [OF False h1]
-        have "e = P thread cs1" .
-        with vt_e.vt_e have "vt ((P thread cs1)# moment t1 s)" by simp
-        from p_pre [OF this] have "thread \<in> runing (moment t1 s)" by simp
-        with runing_ready have "thread \<in> readys (moment t1 s)" by auto
-        with nn2 [rule_format, OF lt12]
-        show ?thesis  by (simp add:readys_def wq_def s_waiting_def, auto)
-      qed
-    } moreover {
-      assume eqt12: "t1 = t2"
-      let ?t3 = "Suc t1"
-      from lt1 have le_t3: "?t3 \<le> length s" by auto
-      from moment_plus [OF this] 
-      obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
-      have lt_t3: "t1 < ?t3" by simp
-      from nn1 [rule_format, OF this] and eq_m
-      have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
-        h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
-      have vt_e: "vt (e#moment t1 s)"
-      proof -
-        from vt_moment
-        have "vt (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      then interpret vt_e: valid_trace_e "moment t1 s" e
-        by (unfold_locales, auto, cases, auto)
-      have ?thesis
-      proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
-        case True
-        from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
-          by auto
-        from vt_e.abs2 [OF True eq_th h2 h1]
-        show ?thesis by auto
-      next
-        case False
-        from vt_e.block_pre [OF False h1]
-        have eq_e1: "e = P thread cs1" .
-        have lt_t3: "t1 < ?t3" by simp
-        with eqt12 have "t2 < ?t3" by simp
-        from nn2 [rule_format, OF this] and eq_m and eqt12
-        have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
-          h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
-        show ?thesis
-        proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
-          case True
-          from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
-            by auto
-          from vt_e and eqt12 have "vt (e#moment t2 s)" by simp 
-          then interpret vt_e2: valid_trace_e "moment t2 s" e
-            by (unfold_locales, auto, cases, auto)
-          from vt_e2.abs2 [OF True eq_th h2 h1]
-          show ?thesis .
-        next
-          case False
-          have "vt (e#moment t2 s)"
-          proof -
-            from vt_moment eqt12
-            have "vt (moment (Suc t2) s)" by auto
-            with eq_m eqt12 show ?thesis by simp
-          qed
-          then interpret vt_e2: valid_trace_e "moment t2 s" e
-            by (unfold_locales, auto, cases, auto)
-          from vt_e2.block_pre [OF False h1]
-          have "e = P thread cs2" .
-          with eq_e1 neq12 show ?thesis by auto
-        qed
-      qed
-    } ultimately show ?thesis by arith
-  qed
-qed
-
-text {*
-  This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
-*}
-
-lemma waiting_unique:
-  assumes "waiting s th cs1"
-  and "waiting s th cs2"
-  shows "cs1 = cs2"
-using waiting_unique_pre assms
-unfolding wq_def s_waiting_def
-by auto
-
-end
-
-(* not used *)
-text {*
-  Every thread can only be blocked on one critical resource, 
-  symmetrically, every critical resource can only be held by one thread. 
-  This fact is much more easier according to our definition. 
-*}
-lemma held_unique:
-  assumes "holding (s::event list) th1 cs"
-  and "holding s th2 cs"
-  shows "th1 = th2"
- by (insert assms, unfold s_holding_def, auto)
-
-
-lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s"
-  apply (induct s, auto)
-  by (case_tac a, auto split:if_splits)
-
-lemma last_set_unique: 
-  "\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk>
-          \<Longrightarrow> th1 = th2"
-  apply (induct s, auto)
-  by (case_tac a, auto split:if_splits dest:last_set_lt)
-
-lemma preced_unique : 
-  assumes pcd_eq: "preced th1 s = preced th2 s"
-  and th_in1: "th1 \<in> threads s"
-  and th_in2: " th2 \<in> threads s"
-  shows "th1 = th2"
-proof -
-  from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def)
-  from last_set_unique [OF this th_in1 th_in2]
-  show ?thesis .
-qed
-
-lemma preced_linorder: 
-  assumes neq_12: "th1 \<noteq> th2"
-  and th_in1: "th1 \<in> threads s"
-  and th_in2: " th2 \<in> threads s"
-  shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
-proof -
-  from preced_unique [OF _ th_in1 th_in2] and neq_12 
-  have "preced th1 s \<noteq> preced th2 s" by auto
-  thus ?thesis by auto
-qed
-
-(* An aux lemma used later *)
-lemma unique_minus:
-  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
-  and xy: "(x, y) \<in> r"
-  and xz: "(x, z) \<in> r^+"
-  and neq: "y \<noteq> z"
-  shows "(y, z) \<in> r^+"
-proof -
- from xz and neq show ?thesis
- proof(induct)
-   case (base ya)
-   have "(x, ya) \<in> r" by fact
-   from unique [OF xy this] have "y = ya" .
-   with base show ?case by auto
- next
-   case (step ya z)
-   show ?case
-   proof(cases "y = ya")
-     case True
-     from step True show ?thesis by simp
-   next
-     case False
-     from step False
-     show ?thesis by auto
-   qed
- qed
-qed
-
-lemma unique_base:
-  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
-  and xy: "(x, y) \<in> r"
-  and xz: "(x, z) \<in> r^+"
-  and neq_yz: "y \<noteq> z"
-  shows "(y, z) \<in> r^+"
-proof -
-  from xz neq_yz show ?thesis
-  proof(induct)
-    case (base ya)
-    from xy unique base show ?case by auto
-  next
-    case (step ya z)
-    show ?case
-    proof(cases "y = ya")
-      case True
-      from True step show ?thesis by auto
-    next
-      case False
-      from False step 
-      have "(y, ya) \<in> r\<^sup>+" by auto
-      with step show ?thesis by auto
-    qed
-  qed
-qed
-
-lemma unique_chain:
-  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
-  and xy: "(x, y) \<in> r^+"
-  and xz: "(x, z) \<in> r^+"
-  and neq_yz: "y \<noteq> z"
-  shows "(y, z) \<in> r^+ \<or> (z, y) \<in> r^+"
-proof -
-  from xy xz neq_yz show ?thesis
-  proof(induct)
-    case (base y)
-    have h1: "(x, y) \<in> r" and h2: "(x, z) \<in> r\<^sup>+" and h3: "y \<noteq> z" using base by auto
-    from unique_base [OF _ h1 h2 h3] and unique show ?case by auto
-  next
-    case (step y za)
-    show ?case
-    proof(cases "y = z")
-      case True
-      from True step show ?thesis by auto
-    next
-      case False
-      from False step have "(y, z) \<in> r\<^sup>+ \<or> (z, y) \<in> r\<^sup>+" by auto
-      thus ?thesis
-      proof
-        assume "(z, y) \<in> r\<^sup>+"
-        with step have "(z, za) \<in> r\<^sup>+" by auto
-        thus ?thesis by auto
-      next
-        assume h: "(y, z) \<in> r\<^sup>+"
-        from step have yza: "(y, za) \<in> r" by simp
-        from step have "za \<noteq> z" by simp
-        from unique_minus [OF _ yza h this] and unique
-        have "(za, z) \<in> r\<^sup>+" by auto
-        thus ?thesis by auto
-      qed
-    qed
-  qed
-qed
-
-text {*
-  The following three lemmas show that @{text "RAG"} does not change
-  by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
-  events, respectively.
-*}
-
-lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
-apply (unfold s_RAG_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-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 RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
-apply (unfold s_RAG_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-
-text {* 
-  The following lemmas are used in the proof of 
-  lemma @{text "step_RAG_v"}, which characterizes how the @{text "RAG"} is changed
-  by @{text "V"}-events. 
-  However, since our model is very concise, such  seemingly obvious lemmas need to be derived from scratch,
-  starting from the model definitions.
-*}
-lemma step_v_hold_inv[elim_format]:
-  "\<And>c t. \<lbrakk>vt (V th cs # s); 
-          \<not> holding (wq s) t c; holding (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> 
-            next_th s th cs t \<and> c = cs"
-proof -
-  fix c t
-  assume vt: "vt (V th cs # s)"
-    and nhd: "\<not> holding (wq s) t c"
-    and hd: "holding (wq (V th cs # s)) t c"
-  show "next_th s th cs t \<and> c = cs"
-  proof(cases "c = cs")
-    case False
-    with nhd hd show ?thesis
-      by (unfold cs_holding_def wq_def, auto simp:Let_def)
-  next
-    case True
-    with step_back_step [OF vt] 
-    have "step s (V th c)" by simp
-    hence "next_th s th cs t"
-    proof(cases)
-      assume "holding s th c"
-      with nhd hd show ?thesis
-        apply (unfold s_holding_def cs_holding_def wq_def next_th_def,
-               auto simp:Let_def split:list.splits if_splits)
-        proof -
-          assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
-          moreover have "\<dots> = set []"
-          proof(rule someI2)
-            show "distinct [] \<and> [] = []" by auto
-          next
-            fix x assume "distinct x \<and> x = []"
-            thus "set x = set []" by auto
-          qed
-          ultimately show False by auto
-        next
-          assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
-          moreover have "\<dots> = set []"
-          proof(rule someI2)
-            show "distinct [] \<and> [] = []" by auto
-          next
-            fix x assume "distinct x \<and> x = []"
-            thus "set x = set []" by auto
-          qed
-          ultimately show False by auto
-        qed
-    qed
-    with True show ?thesis by auto
-  qed
-qed
-
-text {* 
-  The following @{text "step_v_wait_inv"} is also an obvious lemma, which, however, needs to be
-  derived from scratch, which confirms the correctness of the definition of @{text "next_th"}.
-*}
-lemma step_v_wait_inv[elim_format]:
-    "\<And>t c. \<lbrakk>vt (V th cs # s); \<not> waiting (wq (V th cs # s)) t c; waiting (wq s) t c
-           \<rbrakk>
-          \<Longrightarrow> (next_th s th cs t \<and> cs = c)"
-proof -
-  fix t c 
-  assume vt: "vt (V th cs # s)"
-    and nw: "\<not> waiting (wq (V th cs # s)) t c"
-    and wt: "waiting (wq s) t c"
-  from vt interpret vt_v: valid_trace_e s "V th cs" 
-    by  (cases, unfold_locales, simp)
-  show "next_th s th cs t \<and> cs = c"
-  proof(cases "cs = c")
-    case False
-    with nw wt show ?thesis
-      by (auto simp:cs_waiting_def wq_def Let_def)
-  next
-    case True
-    from nw[folded True] wt[folded True]
-    have "next_th s th cs t"
-      apply (unfold next_th_def, auto simp:cs_waiting_def wq_def Let_def split:list.splits)
-    proof -
-      fix a list
-      assume t_in: "t \<in> set list"
-        and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
-        and eq_wq: "wq_fun (schs s) cs = a # list"
-      have " set (SOME q. distinct q \<and> set q = set list) = set list"
-      proof(rule someI2)
-        from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def]
-        show "distinct list \<and> set list = set list" by auto
-      next
-        show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
-          by auto
-      qed
-      with t_ni and t_in show "a = th" by auto
-    next
-      fix a list
-      assume t_in: "t \<in> set list"
-        and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
-        and eq_wq: "wq_fun (schs s) cs = a # list"
-      have " set (SOME q. distinct q \<and> set q = set list) = set list"
-      proof(rule someI2)
-        from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def]
-        show "distinct list \<and> set list = set list" by auto
-      next
-        show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
-          by auto
-      qed
-      with t_ni and t_in show "t = hd (SOME q. distinct q \<and> set q = set list)" by auto
-    next
-      fix a list
-      assume eq_wq: "wq_fun (schs s) cs = a # list"
-      from step_back_step[OF vt]
-      show "a = th"
-      proof(cases)
-        assume "holding s th cs"
-        with eq_wq show ?thesis
-          by (unfold s_holding_def wq_def, auto)
-      qed
-    qed
-    with True show ?thesis by simp
-  qed
-qed
-
-lemma step_v_not_wait[consumes 3]:
-  "\<lbrakk>vt (V th cs # s); next_th s th cs t; waiting (wq (V th cs # s)) t cs\<rbrakk> \<Longrightarrow> False"
-  by (unfold next_th_def cs_waiting_def wq_def, auto simp:Let_def)
-
-lemma step_v_release:
-  "\<lbrakk>vt (V th cs # s); holding (wq (V th cs # s)) th cs\<rbrakk> \<Longrightarrow> False"
-proof -
-  assume vt: "vt (V th cs # s)"
-    and hd: "holding (wq (V th cs # s)) th cs"
-  from vt interpret vt_v: valid_trace_e s "V th cs"
-    by (cases, unfold_locales, simp+)
-  from step_back_step [OF vt] and hd
-  show "False"
-  proof(cases)
-    assume "holding (wq (V th cs # s)) th cs" and "holding s th cs"
-    thus ?thesis
-      apply (unfold s_holding_def wq_def cs_holding_def)
-      apply (auto simp:Let_def split:list.splits)
-    proof -
-      fix list
-      assume eq_wq[folded wq_def]: 
-        "wq_fun (schs s) cs = hd (SOME q. distinct q \<and> set q = set list) # list"
-      and hd_in: "hd (SOME q. distinct q \<and> set q = set list)
-            \<in> set (SOME q. distinct q \<and> set q = set list)"
-      have "set (SOME q. distinct q \<and> set q = set list) = set list"
-      proof(rule someI2)
-        from vt_v.wq_distinct[of cs] and eq_wq
-        show "distinct list \<and> set list = set list" by auto
-      next
-        show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
-          by auto
-      qed
-      moreover have "distinct  (hd (SOME q. distinct q \<and> set q = set list) # list)"
-      proof -
-        from vt_v.wq_distinct[of cs] and eq_wq
-        show ?thesis by auto
-      qed
-      moreover note eq_wq and hd_in
-      ultimately show "False" by auto
-    qed
-  qed
-qed
-
-lemma step_v_get_hold:
-  "\<And>th'. \<lbrakk>vt (V th cs # s); \<not> holding (wq (V th cs # s)) th' cs; next_th s th cs th'\<rbrakk> \<Longrightarrow> False"
-  apply (unfold cs_holding_def next_th_def wq_def,
-         auto simp:Let_def)
-proof -
-  fix rest
-  assume vt: "vt (V th cs # s)"
-    and eq_wq[folded wq_def]: " wq_fun (schs s) cs = th # rest"
-    and nrest: "rest \<noteq> []"
-    and ni: "hd (SOME q. distinct q \<and> set q = set rest)
-            \<notin> set (SOME q. distinct q \<and> set q = set rest)"
-  from vt interpret vt_v: valid_trace_e s "V th cs"
-    by (cases, unfold_locales, simp+)
-  have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-  proof(rule someI2)
-    from vt_v.wq_distinct[of cs] and eq_wq
-    show "distinct rest \<and> set rest = set rest" by auto
-  next
-    fix x assume "distinct x \<and> set x = set rest"
-    hence "set x = set rest" by auto
-    with nrest
-    show "x \<noteq> []" by (case_tac x, auto)
-  qed
-  with ni show "False" by auto
-qed
-
-lemma step_v_release_inv[elim_format]:
-"\<And>c t. \<lbrakk>vt (V th cs # s); \<not> holding (wq (V th cs # s)) t c; holding (wq s) t c\<rbrakk> \<Longrightarrow> 
-  c = cs \<and> t = th"
-  apply (unfold cs_holding_def wq_def, auto simp:Let_def split:if_splits list.splits)
-  proof -
-    fix a list
-    assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list"
-    from step_back_step [OF vt] show "a = th"
-    proof(cases)
-      assume "holding s th cs" with eq_wq
-      show ?thesis
-        by (unfold s_holding_def wq_def, auto)
-    qed
-  next
-    fix a list
-    assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list"
-    from step_back_step [OF vt] show "a = th"
-    proof(cases)
-      assume "holding s th cs" with eq_wq
-      show ?thesis
-        by (unfold s_holding_def wq_def, auto)
-    qed
-  qed
-
-lemma step_v_waiting_mono:
-  "\<And>t c. \<lbrakk>vt (V th cs # s); waiting (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> waiting (wq s) t c"
-proof -
-  fix t c
-  let ?s' = "(V th cs # s)"
-  assume vt: "vt ?s'" 
-    and wt: "waiting (wq ?s') t c"
-  from vt interpret vt_v: valid_trace_e s "V th cs"
-    by (cases, unfold_locales, simp+)
-  show "waiting (wq s) t c"
-  proof(cases "c = cs")
-    case False
-    assume neq_cs: "c \<noteq> cs"
-    hence "waiting (wq ?s') t c = waiting (wq s) t c"
-      by (unfold cs_waiting_def wq_def, auto simp:Let_def)
-    with wt show ?thesis by simp
-  next
-    case True
-    with wt show ?thesis
-      apply (unfold cs_waiting_def wq_def, auto simp:Let_def split:list.splits)
-    proof -
-      fix a list
-      assume not_in: "t \<notin> set list"
-        and is_in: "t \<in> set (SOME q. distinct q \<and> set q = set list)"
-        and eq_wq: "wq_fun (schs s) cs = a # list"
-      have "set (SOME q. distinct q \<and> set q = set list) = set list"
-      proof(rule someI2)
-        from vt_v.wq_distinct [of cs]
-        and eq_wq[folded wq_def]
-        show "distinct list \<and> set list = set list" by auto
-      next
-        fix x assume "distinct x \<and> set x = set list"
-        thus "set x = set list" by auto
-      qed
-      with not_in is_in show "t = a" by auto
-    next
-      fix list
-      assume is_waiting: "waiting (wq (V th cs # s)) t cs"
-      and eq_wq: "wq_fun (schs s) cs = t # list"
-      hence "t \<in> set list"
-        apply (unfold wq_def, auto simp:Let_def cs_waiting_def)
-      proof -
-        assume " t \<in> set (SOME q. distinct q \<and> set q = set list)"
-        moreover have "\<dots> = set list" 
-        proof(rule someI2)
-          from vt_v.wq_distinct [of cs]
-            and eq_wq[folded wq_def]
-          show "distinct list \<and> set list = set list" by auto
-        next
-          fix x assume "distinct x \<and> set x = set list" 
-          thus "set x = set list" by auto
-        qed
-        ultimately show "t \<in> set list" by simp
-      qed
-      with eq_wq and vt_v.wq_distinct [of cs, unfolded wq_def]
-      show False by auto
-    qed
-  qed
-qed
-
-text {* (* ddd *) 
-  The following @{text "step_RAG_v"} lemma charaterizes how @{text "RAG"} is changed
-  with the happening of @{text "V"}-events:
-*}
-lemma step_RAG_v:
-assumes vt:
-  "vt (V th cs#s)"
-shows "
-  RAG (V th cs # s) =
-  RAG s - {(Cs cs, Th th)} -
-  {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-  {(Cs cs, Th th') |th'.  next_th s th cs th'}"
-  apply (insert vt, unfold s_RAG_def) 
-  apply (auto split:if_splits list.splits simp:Let_def)
-  apply (auto elim: step_v_waiting_mono step_v_hold_inv 
-              step_v_release step_v_wait_inv
-              step_v_get_hold step_v_release_inv)
-  apply (erule_tac step_v_not_wait, auto)
-  done
-
-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
-
-
-lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
-  by (unfold s_RAG_def, auto)
-
-context valid_trace
-begin
-
-text {*
-  The following lemma shows that @{text "RAG"} is acyclic.
-  The overall structure is by induction on the formation of @{text "vt s"}
-  and then case analysis on event @{text "e"}, where the non-trivial cases 
-  for those for @{text "V"} and @{text "P"} events.
-*}
-lemma acyclic_RAG:
-  shows "acyclic (RAG s)"
-using vt
-proof(induct)
-  case (vt_cons s e)
-  interpret vt_s: valid_trace s using vt_cons(1)
-    by (unfold_locales, simp)
-  assume ih: "acyclic (RAG s)"
-    and stp: "step s e"
-    and vt: "vt s"
-  show ?case
-  proof(cases e)
-    case (Create th prio)
-    with ih
-    show ?thesis by (simp add:RAG_create_unchanged)
-  next
-    case (Exit th)
-    with ih show ?thesis by (simp add:RAG_exit_unchanged)
-  next
-    case (V th cs)
-    from V vt stp have vtt: "vt (V th cs#s)" by auto
-    from step_RAG_v [OF this]
-    have eq_de: 
-      "RAG (e # s) = 
-      RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-      {(Cs cs, Th th') |th'. next_th s th cs th'}"
-      (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
-    from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
-    from step_back_step [OF vtt]
-    have "step s (V th cs)" .
-    thus ?thesis
-    proof(cases)
-      assume "holding s th cs"
-      hence th_in: "th \<in> set (wq s cs)" and
-        eq_hd: "th = hd (wq s cs)" unfolding s_holding_def wq_def by auto
-      then obtain rest where
-        eq_wq: "wq s cs = th#rest"
-        by (cases "wq s cs", auto)
-      show ?thesis
-      proof(cases "rest = []")
-        case False
-        let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
-        from eq_wq False have eq_D: "?D = {(Cs cs, Th ?th')}"
-          by (unfold next_th_def, auto)
-        let ?E = "(?A - ?B - ?C)"
-        have "(Th ?th', Cs cs) \<notin> ?E\<^sup>*"
-        proof
-          assume "(Th ?th', Cs cs) \<in> ?E\<^sup>*"
-          hence " (Th ?th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
-          from tranclD [OF this]
-          obtain x where th'_e: "(Th ?th', x) \<in> ?E" by blast
-          hence th_d: "(Th ?th', x) \<in> ?A" by simp
-          from RAG_target_th [OF this]
-          obtain cs' where eq_x: "x = Cs cs'" by auto
-          with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp
-          hence wt_th': "waiting s ?th' cs'"
-            unfolding s_RAG_def s_waiting_def cs_waiting_def wq_def by simp
-          hence "cs' = cs"
-          proof(rule vt_s.waiting_unique)
-            from eq_wq vt_s.wq_distinct[of cs]
-            show "waiting s ?th' cs" 
-              apply (unfold s_waiting_def wq_def, auto)
-            proof -
-              assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-                and eq_wq: "wq_fun (schs s) cs = th # rest"
-              have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
-              next
-                fix x assume "distinct x \<and> set x = set rest"
-                with False show "x \<noteq> []" by auto
-              qed
-              hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
-                set (SOME q. distinct q \<and> set q = set rest)" by auto
-              moreover have "\<dots> = set rest" 
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
-              next
-                show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-              qed
-              moreover note hd_in
-              ultimately show "hd (SOME q. distinct q \<and> set q = set rest) = th" by auto
-            next
-              assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-                and eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest"
-              have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" by auto
-              next
-                fix x assume "distinct x \<and> set x = set rest"
-                with False show "x \<noteq> []" by auto
-              qed
-              hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
-                set (SOME q. distinct q \<and> set q = set rest)" by auto
-              moreover have "\<dots> = set rest" 
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" by auto
-              next
-                show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-              qed
-              moreover note hd_in
-              ultimately show False by auto
-            qed
-          qed
-          with th'_e eq_x have "(Th ?th', Cs cs) \<in> ?E" by simp
-          with False
-          show "False" by (auto simp: next_th_def eq_wq)
-        qed
-        with acyclic_insert[symmetric] and ac
-          and eq_de eq_D show ?thesis by auto
-      next
-        case True
-        with eq_wq
-        have eq_D: "?D = {}"
-          by (unfold next_th_def, auto)
-        with eq_de ac
-        show ?thesis by auto
-      qed 
-    qed
-  next
-    case (P th cs)
-    from P vt stp have vtt: "vt (P th cs#s)" by auto
-    from step_RAG_p [OF this] P
-    have "RAG (e # s) = 
-      (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
-      RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
-      by simp
-    moreover have "acyclic ?R"
-    proof(cases "wq s cs = []")
-      case True
-      hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
-      have "(Th th, Cs cs) \<notin> (RAG s)\<^sup>*"
-      proof
-        assume "(Th th, Cs cs) \<in> (RAG s)\<^sup>*"
-        hence "(Th th, Cs cs) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
-        from tranclD2 [OF this]
-        obtain x where "(x, Cs cs) \<in> RAG s" by auto
-        with True show False by (auto simp:s_RAG_def cs_waiting_def)
-      qed
-      with acyclic_insert ih eq_r show ?thesis by auto
-    next
-      case False
-      hence eq_r: "?R =  RAG s \<union> {(Th th, Cs cs)}" by simp
-      have "(Cs cs, Th th) \<notin> (RAG s)\<^sup>*"
-      proof
-        assume "(Cs cs, Th th) \<in> (RAG s)\<^sup>*"
-        hence "(Cs cs, Th th) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
-        moreover from step_back_step [OF vtt] have "step s (P th cs)" .
-        ultimately show False
-        proof -
-          show " \<lbrakk>(Cs cs, Th th) \<in> (RAG s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
-            by (ind_cases "step s (P th cs)", simp)
-        qed
-      qed
-      with acyclic_insert ih eq_r show ?thesis by auto
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (Set thread prio)
-      with ih
-      thm RAG_set_unchanged
-      show ?thesis by (simp add:RAG_set_unchanged)
-    qed
-  next
-    case vt_nil
-    show "acyclic (RAG ([]::state))"
-      by (auto simp: s_RAG_def cs_waiting_def 
-        cs_holding_def wq_def acyclic_def)
-qed
-
-
-lemma finite_RAG:
-  shows "finite (RAG s)"
-proof -
-  from vt show ?thesis
-  proof(induct)
-    case (vt_cons s e)
-    interpret vt_s: valid_trace s using vt_cons(1)
-      by (unfold_locales, simp)
-    assume ih: "finite (RAG s)"
-      and stp: "step s e"
-      and vt: "vt s"
-    show ?case
-    proof(cases e)
-      case (Create th prio)
-      with ih
-      show ?thesis by (simp add:RAG_create_unchanged)
-    next
-      case (Exit th)
-      with ih show ?thesis by (simp add:RAG_exit_unchanged)
-    next
-      case (V th cs)
-      from V vt stp have vtt: "vt (V th cs#s)" by auto
-      from step_RAG_v [OF this]
-      have eq_de: "RAG (e # s) = 
-                   RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-                      {(Cs cs, Th th') |th'. next_th s th cs th'}
-"
-        (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
-      moreover from ih have ac: "finite (?A - ?B - ?C)" by simp
-      moreover have "finite ?D"
-      proof -
-        have "?D = {} \<or> (\<exists> a. ?D = {a})" 
-          by (unfold next_th_def, auto)
-        thus ?thesis
-        proof
-          assume h: "?D = {}"
-          show ?thesis by (unfold h, simp)
-        next
-          assume "\<exists> a. ?D = {a}"
-          thus ?thesis
-            by (metis finite.simps)
-        qed
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (P th cs)
-      from P vt stp have vtt: "vt (P th cs#s)" by auto
-      from step_RAG_p [OF this] P
-      have "RAG (e # s) = 
-              (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
-                                    RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
-        by simp
-      moreover have "finite ?R"
-      proof(cases "wq s cs = []")
-        case True
-        hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
-        with True and ih show ?thesis by auto
-      next
-        case False
-        hence "?R = RAG s \<union> {(Th th, Cs cs)}" by simp
-        with False and ih show ?thesis by auto
-      qed
-      ultimately show ?thesis by auto
-    next
-      case (Set thread prio)
-      with ih
-      show ?thesis by (simp add:RAG_set_unchanged)
-    qed
-  next
-    case vt_nil
-    show "finite (RAG ([]::state))"
-      by (auto simp: s_RAG_def cs_waiting_def 
-                   cs_holding_def wq_def acyclic_def)
-  qed
-qed
-
-text {* Several useful lemmas *}
-
-lemma wf_dep_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
-
-end
-
-lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l"
-  by (induct l, auto)
-
-lemma th_chasing: "(Th th, Cs cs) \<in> RAG (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> RAG s"
-  by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
-
-context valid_trace
-begin
-
-lemma wq_threads: 
-  assumes h: "th \<in> set (wq s cs)"
-  shows "th \<in> threads s"
-proof -
- from vt and h show ?thesis
-  proof(induct arbitrary: th cs)
-    case (vt_cons s e)
-    interpret vt_s: valid_trace s
-      using vt_cons(1) by (unfold_locales, auto)
-    assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s"
-      and stp: "step s e"
-      and vt: "vt s"
-      and h: "th \<in> set (wq (e # s) cs)"
-    show ?case
-    proof(cases e)
-      case (Create th' prio)
-      with ih h show ?thesis
-        by (auto simp:wq_def Let_def)
-    next
-      case (Exit th')
-      with stp ih h show ?thesis
-        apply (auto simp:wq_def Let_def)
-        apply (ind_cases "step s (Exit th')")
-        apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def
-               s_RAG_def s_holding_def cs_holding_def)
-        done
-    next
-      case (V th' cs')
-      show ?thesis
-      proof(cases "cs' = cs")
-        case False
-        with h
-        show ?thesis
-          apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def)
-          by (drule_tac ih, simp)
-      next
-        case True
-        from h
-        show ?thesis
-        proof(unfold V wq_def)
-          assume th_in: "th \<in> set (wq_fun (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l")
-          show "th \<in> threads (V th' cs' # s)"
-          proof(cases "cs = cs'")
-            case False
-            hence "?l = wq_fun (schs s) cs" by (simp add:Let_def)
-            with th_in have " th \<in> set (wq s cs)" 
-              by (fold wq_def, simp)
-            from ih [OF this] show ?thesis by simp
-          next
-            case True
-            show ?thesis
-            proof(cases "wq_fun (schs s) cs'")
-              case Nil
-              with h V show ?thesis
-                apply (auto simp:wq_def Let_def split:if_splits)
-                by (fold wq_def, drule_tac ih, simp)
-            next
-              case (Cons a rest)
-              assume eq_wq: "wq_fun (schs s) cs' = a # rest"
-              with h V show ?thesis
-                apply (auto simp:Let_def wq_def split:if_splits)
-              proof -
-                assume th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
-                have "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
-                proof(rule someI2)
-                  from vt_s.wq_distinct[of cs'] and eq_wq[folded wq_def]
-                  show "distinct rest \<and> set rest = set rest" by auto
-                next
-                  show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
-                    by auto
-                qed
-                with eq_wq th_in have "th \<in> set (wq_fun (schs s) cs')" by auto
-                from ih[OF this[folded wq_def]] show "th \<in> threads s" .
-              next
-                assume th_in: "th \<in> set (wq_fun (schs s) cs)"
-                from ih[OF this[folded wq_def]]
-                show "th \<in> threads s" .
-              qed
-            qed
-          qed
-        qed
-      qed
-    next
-      case (P th' cs')
-      from h stp
-      show ?thesis
-        apply (unfold P wq_def)
-        apply (auto simp:Let_def split:if_splits, fold wq_def)
-        apply (auto intro:ih)
-        apply(ind_cases "step s (P th' cs')")
-        by (unfold runing_def readys_def, auto)
-    next
-      case (Set thread prio)
-      with ih h show ?thesis
-        by (auto simp:wq_def Let_def)
-    qed
-  next
-    case vt_nil
-    thus ?case by (auto simp:wq_def)
-  qed
-qed
-
-lemma range_in: "\<lbrakk>(Th th) \<in> Range (RAG (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
-  apply(unfold s_RAG_def cs_waiting_def cs_holding_def)
-  by (auto intro:wq_threads)
-
-lemma readys_v_eq:
-  assumes neq_th: "th \<noteq> thread"
-  and eq_wq: "wq s cs = thread#rest"
-  and not_in: "th \<notin>  set rest"
-  shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)"
-proof -
-  from assms show ?thesis
-    apply (auto simp:readys_def)
-    apply(simp add:s_waiting_def[folded wq_def])
-    apply (erule_tac x = csa in allE)
-    apply (simp add:s_waiting_def wq_def Let_def split:if_splits)
-    apply (case_tac "csa = cs", simp)
-    apply (erule_tac x = cs in allE)
-    apply(auto simp add: s_waiting_def[folded wq_def] Let_def split: list.splits)
-    apply(auto simp add: wq_def)
-    apply (auto simp:s_waiting_def wq_def Let_def split:list.splits)
-    proof -
-       assume th_nin: "th \<notin> set rest"
-        and th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
-        and eq_wq: "wq_fun (schs s) cs = thread # rest"
-      have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-      proof(rule someI2)
-        from wq_distinct[of cs, unfolded wq_def] and eq_wq[unfolded wq_def]
-        show "distinct rest \<and> set rest = set rest" by auto
-      next
-        show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-      qed
-      with th_nin th_in show False by auto
-    qed
-qed
-
-text {* \noindent
-  The following lemmas shows that: starting from any node in @{text "RAG"}, 
-  by chasing out-going edges, it is always possible to reach a node representing a ready
-  thread. In this lemma, it is the @{text "th'"}.
-*}
-
-lemma chain_building:
-  shows "node \<in> Domain (RAG s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (RAG s)^+)"
-proof -
-  from wf_dep_converse
-  have h: "wf ((RAG s)\<inverse>)" .
-  show ?thesis
-  proof(induct rule:wf_induct [OF h])
-    fix x
-    assume ih [rule_format]: 
-      "\<forall>y. (y, x) \<in> (RAG s)\<inverse> \<longrightarrow> 
-           y \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (RAG s)\<^sup>+)"
-    show "x \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+)"
-    proof
-      assume x_d: "x \<in> Domain (RAG s)"
-      show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+"
-      proof(cases x)
-        case (Th th)
-        from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> RAG s" by (auto simp:s_RAG_def)
-        with Th have x_in_r: "(Cs cs, x) \<in> (RAG s)^-1" by simp
-        from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> RAG s" by blast
-        hence "Cs cs \<in> Domain (RAG s)" by auto
-        from ih [OF x_in_r this] obtain th'
-          where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (RAG s)\<^sup>+" by auto
-        have "(x, Th th') \<in> (RAG s)\<^sup>+" using Th x_in cs_in by auto
-        with th'_ready show ?thesis by auto
-      next
-        case (Cs cs)
-        from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (RAG s)^-1" by (auto simp:s_RAG_def)
-        show ?thesis
-        proof(cases "th' \<in> readys s")
-          case True
-          from True and th'_d show ?thesis by auto
-        next
-          case False
-          from th'_d and range_in  have "th' \<in> threads s" by auto
-          with False have "Th th' \<in> Domain (RAG s)" 
-            by (auto simp:readys_def wq_def s_waiting_def s_RAG_def cs_waiting_def Domain_def)
-          from ih [OF th'_d this]
-          obtain th'' where 
-            th''_r: "th'' \<in> readys s" and 
-            th''_in: "(Th th', Th th'') \<in> (RAG s)\<^sup>+" by auto
-          from th'_d and th''_in 
-          have "(x, Th th'') \<in> (RAG s)\<^sup>+" by auto
-          with th''_r show ?thesis by auto
-        qed
-      qed
-    qed
-  qed
-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 waiting_eq: "waiting s th cs = waiting (wq s) th cs"
-  by  (unfold s_waiting_def cs_waiting_def wq_def, auto)
-
-lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
-  by (unfold s_holding_def wq_def cs_holding_def, simp)
-
-lemma holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2"
-  by (unfold s_holding_def cs_holding_def, auto)
-
-context valid_trace
-begin
-
-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 holding_unique)
-
-end
-
-
-lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
-by (induct rule:trancl_induct, auto)
-
-context valid_trace
-begin
-
-lemma dchain_unique:
-  assumes th1_d: "(n, Th th1) \<in> (RAG s)^+"
-  and th1_r: "th1 \<in> readys s"
-  and th2_d: "(n, Th th2) \<in> (RAG s)^+"
-  and th2_r: "th2 \<in> readys s"
-  shows "th1 = th2"
-proof -
-  { assume neq: "th1 \<noteq> th2"
-    hence "Th th1 \<noteq> Th th2" by simp
-    from unique_chain [OF _ th1_d th2_d this] and unique_RAG 
-    have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG s)\<^sup>+" by auto
-    hence "False"
-    proof
-      assume "(Th th1, Th th2) \<in> (RAG s)\<^sup>+"
-      from trancl_split [OF this]
-      obtain n where dd: "(Th th1, n) \<in> RAG s" by auto
-      then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
-      from dd eq_n have "th1 \<notin> readys s"
-        by (auto simp:readys_def s_RAG_def wq_def s_waiting_def cs_waiting_def)
-      with th1_r show ?thesis by auto
-    next
-      assume "(Th th2, Th th1) \<in> (RAG s)\<^sup>+"
-      from trancl_split [OF this]
-      obtain n where dd: "(Th th2, n) \<in> RAG s" by auto
-      then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
-      from dd eq_n have "th2 \<notin> readys s"
-        by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def)
-      with th2_r show ?thesis by auto
-    qed
-  } thus ?thesis by auto
-qed
-
-end
-             
-
-lemma step_holdents_p_add:
-  assumes vt: "vt (P th cs#s)"
-  and "wq s cs = []"
-  shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
-proof -
-  from assms show ?thesis
-  unfolding  holdents_test step_RAG_p[OF vt] by (auto)
-qed
-
-lemma step_holdents_p_eq:
-  assumes vt: "vt (P th cs#s)"
-  and "wq s cs \<noteq> []"
-  shows "holdents (P th cs#s) th = holdents s th"
-proof -
-  from assms show ?thesis
-  unfolding  holdents_test step_RAG_p[OF vt] by auto
-qed
-
-
-lemma (in valid_trace) finite_holding :
-  shows "finite (holdents s th)"
-proof -
-  let ?F = "\<lambda> (x, y). the_cs x"
-  from finite_RAG 
-  have "finite (RAG s)" .
-  hence "finite (?F `(RAG s))" by simp
-  moreover have "{cs . (Cs cs, Th th) \<in> RAG s} \<subseteq> \<dots>" 
-  proof -
-    { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
-      fix x assume "(Cs x, Th th) \<in> RAG s"
-      hence "?F (Cs x, Th th) \<in> ?F `(RAG s)" by (rule h)
-      moreover have "?F (Cs x, Th th) = x" by simp
-      ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` RAG s" by simp 
-    } thus ?thesis by auto
-  qed
-  ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset)
-qed
-
-lemma cntCS_v_dec: 
-  assumes vtv: "vt (V thread cs#s)"
-  shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
-proof -
-  from vtv interpret vt_s: valid_trace s
-    by (cases, unfold_locales, simp)
-  from vtv interpret vt_v: valid_trace "V thread cs#s"
-     by (unfold_locales, simp)
-  from step_back_step[OF vtv]
-  have cs_in: "cs \<in> holdents s thread" 
-    apply (cases, unfold holdents_test s_RAG_def, simp)
-    by (unfold cs_holding_def s_holding_def wq_def, auto)
-  moreover have cs_not_in: 
-    "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
-    apply (insert vt_s.wq_distinct[of cs])
-    apply (unfold holdents_test, unfold step_RAG_v[OF vtv],
-            auto simp:next_th_def)
-  proof -
-    fix rest
-    assume dst: "distinct (rest::thread list)"
-      and ne: "rest \<noteq> []"
-    and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-    moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-    qed
-    ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
-                     set (SOME q. distinct q \<and> set q = set rest)" by simp
-    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      fix x assume " distinct x \<and> set x = set rest" with ne
-      show "x \<noteq> []" by auto
-    qed
-    ultimately 
-    show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
-      by auto
-  next
-    fix rest
-    assume dst: "distinct (rest::thread list)"
-      and ne: "rest \<noteq> []"
-    and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-    moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-    qed
-    ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
-                     set (SOME q. distinct q \<and> set q = set rest)" by simp
-    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      fix x assume " distinct x \<and> set x = set rest" with ne
-      show "x \<noteq> []" by auto
-    qed
-    ultimately show "False" by auto 
-  qed
-  ultimately 
-  have "holdents s thread = insert cs (holdents (V thread cs#s) thread)"
-    by auto
-  moreover have "card \<dots> = 
-                    Suc (card ((holdents (V thread cs#s) thread) - {cs}))"
-  proof(rule card_insert)
-    from vt_v.finite_holding
-    show " finite (holdents (V thread cs # s) thread)" .
-  qed
-  moreover from cs_not_in 
-  have "cs \<notin> (holdents (V thread cs#s) thread)" by auto
-  ultimately show ?thesis by (simp add:cntCS_def)
-qed 
-
-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_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 isP_E:
-  assumes "isP e"
-  obtains cs where "e = P (actor e) cs"
-  using assms by (cases e, auto)
-
-lemma isV_E:
-  assumes "isV e"
-  obtains cs where "e = V (actor e) cs"
-  using assms by (cases e, auto) (* ccc *)
-
-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)
-
-context valid_trace
-begin
-
-text {* (* ddd *) \noindent
-  The relationship between @{text "cntP"}, @{text "cntV"} and @{text "cntCS"} 
-  of one particular thread. 
-*} 
-
-lemma cnp_cnv_cncs:
-  shows "cntP s th = cntV s th + (if (th \<in> readys s \<or> th \<notin> threads s) 
-                                       then cntCS s th else cntCS s th + 1)"
-proof -
-  from vt show ?thesis
-  proof(induct arbitrary:th)
-    case (vt_cons s e)
-    interpret vt_s: valid_trace s using vt_cons(1) by (unfold_locales, simp)
-    assume vt: "vt s"
-    and ih: "\<And>th. cntP s th  = cntV s th +
-               (if (th \<in> readys s \<or> th \<notin> threads s) then cntCS s th else cntCS s th + 1)"
-    and stp: "step s e"
-    from stp show ?case
-    proof(cases)
-      case (thread_create thread prio)
-      assume eq_e: "e = Create thread prio"
-        and not_in: "thread \<notin> threads s"
-      show ?thesis
-      proof -
-        { fix cs 
-          assume "thread \<in> set (wq s cs)"
-          from vt_s.wq_threads [OF this] have "thread \<in> threads s" .
-          with not_in have "False" by simp
-        } with eq_e have eq_readys: "readys (e#s) = readys s \<union> {thread}"
-          by (auto simp:readys_def threads.simps s_waiting_def 
-            wq_def cs_waiting_def Let_def)
-        from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
-        from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
-        have eq_cncs: "cntCS (e#s) th = cntCS s th"
-          unfolding cntCS_def holdents_test
-          by (simp add:RAG_create_unchanged eq_e)
-        { assume "th \<noteq> thread"
-          with eq_readys eq_e
-          have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
-                      (th \<in> readys (s) \<or> th \<notin> threads (s))" 
-            by (simp add:threads.simps)
-          with eq_cnp eq_cnv eq_cncs ih not_in
-          have ?thesis by simp
-        } moreover {
-          assume eq_th: "th = thread"
-          with not_in ih have " cntP s th  = cntV s th + cntCS s th" by simp
-          moreover from eq_th and eq_readys have "th \<in> readys (e#s)" by simp
-          moreover note eq_cnp eq_cnv eq_cncs
-          ultimately have ?thesis by auto
-        } ultimately show ?thesis by blast
-      qed
-    next
-      case (thread_exit thread)
-      assume eq_e: "e = Exit thread" 
-      and is_runing: "thread \<in> runing s"
-      and no_hold: "holdents s thread = {}"
-      from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
-      from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
-      have eq_cncs: "cntCS (e#s) th = cntCS s th"
-        unfolding cntCS_def holdents_test
-        by (simp add:RAG_exit_unchanged eq_e)
-      { assume "th \<noteq> thread"
-        with eq_e
-        have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
-          (th \<in> readys (s) \<or> th \<notin> threads (s))" 
-          apply (simp add:threads.simps readys_def)
-          apply (subst s_waiting_def)
-          apply (simp add:Let_def)
-          apply (subst s_waiting_def, simp)
-          done
-        with eq_cnp eq_cnv eq_cncs ih
-        have ?thesis by simp
-      } moreover {
-        assume eq_th: "th = thread"
-        with ih is_runing have " cntP s th = cntV s th + cntCS s th" 
-          by (simp add:runing_def)
-        moreover from eq_th eq_e have "th \<notin> threads (e#s)"
-          by simp
-        moreover note eq_cnp eq_cnv eq_cncs
-        ultimately have ?thesis by auto
-      } ultimately show ?thesis by blast
-    next
-      case (thread_P thread cs)
-      assume eq_e: "e = P thread cs"
-        and is_runing: "thread \<in> runing s"
-        and no_dep: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+"
-      from thread_P vt stp ih  have vtp: "vt (P thread cs#s)" by auto
-      then interpret vt_p: valid_trace "(P thread cs#s)"
-        by (unfold_locales, simp)
-      show ?thesis 
-      proof -
-        { have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast
-          assume neq_th: "th \<noteq> thread"
-          with eq_e
-          have eq_readys: "(th \<in> readys (e#s)) = (th \<in> readys (s))"
-            apply (simp add:readys_def s_waiting_def wq_def Let_def)
-            apply (rule_tac hh)
-             apply (intro iffI allI, clarify)
-            apply (erule_tac x = csa in allE, auto)
-            apply (subgoal_tac "wq_fun (schs s) cs \<noteq> []", auto)
-            apply (erule_tac x = cs in allE, auto)
-            by (case_tac "(wq_fun (schs s) cs)", auto)
-          moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th"
-            apply (simp add:cntCS_def holdents_test)
-            by (unfold  step_RAG_p [OF vtp], auto)
-          moreover from eq_e neq_th have "cntP (e # s) th = cntP s th"
-            by (simp add:cntP_def count_def)
-          moreover from eq_e neq_th have "cntV (e#s) th = cntV s th"
-            by (simp add:cntV_def count_def)
-          moreover from eq_e neq_th have "threads (e#s) = threads s" by simp
-          moreover note ih [of th] 
-          ultimately have ?thesis by simp
-        } moreover {
-          assume eq_th: "th = thread"
-          have ?thesis
-          proof -
-            from eq_e eq_th have eq_cnp: "cntP (e # s) th  = 1 + (cntP s th)" 
-              by (simp add:cntP_def count_def)
-            from eq_e eq_th have eq_cnv: "cntV (e#s) th = cntV s th"
-              by (simp add:cntV_def count_def)
-            show ?thesis
-            proof (cases "wq s cs = []")
-              case True
-              with is_runing
-              have "th \<in> readys (e#s)"
-                apply (unfold eq_e wq_def, unfold readys_def s_RAG_def)
-                apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def)
-                by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def)
-              moreover have "cntCS (e # s) th = 1 + cntCS s th"
-              proof -
-                have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> RAG s} =
-                  Suc (card {cs. (Cs cs, Th thread) \<in> RAG s})" (is "card ?L = Suc (card ?R)")
-                proof -
-                  have "?L = insert cs ?R" by auto
-                  moreover have "card \<dots> = Suc (card (?R - {cs}))" 
-                  proof(rule card_insert)
-                    from vt_s.finite_holding [of thread]
-                    show " finite {cs. (Cs cs, Th thread) \<in> RAG s}"
-                      by (unfold holdents_test, simp)
-                  qed
-                  moreover have "?R - {cs} = ?R"
-                  proof -
-                    have "cs \<notin> ?R"
-                    proof
-                      assume "cs \<in> {cs. (Cs cs, Th thread) \<in> RAG s}"
-                      with no_dep show False by auto
-                    qed
-                    thus ?thesis by auto
-                  qed
-                  ultimately show ?thesis by auto
-                qed
-                thus ?thesis
-                  apply (unfold eq_e eq_th cntCS_def)
-                  apply (simp add: holdents_test)
-                  by (unfold step_RAG_p [OF vtp], auto simp:True)
-              qed
-              moreover from is_runing have "th \<in> readys s"
-                by (simp add:runing_def eq_th)
-              moreover note eq_cnp eq_cnv ih [of th]
-              ultimately show ?thesis by auto
-            next
-              case False
-              have eq_wq: "wq (e#s) cs = wq s cs @ [th]"
-                    by (unfold eq_th eq_e wq_def, auto simp:Let_def)
-              have "th \<notin> readys (e#s)"
-              proof
-                assume "th \<in> readys (e#s)"
-                hence "\<forall>cs. \<not> waiting (e # s) th cs" by (simp add:readys_def)
-                from this[rule_format, of cs] have " \<not> waiting (e # s) th cs" .
-                hence "th \<in> set (wq (e#s) cs) \<Longrightarrow> th = hd (wq (e#s) cs)" 
-                  by (simp add:s_waiting_def wq_def)
-                moreover from eq_wq have "th \<in> set (wq (e#s) cs)" by auto
-                ultimately have "th = hd (wq (e#s) cs)" by blast
-                with eq_wq have "th = hd (wq s cs @ [th])" by simp
-                hence "th = hd (wq s cs)" using False by auto
-                with False eq_wq vt_p.wq_distinct [of cs]
-                show False by (fold eq_e, auto)
-              qed
-              moreover from is_runing have "th \<in> threads (e#s)" 
-                by (unfold eq_e, auto simp:runing_def readys_def eq_th)
-              moreover have "cntCS (e # s) th = cntCS s th"
-                apply (unfold cntCS_def holdents_test eq_e step_RAG_p[OF vtp])
-                by (auto simp:False)
-              moreover note eq_cnp eq_cnv ih[of th]
-              moreover from is_runing have "th \<in> readys s"
-                by (simp add:runing_def eq_th)
-              ultimately show ?thesis by auto
-            qed
-          qed
-        } ultimately show ?thesis by blast
-      qed
-    next
-      case (thread_V thread cs)
-      from assms vt stp ih thread_V have vtv: "vt (V thread cs # s)" by auto
-      then interpret vt_v: valid_trace "(V thread cs # s)" by (unfold_locales, simp)
-      assume eq_e: "e = V thread cs"
-        and is_runing: "thread \<in> runing s"
-        and hold: "holding s thread cs"
-      from hold obtain rest 
-        where eq_wq: "wq s cs = thread # rest"
-        by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
-      have eq_threads: "threads (e#s) = threads s" by (simp add: eq_e)
-      have eq_set: "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-      proof(rule someI2)
-        from vt_v.wq_distinct[of cs] and eq_wq
-        show "distinct rest \<and> set rest = set rest"
-          by (metis distinct.simps(2) vt_s.wq_distinct)
-      next
-        show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
-          by auto
-      qed
-      show ?thesis
-      proof -
-        { assume eq_th: "th = thread"
-          from eq_th have eq_cnp: "cntP (e # s) th = cntP s th"
-            by (unfold eq_e, simp add:cntP_def count_def)
-          moreover from eq_th have eq_cnv: "cntV (e#s) th = 1 + cntV s th"
-            by (unfold eq_e, simp add:cntV_def count_def)
-          moreover from cntCS_v_dec [OF vtv] 
-          have "cntCS (e # s) thread + 1 = cntCS s thread"
-            by (simp add:eq_e)
-          moreover from is_runing have rd_before: "thread \<in> readys s"
-            by (unfold runing_def, simp)
-          moreover have "thread \<in> readys (e # s)"
-          proof -
-            from is_runing
-            have "thread \<in> threads (e#s)" 
-              by (unfold eq_e, auto simp:runing_def readys_def)
-            moreover have "\<forall> cs1. \<not> waiting (e#s) thread cs1"
-            proof
-              fix cs1
-              { assume eq_cs: "cs1 = cs" 
-                have "\<not> waiting (e # s) thread cs1"
-                proof -
-                  from eq_wq
-                  have "thread \<notin> set (wq (e#s) cs1)"
-                    apply(unfold eq_e wq_def eq_cs s_holding_def)
-                    apply (auto simp:Let_def)
-                  proof -
-                    assume "thread \<in> set (SOME q. distinct q \<and> set q = set rest)"
-                    with eq_set have "thread \<in> set rest" by simp
-                    with vt_v.wq_distinct[of cs]
-                    and eq_wq show False
-                        by (metis distinct.simps(2) vt_s.wq_distinct)
-                  qed
-                  thus ?thesis by (simp add:wq_def s_waiting_def)
-                qed
-              } moreover {
-                assume neq_cs: "cs1 \<noteq> cs"
-                  have "\<not> waiting (e # s) thread cs1" 
-                  proof -
-                    from wq_v_neq [OF neq_cs[symmetric]]
-                    have "wq (V thread cs # s) cs1 = wq s cs1" .
-                    moreover have "\<not> waiting s thread cs1" 
-                    proof -
-                      from runing_ready and is_runing
-                      have "thread \<in> readys s" by auto
-                      thus ?thesis by (simp add:readys_def)
-                    qed
-                    ultimately show ?thesis 
-                      by (auto simp:wq_def s_waiting_def eq_e)
-                  qed
-              } ultimately show "\<not> waiting (e # s) thread cs1" by blast
-            qed
-            ultimately show ?thesis by (simp add:readys_def)
-          qed
-          moreover note eq_th ih
-          ultimately have ?thesis by auto
-        } moreover {
-          assume neq_th: "th \<noteq> thread"
-          from neq_th eq_e have eq_cnp: "cntP (e # s) th = cntP s th" 
-            by (simp add:cntP_def count_def)
-          from neq_th eq_e have eq_cnv: "cntV (e # s) th = cntV s th" 
-            by (simp add:cntV_def count_def)
-          have ?thesis
-          proof(cases "th \<in> set rest")
-            case False
-            have "(th \<in> readys (e # s)) = (th \<in> readys s)"
-              apply (insert step_back_vt[OF vtv])
-              by (simp add: False eq_e eq_wq neq_th vt_s.readys_v_eq)
-            moreover have "cntCS (e#s) th = cntCS s th"
-              apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
-              proof -
-                have "{csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
-                      {cs. (Cs cs, Th th) \<in> RAG s}"
-                proof -
-                  from False eq_wq
-                  have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> RAG s"
-                    apply (unfold next_th_def, auto)
-                  proof -
-                    assume ne: "rest \<noteq> []"
-                      and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-                      and eq_wq: "wq s cs = thread # rest"
-                    from eq_set ni have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
-                                  set (SOME q. distinct q \<and> set q = set rest)
-                                  " by simp
-                    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-                    proof(rule someI2)
-                      from vt_s.wq_distinct[ of cs] and eq_wq
-                      show "distinct rest \<and> set rest = set rest" by auto
-                    next
-                      fix x assume "distinct x \<and> set x = set rest"
-                      with ne show "x \<noteq> []" by auto
-                    qed
-                    ultimately show 
-                      "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
-                      by auto
-                  qed    
-                  thus ?thesis by auto
-                qed
-                thus "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
-                             card {cs. (Cs cs, Th th) \<in> RAG s}" by simp 
-              qed
-            moreover note ih eq_cnp eq_cnv eq_threads
-            ultimately show ?thesis by auto
-          next
-            case True
-            assume th_in: "th \<in> set rest"
-            show ?thesis
-            proof(cases "next_th s thread cs th")
-              case False
-              with eq_wq and th_in have 
-                neq_hd: "th \<noteq> hd (SOME q. distinct q \<and> set q = set rest)" (is "th \<noteq> hd ?rest")
-                by (auto simp:next_th_def)
-              have "(th \<in> readys (e # s)) = (th \<in> readys s)"
-              proof -
-                from eq_wq and th_in
-                have "\<not> th \<in> readys s"
-                  apply (auto simp:readys_def s_waiting_def)
-                  apply (rule_tac x = cs in exI, auto)
-                  by (insert vt_s.wq_distinct[of cs], auto simp add: wq_def)
-                moreover 
-                from eq_wq and th_in and neq_hd
-                have "\<not> (th \<in> readys (e # s))"
-                  apply (auto simp:readys_def s_waiting_def eq_e wq_def Let_def split:list.splits)
-                  by (rule_tac x = cs in exI, auto simp:eq_set)
-                ultimately show ?thesis by auto
-              qed
-              moreover have "cntCS (e#s) th = cntCS s th" 
-              proof -
-                from eq_wq and  th_in and neq_hd
-                have "(holdents (e # s) th) = (holdents s th)"
-                  apply (unfold eq_e step_RAG_v[OF vtv], 
-                         auto simp:next_th_def eq_set s_RAG_def holdents_test wq_def
-                                   Let_def cs_holding_def)
-                  by (insert vt_s.wq_distinct[of cs], auto simp:wq_def)
-                thus ?thesis by (simp add:cntCS_def)
-              qed
-              moreover note ih eq_cnp eq_cnv eq_threads
-              ultimately show ?thesis by auto
-            next
-              case True
-              let ?rest = " (SOME q. distinct q \<and> set q = set rest)"
-              let ?t = "hd ?rest"
-              from True eq_wq th_in neq_th
-              have "th \<in> readys (e # s)"
-                apply (auto simp:eq_e readys_def s_waiting_def wq_def
-                        Let_def next_th_def)
-              proof -
-                assume eq_wq: "wq_fun (schs s) cs = thread # rest"
-                  and t_in: "?t \<in> set rest"
-                show "?t \<in> threads s"
-                proof(rule vt_s.wq_threads)
-                  from eq_wq and t_in
-                  show "?t \<in> set (wq s cs)" by (auto simp:wq_def)
-                qed
-              next
-                fix csa
-                assume eq_wq: "wq_fun (schs s) cs = thread # rest"
-                  and t_in: "?t \<in> set rest"
-                  and neq_cs: "csa \<noteq> cs"
-                  and t_in': "?t \<in>  set (wq_fun (schs s) csa)"
-                show "?t = hd (wq_fun (schs s) csa)"
-                proof -
-                  { assume neq_hd': "?t \<noteq> hd (wq_fun (schs s) csa)"
-                    from vt_s.wq_distinct[of cs] and 
-                    eq_wq[folded wq_def] and t_in eq_wq
-                    have "?t \<noteq> thread" by auto
-                    with eq_wq and t_in
-                    have w1: "waiting s ?t cs"
-                      by (auto simp:s_waiting_def wq_def)
-                    from t_in' neq_hd'
-                    have w2: "waiting s ?t csa"
-                      by (auto simp:s_waiting_def wq_def)
-                    from vt_s.waiting_unique[OF w1 w2]
-                    and neq_cs have "False" by auto
-                  } thus ?thesis by auto
-                qed
-              qed
-              moreover have "cntP s th = cntV s th + cntCS s th + 1"
-              proof -
-                have "th \<notin> readys s" 
-                proof -
-                  from True eq_wq neq_th th_in
-                  show ?thesis
-                    apply (unfold readys_def s_waiting_def, auto)
-                    by (rule_tac x = cs in exI, auto simp add: wq_def)
-                qed
-                moreover have "th \<in> threads s"
-                proof -
-                  from th_in eq_wq
-                  have "th \<in> set (wq s cs)" by simp
-                  from vt_s.wq_threads [OF this] 
-                  show ?thesis .
-                qed
-                ultimately show ?thesis using ih by auto
-              qed
-              moreover from True neq_th have "cntCS (e # s) th = 1 + cntCS s th"
-                apply (unfold cntCS_def holdents_test eq_e step_RAG_v[OF vtv], auto)
-              proof -
-                show "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs} =
-                               Suc (card {cs. (Cs cs, Th th) \<in> RAG s})"
-                  (is "card ?A = Suc (card ?B)")
-                proof -
-                  have "?A = insert cs ?B" by auto
-                  hence "card ?A = card (insert cs ?B)" by simp
-                  also have "\<dots> = Suc (card ?B)"
-                  proof(rule card_insert_disjoint)
-                    have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` RAG s)" 
-                      apply (auto simp:image_def)
-                      by (rule_tac x = "(Cs x, Th th)" in bexI, auto)
-                    with vt_s.finite_RAG
-                    show "finite {cs. (Cs cs, Th th) \<in> RAG s}" by (auto intro:finite_subset)
-                  next
-                    show "cs \<notin> {cs. (Cs cs, Th th) \<in> RAG s}"
-                    proof
-                      assume "cs \<in> {cs. (Cs cs, Th th) \<in> RAG s}"
-                      hence "(Cs cs, Th th) \<in> RAG s" by simp
-                      with True neq_th eq_wq show False
-                        by (auto simp:next_th_def s_RAG_def cs_holding_def)
-                    qed
-                  qed
-                  finally show ?thesis .
-                qed
-              qed
-              moreover note eq_cnp eq_cnv
-              ultimately show ?thesis by simp
-            qed
-          qed
-        } ultimately show ?thesis by blast
-      qed
-    next
-      case (thread_set thread prio)
-      assume eq_e: "e = Set thread prio"
-        and is_runing: "thread \<in> runing s"
-      show ?thesis
-      proof -
-        from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
-        from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
-        have eq_cncs: "cntCS (e#s) th = cntCS s th"
-          unfolding cntCS_def holdents_test
-          by (simp add:RAG_set_unchanged eq_e)
-        from eq_e have eq_readys: "readys (e#s) = readys s" 
-          by (simp add:readys_def cs_waiting_def s_waiting_def wq_def,
-                  auto simp:Let_def)
-        { assume "th \<noteq> thread"
-          with eq_readys eq_e
-          have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
-                      (th \<in> readys (s) \<or> th \<notin> threads (s))" 
-            by (simp add:threads.simps)
-          with eq_cnp eq_cnv eq_cncs ih is_runing
-          have ?thesis by simp
-        } moreover {
-          assume eq_th: "th = thread"
-          with is_runing ih have " cntP s th  = cntV s th + cntCS s th" 
-            by (unfold runing_def, auto)
-          moreover from eq_th and eq_readys is_runing have "th \<in> readys (e#s)"
-            by (simp add:runing_def)
-          moreover note eq_cnp eq_cnv eq_cncs
-          ultimately have ?thesis by auto
-        } ultimately show ?thesis by blast
-      qed   
-    qed
-  next
-    case vt_nil
-    show ?case 
-      by (unfold cntP_def cntV_def cntCS_def, 
-        auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
-  qed
-qed
-
-lemma not_thread_cncs:
-  assumes not_in: "th \<notin> threads s" 
-  shows "cntCS s th = 0"
-proof -
-  from vt not_in show ?thesis
-  proof(induct arbitrary:th)
-    case (vt_cons s e th)
-    interpret vt_s: valid_trace s using vt_cons(1)
-       by (unfold_locales, simp)
-    assume vt: "vt s"
-      and ih: "\<And>th. th \<notin> threads s \<Longrightarrow> cntCS s th = 0"
-      and stp: "step s e"
-      and not_in: "th \<notin> threads (e # s)"
-    from stp show ?case
-    proof(cases)
-      case (thread_create thread prio)
-      assume eq_e: "e = Create thread prio"
-        and not_in': "thread \<notin> threads s"
-      have "cntCS (e # s) th = cntCS s th"
-        apply (unfold eq_e cntCS_def holdents_test)
-        by (simp add:RAG_create_unchanged)
-      moreover have "th \<notin> threads s" 
-      proof -
-        from not_in eq_e show ?thesis by simp
-      qed
-      moreover note ih ultimately show ?thesis by auto
-    next
-      case (thread_exit thread)
-      assume eq_e: "e = Exit thread"
-      and nh: "holdents s thread = {}"
-      have eq_cns: "cntCS (e # s) th = cntCS s th"
-        apply (unfold eq_e cntCS_def holdents_test)
-        by (simp add:RAG_exit_unchanged)
-      show ?thesis
-      proof(cases "th = thread")
-        case True
-        have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True)
-        with eq_cns show ?thesis by simp
-      next
-        case False
-        with not_in and eq_e
-        have "th \<notin> threads s" by simp
-        from ih[OF this] and eq_cns show ?thesis by simp
-      qed
-    next
-      case (thread_P thread cs)
-      assume eq_e: "e = P thread cs"
-      and is_runing: "thread \<in> runing s"
-      from assms thread_P ih vt stp thread_P have vtp: "vt (P thread cs#s)" by auto
-      have neq_th: "th \<noteq> thread" 
-      proof -
-        from not_in eq_e have "th \<notin> threads s" by simp
-        moreover from is_runing have "thread \<in> threads s"
-          by (simp add:runing_def readys_def)
-        ultimately show ?thesis by auto
-      qed
-      hence "cntCS (e # s) th  = cntCS s th "
-        apply (unfold cntCS_def holdents_test eq_e)
-        by (unfold step_RAG_p[OF vtp], auto)
-      moreover have "cntCS s th = 0"
-      proof(rule ih)
-        from not_in eq_e show "th \<notin> threads s" by simp
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (thread_V thread cs)
-      assume eq_e: "e = V thread cs"
-        and is_runing: "thread \<in> runing s"
-        and hold: "holding s thread cs"
-      have neq_th: "th \<noteq> thread" 
-      proof -
-        from not_in eq_e have "th \<notin> threads s" by simp
-        moreover from is_runing have "thread \<in> threads s"
-          by (simp add:runing_def readys_def)
-        ultimately show ?thesis by auto
-      qed
-      from assms thread_V vt stp ih 
-      have vtv: "vt (V thread cs#s)" by auto
-      then interpret vt_v: valid_trace "(V thread cs#s)"
-        by (unfold_locales, simp)
-      from hold obtain rest 
-        where eq_wq: "wq s cs = thread # rest"
-        by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
-      from not_in eq_e eq_wq
-      have "\<not> next_th s thread cs th"
-        apply (auto simp:next_th_def)
-      proof -
-        assume ne: "rest \<noteq> []"
-          and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> threads s" (is "?t \<notin> threads s")
-        have "?t \<in> set rest"
-        proof(rule someI2)
-          from vt_v.wq_distinct[of cs] and eq_wq
-          show "distinct rest \<and> set rest = set rest"
-            by (metis distinct.simps(2) vt_s.wq_distinct) 
-        next
-          fix x assume "distinct x \<and> set x = set rest" with ne
-          show "hd x \<in> set rest" by (cases x, auto)
-        qed
-        with eq_wq have "?t \<in> set (wq s cs)" by simp
-        from vt_s.wq_threads[OF this] and ni
-        show False
-          using `hd (SOME q. distinct q \<and> set q = set rest) \<in> set (wq s cs)` 
-            ni vt_s.wq_threads by blast 
-      qed
-      moreover note neq_th eq_wq
-      ultimately have "cntCS (e # s) th  = cntCS s th"
-        by (unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
-      moreover have "cntCS s th = 0"
-      proof(rule ih)
-        from not_in eq_e show "th \<notin> threads s" by simp
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (thread_set thread prio)
-      print_facts
-      assume eq_e: "e = Set thread prio"
-        and is_runing: "thread \<in> runing s"
-      from not_in and eq_e have "th \<notin> threads s" by auto
-      from ih [OF this] and eq_e
-      show ?thesis 
-        apply (unfold eq_e cntCS_def holdents_test)
-        by (simp add:RAG_set_unchanged)
-    qed
-    next
-      case vt_nil
-      show ?case
-      by (unfold cntCS_def, 
-        auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
-  qed
-qed
-
-end
-
-lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs"
-  by (auto simp:s_waiting_def cs_waiting_def wq_def)
-
-context valid_trace
-begin
-
-lemma dm_RAG_threads:
-  assumes in_dom: "(Th th) \<in> Domain (RAG s)"
-  shows "th \<in> threads s"
-proof -
-  from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
-  moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
-  ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
-  hence "th \<in> set (wq s cs)"
-    by (unfold s_RAG_def, auto simp:cs_waiting_def)
-  from wq_threads [OF this] show ?thesis .
-qed
-
-end
-
-lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
-unfolding cp_def wq_def
-apply(induct s rule: schs.induct)
-thm cpreced_initial
-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
-
-context valid_trace
-begin
-
-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
-    apply(simp)
-    done
-  hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1)) =
-                 Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependants (wq s) th2))"
-    (is "Max (?f ` ?A) = Max (?f ` ?B)")
-    unfolding cp_eq_cpreced 
-    unfolding cpreced_def .
-  obtain th1' where th1_in: "th1' \<in> ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)"
-  proof -
-    have h1: "finite (?f ` ?A)"
-    proof -
-      have "finite ?A" 
-      proof -
-        have "finite (dependants (wq s) th1)"
-        proof-
-          have "finite {th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+}"
-          proof -
-            let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
-              apply (auto simp:image_def)
-              by (rule_tac x = "(Th x, Th th1)" in bexI, auto)
-            moreover have "finite \<dots>"
-            proof -
-              from finite_RAG have "finite (RAG s)" .
-              hence "finite ((RAG (wq s))\<^sup>+)"
-                apply (unfold finite_trancl)
-                by (auto simp: s_RAG_def cs_RAG_def wq_def)
-              thus ?thesis by auto
-            qed
-            ultimately show ?thesis by (auto intro:finite_subset)
-          qed
-          thus ?thesis by (simp add:cs_dependants_def)
-        qed
-        thus ?thesis by simp
-      qed
-      thus ?thesis by auto
-    qed
-    moreover have h2: "(?f ` ?A) \<noteq> {}"
-    proof -
-      have "?A \<noteq> {}" by simp
-      thus ?thesis by simp
-    qed
-    from Max_in [OF h1 h2]
-    have "Max (?f ` ?A) \<in> (?f ` ?A)" .
-    thus ?thesis 
-      thm cpreced_def
-      unfolding cpreced_def[symmetric] 
-      unfolding cp_eq_cpreced[symmetric] 
-      unfolding cpreced_def 
-      using that[intro] by (auto)
-  qed
-  obtain th2' where th2_in: "th2' \<in> ?B" and eq_f_th2: "?f th2' = Max (?f ` ?B)"
-  proof -
-    have h1: "finite (?f ` ?B)"
-    proof -
-      have "finite ?B" 
-      proof -
-        have "finite (dependants (wq s) th2)"
-        proof-
-          have "finite {th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+}"
-          proof -
-            let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
-              apply (auto simp:image_def)
-              by (rule_tac x = "(Th x, Th th2)" in bexI, auto)
-            moreover have "finite \<dots>"
-            proof -
-              from finite_RAG have "finite (RAG s)" .
-              hence "finite ((RAG (wq s))\<^sup>+)"
-                apply (unfold finite_trancl)
-                by (auto simp: s_RAG_def cs_RAG_def wq_def)
-              thus ?thesis by auto
-            qed
-            ultimately show ?thesis by (auto intro:finite_subset)
-          qed
-          thus ?thesis by (simp add:cs_dependants_def)
-        qed
-        thus ?thesis by simp
-      qed
-      thus ?thesis by auto
-    qed
-    moreover have h2: "(?f ` ?B) \<noteq> {}"
-    proof -
-      have "?B \<noteq> {}" by simp
-      thus ?thesis by simp
-    qed
-    from Max_in [OF h1 h2]
-    have "Max (?f ` ?B) \<in> (?f ` ?B)" .
-    thus ?thesis by (auto intro:that)
-  qed
-  from eq_f_th1 eq_f_th2 eq_max 
-  have eq_preced: "preced th1' s = preced th2' s" by auto
-  hence eq_th12: "th1' = th2'"
-  proof (rule preced_unique)
-    from th1_in have "th1' = th1 \<or> (th1' \<in> dependants (wq s) th1)" by simp
-    thus "th1' \<in> threads s"
-    proof
-      assume "th1' \<in> dependants (wq s) th1"
-      hence "(Th th1') \<in> Domain ((RAG s)^+)"
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "(Th th1') \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      from dm_RAG_threads[OF this] show ?thesis .
-    next
-      assume "th1' = th1"
-      with runing_1 show ?thesis
-        by (unfold runing_def readys_def, auto)
-    qed
-  next
-    from th2_in have "th2' = th2 \<or> (th2' \<in> dependants (wq s) th2)" by simp
-    thus "th2' \<in> threads s"
-    proof
-      assume "th2' \<in> dependants (wq s) th2"
-      hence "(Th th2') \<in> Domain ((RAG s)^+)"
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "(Th th2') \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      from dm_RAG_threads[OF this] show ?thesis .
-    next
-      assume "th2' = th2"
-      with runing_2 show ?thesis
-        by (unfold runing_def readys_def, auto)
-    qed
-  qed
-  from th1_in have "th1' = th1 \<or> th1' \<in> dependants (wq s) th1" by simp
-  thus ?thesis
-  proof
-    assume eq_th': "th1' = th1"
-    from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
-    thus ?thesis
-    proof
-      assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
-    next
-      assume "th2' \<in> dependants (wq s) th2"
-      with eq_th12 eq_th' have "th1 \<in> dependants (wq s) th2" by simp
-      hence "(Th th1, Th th2) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      hence "Th th1 \<in> Domain ((RAG s)^+)" 
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "Th th1 \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      then obtain n where d: "(Th th1, n) \<in> RAG s" by (auto simp:Domain_def)
-      from RAG_target_th [OF this]
-      obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th1, Cs cs') \<in> RAG s" by simp
-      with runing_1 have "False"
-        apply (unfold runing_def readys_def s_RAG_def)
-        by (auto simp:eq_waiting)
-      thus ?thesis by simp
-    qed
-  next
-    assume th1'_in: "th1' \<in> dependants (wq s) th1"
-    from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
-    thus ?thesis 
-    proof
-      assume "th2' = th2"
-      with th1'_in eq_th12 have "th2 \<in> dependants (wq s) th1" by simp
-      hence "(Th th2, Th th1) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      hence "Th th2 \<in> Domain ((RAG s)^+)" 
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "Th th2 \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      then obtain n where d: "(Th th2, n) \<in> RAG s" by (auto simp:Domain_def)
-      from RAG_target_th [OF this]
-      obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th2, Cs cs') \<in> RAG s" by simp
-      with runing_2 have "False"
-        apply (unfold runing_def readys_def s_RAG_def)
-        by (auto simp:eq_waiting)
-      thus ?thesis by simp
-    next
-      assume "th2' \<in> dependants (wq s) th2"
-      with eq_th12 have "th1' \<in> dependants (wq s) th2" by simp
-      hence h1: "(Th th1', Th th2) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      from th1'_in have h2: "(Th th1', Th th1) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      show ?thesis
-      proof(rule dchain_unique[OF h1 _ h2, symmetric])
-        from runing_1 show "th1 \<in> readys s" by (simp add:runing_def)
-        from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) 
-      qed
-    qed
-  qed
-qed
-
-
-lemma "card (runing s) \<le> 1"
-apply(subgoal_tac "finite (runing s)")
-prefer 2
-apply (metis finite_nat_set_iff_bounded lessI runing_unique)
-apply(rule ccontr)
-apply(simp)
-apply(case_tac "Suc (Suc 0) \<le> card (runing s)")
-apply(subst (asm) card_le_Suc_iff)
-apply(simp)
-apply(auto)[1]
-apply (metis insertCI runing_unique)
-apply(auto) 
-done
-
-end
-
-
-lemma create_pre:
-  assumes stp: "step s e"
-  and not_in: "th \<notin> threads s"
-  and is_in: "th \<in> threads (e#s)"
-  obtains prio where "e = Create th prio"
-proof -
-  from assms  
-  show ?thesis
-  proof(cases)
-    case (thread_create thread prio)
-    with is_in not_in have "e = Create th prio" by simp
-    from that[OF this] show ?thesis .
-  next
-    case (thread_exit thread)
-    with assms show ?thesis by (auto intro!:that)
-  next
-    case (thread_P thread)
-    with assms show ?thesis by (auto intro!:that)
-  next
-    case (thread_V thread)
-    with assms show ?thesis by (auto intro!:that)
-  next 
-    case (thread_set thread)
-    with assms show ?thesis by (auto intro!:that)
-  qed
-qed
-
-
-context valid_trace
-begin
-
-lemma cnp_cnv_eq:
-  assumes "th \<notin> threads s"
-  shows "cntP s th = cntV s th"
-  using assms
-  using cnp_cnv_cncs not_thread_cncs by auto
-
-end
-
-
-lemma eq_RAG: 
-  "RAG (wq s) = RAG s"
-by (unfold cs_RAG_def s_RAG_def, auto)
-
-context valid_trace
-begin
-
-lemma count_eq_dependants:
-  assumes eq_pv: "cntP s th = cntV s th"
-  shows "dependants (wq s) th = {}"
-proof -
-  from cnp_cnv_cncs and eq_pv
-  have "cntCS s th = 0" 
-    by (auto split:if_splits)
-  moreover have "finite {cs. (Cs cs, Th th) \<in> RAG s}"
-  proof -
-    from finite_holding[of th] show ?thesis
-      by (simp add:holdents_test)
-  qed
-  ultimately have h: "{cs. (Cs cs, Th th) \<in> RAG s} = {}"
-    by (unfold cntCS_def holdents_test cs_dependants_def, auto)
-  show ?thesis
-  proof(unfold cs_dependants_def)
-    { assume "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}"
-      then obtain th' where "(Th th', Th th) \<in> (RAG (wq s))\<^sup>+" by auto
-      hence "False"
-      proof(cases)
-        assume "(Th th', Th th) \<in> RAG (wq s)"
-        thus "False" by (auto simp:cs_RAG_def)
-      next
-        fix c
-        assume "(c, Th th) \<in> RAG (wq s)"
-        with h and eq_RAG show "False"
-          by (cases c, auto simp:cs_RAG_def)
-      qed
-    } thus "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} = {}" by auto
-  qed
-qed
-
-lemma dependants_threads:
-  shows "dependants (wq s) th \<subseteq> threads s"
-proof
-  { fix th th'
-    assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (RAG (wq s))\<^sup>+}"
-    have "Th th \<in> Domain (RAG s)"
-    proof -
-      from h obtain th' where "(Th th, Th th') \<in> (RAG (wq s))\<^sup>+" by auto
-      hence "(Th th) \<in> Domain ( (RAG (wq s))\<^sup>+)" by (auto simp:Domain_def)
-      with trancl_domain have "(Th th) \<in> Domain (RAG (wq s))" by simp
-      thus ?thesis using eq_RAG by simp
-    qed
-    from dm_RAG_threads[OF this]
-    have "th \<in> threads s" .
-  } note hh = this
-  fix th1 
-  assume "th1 \<in> dependants (wq s) th"
-  hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (RAG (wq s))\<^sup>+}"
-    by (unfold cs_dependants_def, simp)
-  from hh [OF this] show "th1 \<in> threads s" .
-qed
-
-lemma finite_threads:
-  shows "finite (threads s)"
-using vt by (induct) (auto elim: step.cases)
-
-end
-
-lemma Max_f_mono:
-  assumes seq: "A \<subseteq> B"
-  and np: "A \<noteq> {}"
-  and fnt: "finite B"
-  shows "Max (f ` A) \<le> Max (f ` B)"
-proof(rule Max_mono)
-  from seq show "f ` A \<subseteq> f ` B" by auto
-next
-  from np show "f ` A \<noteq> {}" by auto
-next
-  from fnt and seq show "finite (f ` B)" by auto
-qed
-
-context valid_trace
-begin
-
-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 le_cp:
-  shows "preced th s \<le> cp s th"
-proof(unfold cp_eq_cpreced preced_def cpreced_def, simp)
-  show "Prc (priority th s) (last_set th s)
-    \<le> Max (insert (Prc (priority th s) (last_set th s))
-            ((\<lambda>th. Prc (priority th s) (last_set th s)) ` dependants (wq s) th))"
-    (is "?l \<le> Max (insert ?l ?A)")
-  proof(cases "?A = {}")
-    case False
-    have "finite ?A" (is "finite (?f ` ?B)")
-    proof -
-      have "finite ?B" 
-      proof-
-        have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}"
-        proof -
-          let ?F = "\<lambda> (x, y). the_th x"
-          have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
-            apply (auto simp:image_def)
-            by (rule_tac x = "(Th x, Th th)" in bexI, auto)
-          moreover have "finite \<dots>"
-          proof -
-            from finite_RAG have "finite (RAG s)" .
-            hence "finite ((RAG (wq s))\<^sup>+)"
-              apply (unfold finite_trancl)
-              by (auto simp: s_RAG_def cs_RAG_def wq_def)
-            thus ?thesis by auto
-          qed
-          ultimately show ?thesis by (auto intro:finite_subset)
-        qed
-        thus ?thesis by (simp add:cs_dependants_def)
-      qed
-      thus ?thesis by simp
-    qed
-    from Max_insert [OF this False, of ?l] show ?thesis by auto
-  next
-    case True
-    thus ?thesis by auto
-  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_readys_threads_pre:
-  assumes np: "threads s \<noteq> {}"
-  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
-proof(unfold max_cp_eq)
-  show "Max (cp s ` readys s) = Max ((\<lambda>th. preced th s) ` threads s)"
-  proof -
-    let ?p = "Max ((\<lambda>th. preced th s) ` threads s)" 
-    let ?f = "(\<lambda>th. preced th s)"
-    have "?p \<in> ((\<lambda>th. preced th s) ` threads s)"
-    proof(rule Max_in)
-      from finite_threads show "finite (?f ` threads s)" by simp
-    next
-      from np show "?f ` threads s \<noteq> {}" by simp
-    qed
-    then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
-      by (auto simp:Image_def)
-    from th_chain_to_ready [OF tm_in]
-    have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" .
-    thus ?thesis
-    proof
-      assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ "
-      then obtain th' where th'_in: "th' \<in> readys s" 
-        and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto
-      have "cp s th' = ?f tm"
-      proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
-        from dependants_threads finite_threads
-        show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" 
-          by (auto intro:finite_subset)
-      next
-        fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
-        from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" .
-        moreover have "p \<le> \<dots>"
-        proof(rule Max_ge)
-          from finite_threads
-          show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        next
-          from p_in and th'_in and dependants_threads[of th']
-          show "p \<in> (\<lambda>th. preced th s) ` threads s"
-            by (auto simp:readys_def)
-        qed
-        ultimately show "p \<le> preced tm s" by auto
-      next
-        show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
-        proof -
-          from tm_chain
-          have "tm \<in> dependants (wq s) th'"
-            by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto)
-          thus ?thesis by auto
-        qed
-      qed
-      with tm_max
-      have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
-      show ?thesis
-      proof (fold h, rule Max_eqI)
-        fix q 
-        assume "q \<in> cp s ` readys s"
-        then obtain th1 where th1_in: "th1 \<in> readys s"
-          and eq_q: "q = cp s th1" by auto
-        show "q \<le> cp s th'"
-          apply (unfold h eq_q)
-          apply (unfold cp_eq_cpreced cpreced_def)
-          apply (rule Max_mono)
-        proof -
-          from dependants_threads [of th1] th1_in
-          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<subseteq> 
-                 (\<lambda>th. preced th s) ` threads s"
-            by (auto simp:readys_def)
-        next
-          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}" by simp
-        next
-          from finite_threads 
-          show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        qed
-      next
-        from finite_threads
-        show "finite (cp s ` readys s)" by (auto simp:readys_def)
-      next
-        from th'_in
-        show "cp s th' \<in> cp s ` readys s" by simp
-      qed
-    next
-      assume tm_ready: "tm \<in> readys s"
-      show ?thesis
-      proof(fold tm_max)
-        have cp_eq_p: "cp s tm = preced tm s"
-        proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
-          fix y 
-          assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
-          show "y \<le> preced tm s"
-          proof -
-            { fix y'
-              assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependants (wq s) tm)"
-              have "y' \<le> preced tm s"
-              proof(unfold tm_max, rule Max_ge)
-                from hy' dependants_threads[of tm]
-                show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
-              next
-                from finite_threads
-                show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-              qed
-            } with hy show ?thesis by auto
-          qed
-        next
-          from dependants_threads[of tm] finite_threads
-          show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm))"
-            by (auto intro:finite_subset)
-        next
-          show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
-            by simp
-        qed 
-        moreover have "Max (cp s ` readys s) = cp s tm"
-        proof(rule Max_eqI)
-          from tm_ready show "cp s tm \<in> cp s ` readys s" by simp
-        next
-          from finite_threads
-          show "finite (cp s ` readys s)" by (auto simp:readys_def)
-        next
-          fix y assume "y \<in> cp s ` readys s"
-          then obtain th1 where th1_readys: "th1 \<in> readys s"
-            and h: "y = cp s th1" by auto
-          show "y \<le> cp s tm"
-            apply(unfold cp_eq_p h)
-            apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono)
-          proof -
-            from finite_threads
-            show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-          next
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}"
-              by simp
-          next
-            from dependants_threads[of th1] th1_readys
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) 
-                    \<subseteq> (\<lambda>th. preced th s) ` threads s"
-              by (auto simp:readys_def)
-          qed
-        qed
-        ultimately show " Max (cp s ` readys s) = preced tm s" by simp
-      qed 
-    qed
-  qed
-qed
-
-text {* (* ccc *) \noindent
-  Since the current precedence of the threads in ready queue will always be boosted,
-  there must be one inside it has the maximum precedence of the whole system. 
-*}
-lemma max_cp_readys_threads:
-  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
-proof(cases "threads s = {}")
-  case True
-  thus ?thesis 
-    by (auto simp:readys_def)
-next
-  case False
-  show ?thesis by (rule max_cp_readys_threads_pre[OF False])
-qed
-
-end
-
-lemma eq_holding: "holding (wq s) th cs = holding s th cs"
-  apply (unfold s_holding_def cs_holding_def wq_def, simp)
-  done
-
-lemma f_image_eq:
-  assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
-  shows "f ` A = g ` A"
-proof
-  show "f ` A \<subseteq> g ` A"
-    by(rule image_subsetI, auto intro:h)
-next
-  show "g ` A \<subseteq> f ` A"
-   by (rule image_subsetI, auto intro:h[symmetric])
-qed
-
-
-definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
-  where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
-
-
-lemma detached_test:
-  shows "detached s th = (Th th \<notin> Field (RAG s))"
-apply(simp add: detached_def Field_def)
-apply(simp add: s_RAG_def)
-apply(simp add: s_holding_abv s_waiting_abv)
-apply(simp add: Domain_iff Range_iff)
-apply(simp add: wq_def)
-apply(auto)
-done
-
-context valid_trace
-begin
-
-lemma detached_intro:
-  assumes eq_pv: "cntP s th = cntV s th"
-  shows "detached s th"
-proof -
- from cnp_cnv_cncs
-  have eq_cnt: "cntP s th =
-    cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
-  hence cncs_zero: "cntCS s th = 0"
-    by (auto simp:eq_pv split:if_splits)
-  with eq_cnt
-  have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:eq_pv)
-  thus ?thesis
-  proof
-    assume "th \<notin> threads s"
-    with range_in dm_RAG_threads
-    show ?thesis
-      by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def Domain_iff Range_iff)
-  next
-    assume "th \<in> readys s"
-    moreover have "Th th \<notin> Range (RAG s)"
-    proof -
-      from card_0_eq [OF finite_holding] and cncs_zero
-      have "holdents s th = {}"
-        by (simp add:cntCS_def)
-      thus ?thesis
-        apply(auto simp:holdents_test)
-        apply(case_tac a)
-        apply(auto simp:holdents_test s_RAG_def)
-        done
-    qed
-    ultimately show ?thesis
-      by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def readys_def)
-  qed
-qed
-
-lemma detached_elim:
-  assumes dtc: "detached s th"
-  shows "cntP s th = cntV s th"
-proof -
-  from cnp_cnv_cncs
-  have eq_pv: " cntP s th =
-    cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
-  have cncs_z: "cntCS s th = 0"
-  proof -
-    from dtc have "holdents s th = {}"
-      unfolding detached_def holdents_test s_RAG_def
-      by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
-    thus ?thesis by (auto simp:cntCS_def)
-  qed
-  show ?thesis
-  proof(cases "th \<in> threads s")
-    case True
-    with dtc 
-    have "th \<in> readys s"
-      by (unfold readys_def detached_def Field_def Domain_def Range_def, 
-           auto simp:eq_waiting s_RAG_def)
-    with cncs_z and eq_pv show ?thesis by simp
-  next
-    case False
-    with cncs_z and eq_pv show ?thesis by simp
-  qed
-qed
-
-lemma detached_eq:
-  shows "(detached s th) = (cntP s th = cntV s th)"
-  by (insert vt, auto intro:detached_intro detached_elim)
-
-end
-
-text {* 
-  The lemmas in this .thy file are all obvious lemmas, however, they still needs to be derived
-  from the concise and miniature model of PIP given in PrioGDef.thy.
-*}
-
-lemma eq_dependants: "dependants (wq s) = dependants s"
-  by (simp add: s_dependants_abv wq_def)
-
-lemma next_th_unique: 
-  assumes nt1: "next_th s th cs th1"
-  and nt2: "next_th s th cs th2"
-  shows "th1 = th2"
-using assms by (unfold next_th_def, auto)
-
-lemma birth_time_lt:  "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
-  apply (induct s, simp)
-proof -
-  fix a s
-  assume ih: "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
-    and eq_as: "a # s \<noteq> []"
-  show "last_set th (a # s) < length (a # s)"
-  proof(cases "s \<noteq> []")
-    case False
-    from False show ?thesis
-      by (cases a, auto simp:last_set.simps)
-  next
-    case True
-    from ih [OF True] show ?thesis
-      by (cases a, auto simp:last_set.simps)
-  qed
-qed
-
-lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
-  by (induct s, auto simp:threads.simps)
-
-lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
-  apply (drule_tac th_in_ne)
-  by (unfold preced_def, auto intro: birth_time_lt)
-
-lemma inj_the_preced: 
-  "inj_on (the_preced s) (threads s)"
-  by (metis inj_onI preced_unique the_preced_def)
-
-lemma tRAG_alt_def: 
-  "tRAG s = {(Th th1, Th th2) | th1 th2. 
-                  \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
- by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
-
-lemma tRAG_Field:
-  "Field (tRAG s) \<subseteq> Field (RAG s)"
-  by (unfold tRAG_alt_def Field_def, auto)
-
-lemma tRAG_ancestorsE:
-  assumes "x \<in> ancestors (tRAG s) u"
-  obtains th where "x = Th th"
-proof -
-  from assms have "(u, x) \<in> (tRAG s)^+" 
-      by (unfold ancestors_def, auto)
-  from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
-  then obtain th where "x = Th th"
-    by (unfold tRAG_alt_def, auto)
-  from that[OF this] show ?thesis .
-qed
-
-lemma tRAG_mono:
-  assumes "RAG s' \<subseteq> RAG s"
-  shows "tRAG s' \<subseteq> tRAG s"
-  using assms 
-  by (unfold tRAG_alt_def, auto)
-
-lemma holding_next_thI:
-  assumes "holding s th cs"
-  and "length (wq s cs) > 1"
-  obtains th' where "next_th s th cs th'"
-proof -
-  from assms(1)[folded eq_holding, unfolded cs_holding_def]
-  have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" .
-  then obtain rest where h1: "wq s cs = th#rest" 
-    by (cases "wq s cs", auto)
-  with assms(2) have h2: "rest \<noteq> []" by auto
-  let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
-  have "next_th s th cs ?th'" using  h1(1) h2 
-    by (unfold next_th_def, auto)
-  from that[OF this] show ?thesis .
-qed
-
-lemma RAG_tRAG_transfer:
-  assumes "vt s'"
-  assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
-  and "(Cs cs, Th th'') \<in> RAG s'"
-  shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
-proof -
-  interpret vt_s': valid_trace "s'" using assms(1)
-    by (unfold_locales, simp)
-  interpret rtree: rtree "RAG s'"
-  proof
-  show "single_valued (RAG s')"
-  apply (intro_locales)
-    by (unfold single_valued_def, 
-        auto intro:vt_s'.unique_RAG)
-
-  show "acyclic (RAG s')"
-     by (rule vt_s'.acyclic_RAG)
-  qed
-  { fix n1 n2
-    assume "(n1, n2) \<in> ?L"
-    from this[unfolded tRAG_alt_def]
-    obtain th1 th2 cs' where 
-      h: "n1 = Th th1" "n2 = Th th2" 
-         "(Th th1, Cs cs') \<in> RAG s"
-         "(Cs cs', Th th2) \<in> RAG s" by auto
-    from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto
-    from h(3) and assms(2) 
-    have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
-          (Th th1, Cs cs') \<in> RAG s'" by auto
-    hence "(n1, n2) \<in> ?R"
-    proof
-      assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
-      hence eq_th1: "th1 = th" by simp
-      moreover have "th2 = th''"
-      proof -
-        from h1 have "cs' = cs" by simp
-        from assms(3) cs_in[unfolded this] rtree.sgv
-        show ?thesis
-          by (unfold single_valued_def, auto)
-      qed
-      ultimately show ?thesis using h(1,2) by auto
-    next
-      assume "(Th th1, Cs cs') \<in> RAG s'"
-      with cs_in have "(Th th1, Th th2) \<in> tRAG s'"
-        by (unfold tRAG_alt_def, auto)
-      from this[folded h(1, 2)] show ?thesis by auto
-    qed
-  } moreover {
-    fix n1 n2
-    assume "(n1, n2) \<in> ?R"
-    hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto
-    hence "(n1, n2) \<in> ?L" 
-    proof
-      assume "(n1, n2) \<in> tRAG s'"
-      moreover have "... \<subseteq> ?L"
-      proof(rule tRAG_mono)
-        show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto)
-      qed
-      ultimately show ?thesis by auto
-    next
-      assume eq_n: "(n1, n2) = (Th th, Th th'')"
-      from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto
-      moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto
-      ultimately show ?thesis 
-        by (unfold eq_n tRAG_alt_def, auto)
-    qed
-  } ultimately show ?thesis by auto
-qed
-
-context valid_trace
-begin
-
-lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt]
-
-end
-
-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
-
-lemma cp_gen_alt_def:
-  "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
-    by (auto simp:cp_gen_def)
-
-lemma tRAG_nodeE:
-  assumes "(n1, n2) \<in> tRAG s"
-  obtains th1 th2 where "n1 = Th th1" "n2 = Th th2"
-  using assms
-  by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def)
-
-lemma subtree_nodeE:
-  assumes "n \<in> subtree (tRAG s) (Th th)"
-  obtains th1 where "n = Th th1"
-proof -
-  show ?thesis
-  proof(rule subtreeE[OF assms])
-    assume "n = Th th"
-    from that[OF this] show ?thesis .
-  next
-    assume "Th th \<in> ancestors (tRAG s) n"
-    hence "(n, Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def)
-    hence "\<exists> th1. n = Th th1"
-    proof(induct)
-      case (base y)
-      from tRAG_nodeE[OF this] show ?case by metis
-    next
-      case (step y z)
-      thus ?case by auto
-    qed
-    with that show ?thesis by auto
-  qed
-qed
-
-lemma tRAG_star_RAG: "(tRAG s)^* \<subseteq> (RAG s)^*"
-proof -
-  have "(wRAG s O hRAG s)^* \<subseteq> (RAG s O RAG s)^*" 
-    by (rule rtrancl_mono, auto simp:RAG_split)
-  also have "... \<subseteq> ((RAG s)^*)^*"
-    by (rule rtrancl_mono, auto)
-  also have "... = (RAG s)^*" by simp
-  finally show ?thesis by (unfold tRAG_def, simp)
-qed
-
-lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x"
-proof -
-  { fix a
-    assume "a \<in> subtree (tRAG s) x"
-    hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def)
-    with tRAG_star_RAG[of s]
-    have "(a, x) \<in> (RAG s)^*" by auto
-    hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def)
-  } thus ?thesis by auto
-qed
-
-lemma tRAG_trancl_eq:
-   "{th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
-    {th'. (Th th', Th th)  \<in> (RAG s)^+}"
-   (is "?L = ?R")
-proof -
-  { fix th'
-    assume "th' \<in> ?L"
-    hence "(Th th', Th th) \<in> (tRAG s)^+" by auto
-    from tranclD[OF this]
-    obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto
-    from tRAG_subtree_RAG[of s] and this(2)
-    have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) 
-    moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto 
-    ultimately have "th' \<in> ?R"  by auto 
-  } moreover 
-  { fix th'
-    assume "th' \<in> ?R"
-    hence "(Th th', Th th) \<in> (RAG s)^+" by (auto)
-    from plus_rpath[OF this]
-    obtain xs where rp: "rpath (RAG s) (Th th') xs (Th th)" "xs \<noteq> []" by auto
-    hence "(Th th', Th th) \<in> (tRAG s)^+"
-    proof(induct xs arbitrary:th' th rule:length_induct)
-      case (1 xs th' th)
-      then obtain x1 xs1 where Cons1: "xs = x1#xs1" by (cases xs, auto)
-      show ?case
-      proof(cases "xs1")
-        case Nil
-        from 1(2)[unfolded Cons1 Nil]
-        have rp: "rpath (RAG s) (Th th') [x1] (Th th)" .
-        hence "(Th th', x1) \<in> (RAG s)" by (cases, simp)
-        then obtain cs where "x1 = Cs cs" 
-              by (unfold s_RAG_def, auto)
-        from rpath_nnl_lastE[OF rp[unfolded this]]
-        show ?thesis by auto
-      next
-        case (Cons x2 xs2)
-        from 1(2)[unfolded Cons1[unfolded this]]
-        have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" .
-        from rpath_edges_on[OF this]
-        have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" .
-        have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)"
-            by (simp add: edges_on_unfold)
-        with eds have rg1: "(Th th', x1) \<in> RAG s" by auto
-        then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto)
-        have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)"
-            by (simp add: edges_on_unfold)
-        from this eds
-        have rg2: "(x1, x2) \<in> RAG s" by auto
-        from this[unfolded eq_x1] 
-        obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto)
-        from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2]
-        have rt1: "(Th th', Th th1) \<in> tRAG s" by (unfold tRAG_alt_def, auto)
-        from rp have "rpath (RAG s) x2 xs2 (Th th)"
-           by  (elim rpath_ConsE, simp)
-        from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" .
-        show ?thesis
-        proof(cases "xs2 = []")
-          case True
-          from rpath_nilE[OF rp'[unfolded this]]
-          have "th1 = th" by auto
-          from rt1[unfolded this] show ?thesis by auto
-        next
-          case False
-          from 1(1)[rule_format, OF _ rp' this, unfolded Cons1 Cons]
-          have "(Th th1, Th th) \<in> (tRAG s)\<^sup>+" by simp
-          with rt1 show ?thesis by auto
-        qed
-      qed
-    qed
-    hence "th' \<in> ?L" by auto
-  } ultimately show ?thesis by blast
-qed
-
-lemma tRAG_trancl_eq_Th:
-   "{Th th' | th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
-    {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}"
-    using tRAG_trancl_eq by auto
-
-lemma dependants_alt_def:
-  "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
-  by (metis eq_RAG s_dependants_def tRAG_trancl_eq)
-  
-context valid_trace
-begin
-
-lemma count_eq_tRAG_plus:
-  assumes "cntP s th = cntV s th"
-  shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
-  using assms count_eq_dependants dependants_alt_def eq_dependants by auto 
-
-lemma count_eq_RAG_plus:
-  assumes "cntP s th = cntV s th"
-  shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
-  using assms count_eq_dependants cs_dependants_def eq_RAG by auto
-
-lemma count_eq_RAG_plus_Th:
-  assumes "cntP s th = cntV s th"
-  shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
-  using count_eq_RAG_plus[OF assms] by auto
-
-lemma count_eq_tRAG_plus_Th:
-  assumes "cntP s th = cntV s th"
-  shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
-   using count_eq_tRAG_plus[OF assms] by auto
-
-end
-
-lemma tRAG_subtree_eq: 
-   "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
-   (is "?L = ?R")
-proof -
-  { fix n 
-    assume h: "n \<in> ?L"
-    hence "n \<in> ?R"
-    by (smt mem_Collect_eq subsetCE subtree_def subtree_nodeE tRAG_subtree_RAG) 
-  } moreover {
-    fix n
-    assume "n \<in> ?R"
-    then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*"
-      by (auto simp:subtree_def)
-    from rtranclD[OF this(2)]
-    have "n \<in> ?L"
-    proof
-      assume "Th th' \<noteq> Th th \<and> (Th th', Th th) \<in> (RAG s)\<^sup>+"
-      with h have "n \<in> {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}" by auto
-      thus ?thesis using subtree_def tRAG_trancl_eq by fastforce
-    qed (insert h, auto simp:subtree_def)
-  } ultimately show ?thesis by auto
-qed
-
-lemma threads_set_eq: 
-   "the_thread ` (subtree (tRAG s) (Th th)) = 
-                  {th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R")
-   by (auto intro:rev_image_eqI simp:tRAG_subtree_eq)
-
-lemma cp_alt_def1: 
-  "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))"
-proof -
-  have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) =
-       ((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))"
-       by auto
-  thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto)
-qed
-
-lemma cp_gen_def_cond: 
-  assumes "x = Th th"
-  shows "cp s th = cp_gen s (Th th)"
-by (unfold cp_alt_def1 cp_gen_def, simp)
-
-lemma cp_gen_over_set:
-  assumes "\<forall> x \<in> A. \<exists> th. x = Th th"
-  shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A"
-proof(rule f_image_eq)
-  fix a
-  assume "a \<in> A"
-  from assms[rule_format, OF this]
-  obtain th where eq_a: "a = Th th" by auto
-  show "cp_gen s a = (cp s \<circ> the_thread) a"
-    by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
-qed
-
-
-context valid_trace
-begin
-
-lemma RAG_threads:
-  assumes "(Th th) \<in> Field (RAG s)"
-  shows "th \<in> threads s"
-  using assms
-  by (metis Field_def UnE dm_RAG_threads range_in vt)
-
-lemma subtree_tRAG_thread:
-  assumes "th \<in> threads s"
-  shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R")
-proof -
-  have "?L = {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
-    by (unfold tRAG_subtree_eq, simp)
-  also have "... \<subseteq> ?R"
-  proof
-    fix x
-    assume "x \<in> {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
-    then obtain th' where h: "x = Th th'" "Th th' \<in> subtree (RAG s) (Th th)" by auto
-    from this(2)
-    show "x \<in> ?R"
-    proof(cases rule:subtreeE)
-      case 1
-      thus ?thesis by (simp add: assms h(1)) 
-    next
-      case 2
-      thus ?thesis by (metis ancestors_Field dm_RAG_threads h(1) image_eqI) 
-    qed
-  qed
-  finally show ?thesis .
-qed
-
-lemma readys_root:
-  assumes "th \<in> readys s"
-  shows "root (RAG s) (Th th)"
-proof -
-  { fix x
-    assume "x \<in> ancestors (RAG s) (Th th)"
-    hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def)
-    from tranclD[OF this]
-    obtain z where "(Th th, z) \<in> RAG s" by auto
-    with assms(1) have False
-         apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def)
-         by (fold wq_def, blast)
-  } thus ?thesis by (unfold root_def, auto)
-qed
-
-lemma readys_in_no_subtree:
-  assumes "th \<in> readys s"
-  and "th' \<noteq> th"
-  shows "Th th \<notin> subtree (RAG s) (Th th')" 
-proof
-   assume "Th th \<in> subtree (RAG s) (Th th')"
-   thus False
-   proof(cases rule:subtreeE)
-      case 1
-      with assms show ?thesis by auto
-   next
-      case 2
-      with readys_root[OF assms(1)]
-      show ?thesis by (auto simp:root_def)
-   qed
-qed
-
-lemma not_in_thread_isolated:
-  assumes "th \<notin> threads s"
-  shows "(Th th) \<notin> Field (RAG s)"
-proof
-  assume "(Th th) \<in> Field (RAG s)"
-  with dm_RAG_threads and range_in assms
-  show False by (unfold Field_def, blast)
-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 holding_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 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
-
-sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
-proof -
-  have "fsubtree (tRAG s)"
-  proof -
-    have "fbranch (tRAG s)"
-    proof(unfold tRAG_def, rule fbranch_compose)
-        show "fbranch (wRAG s)"
-        proof(rule finite_fbranchI)
-           from finite_RAG show "finite (wRAG s)"
-           by (unfold RAG_split, auto)
-        qed
-    next
-        show "fbranch (hRAG s)"
-        proof(rule finite_fbranchI)
-           from finite_RAG 
-           show "finite (hRAG s)" by (unfold RAG_split, auto)
-        qed
-    qed
-    moreover have "wf (tRAG s)"
-    proof(rule wf_subset)
-      show "wf (RAG s O RAG s)" using wf_RAG
-        by (fold wf_comp_self, simp)
-    next
-      show "tRAG s \<subseteq> (RAG s O RAG s)"
-        by (unfold tRAG_alt_def, auto)
-    qed
-    ultimately show ?thesis
-      by (unfold fsubtree_def fsubtree_axioms_def,auto)
-  qed
-  from this[folded tRAG_def] show "fsubtree (tRAG s)" .
-qed
-
-lemma Max_UNION: 
-  assumes "finite A"
-  and "A \<noteq> {}"
-  and "\<forall> M \<in> f ` A. finite M"
-  and "\<forall> M \<in> f ` A. M \<noteq> {}"
-  shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R")
-  using assms[simp]
-proof -
-  have "?L = Max (\<Union>(f ` A))"
-    by (fold Union_image_eq, simp)
-  also have "... = ?R"
-    by (subst Max_Union, simp+)
-  finally show ?thesis .
-qed
-
-lemma max_Max_eq:
-  assumes "finite A"
-    and "A \<noteq> {}"
-    and "x = y"
-  shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R")
-proof -
-  have "?R = Max (insert y A)" by simp
-  also from assms have "... = ?L"
-      by (subst Max.insert, simp+)
-  finally show ?thesis by simp
-qed
-
-context valid_trace
-begin
-
-(* ddd *)
-lemma cp_gen_rec:
-  assumes "x = Th th"
-  shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
-proof(cases "children (tRAG s) x = {}")
-  case True
-  show ?thesis
-    by (unfold True cp_gen_def subtree_children, simp add:assms)
-next
-  case False
-  hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
-  note fsbttRAGs.finite_subtree[simp]
-  have [simp]: "finite (children (tRAG s) x)"
-     by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], 
-            rule children_subtree)
-  { fix r x
-    have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
-  } note this[simp]
-  have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
-  proof -
-    from False obtain q where "q \<in> children (tRAG s) x" by blast
-    moreover have "subtree (tRAG s) q \<noteq> {}" by simp
-    ultimately show ?thesis by blast
-  qed
-  have h: "Max ((the_preced s \<circ> the_thread) `
-                ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
-        Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
-                     (is "?L = ?R")
-  proof -
-    let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L
-    let "Max (_ \<union> (?h ` ?B))" = ?R
-    let ?L1 = "?f ` \<Union>(?g ` ?B)"
-    have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)"
-    proof -
-      have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp
-      also have "... =  (\<Union> x \<in> ?B. ?f ` (?g x))" by auto
-      finally have "Max ?L1 = Max ..." by simp
-      also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)"
-        by (subst Max_UNION, simp+)
-      also have "... = Max (cp_gen s ` children (tRAG s) x)"
-          by (unfold image_comp cp_gen_alt_def, simp)
-      finally show ?thesis .
-    qed
-    show ?thesis
-    proof -
-      have "?L = Max (?f ` ?A \<union> ?L1)" by simp
-      also have "... = max (the_preced s (the_thread x)) (Max ?L1)"
-            by (subst Max_Un, simp+)
-      also have "... = max (?f x) (Max (?h ` ?B))"
-        by (unfold eq_Max_L1, simp)
-      also have "... =?R"
-        by (rule max_Max_eq, (simp)+, unfold assms, simp)
-      finally show ?thesis .
-    qed
-  qed  thus ?thesis 
-          by (fold h subtree_children, unfold cp_gen_def, simp) 
-qed
-
-lemma cp_rec:
-  "cp s th = Max ({the_preced s th} \<union> 
-                     (cp s o the_thread) ` children (tRAG s) (Th th))"
-proof -
-  have "Th th = Th th" by simp
-  note h =  cp_gen_def_cond[OF this] cp_gen_rec[OF this]
-  show ?thesis 
-  proof -
-    have "cp_gen s ` children (tRAG s) (Th th) = 
-                (cp s \<circ> the_thread) ` children (tRAG s) (Th th)"
-    proof(rule cp_gen_over_set)
-      show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th"
-        by (unfold tRAG_alt_def, auto simp:children_def)
-    qed
-    thus ?thesis by (subst (1) h(1), unfold h(2), simp)
-  qed
-qed
-
-end
-
-(* keep *)
-lemma next_th_holding:
-  assumes vt: "vt s"
-  and nxt: "next_th s th cs th'"
-  shows "holding (wq s) th cs"
-proof -
-  from nxt[unfolded next_th_def]
-  obtain rest where h: "wq s cs = th # rest"
-                       "rest \<noteq> []" 
-                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
-  thus ?thesis
-    by (unfold cs_holding_def, auto)
-qed
-
-context valid_trace
-begin
-
-lemma next_th_waiting:
-  assumes nxt: "next_th s th cs th'"
-  shows "waiting (wq s) th' cs"
-proof -
-  from nxt[unfolded next_th_def]
-  obtain rest where h: "wq s cs = th # rest"
-                       "rest \<noteq> []" 
-                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
-  from wq_distinct[of cs, unfolded h]
-  have dst: "distinct (th # rest)" .
-  have in_rest: "th' \<in> set rest"
-  proof(unfold h, rule someI2)
-    show "distinct rest \<and> set rest = set rest" using dst by auto
-  next
-    fix x assume "distinct x \<and> set x = set rest"
-    with h(2)
-    show "hd x \<in> set (rest)" by (cases x, auto)
-  qed
-  hence "th' \<in> set (wq s cs)" by (unfold h(1), auto)
-  moreover have "th' \<noteq> hd (wq s cs)"
-    by (unfold h(1), insert in_rest dst, auto)
-  ultimately show ?thesis by (auto simp:cs_waiting_def)
-qed
-
-lemma next_th_RAG:
-  assumes nxt: "next_th (s::event list) th cs th'"
-  shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s"
-  using vt assms next_th_holding next_th_waiting
-  by (unfold s_RAG_def, simp)
-
-end
-
--- {* A useless definition *}
-definition cps:: "state \<Rightarrow> (thread \<times> precedence) set"
-where "cps s = {(th, cp s th) | th . th \<in> threads s}"
-
-
-find_theorems "waiting" holding
-context valid_trace
-begin
-
-find_theorems "waiting" holding
-
-end
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/.hgignore	Fri Jan 29 09:46:07 2016 +0800
@@ -0,0 +1,6 @@
+syntax: glob
+*.*~
+log
+*.patch
+*.hg
+*.*.orig
--- a/Correctness.thy	Thu Jan 28 21:14:17 2016 +0800
+++ b/Correctness.thy	Fri Jan 29 09:46:07 2016 +0800
@@ -694,7 +694,7 @@
   characterisationof the blocking thread @{text "th'"}:
 
 *}
-
+                  
 lemma runing_inversion: (* ddd, one of the main lemmas to present *)
   assumes runing': "th' \<in> runing (t@s)"
   and neq_th: "th' \<noteq> th"
--- a/CpsG - 副本.thy~	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,4410 +0,0 @@
-theory CpsG
-imports PIPDefs 
-begin
-
-lemma Max_f_mono:
-  assumes seq: "A \<subseteq> B"
-  and np: "A \<noteq> {}"
-  and fnt: "finite B"
-  shows "Max (f ` A) \<le> Max (f ` B)"
-proof(rule Max_mono)
-  from seq show "f ` A \<subseteq> f ` B" by auto
-next
-  from np show "f ` A \<noteq> {}" by auto
-next
-  from fnt and seq show "finite (f ` B)" by auto
-qed
-
-(* I am going to use this file as a start point to retrofiting 
-   PIPBasics.thy, which is originally called CpsG.ghy *)
-
-locale valid_trace = 
-  fixes s
-  assumes vt : "vt s"
-
-locale valid_trace_e = valid_trace +
-  fixes e
-  assumes vt_e: "vt (e#s)"
-begin
-
-lemma pip_e: "PIP s e"
-  using vt_e by (cases, simp)  
-
-end
-
-locale valid_trace_create = valid_trace_e + 
-  fixes th prio
-  assumes is_create: "e = Create th prio"
-
-locale valid_trace_exit = valid_trace_e + 
-  fixes th
-  assumes is_exit: "e = Exit th"
-
-locale valid_trace_p = valid_trace_e + 
-  fixes th cs
-  assumes is_p: "e = P th cs"
-
-locale valid_trace_v = valid_trace_e + 
-  fixes th cs
-  assumes is_v: "e = V th cs"
-begin
-  definition "rest = tl (wq s cs)"
-  definition "wq' = (SOME q. distinct q \<and> set q = set rest)"
-end
-
-locale valid_trace_v_n = valid_trace_v +
-  assumes rest_nnl: "rest \<noteq> []"
-
-locale valid_trace_v_e = valid_trace_v +
-  assumes rest_nil: "rest = []"
-
-locale valid_trace_set= valid_trace_e + 
-  fixes th prio
-  assumes is_set: "e = Set th prio"
-
-context valid_trace
-begin
-
-lemma ind [consumes 0, case_names Nil Cons, induct type]:
-  assumes "PP []"
-     and "(\<And>s e. valid_trace_e s e \<Longrightarrow>
-                   PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))"
-     shows "PP s"
-proof(induct rule:vt.induct[OF vt, case_names Init Step])
-  case Init
-  from assms(1) show ?case .
-next
-  case (Step s e)
-  show ?case
-  proof(rule assms(2))
-    show "valid_trace_e s e" using Step by (unfold_locales, auto)
-  next
-    show "PP s" using Step by simp
-  next
-    show "PIP s e" using Step by simp
-  qed
-qed
-
-end
-
-
-lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
-  by  (unfold s_waiting_def cs_waiting_def wq_def, auto)
-
-lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
-  by (unfold s_holding_def wq_def cs_holding_def, simp)
-
-lemma runing_ready: 
-  shows "runing s \<subseteq> readys s"
-  unfolding runing_def readys_def
-  by auto 
-
-lemma readys_threads:
-  shows "readys s \<subseteq> threads s"
-  unfolding readys_def
-  by auto
-
-lemma wq_v_neq [simp]:
-   "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
-  by (auto simp:wq_def Let_def cp_def split:list.splits)
-
-lemma runing_head:
-  assumes "th \<in> runing s"
-  and "th \<in> set (wq_fun (schs s) cs)"
-  shows "th = hd (wq_fun (schs s) cs)"
-  using assms
-  by (simp add:runing_def readys_def s_waiting_def wq_def)
-
-context valid_trace
-begin
-
-lemma runing_wqE:
-  assumes "th \<in> runing s"
-  and "th \<in> set (wq s cs)"
-  obtains rest where "wq s cs = th#rest"
-proof -
-  from assms(2) obtain th' rest where eq_wq: "wq s cs = th'#rest"
-    by (meson list.set_cases)
-  have "th' = th"
-  proof(rule ccontr)
-    assume "th' \<noteq> th"
-    hence "th \<noteq> hd (wq s cs)" using eq_wq by auto 
-    with assms(2)
-    have "waiting s th cs" 
-      by (unfold s_waiting_def, fold wq_def, auto)
-    with assms show False 
-      by (unfold runing_def readys_def, auto)
-  qed
-  with eq_wq that show ?thesis by metis
-qed
-
-end
-
-context valid_trace_p
-begin
-
-lemma wq_neq_simp [simp]:
-  assumes "cs' \<noteq> cs"
-  shows "wq (e#s) cs' = wq s cs'"
-    using assms unfolding is_p wq_def
-  by (auto simp:Let_def)
-
-lemma runing_th_s:
-  shows "th \<in> runing s"
-proof -
-  from pip_e[unfolded is_p]
-  show ?thesis by (cases, simp)
-qed
-
-lemma th_not_waiting: 
-  "\<not> waiting s th c"
-proof -
-  have "th \<in> readys s"
-    using runing_ready runing_th_s by blast 
-  thus ?thesis
-    by (unfold readys_def, auto)
-qed
-
-lemma waiting_neq_th: 
-  assumes "waiting s t c"
-  shows "t \<noteq> th"
-  using assms using th_not_waiting by blast 
-
-lemma th_not_in_wq: 
-  shows "th \<notin> set (wq s cs)"
-proof
-  assume otherwise: "th \<in> set (wq s cs)"
-  from runing_wqE[OF runing_th_s this]
-  obtain rest where eq_wq: "wq s cs = th#rest" by blast
-  with otherwise
-  have "holding s th cs"
-    by (unfold s_holding_def, fold wq_def, simp)
-  hence cs_th_RAG: "(Cs cs, Th th) \<in> RAG s"
-    by (unfold s_RAG_def, fold holding_eq, auto)
-  from pip_e[unfolded is_p]
-  show False
-  proof(cases)
-    case (thread_P)
-    with cs_th_RAG show ?thesis by auto
-  qed
-qed
-
-lemma wq_es_cs: 
-  "wq (e#s) cs =  wq s cs @ [th]"
-  by (unfold is_p wq_def, auto simp:Let_def)
-
-lemma wq_distinct_kept:
-  assumes "distinct (wq s cs')"
-  shows "distinct (wq (e#s) cs')"
-proof(cases "cs' = cs")
-  case True
-  show ?thesis using True assms th_not_in_wq
-    by (unfold True wq_es_cs, auto)
-qed (insert assms, simp)
-
-end
-
-
-context valid_trace_v
-begin
-
-lemma wq_neq_simp [simp]:
-  assumes "cs' \<noteq> cs"
-  shows "wq (e#s) cs' = wq s cs'"
-    using assms unfolding is_v wq_def
-  by (auto simp:Let_def)
-
-lemma runing_th_s:
-  shows "th \<in> runing s"
-proof -
-  from pip_e[unfolded is_v]
-  show ?thesis by (cases, simp)
-qed
-
-lemma th_not_waiting: 
-  "\<not> waiting s th c"
-proof -
-  have "th \<in> readys s"
-    using runing_ready runing_th_s by blast 
-  thus ?thesis
-    by (unfold readys_def, auto)
-qed
-
-lemma waiting_neq_th: 
-  assumes "waiting s t c"
-  shows "t \<noteq> th"
-  using assms using th_not_waiting by blast 
-
-lemma wq_s_cs:
-  "wq s cs = th#rest"
-proof -
-  from pip_e[unfolded is_v]
-  show ?thesis
-  proof(cases)
-    case (thread_V)
-    from this(2) show ?thesis
-      by (unfold rest_def s_holding_def, fold wq_def,
-                 metis empty_iff list.collapse list.set(1))
-  qed
-qed
-
-lemma wq_es_cs:
-  "wq (e#s) cs = wq'"
- using wq_s_cs[unfolded wq_def]
- by (auto simp:Let_def wq_def rest_def wq'_def is_v, simp) 
-
-lemma wq_distinct_kept:
-  assumes "distinct (wq s cs')"
-  shows "distinct (wq (e#s) cs')"
-proof(cases "cs' = cs")
-  case True
-  show ?thesis
-  proof(unfold True wq_es_cs wq'_def, rule someI2)
-    show "distinct rest \<and> set rest = set rest"
-        using assms[unfolded True wq_s_cs] by auto
-  qed simp
-qed (insert assms, simp)
-
-end
-
-context valid_trace
-begin
-
-lemma actor_inv: 
-  assumes "PIP s e"
-  and "\<not> isCreate e"
-  shows "actor e \<in> runing s"
-  using assms
-  by (induct, auto)
-
-lemma isP_E:
-  assumes "isP e"
-  obtains cs where "e = P (actor e) cs"
-  using assms by (cases e, auto)
-
-lemma isV_E:
-  assumes "isV e"
-  obtains cs where "e = V (actor e) cs"
-  using assms by (cases e, auto) 
-
-lemma wq_distinct: "distinct (wq s cs)"
-proof(induct rule:ind)
-  case (Cons s e)
-  interpret vt_e: valid_trace_e s e using Cons by simp
-  show ?case 
-  proof(cases e)
-    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.wq_distinct_kept) 
-  qed
-qed (unfold wq_def Let_def, simp)
-
-end
-
-context valid_trace_e
-begin
-
-text {*
-  The following lemma shows that only the @{text "P"}
-  operation can add new thread into waiting queues. 
-  Such kind of lemmas are very obvious, but need to be checked formally.
-  This is a kind of confirmation that our modelling is correct.
-*}
-
-lemma wq_in_inv: 
-  assumes s_ni: "thread \<notin> set (wq s cs)"
-  and s_i: "thread \<in> set (wq (e#s) cs)"
-  shows "e = P thread cs"
-proof(cases e)
-  -- {* This is the only non-trivial case: *}
-  case (V th cs1)
-  have False
-  proof(cases "cs1 = cs")
-    case True
-    show ?thesis
-    proof(cases "(wq s cs1)")
-      case (Cons w_hd w_tl)
-      have "set (wq (e#s) cs) \<subseteq> set (wq s cs)"
-      proof -
-        have "(wq (e#s) cs) = (SOME q. distinct q \<and> set q = set w_tl)"
-          using  Cons V by (auto simp:wq_def Let_def True split:if_splits)
-        moreover have "set ... \<subseteq> set (wq s cs)"
-        proof(rule someI2)
-          show "distinct w_tl \<and> set w_tl = set w_tl"
-            by (metis distinct.simps(2) local.Cons wq_distinct)
-        qed (insert Cons True, auto)
-        ultimately show ?thesis by simp
-      qed
-      with assms show ?thesis by auto
-    qed (insert assms V True, auto simp:wq_def Let_def split:if_splits)
-  qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
-  thus ?thesis by auto
-qed (insert assms, auto simp:wq_def Let_def split:if_splits)
-
-lemma wq_out_inv: 
-  assumes s_in: "thread \<in> set (wq s cs)"
-  and s_hd: "thread = hd (wq s cs)"
-  and s_i: "thread \<noteq> hd (wq (e#s) cs)"
-  shows "e = V thread cs"
-proof(cases e)
--- {* There are only two non-trivial cases: *}
-  case (V th cs1)
-  show ?thesis
-  proof(cases "cs1 = cs")
-    case True
-    have "PIP s (V th cs)" using pip_e[unfolded V[unfolded True]] .
-    thus ?thesis
-    proof(cases)
-      case (thread_V)
-      moreover have "th = thread" using thread_V(2) s_hd
-          by (unfold s_holding_def wq_def, simp)
-      ultimately show ?thesis using V True by simp
-    qed
-  qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
-next
-  case (P th cs1)
-  show ?thesis
-  proof(cases "cs1 = cs")
-    case True
-    with P have "wq (e#s) cs = wq_fun (schs s) cs @ [th]"
-      by (auto simp:wq_def Let_def split:if_splits)
-    with s_i s_hd s_in have False
-      by (metis empty_iff hd_append2 list.set(1) wq_def) 
-    thus ?thesis by simp
-  qed (insert assms P, auto simp:wq_def Let_def split:if_splits)
-qed (insert assms, auto simp:wq_def Let_def split:if_splits)
-
-end
-
-
-context valid_trace
-begin
-
-end
-
-
-
-context valid_trace
-begin
-
-
-text {* (* ddd *)
-  The nature of the work is like this: since it starts from a very simple and basic 
-  model, even intuitively very `basic` and `obvious` properties need to derived from scratch.
-  For instance, the fact 
-  that one thread can not be blocked by two critical resources at the same time
-  is obvious, because only running threads can make new requests, if one is waiting for 
-  a critical resource and get blocked, it can not make another resource request and get 
-  blocked the second time (because it is not running). 
-
-  To derive this fact, one needs to prove by contraction and 
-  reason about time (or @{text "moement"}). The reasoning is based on a generic theorem
-  named @{text "p_split"}, which is about status changing along the time axis. It says if 
-  a condition @{text "Q"} is @{text "True"} at a state @{text "s"},
-  but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} 
-  in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history 
-  of events leading to it), such that @{text "Q"} switched 
-  from being @{text "False"} to @{text "True"} and kept being @{text "True"}
-  till the last moment of @{text "s"}.
-
-  Suppose a thread @{text "th"} is blocked
-  on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, 
-  since no thread is blocked at the very beginning, by applying 
-  @{text "p_split"} to these two blocking facts, there exist 
-  two moments @{text "t1"} and @{text "t2"}  in @{text "s"}, such that 
-  @{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} 
-  and kept on blocked on them respectively ever since.
- 
-  Without lost of generality, we assume @{text "t1"} is earlier than @{text "t2"}.
-  However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
-  in blocked state at moment @{text "t2"} and could not
-  make any request and get blocked the second time: Contradiction.
-*}
-
-lemma waiting_unique_pre: (* ddd *)
-  assumes h11: "thread \<in> set (wq s cs1)"
-  and h12: "thread \<noteq> hd (wq s cs1)"
-  assumes h21: "thread \<in> set (wq s cs2)"
-  and h22: "thread \<noteq> hd (wq s cs2)"
-  and neq12: "cs1 \<noteq> cs2"
-  shows "False"
-proof -
-  let "?Q" = "\<lambda> cs s. thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
-  from h11 and h12 have q1: "?Q cs1 s" by simp
-  from h21 and h22 have q2: "?Q cs2 s" by simp
-  have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
-  have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
-  from p_split [of "?Q cs1", OF q1 nq1]
-  obtain t1 where lt1: "t1 < length s"
-    and np1: "\<not> ?Q cs1 (moment t1 s)"
-    and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))" by auto
-  from p_split [of "?Q cs2", OF q2 nq2]
-  obtain t2 where lt2: "t2 < length s"
-    and np2: "\<not> ?Q cs2 (moment t2 s)"
-    and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))" by auto
-  { fix s cs
-    assume q: "?Q cs s"
-    have "thread \<notin> runing s"
-    proof
-      assume "thread \<in> runing s"
-      hence " \<forall>cs. \<not> (thread \<in> set (wq_fun (schs s) cs) \<and> 
-                 thread \<noteq> hd (wq_fun (schs s) cs))"
-        by (unfold runing_def s_waiting_def readys_def, auto)
-      from this[rule_format, of cs] q 
-      show False by (simp add: wq_def) 
-    qed
-  } note q_not_runing = this
-  { fix t1 t2 cs1 cs2
-    assume  lt1: "t1 < length s"
-    and np1: "\<not> ?Q cs1 (moment t1 s)"
-    and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))"
-    and lt2: "t2 < length s"
-    and np2: "\<not> ?Q cs2 (moment t2 s)"
-    and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))"
-    and lt12: "t1 < t2"
-    let ?t3 = "Suc t2"
-    from lt2 have le_t3: "?t3 \<le> length s" by auto
-    from moment_plus [OF this] 
-    obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
-    have "t2 < ?t3" by simp
-    from nn2 [rule_format, OF this] and eq_m
-    have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
-         h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
-    have "vt (e#moment t2 s)"
-    proof -
-      from vt_moment 
-      have "vt (moment ?t3 s)" .
-      with eq_m show ?thesis by simp
-    qed
-    then interpret vt_e: valid_trace_e "moment t2 s" "e"
-        by (unfold_locales, auto, cases, simp)
-    have ?thesis
-    proof -
-      have "thread \<in> runing (moment t2 s)"
-      proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
-        case True
-        have "e = V thread cs2"
-        proof -
-          have eq_th: "thread = hd (wq (moment t2 s) cs2)" 
-              using True and np2  by auto 
-          from vt_e.wq_out_inv[OF True this h2]
-          show ?thesis .
-        qed
-        thus ?thesis using vt_e.actor_inv[OF vt_e.pip_e] by auto
-      next
-        case False
-        have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] .
-        with vt_e.actor_inv[OF vt_e.pip_e]
-        show ?thesis by auto
-      qed
-      moreover have "thread \<notin> runing (moment t2 s)"
-        by (rule q_not_runing[OF nn1[rule_format, OF lt12]])
-      ultimately show ?thesis by simp
-    qed
-  } note lt_case = this
-  show ?thesis
-  proof -
-    { assume "t1 < t2"
-      from lt_case[OF lt1 np1 nn1 lt2 np2 nn2 this]
-      have ?thesis .
-    } moreover {
-      assume "t2 < t1"
-      from lt_case[OF lt2 np2 nn2 lt1 np1 nn1 this]
-      have ?thesis .
-    } moreover {
-      assume eq_12: "t1 = t2"
-      let ?t3 = "Suc t2"
-      from lt2 have le_t3: "?t3 \<le> length s" by auto
-      from moment_plus [OF this] 
-      obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
-      have lt_2: "t2 < ?t3" by simp
-      from nn2 [rule_format, OF this] and eq_m
-      have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
-           h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
-      from nn1[rule_format, OF lt_2[folded eq_12]] eq_m[folded eq_12]
-      have g1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
-           g2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
-      have "vt (e#moment t2 s)"
-      proof -
-        from vt_moment 
-        have "vt (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      then interpret vt_e: valid_trace_e "moment t2 s" "e"
-          by (unfold_locales, auto, cases, simp)
-      have "e = V thread cs2 \<or> e = P thread cs2"
-      proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
-        case True
-        have "e = V thread cs2"
-        proof -
-          have eq_th: "thread = hd (wq (moment t2 s) cs2)" 
-              using True and np2  by auto 
-          from vt_e.wq_out_inv[OF True this h2]
-          show ?thesis .
-        qed
-        thus ?thesis by auto
-      next
-        case False
-        have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] .
-        thus ?thesis by auto
-      qed
-      moreover have "e = V thread cs1 \<or> e = P thread cs1"
-      proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
-        case True
-        have eq_th: "thread = hd (wq (moment t1 s) cs1)" 
-              using True and np1  by auto 
-        from vt_e.wq_out_inv[folded eq_12, OF True this g2]
-        have "e = V thread cs1" .
-        thus ?thesis by auto
-      next
-        case False
-        have "e = P thread cs1" using vt_e.wq_in_inv[folded eq_12, OF False g1] .
-        thus ?thesis by auto
-      qed
-      ultimately have ?thesis using neq12 by auto
-    } ultimately show ?thesis using nat_neq_iff by blast 
-  qed
-qed
-
-text {*
-  This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
-*}
-
-lemma waiting_unique:
-  assumes "waiting s th cs1"
-  and "waiting s th cs2"
-  shows "cs1 = cs2"
-  using waiting_unique_pre assms
-  unfolding wq_def s_waiting_def
-  by auto
-
-end
-
-(* not used *)
-text {*
-  Every thread can only be blocked on one critical resource, 
-  symmetrically, every critical resource can only be held by one thread. 
-  This fact is much more easier according to our definition. 
-*}
-lemma held_unique:
-  assumes "holding (s::event list) th1 cs"
-  and "holding s th2 cs"
-  shows "th1 = th2"
- by (insert assms, unfold s_holding_def, auto)
-
-
-lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s"
-  apply (induct s, auto)
-  by (case_tac a, auto split:if_splits)
-
-lemma last_set_unique: 
-  "\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk>
-          \<Longrightarrow> th1 = th2"
-  apply (induct s, auto)
-  by (case_tac a, auto split:if_splits dest:last_set_lt)
-
-lemma preced_unique : 
-  assumes pcd_eq: "preced th1 s = preced th2 s"
-  and th_in1: "th1 \<in> threads s"
-  and th_in2: " th2 \<in> threads s"
-  shows "th1 = th2"
-proof -
-  from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def)
-  from last_set_unique [OF this th_in1 th_in2]
-  show ?thesis .
-qed
-                      
-lemma preced_linorder: 
-  assumes neq_12: "th1 \<noteq> th2"
-  and th_in1: "th1 \<in> threads s"
-  and th_in2: " th2 \<in> threads s"
-  shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
-proof -
-  from preced_unique [OF _ th_in1 th_in2] and neq_12 
-  have "preced th1 s \<noteq> preced th2 s" by auto
-  thus ?thesis by auto
-qed
-
-(* An aux lemma used later *) 
-lemma unique_minus:
-  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
-  and xy: "(x, y) \<in> r"
-  and xz: "(x, z) \<in> r^+"
-  and neq: "y \<noteq> z"
-  shows "(y, z) \<in> r^+"
-proof -
- from xz and neq show ?thesis
- proof(induct)
-   case (base ya)
-   have "(x, ya) \<in> r" by fact
-   from unique [OF xy this] have "y = ya" .
-   with base show ?case by auto
- next
-   case (step ya z)
-   show ?case
-   proof(cases "y = ya")
-     case True
-     from step True show ?thesis by simp
-   next
-     case False
-     from step False
-     show ?thesis by auto
-   qed
- qed
-qed
-
-lemma unique_base:
-  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
-  and xy: "(x, y) \<in> r"
-  and xz: "(x, z) \<in> r^+"
-  and neq_yz: "y \<noteq> z"
-  shows "(y, z) \<in> r^+"
-proof -
-  from xz neq_yz show ?thesis
-  proof(induct)
-    case (base ya)
-    from xy unique base show ?case by auto
-  next
-    case (step ya z)
-    show ?case
-    proof(cases "y = ya")
-      case True
-      from True step show ?thesis by auto
-    next
-      case False
-      from False step 
-      have "(y, ya) \<in> r\<^sup>+" by auto
-      with step show ?thesis by auto
-    qed
-  qed
-qed
-
-lemma unique_chain:
-  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
-  and xy: "(x, y) \<in> r^+"
-  and xz: "(x, z) \<in> r^+"
-  and neq_yz: "y \<noteq> z"
-  shows "(y, z) \<in> r^+ \<or> (z, y) \<in> r^+"
-proof -
-  from xy xz neq_yz show ?thesis
-  proof(induct)
-    case (base y)
-    have h1: "(x, y) \<in> r" and h2: "(x, z) \<in> r\<^sup>+" and h3: "y \<noteq> z" using base by auto
-    from unique_base [OF _ h1 h2 h3] and unique show ?case by auto
-  next
-    case (step y za)
-    show ?case
-    proof(cases "y = z")
-      case True
-      from True step show ?thesis by auto
-    next
-      case False
-      from False step have "(y, z) \<in> r\<^sup>+ \<or> (z, y) \<in> r\<^sup>+" by auto
-      thus ?thesis
-      proof
-        assume "(z, y) \<in> r\<^sup>+"
-        with step have "(z, za) \<in> r\<^sup>+" by auto
-        thus ?thesis by auto
-      next
-        assume h: "(y, z) \<in> r\<^sup>+"
-        from step have yza: "(y, za) \<in> r" by simp
-        from step have "za \<noteq> z" by simp
-        from unique_minus [OF _ yza h this] and unique
-        have "(za, z) \<in> r\<^sup>+" by auto
-        thus ?thesis by auto
-      qed
-    qed
-  qed
-qed
-
-text {*
-  The following three lemmas show that @{text "RAG"} does not change
-  by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
-  events, respectively.
-*}
-
-lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
-apply (unfold s_RAG_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-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 RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
-apply (unfold s_RAG_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-
-
-context valid_trace_v
-begin
-
-
-lemma distinct_rest: "distinct rest"
-  by (simp add: distinct_tl rest_def wq_distinct)
-
-definition "wq' = (SOME q. distinct q \<and> set q = set rest)"
-
-lemma runing_th_s:
-  shows "th \<in> runing s"
-proof -
-  from pip_e[unfolded is_v]
-  show ?thesis by (cases, simp)
-qed
-
-lemma holding_cs_eq_th:
-  assumes "holding s t cs"
-  shows "t = th"
-proof -
-  from pip_e[unfolded is_v]
-  show ?thesis
-  proof(cases)
-    case (thread_V)
-    from held_unique[OF this(2) assms]
-    show ?thesis by simp
-  qed
-qed
-
-lemma th_not_waiting: 
-  "\<not> waiting s th c"
-proof -
-  have "th \<in> readys s"
-    using runing_ready runing_th_s by blast 
-  thus ?thesis
-    by (unfold readys_def, auto)
-qed
-
-lemma waiting_neq_th: 
-  assumes "waiting s t c"
-  shows "t \<noteq> th"
-  using assms using th_not_waiting by blast 
-
-lemma wq_s_cs:
-  "wq s cs = th#rest"
-proof -
-  from pip_e[unfolded is_v]
-  show ?thesis
-  proof(cases)
-    case (thread_V)
-    from this(2) show ?thesis
-      by (unfold rest_def s_holding_def, fold wq_def,
-                 metis empty_iff list.collapse list.set(1))
-  qed
-qed
-
-lemma wq_es_cs:
-  "wq (e#s) cs = wq'"
- using wq_s_cs[unfolded wq_def]
- by (auto simp:Let_def wq_def rest_def wq'_def is_v, simp) 
-
-lemma distinct_wq': "distinct wq'"
-  by (metis (mono_tags, lifting) distinct_rest  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) 
-
-lemma neq_t_th: 
-  assumes "waiting (e#s) t c"
-  shows "t \<noteq> th"
-proof
-  assume otherwise: "t = th"
-  show False
-  proof(cases "c = cs")
-    case True
-    have "t \<in> set wq'" 
-     using assms[unfolded True s_waiting_def, folded wq_def, unfolded wq_es_cs]
-     by simp 
-    from th'_in_inv[OF this] have "t \<in> set rest" .
-    with wq_s_cs[folded otherwise] wq_distinct[of cs]
-    show ?thesis by simp
-  next
-    case False
-    have "wq (e#s) c = wq s c" using False
-        by (unfold is_v, simp)
-    hence "waiting s t c" using assms 
-        by (simp add: cs_waiting_def waiting_eq)
-    hence "t \<notin> readys s" by (unfold readys_def, auto)
-    hence "t \<notin> runing s" using runing_ready by auto 
-    with runing_th_s[folded otherwise] show ?thesis by auto
-  qed
-qed
-
-lemma waiting_esI1:
-  assumes "waiting s t c"
-      and "c \<noteq> cs" 
-  shows "waiting (e#s) t c" 
-proof -
-  have "wq (e#s) c = wq s c" 
-    using assms(2) is_v by auto
-  with assms(1) show ?thesis 
-    using cs_waiting_def waiting_eq by auto 
-qed
-
-lemma holding_esI2:
-  assumes "c \<noteq> cs" 
-  and "holding s t c"
-  shows "holding (e#s) t c"
-proof -
-  from assms(1) have "wq (e#s) c = wq s c" using is_v by auto
-  from assms(2)[unfolded s_holding_def, folded wq_def, 
-                folded this, unfolded wq_def, folded s_holding_def]
-  show ?thesis .
-qed
-
-lemma holding_esI1:
-  assumes "holding s t c"
-  and "t \<noteq> th"
-  shows "holding (e#s) t c"
-proof -
-  have "c \<noteq> cs" using assms using holding_cs_eq_th by blast 
-  from holding_esI2[OF this assms(1)]
-  show ?thesis .
-qed
-
-end
-
-context valid_trace_v_n
-begin
-
-lemma neq_wq': "wq' \<noteq> []" 
-proof (unfold wq'_def, rule someI2)
-  show "distinct rest \<and> set rest = set rest"
-    by (simp add: distinct_rest) 
-next
-  fix x
-  assume " distinct x \<and> set x = set rest" 
-  thus "x \<noteq> []" using rest_nnl by auto
-qed 
-
-definition "taker = hd wq'"
-
-definition "rest' = tl wq'"
-
-lemma eq_wq': "wq' = taker # rest'"
-  by (simp add: neq_wq' rest'_def taker_def)
-
-lemma next_th_taker: 
-  shows "next_th s th cs taker"
-  using rest_nnl taker_def wq'_def wq_s_cs 
-  by (auto simp:next_th_def)
-
-lemma taker_unique: 
-  assumes "next_th s th cs taker'"
-  shows "taker' = taker"
-proof -
-  from assms
-  obtain rest' where 
-    h: "wq s cs = th # rest'" 
-       "taker' = hd (SOME q. distinct q \<and> set q = set rest')"
-          by (unfold next_th_def, auto)
-  with wq_s_cs have "rest' = rest" by auto
-  thus ?thesis using h(2) taker_def wq'_def by auto 
-qed
-
-lemma waiting_set_eq:
-  "{(Th th', Cs cs) |th'. next_th s th cs th'} = {(Th taker, Cs cs)}"
-  by (smt all_not_in_conv bot.extremum insertI1 insert_subset 
-      mem_Collect_eq next_th_taker subsetI subset_antisym taker_def taker_unique)
-
-lemma holding_set_eq:
-  "{(Cs cs, Th th') |th'.  next_th s th cs th'} = {(Cs cs, Th taker)}"
-  using next_th_taker taker_def waiting_set_eq 
-  by fastforce
-   
-lemma holding_taker:
-  shows "holding (e#s) taker cs"
-    by (unfold s_holding_def, fold wq_def, unfold wq_es_cs, 
-        auto simp:neq_wq' taker_def)
-
-lemma waiting_esI2:
-  assumes "waiting s t cs"
-      and "t \<noteq> taker"
-  shows "waiting (e#s) t cs" 
-proof -
-  have "t \<in> set wq'" 
-  proof(unfold wq'_def, rule someI2)
-    show "distinct rest \<and> set rest = set rest"
-          by (simp add: distinct_rest)
-  next
-    fix x
-    assume "distinct x \<and> set x = set rest"
-    moreover have "t \<in> set rest"
-        using assms(1) cs_waiting_def waiting_eq wq_s_cs by auto 
-    ultimately show "t \<in> set x" by simp
-  qed
-  moreover have "t \<noteq> hd wq'"
-    using assms(2) taker_def by auto 
-  ultimately show ?thesis
-    by (unfold s_waiting_def, fold wq_def, unfold wq_es_cs, simp)
-qed
-
-lemma waiting_esE:
-  assumes "waiting (e#s) t c" 
-  obtains "c \<noteq> cs" "waiting s t c"
-     |    "c = cs" "t \<noteq> taker" "waiting s t cs" "t \<in> set rest'"
-proof(cases "c = cs")
-  case False
-  hence "wq (e#s) c = wq s c" using is_v by auto
-  with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto 
-  from that(1)[OF False this] show ?thesis .
-next
-  case True
-  from assms[unfolded s_waiting_def True, folded wq_def, unfolded wq_es_cs]
-  have "t \<noteq> hd wq'" "t \<in> set wq'" by auto
-  hence "t \<noteq> taker" by (simp add: taker_def) 
-  moreover hence "t \<noteq> th" using assms neq_t_th by blast 
-  moreover have "t \<in> set rest" by (simp add: `t \<in> set wq'` th'_in_inv) 
-  ultimately have "waiting s t cs"
-    by (metis cs_waiting_def list.distinct(2) list.sel(1) 
-                list.set_sel(2) rest_def waiting_eq wq_s_cs)  
-  show ?thesis using that(2)
-  using True `t \<in> set wq'` `t \<noteq> taker` `waiting s t cs` eq_wq' by auto   
-qed
-
-lemma holding_esI1:
-  assumes "c = cs"
-  and "t = taker"
-  shows "holding (e#s) t c"
-  by (unfold assms, simp add: holding_taker)
-
-lemma holding_esE:
-  assumes "holding (e#s) t c" 
-  obtains "c = cs" "t = taker"
-      | "c \<noteq> cs" "holding s t c"
-proof(cases "c = cs")
-  case True
-  from assms[unfolded True, unfolded s_holding_def, 
-             folded wq_def, unfolded wq_es_cs]
-  have "t = taker" by (simp add: taker_def) 
-  from that(1)[OF True this] show ?thesis .
-next
-  case False
-  hence "wq (e#s) c = wq s c" using is_v by auto
-  from assms[unfolded s_holding_def, folded wq_def, 
-             unfolded this, unfolded wq_def, folded s_holding_def]
-  have "holding s t c"  .
-  from that(2)[OF False this] show ?thesis .
-qed
-
-end 
-
-
-context valid_trace_v_n
-begin
-
-lemma nil_wq': "wq' = []" 
-proof (unfold wq'_def, rule someI2)
-  show "distinct rest \<and> set rest = set rest"
-    by (simp add: distinct_rest) 
-next
-  fix x
-  assume " distinct x \<and> set x = set rest" 
-  thus "x = []" using rest_nil by auto
-qed 
-
-lemma no_taker: 
-  assumes "next_th s th cs taker"
-  shows "False"
-proof -
-  from assms[unfolded next_th_def]
-  obtain rest' where "wq s cs = th # rest'" "rest' \<noteq> []"
-    by auto
-  thus ?thesis using rest_def rest_nil by auto 
-qed
-
-lemma waiting_set_eq:
-  "{(Th th', Cs cs) |th'. next_th s th cs th'} = {}"
-  using no_taker by auto
-
-lemma holding_set_eq:
-  "{(Cs cs, Th th') |th'.  next_th s th cs th'} = {}"
-  using no_taker by auto
-   
-lemma no_holding:
-  assumes "holding (e#s) taker cs"
-  shows False
-proof -
-  from wq_es_cs[unfolded nil_wq']
-  have " wq (e # s) cs = []" .
-  from assms[unfolded s_holding_def, folded wq_def, unfolded this]
-  show ?thesis by auto
-qed
-
-lemma no_waiting:
-  assumes "waiting (e#s) t cs"
-  shows False
-proof -
-  from wq_es_cs[unfolded nil_wq']
-  have " wq (e # s) cs = []" .
-  from assms[unfolded s_waiting_def, folded wq_def, unfolded this]
-  show ?thesis by auto
-qed
-
-lemma waiting_esI2:
-  assumes "waiting s t c"
-  shows "waiting (e#s) t c"
-proof -
-  have "c \<noteq> cs" using assms
-    using cs_waiting_def rest_nil waiting_eq wq_s_cs by auto 
-  from waiting_esI1[OF assms this]
-  show ?thesis .
-qed
-
-lemma waiting_esE:
-  assumes "waiting (e#s) t c" 
-  obtains "c \<noteq> cs" "waiting s t c"
-proof(cases "c = cs")
-  case False
-  hence "wq (e#s) c = wq s c" using is_v by auto
-  with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto 
-  from that(1)[OF False this] show ?thesis .
-next
-  case True
-  from no_waiting[OF assms[unfolded True]]
-  show ?thesis by auto
-qed
-
-lemma holding_esE:
-  assumes "holding (e#s) t c" 
-  obtains "c \<noteq> cs" "holding s t c"
-proof(cases "c = cs")
-  case True
-  from no_holding[OF assms[unfolded True]] 
-  show ?thesis by auto
-next
-  case False
-  hence "wq (e#s) c = wq s c" using is_v by auto
-  from assms[unfolded s_holding_def, folded wq_def, 
-             unfolded this, unfolded wq_def, folded s_holding_def]
-  have "holding s t c"  .
-  from that[OF False this] show ?thesis .
-qed
-
-end (* ccc *)
-
-lemma rel_eqI:
-  assumes "\<And> x y. (x,y) \<in> A \<Longrightarrow> (x,y) \<in> B"
-  and "\<And> x y. (x,y) \<in> B \<Longrightarrow> (x, y) \<in> A"
-  shows "A = B"
-  using assms by auto
-
-lemma in_RAG_E:
-  assumes "(n1, n2) \<in> RAG (s::state)"
-  obtains (waiting) th cs where "n1 = Th th" "n2 = Cs cs" "waiting s th cs"
-      | (holding) th cs where "n1 = Cs cs" "n2 = Th th" "holding s th cs"
-  using assms[unfolded s_RAG_def, folded waiting_eq holding_eq]
-  by auto
-  
-context valid_trace_v
-begin
-
-lemma RAG_es:
-  "RAG (e # s) =
-   RAG s - {(Cs cs, Th th)} -
-     {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-     {(Cs cs, Th th') |th'.  next_th s th cs th'}" (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')
-    show ?thesis
-    proof(cases "rest = []")
-      case False
-      interpret h_n: valid_trace_v_n s e th cs
-        by (unfold_locales, insert False, simp)
-      from waiting(3)
-      show ?thesis
-      proof(cases rule:h_n.waiting_esE)
-        case 1
-        with waiting(1,2)
-        show ?thesis
-        by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
-             fold waiting_eq, auto)
-      next
-        case 2
-        with waiting(1,2)
-        show ?thesis
-         by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
-             fold waiting_eq, auto)
-      qed
-    next
-      case True
-      interpret h_e: valid_trace_v_e s e th cs
-        by (unfold_locales, insert True, simp)
-      from waiting(3)
-      show ?thesis
-      proof(cases rule:h_e.waiting_esE)
-        case 1
-        with waiting(1,2)
-        show ?thesis
-        by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, 
-             fold waiting_eq, auto)
-      qed
-    qed
-  next
-    case (holding th' cs')
-    show ?thesis
-    proof(cases "rest = []")
-      case False
-      interpret h_n: valid_trace_v_n s e th cs
-        by (unfold_locales, insert False, simp)
-      from holding(3)
-      show ?thesis
-      proof(cases rule:h_n.holding_esE)
-        case 1
-        with holding(1,2)
-        show ?thesis
-        by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
-             fold waiting_eq, auto)
-      next
-        case 2
-        with holding(1,2)
-        show ?thesis
-         by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
-             fold holding_eq, auto)
-      qed
-    next
-      case True
-      interpret h_e: valid_trace_v_e s e th cs
-        by (unfold_locales, insert True, simp)
-      from holding(3)
-      show ?thesis
-      proof(cases rule:h_e.holding_esE)
-        case 1
-        with holding(1,2)
-        show ?thesis
-        by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, 
-             fold holding_eq, auto)
-      qed
-    qed
-  qed
-next
-  fix n1 n2
-  assume h: "(n1, n2) \<in> ?R"
-  show "(n1, n2) \<in> ?L"
-  proof(cases "rest = []")
-    case False
-    interpret h_n: valid_trace_v_n s e th cs
-        by (unfold_locales, insert False, simp)
-    from h[unfolded h_n.waiting_set_eq h_n.holding_set_eq]
-    have "((n1, n2) \<in> RAG s \<and> (n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th)
-                            \<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)) \<or> 
-          (n2 = Th h_n.taker \<and> n1 = Cs cs)" 
-      by auto
-   thus ?thesis
-   proof
-      assume "n2 = Th h_n.taker \<and> n1 = Cs cs"
-      with h_n.holding_taker
-      show ?thesis 
-        by (unfold s_RAG_def, fold holding_eq, auto)
-   next
-    assume h: "(n1, n2) \<in> RAG s \<and>
-        (n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th) \<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)"
-    hence "(n1, n2) \<in> RAG s" by simp
-    thus ?thesis
-    proof(cases rule:in_RAG_E)
-      case (waiting th' cs')
-      from h and this(1,2)
-      have "th' \<noteq> h_n.taker \<or> cs' \<noteq> cs" by auto
-      hence "waiting (e#s) th' cs'" 
-      proof
-        assume "cs' \<noteq> cs"
-        from waiting_esI1[OF waiting(3) this] 
-        show ?thesis .
-      next
-        assume neq_th': "th' \<noteq> h_n.taker"
-        show ?thesis
-        proof(cases "cs' = cs")
-          case False
-          from waiting_esI1[OF waiting(3) this] 
-          show ?thesis .
-        next
-          case True
-          from h_n.waiting_esI2[OF waiting(3)[unfolded True] neq_th', folded True]
-          show ?thesis .
-        qed
-      qed
-      thus ?thesis using waiting(1,2)
-        by (unfold s_RAG_def, fold waiting_eq, auto)
-    next
-      case (holding th' cs')
-      from h this(1,2)
-      have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto
-      hence "holding (e#s) th' cs'"
-      proof
-        assume "cs' \<noteq> cs"
-        from holding_esI2[OF this holding(3)] 
-        show ?thesis .
-      next
-        assume "th' \<noteq> th"
-        from holding_esI1[OF holding(3) this]
-        show ?thesis .
-      qed
-      thus ?thesis using holding(1,2)
-        by (unfold s_RAG_def, fold holding_eq, auto)
-    qed
-   qed
- next
-   case True
-   interpret h_e: valid_trace_v_e s e th cs
-        by (unfold_locales, insert True, simp)
-   from h[unfolded h_e.waiting_set_eq h_e.holding_set_eq]
-   have h_s: "(n1, n2) \<in> RAG s" "(n1, n2) \<noteq> (Cs cs, Th th)" 
-      by auto
-   from h_s(1)
-   show ?thesis
-   proof(cases rule:in_RAG_E)
-    case (waiting th' cs')
-    from h_e.waiting_esI2[OF this(3)]
-    show ?thesis using waiting(1,2)
-      by (unfold s_RAG_def, fold waiting_eq, auto)
-   next
-    case (holding th' cs')
-    with h_s(2)
-    have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto
-    thus ?thesis
-    proof
-      assume neq_cs: "cs' \<noteq> cs"
-      from holding_esI2[OF this holding(3)]
-      show ?thesis using holding(1,2)
-        by (unfold s_RAG_def, fold holding_eq, auto)
-    next
-      assume "th' \<noteq> th"
-      from holding_esI1[OF holding(3) this]
-      show ?thesis using holding(1,2)
-        by (unfold s_RAG_def, fold holding_eq, auto)
-    qed
-   qed
- qed
-qed
-
-end
-
-
-
-context valid_trace
-begin
-
-lemma finite_threads:
-  shows "finite (threads s)"
-using vt by (induct) (auto elim: step.cases)
-
-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)
-
-lemma wq_threads: 
-  assumes h: "th \<in> set (wq s cs)"
-  shows "th \<in> threads s"
-
-
-lemma wq_threads: 
-  assumes h: "th \<in> set (wq s cs)"
-  shows "th \<in> threads s"
-proof -
- from vt and h show ?thesis
-  proof(induct arbitrary: th cs)
-    case (vt_cons s e)
-    interpret vt_s: valid_trace s
-      using vt_cons(1) by (unfold_locales, auto)
-    assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s"
-      and stp: "step s e"
-      and vt: "vt s"
-      and h: "th \<in> set (wq (e # s) cs)"
-    show ?case
-    proof(cases e)
-      case (Create th' prio)
-      with ih h show ?thesis
-        by (auto simp:wq_def Let_def)
-    next
-      case (Exit th')
-      with stp ih h show ?thesis
-        apply (auto simp:wq_def Let_def)
-        apply (ind_cases "step s (Exit th')")
-        apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def
-               s_RAG_def s_holding_def cs_holding_def)
-        done
-    next
-      case (V th' cs')
-      show ?thesis
-      proof(cases "cs' = cs")
-        case False
-        with h
-        show ?thesis
-          apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def)
-          by (drule_tac ih, simp)
-      next
-        case True
-        from h
-        show ?thesis
-        proof(unfold V wq_def)
-          assume th_in: "th \<in> set (wq_fun (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l")
-          show "th \<in> threads (V th' cs' # s)"
-          proof(cases "cs = cs'")
-            case False
-            hence "?l = wq_fun (schs s) cs" by (simp add:Let_def)
-            with th_in have " th \<in> set (wq s cs)" 
-              by (fold wq_def, simp)
-            from ih [OF this] show ?thesis by simp
-          next
-            case True
-            show ?thesis
-            proof(cases "wq_fun (schs s) cs'")
-              case Nil
-              with h V show ?thesis
-                apply (auto simp:wq_def Let_def split:if_splits)
-                by (fold wq_def, drule_tac ih, simp)
-            next
-              case (Cons a rest)
-              assume eq_wq: "wq_fun (schs s) cs' = a # rest"
-              with h V show ?thesis
-                apply (auto simp:Let_def wq_def split:if_splits)
-              proof -
-                assume th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
-                have "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
-                proof(rule someI2)
-                  from vt_s.wq_distinct[of cs'] and eq_wq[folded wq_def]
-                  show "distinct rest \<and> set rest = set rest" by auto
-                next
-                  show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
-                    by auto
-                qed
-                with eq_wq th_in have "th \<in> set (wq_fun (schs s) cs')" by auto
-                from ih[OF this[folded wq_def]] show "th \<in> threads s" .
-              next
-                assume th_in: "th \<in> set (wq_fun (schs s) cs)"
-                from ih[OF this[folded wq_def]]
-                show "th \<in> threads s" .
-              qed
-            qed
-          qed
-        qed
-      qed
-    next
-      case (P th' cs')
-      from h stp
-      show ?thesis
-        apply (unfold P wq_def)
-        apply (auto simp:Let_def split:if_splits, fold wq_def)
-        apply (auto intro:ih)
-        apply(ind_cases "step s (P th' cs')")
-        by (unfold runing_def readys_def, auto)
-    next
-      case (Set thread prio)
-      with ih h show ?thesis
-        by (auto simp:wq_def Let_def)
-    qed
-  next
-    case vt_nil
-    thus ?case by (auto simp:wq_def)
-  qed
-qed
-
-lemma dm_RAG_threads:
-  assumes in_dom: "(Th th) \<in> Domain (RAG s)"
-  shows "th \<in> threads s"
-proof -
-  from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
-  moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
-  ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
-  hence "th \<in> set (wq s cs)"
-    by (unfold s_RAG_def, auto simp:cs_waiting_def)
-  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 le_cp:
-  shows "preced th s \<le> cp s th"
-proof(unfold cp_eq_cpreced preced_def cpreced_def, simp)
-  show "Prc (priority th s) (last_set th s)
-    \<le> Max (insert (Prc (priority th s) (last_set th s))
-            ((\<lambda>th. Prc (priority th s) (last_set th s)) ` dependants (wq s) th))"
-    (is "?l \<le> Max (insert ?l ?A)")
-  proof(cases "?A = {}")
-    case False
-    have "finite ?A" (is "finite (?f ` ?B)")
-    proof -
-      have "finite ?B" 
-      proof-
-        have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}"
-        proof -
-          let ?F = "\<lambda> (x, y). the_th x"
-          have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
-            apply (auto simp:image_def)
-            by (rule_tac x = "(Th x, Th th)" in bexI, auto)
-          moreover have "finite \<dots>"
-          proof -
-            from finite_RAG have "finite (RAG s)" .
-            hence "finite ((RAG (wq s))\<^sup>+)"
-              apply (unfold finite_trancl)
-              by (auto simp: s_RAG_def cs_RAG_def wq_def)
-            thus ?thesis by auto
-          qed
-          ultimately show ?thesis by (auto intro:finite_subset)
-        qed
-        thus ?thesis by (simp add:cs_dependants_def)
-      qed
-      thus ?thesis by simp
-    qed
-    from Max_insert [OF this False, of ?l] show ?thesis by auto
-  next
-    case True
-    thus ?thesis by auto
-  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 
-
-end
-
-lemma preced_v [simp]: "preced th' (V th cs#s) = preced th' s"
-  by (unfold preced_def, simp)
-
-lemma the_preced_v[simp]: "the_preced (V th cs#s) = the_preced s"
-proof
-  fix th'
-  show "the_preced (V th cs # s) th' = the_preced s th'"
-    by (unfold the_preced_def preced_def, simp)
-qed
-
-lemma step_RAG_v: 
-assumes vt:
-  "vt (V th cs#s)"
-shows "
-  RAG (V th cs # s) =
-  RAG s - {(Cs cs, Th th)} -
-  {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-  {(Cs cs, Th th') |th'.  next_th s th cs th'}" (is "?L = ?R")
-proof -
-  interpret vt_v: valid_trace_v s "V th cs"
-    using assms step_back_vt by (unfold_locales, auto) 
-  show ?thesis using vt_v.RAG_es .
-qed
-
-
-
-
-
-text {* (* ddd *) 
-  The following @{text "step_RAG_v"} lemma charaterizes how @{text "RAG"} is changed
-  with the happening of @{text "V"}-events:
-*}
-lemma step_RAG_v:
-assumes vt:
-  "vt (V th cs#s)"
-shows "
-  RAG (V th cs # s) =
-  RAG s - {(Cs cs, Th th)} -
-  {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-  {(Cs cs, Th th') |th'.  next_th s th cs th'}"
-  apply (insert vt, unfold s_RAG_def) 
-  apply (auto split:if_splits list.splits simp:Let_def)
-  apply (auto elim: step_v_waiting_mono step_v_hold_inv
-              step_v_release step_v_wait_inv
-              step_v_get_hold step_v_release_inv)
-  apply (erule_tac step_v_not_wait, auto)
-  done
-
-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
-
-
-lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
-  by (unfold s_RAG_def, auto)
-
-context valid_trace
-begin
-
-text {*
-  The following lemma shows that @{text "RAG"} is acyclic.
-  The overall structure is by induction on the formation of @{text "vt s"}
-  and then case analysis on event @{text "e"}, where the non-trivial cases 
-  for those for @{text "V"} and @{text "P"} events.
-*}
-lemma acyclic_RAG:
-  shows "acyclic (RAG s)"
-using vt
-proof(induct)
-  case (vt_cons s e)
-  interpret vt_s: valid_trace s using vt_cons(1)
-    by (unfold_locales, simp)
-  assume ih: "acyclic (RAG s)"
-    and stp: "step s e"
-    and vt: "vt s"
-  show ?case
-  proof(cases e)
-    case (Create th prio)
-    with ih
-    show ?thesis by (simp add:RAG_create_unchanged)
-  next
-    case (Exit th)
-    with ih show ?thesis by (simp add:RAG_exit_unchanged)
-  next
-    case (V th cs)
-    from V vt stp have vtt: "vt (V th cs#s)" by auto
-    from step_RAG_v [OF this]
-    have eq_de: 
-      "RAG (e # s) = 
-      RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-      {(Cs cs, Th th') |th'. next_th s th cs th'}"
-      (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
-    from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
-    from step_back_step [OF vtt]
-    have "step s (V th cs)" .
-    thus ?thesis
-    proof(cases)
-      assume "holding s th cs"
-      hence th_in: "th \<in> set (wq s cs)" and
-        eq_hd: "th = hd (wq s cs)" unfolding s_holding_def wq_def by auto
-      then obtain rest where
-        eq_wq: "wq s cs = th#rest"
-        by (cases "wq s cs", auto)
-      show ?thesis
-      proof(cases "rest = []")
-        case False
-        let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
-        from eq_wq False have eq_D: "?D = {(Cs cs, Th ?th')}"
-          by (unfold next_th_def, auto)
-        let ?E = "(?A - ?B - ?C)"
-        have "(Th ?th', Cs cs) \<notin> ?E\<^sup>*"
-        proof
-          assume "(Th ?th', Cs cs) \<in> ?E\<^sup>*"
-          hence " (Th ?th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
-          from tranclD [OF this]
-          obtain x where th'_e: "(Th ?th', x) \<in> ?E" by blast
-          hence th_d: "(Th ?th', x) \<in> ?A" by simp
-          from RAG_target_th [OF this]
-          obtain cs' where eq_x: "x = Cs cs'" by auto
-          with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp
-          hence wt_th': "waiting s ?th' cs'"
-            unfolding s_RAG_def s_waiting_def cs_waiting_def wq_def by simp
-          hence "cs' = cs"
-          proof(rule vt_s.waiting_unique)
-            from eq_wq vt_s.wq_distinct[of cs]
-            show "waiting s ?th' cs" 
-              apply (unfold s_waiting_def wq_def, auto)
-            proof -
-              assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-                and eq_wq: "wq_fun (schs s) cs = th # rest"
-              have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
-              next
-                fix x assume "distinct x \<and> set x = set rest"
-                with False show "x \<noteq> []" by auto
-              qed
-              hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
-                set (SOME q. distinct q \<and> set q = set rest)" by auto
-              moreover have "\<dots> = set rest" 
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
-              next
-                show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-              qed
-              moreover note hd_in
-              ultimately show "hd (SOME q. distinct q \<and> set q = set rest) = th" by auto
-            next
-              assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-                and eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest"
-              have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" by auto
-              next
-                fix x assume "distinct x \<and> set x = set rest"
-                with False show "x \<noteq> []" by auto
-              qed
-              hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
-                set (SOME q. distinct q \<and> set q = set rest)" by auto
-              moreover have "\<dots> = set rest" 
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" by auto
-              next
-                show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-              qed
-              moreover note hd_in
-              ultimately show False by auto
-            qed
-          qed
-          with th'_e eq_x have "(Th ?th', Cs cs) \<in> ?E" by simp
-          with False
-          show "False" by (auto simp: next_th_def eq_wq)
-        qed
-        with acyclic_insert[symmetric] and ac
-          and eq_de eq_D show ?thesis by auto
-      next
-        case True
-        with eq_wq
-        have eq_D: "?D = {}"
-          by (unfold next_th_def, auto)
-        with eq_de ac
-        show ?thesis by auto
-      qed 
-    qed
-  next
-    case (P th cs)
-    from P vt stp have vtt: "vt (P th cs#s)" by auto
-    from step_RAG_p [OF this] P
-    have "RAG (e # s) = 
-      (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
-      RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
-      by simp
-    moreover have "acyclic ?R"
-    proof(cases "wq s cs = []")
-      case True
-      hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
-      have "(Th th, Cs cs) \<notin> (RAG s)\<^sup>*"
-      proof
-        assume "(Th th, Cs cs) \<in> (RAG s)\<^sup>*"
-        hence "(Th th, Cs cs) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
-        from tranclD2 [OF this]
-        obtain x where "(x, Cs cs) \<in> RAG s" by auto
-        with True show False by (auto simp:s_RAG_def cs_waiting_def)
-      qed
-      with acyclic_insert ih eq_r show ?thesis by auto
-    next
-      case False
-      hence eq_r: "?R =  RAG s \<union> {(Th th, Cs cs)}" by simp
-      have "(Cs cs, Th th) \<notin> (RAG s)\<^sup>*"
-      proof
-        assume "(Cs cs, Th th) \<in> (RAG s)\<^sup>*"
-        hence "(Cs cs, Th th) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
-        moreover from step_back_step [OF vtt] have "step s (P th cs)" .
-        ultimately show False
-        proof -
-          show " \<lbrakk>(Cs cs, Th th) \<in> (RAG s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
-            by (ind_cases "step s (P th cs)", simp)
-        qed
-      qed
-      with acyclic_insert ih eq_r show ?thesis by auto
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (Set thread prio)
-      with ih
-      thm RAG_set_unchanged
-      show ?thesis by (simp add:RAG_set_unchanged)
-    qed
-  next
-    case vt_nil
-    show "acyclic (RAG ([]::state))"
-      by (auto simp: s_RAG_def cs_waiting_def 
-        cs_holding_def wq_def acyclic_def)
-qed
-
-
-lemma finite_RAG:
-  shows "finite (RAG s)"
-proof -
-  from vt show ?thesis
-  proof(induct)
-    case (vt_cons s e)
-    interpret vt_s: valid_trace s using vt_cons(1)
-      by (unfold_locales, simp)
-    assume ih: "finite (RAG s)"
-      and stp: "step s e"
-      and vt: "vt s"
-    show ?case
-    proof(cases e)
-      case (Create th prio)
-      with ih
-      show ?thesis by (simp add:RAG_create_unchanged)
-    next
-      case (Exit th)
-      with ih show ?thesis by (simp add:RAG_exit_unchanged)
-    next
-      case (V th cs)
-      from V vt stp have vtt: "vt (V th cs#s)" by auto
-      from step_RAG_v [OF this]
-      have eq_de: "RAG (e # s) = 
-                   RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-                      {(Cs cs, Th th') |th'. next_th s th cs th'}
-"
-        (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
-      moreover from ih have ac: "finite (?A - ?B - ?C)" by simp
-      moreover have "finite ?D"
-      proof -
-        have "?D = {} \<or> (\<exists> a. ?D = {a})" 
-          by (unfold next_th_def, auto)
-        thus ?thesis
-        proof
-          assume h: "?D = {}"
-          show ?thesis by (unfold h, simp)
-        next
-          assume "\<exists> a. ?D = {a}"
-          thus ?thesis
-            by (metis finite.simps)
-        qed
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (P th cs)
-      from P vt stp have vtt: "vt (P th cs#s)" by auto
-      from step_RAG_p [OF this] P
-      have "RAG (e # s) = 
-              (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
-                                    RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
-        by simp
-      moreover have "finite ?R"
-      proof(cases "wq s cs = []")
-        case True
-        hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
-        with True and ih show ?thesis by auto
-      next
-        case False
-        hence "?R = RAG s \<union> {(Th th, Cs cs)}" by simp
-        with False and ih show ?thesis by auto
-      qed
-      ultimately show ?thesis by auto
-    next
-      case (Set thread prio)
-      with ih
-      show ?thesis by (simp add:RAG_set_unchanged)
-    qed
-  next
-    case vt_nil
-    show "finite (RAG ([]::state))"
-      by (auto simp: s_RAG_def cs_waiting_def 
-                   cs_holding_def wq_def acyclic_def)
-  qed
-qed
-
-text {* Several useful lemmas *}
-
-lemma wf_dep_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
-
-end
-
-lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l"
-  by (induct l, auto)
-
-lemma th_chasing: "(Th th, Cs cs) \<in> RAG (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> RAG s"
-  by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
-
-context valid_trace
-begin
-
-lemma wq_threads: 
-  assumes h: "th \<in> set (wq s cs)"
-  shows "th \<in> threads s"
-proof -
- from vt and h show ?thesis
-  proof(induct arbitrary: th cs)
-    case (vt_cons s e)
-    interpret vt_s: valid_trace s
-      using vt_cons(1) by (unfold_locales, auto)
-    assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s"
-      and stp: "step s e"
-      and vt: "vt s"
-      and h: "th \<in> set (wq (e # s) cs)"
-    show ?case
-    proof(cases e)
-      case (Create th' prio)
-      with ih h show ?thesis
-        by (auto simp:wq_def Let_def)
-    next
-      case (Exit th')
-      with stp ih h show ?thesis
-        apply (auto simp:wq_def Let_def)
-        apply (ind_cases "step s (Exit th')")
-        apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def
-               s_RAG_def s_holding_def cs_holding_def)
-        done
-    next
-      case (V th' cs')
-      show ?thesis
-      proof(cases "cs' = cs")
-        case False
-        with h
-        show ?thesis
-          apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def)
-          by (drule_tac ih, simp)
-      next
-        case True
-        from h
-        show ?thesis
-        proof(unfold V wq_def)
-          assume th_in: "th \<in> set (wq_fun (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l")
-          show "th \<in> threads (V th' cs' # s)"
-          proof(cases "cs = cs'")
-            case False
-            hence "?l = wq_fun (schs s) cs" by (simp add:Let_def)
-            with th_in have " th \<in> set (wq s cs)" 
-              by (fold wq_def, simp)
-            from ih [OF this] show ?thesis by simp
-          next
-            case True
-            show ?thesis
-            proof(cases "wq_fun (schs s) cs'")
-              case Nil
-              with h V show ?thesis
-                apply (auto simp:wq_def Let_def split:if_splits)
-                by (fold wq_def, drule_tac ih, simp)
-            next
-              case (Cons a rest)
-              assume eq_wq: "wq_fun (schs s) cs' = a # rest"
-              with h V show ?thesis
-                apply (auto simp:Let_def wq_def split:if_splits)
-              proof -
-                assume th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
-                have "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
-                proof(rule someI2)
-                  from vt_s.wq_distinct[of cs'] and eq_wq[folded wq_def]
-                  show "distinct rest \<and> set rest = set rest" by auto
-                next
-                  show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
-                    by auto
-                qed
-                with eq_wq th_in have "th \<in> set (wq_fun (schs s) cs')" by auto
-                from ih[OF this[folded wq_def]] show "th \<in> threads s" .
-              next
-                assume th_in: "th \<in> set (wq_fun (schs s) cs)"
-                from ih[OF this[folded wq_def]]
-                show "th \<in> threads s" .
-              qed
-            qed
-          qed
-        qed
-      qed
-    next
-      case (P th' cs')
-      from h stp
-      show ?thesis
-        apply (unfold P wq_def)
-        apply (auto simp:Let_def split:if_splits, fold wq_def)
-        apply (auto intro:ih)
-        apply(ind_cases "step s (P th' cs')")
-        by (unfold runing_def readys_def, auto)
-    next
-      case (Set thread prio)
-      with ih h show ?thesis
-        by (auto simp:wq_def Let_def)
-    qed
-  next
-    case vt_nil
-    thus ?case by (auto simp:wq_def)
-  qed
-qed
-
-lemma range_in: "\<lbrakk>(Th th) \<in> Range (RAG (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
-  apply(unfold s_RAG_def cs_waiting_def cs_holding_def)
-  by (auto intro:wq_threads)
-
-lemma readys_v_eq:
-  assumes neq_th: "th \<noteq> thread"
-  and eq_wq: "wq s cs = thread#rest"
-  and not_in: "th \<notin>  set rest"
-  shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)"
-proof -
-  from assms show ?thesis
-    apply (auto simp:readys_def)
-    apply(simp add:s_waiting_def[folded wq_def])
-    apply (erule_tac x = csa in allE)
-    apply (simp add:s_waiting_def wq_def Let_def split:if_splits)
-    apply (case_tac "csa = cs", simp)
-    apply (erule_tac x = cs in allE)
-    apply(auto simp add: s_waiting_def[folded wq_def] Let_def split: list.splits)
-    apply(auto simp add: wq_def)
-    apply (auto simp:s_waiting_def wq_def Let_def split:list.splits)
-    proof -
-       assume th_nin: "th \<notin> set rest"
-        and th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
-        and eq_wq: "wq_fun (schs s) cs = thread # rest"
-      have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-      proof(rule someI2)
-        from wq_distinct[of cs, unfolded wq_def] and eq_wq[unfolded wq_def]
-        show "distinct rest \<and> set rest = set rest" by auto
-      next
-        show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-      qed
-      with th_nin th_in show False by auto
-    qed
-qed
-
-text {* \noindent
-  The following lemmas shows that: starting from any node in @{text "RAG"}, 
-  by chasing out-going edges, it is always possible to reach a node representing a ready
-  thread. In this lemma, it is the @{text "th'"}.
-*}
-
-lemma chain_building:
-  shows "node \<in> Domain (RAG s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (RAG s)^+)"
-proof -
-  from wf_dep_converse
-  have h: "wf ((RAG s)\<inverse>)" .
-  show ?thesis
-  proof(induct rule:wf_induct [OF h])
-    fix x
-    assume ih [rule_format]: 
-      "\<forall>y. (y, x) \<in> (RAG s)\<inverse> \<longrightarrow> 
-           y \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (RAG s)\<^sup>+)"
-    show "x \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+)"
-    proof
-      assume x_d: "x \<in> Domain (RAG s)"
-      show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+"
-      proof(cases x)
-        case (Th th)
-        from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> RAG s" by (auto simp:s_RAG_def)
-        with Th have x_in_r: "(Cs cs, x) \<in> (RAG s)^-1" by simp
-        from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> RAG s" by blast
-        hence "Cs cs \<in> Domain (RAG s)" by auto
-        from ih [OF x_in_r this] obtain th'
-          where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (RAG s)\<^sup>+" by auto
-        have "(x, Th th') \<in> (RAG s)\<^sup>+" using Th x_in cs_in by auto
-        with th'_ready show ?thesis by auto
-      next
-        case (Cs cs)
-        from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (RAG s)^-1" by (auto simp:s_RAG_def)
-        show ?thesis
-        proof(cases "th' \<in> readys s")
-          case True
-          from True and th'_d show ?thesis by auto
-        next
-          case False
-          from th'_d and range_in  have "th' \<in> threads s" by auto
-          with False have "Th th' \<in> Domain (RAG s)" 
-            by (auto simp:readys_def wq_def s_waiting_def s_RAG_def cs_waiting_def Domain_def)
-          from ih [OF th'_d this]
-          obtain th'' where 
-            th''_r: "th'' \<in> readys s" and 
-            th''_in: "(Th th', Th th'') \<in> (RAG s)\<^sup>+" by auto
-          from th'_d and th''_in 
-          have "(x, Th th'') \<in> (RAG s)\<^sup>+" by auto
-          with th''_r show ?thesis by auto
-        qed
-      qed
-    qed
-  qed
-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 holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2"
-  by (unfold s_holding_def cs_holding_def, auto)
-
-context valid_trace
-begin
-
-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 holding_unique)
-
-end
-
-
-lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
-by (induct rule:trancl_induct, auto)
-
-context valid_trace
-begin
-
-lemma dchain_unique:
-  assumes th1_d: "(n, Th th1) \<in> (RAG s)^+"
-  and th1_r: "th1 \<in> readys s"
-  and th2_d: "(n, Th th2) \<in> (RAG s)^+"
-  and th2_r: "th2 \<in> readys s"
-  shows "th1 = th2"
-proof -
-  { assume neq: "th1 \<noteq> th2"
-    hence "Th th1 \<noteq> Th th2" by simp
-    from unique_chain [OF _ th1_d th2_d this] and unique_RAG 
-    have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG s)\<^sup>+" by auto
-    hence "False"
-    proof
-      assume "(Th th1, Th th2) \<in> (RAG s)\<^sup>+"
-      from trancl_split [OF this]
-      obtain n where dd: "(Th th1, n) \<in> RAG s" by auto
-      then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
-      from dd eq_n have "th1 \<notin> readys s"
-        by (auto simp:readys_def s_RAG_def wq_def s_waiting_def cs_waiting_def)
-      with th1_r show ?thesis by auto
-    next
-      assume "(Th th2, Th th1) \<in> (RAG s)\<^sup>+"
-      from trancl_split [OF this]
-      obtain n where dd: "(Th th2, n) \<in> RAG s" by auto
-      then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
-      from dd eq_n have "th2 \<notin> readys s"
-        by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def)
-      with th2_r show ?thesis by auto
-    qed
-  } thus ?thesis by auto
-qed
-
-end
-             
-
-lemma step_holdents_p_add:
-  assumes vt: "vt (P th cs#s)"
-  and "wq s cs = []"
-  shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
-proof -
-  from assms show ?thesis
-  unfolding  holdents_test step_RAG_p[OF vt] by (auto)
-qed
-
-lemma step_holdents_p_eq:
-  assumes vt: "vt (P th cs#s)"
-  and "wq s cs \<noteq> []"
-  shows "holdents (P th cs#s) th = holdents s th"
-proof -
-  from assms show ?thesis
-  unfolding  holdents_test step_RAG_p[OF vt] by auto
-qed
-
-
-lemma (in valid_trace) finite_holding :
-  shows "finite (holdents s th)"
-proof -
-  let ?F = "\<lambda> (x, y). the_cs x"
-  from finite_RAG 
-  have "finite (RAG s)" .
-  hence "finite (?F `(RAG s))" by simp
-  moreover have "{cs . (Cs cs, Th th) \<in> RAG s} \<subseteq> \<dots>" 
-  proof -
-    { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
-      fix x assume "(Cs x, Th th) \<in> RAG s"
-      hence "?F (Cs x, Th th) \<in> ?F `(RAG s)" by (rule h)
-      moreover have "?F (Cs x, Th th) = x" by simp
-      ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` RAG s" by simp 
-    } thus ?thesis by auto
-  qed
-  ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset)
-qed
-
-lemma cntCS_v_dec: 
-  assumes vtv: "vt (V thread cs#s)"
-  shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
-proof -
-  from vtv interpret vt_s: valid_trace s
-    by (cases, unfold_locales, simp)
-  from vtv interpret vt_v: valid_trace "V thread cs#s"
-     by (unfold_locales, simp)
-  from step_back_step[OF vtv]
-  have cs_in: "cs \<in> holdents s thread" 
-    apply (cases, unfold holdents_test s_RAG_def, simp)
-    by (unfold cs_holding_def s_holding_def wq_def, auto)
-  moreover have cs_not_in: 
-    "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
-    apply (insert vt_s.wq_distinct[of cs])
-    apply (unfold holdents_test, unfold step_RAG_v[OF vtv],
-            auto simp:next_th_def)
-  proof -
-    fix rest
-    assume dst: "distinct (rest::thread list)"
-      and ne: "rest \<noteq> []"
-    and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-    moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-    qed
-    ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
-                     set (SOME q. distinct q \<and> set q = set rest)" by simp
-    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      fix x assume " distinct x \<and> set x = set rest" with ne
-      show "x \<noteq> []" by auto
-    qed
-    ultimately 
-    show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
-      by auto
-  next
-    fix rest
-    assume dst: "distinct (rest::thread list)"
-      and ne: "rest \<noteq> []"
-    and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-    moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-    qed
-    ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
-                     set (SOME q. distinct q \<and> set q = set rest)" by simp
-    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      fix x assume " distinct x \<and> set x = set rest" with ne
-      show "x \<noteq> []" by auto
-    qed
-    ultimately show "False" by auto 
-  qed
-  ultimately 
-  have "holdents s thread = insert cs (holdents (V thread cs#s) thread)"
-    by auto
-  moreover have "card \<dots> = 
-                    Suc (card ((holdents (V thread cs#s) thread) - {cs}))"
-  proof(rule card_insert)
-    from vt_v.finite_holding
-    show " finite (holdents (V thread cs # s) thread)" .
-  qed
-  moreover from cs_not_in 
-  have "cs \<notin> (holdents (V thread cs#s) thread)" by auto
-  ultimately show ?thesis by (simp add:cntCS_def)
-qed 
-
-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_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)
-
-context valid_trace
-begin
-
-text {* (* ddd *) \noindent
-  The relationship between @{text "cntP"}, @{text "cntV"} and @{text "cntCS"} 
-  of one particular thread. 
-*} 
-
-lemma cnp_cnv_cncs:
-  shows "cntP s th = cntV s th + (if (th \<in> readys s \<or> th \<notin> threads s) 
-                                       then cntCS s th else cntCS s th + 1)"
-proof -
-  from vt show ?thesis
-  proof(induct arbitrary:th)
-    case (vt_cons s e)
-    interpret vt_s: valid_trace s using vt_cons(1) by (unfold_locales, simp)
-    assume vt: "vt s"
-    and ih: "\<And>th. cntP s th  = cntV s th +
-               (if (th \<in> readys s \<or> th \<notin> threads s) then cntCS s th else cntCS s th + 1)"
-    and stp: "step s e"
-    from stp show ?case
-    proof(cases)
-      case (thread_create thread prio)
-      assume eq_e: "e = Create thread prio"
-        and not_in: "thread \<notin> threads s"
-      show ?thesis
-      proof -
-        { fix cs 
-          assume "thread \<in> set (wq s cs)"
-          from vt_s.wq_threads [OF this] have "thread \<in> threads s" .
-          with not_in have "False" by simp
-        } with eq_e have eq_readys: "readys (e#s) = readys s \<union> {thread}"
-          by (auto simp:readys_def threads.simps s_waiting_def 
-            wq_def cs_waiting_def Let_def)
-        from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
-        from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
-        have eq_cncs: "cntCS (e#s) th = cntCS s th"
-          unfolding cntCS_def holdents_test
-          by (simp add:RAG_create_unchanged eq_e)
-        { assume "th \<noteq> thread"
-          with eq_readys eq_e
-          have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
-                      (th \<in> readys (s) \<or> th \<notin> threads (s))" 
-            by (simp add:threads.simps)
-          with eq_cnp eq_cnv eq_cncs ih not_in
-          have ?thesis by simp
-        } moreover {
-          assume eq_th: "th = thread"
-          with not_in ih have " cntP s th  = cntV s th + cntCS s th" by simp
-          moreover from eq_th and eq_readys have "th \<in> readys (e#s)" by simp
-          moreover note eq_cnp eq_cnv eq_cncs
-          ultimately have ?thesis by auto
-        } ultimately show ?thesis by blast
-      qed
-    next
-      case (thread_exit thread)
-      assume eq_e: "e = Exit thread" 
-      and is_runing: "thread \<in> runing s"
-      and no_hold: "holdents s thread = {}"
-      from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
-      from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
-      have eq_cncs: "cntCS (e#s) th = cntCS s th"
-        unfolding cntCS_def holdents_test
-        by (simp add:RAG_exit_unchanged eq_e)
-      { assume "th \<noteq> thread"
-        with eq_e
-        have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
-          (th \<in> readys (s) \<or> th \<notin> threads (s))" 
-          apply (simp add:threads.simps readys_def)
-          apply (subst s_waiting_def)
-          apply (simp add:Let_def)
-          apply (subst s_waiting_def, simp)
-          done
-        with eq_cnp eq_cnv eq_cncs ih
-        have ?thesis by simp
-      } moreover {
-        assume eq_th: "th = thread"
-        with ih is_runing have " cntP s th = cntV s th + cntCS s th" 
-          by (simp add:runing_def)
-        moreover from eq_th eq_e have "th \<notin> threads (e#s)"
-          by simp
-        moreover note eq_cnp eq_cnv eq_cncs
-        ultimately have ?thesis by auto
-      } ultimately show ?thesis by blast
-    next
-      case (thread_P thread cs)
-      assume eq_e: "e = P thread cs"
-        and is_runing: "thread \<in> runing s"
-        and no_dep: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+"
-      from thread_P vt stp ih  have vtp: "vt (P thread cs#s)" by auto
-      then interpret vt_p: valid_trace "(P thread cs#s)"
-        by (unfold_locales, simp)
-      show ?thesis 
-      proof -
-        { have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast
-          assume neq_th: "th \<noteq> thread"
-          with eq_e
-          have eq_readys: "(th \<in> readys (e#s)) = (th \<in> readys (s))"
-            apply (simp add:readys_def s_waiting_def wq_def Let_def)
-            apply (rule_tac hh)
-             apply (intro iffI allI, clarify)
-            apply (erule_tac x = csa in allE, auto)
-            apply (subgoal_tac "wq_fun (schs s) cs \<noteq> []", auto)
-            apply (erule_tac x = cs in allE, auto)
-            by (case_tac "(wq_fun (schs s) cs)", auto)
-          moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th"
-            apply (simp add:cntCS_def holdents_test)
-            by (unfold  step_RAG_p [OF vtp], auto)
-          moreover from eq_e neq_th have "cntP (e # s) th = cntP s th"
-            by (simp add:cntP_def count_def)
-          moreover from eq_e neq_th have "cntV (e#s) th = cntV s th"
-            by (simp add:cntV_def count_def)
-          moreover from eq_e neq_th have "threads (e#s) = threads s" by simp
-          moreover note ih [of th] 
-          ultimately have ?thesis by simp
-        } moreover {
-          assume eq_th: "th = thread"
-          have ?thesis
-          proof -
-            from eq_e eq_th have eq_cnp: "cntP (e # s) th  = 1 + (cntP s th)" 
-              by (simp add:cntP_def count_def)
-            from eq_e eq_th have eq_cnv: "cntV (e#s) th = cntV s th"
-              by (simp add:cntV_def count_def)
-            show ?thesis
-            proof (cases "wq s cs = []")
-              case True
-              with is_runing
-              have "th \<in> readys (e#s)"
-                apply (unfold eq_e wq_def, unfold readys_def s_RAG_def)
-                apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def)
-                by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def)
-              moreover have "cntCS (e # s) th = 1 + cntCS s th"
-              proof -
-                have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> RAG s} =
-                  Suc (card {cs. (Cs cs, Th thread) \<in> RAG s})" (is "card ?L = Suc (card ?R)")
-                proof -
-                  have "?L = insert cs ?R" by auto
-                  moreover have "card \<dots> = Suc (card (?R - {cs}))" 
-                  proof(rule card_insert)
-                    from vt_s.finite_holding [of thread]
-                    show " finite {cs. (Cs cs, Th thread) \<in> RAG s}"
-                      by (unfold holdents_test, simp)
-                  qed
-                  moreover have "?R - {cs} = ?R"
-                  proof -
-                    have "cs \<notin> ?R"
-                    proof
-                      assume "cs \<in> {cs. (Cs cs, Th thread) \<in> RAG s}"
-                      with no_dep show False by auto
-                    qed
-                    thus ?thesis by auto
-                  qed
-                  ultimately show ?thesis by auto
-                qed
-                thus ?thesis
-                  apply (unfold eq_e eq_th cntCS_def)
-                  apply (simp add: holdents_test)
-                  by (unfold step_RAG_p [OF vtp], auto simp:True)
-              qed
-              moreover from is_runing have "th \<in> readys s"
-                by (simp add:runing_def eq_th)
-              moreover note eq_cnp eq_cnv ih [of th]
-              ultimately show ?thesis by auto
-            next
-              case False
-              have eq_wq: "wq (e#s) cs = wq s cs @ [th]"
-                    by (unfold eq_th eq_e wq_def, auto simp:Let_def)
-              have "th \<notin> readys (e#s)"
-              proof
-                assume "th \<in> readys (e#s)"
-                hence "\<forall>cs. \<not> waiting (e # s) th cs" by (simp add:readys_def)
-                from this[rule_format, of cs] have " \<not> waiting (e # s) th cs" .
-                hence "th \<in> set (wq (e#s) cs) \<Longrightarrow> th = hd (wq (e#s) cs)" 
-                  by (simp add:s_waiting_def wq_def)
-                moreover from eq_wq have "th \<in> set (wq (e#s) cs)" by auto
-                ultimately have "th = hd (wq (e#s) cs)" by blast
-                with eq_wq have "th = hd (wq s cs @ [th])" by simp
-                hence "th = hd (wq s cs)" using False by auto
-                with False eq_wq vt_p.wq_distinct [of cs]
-                show False by (fold eq_e, auto)
-              qed
-              moreover from is_runing have "th \<in> threads (e#s)" 
-                by (unfold eq_e, auto simp:runing_def readys_def eq_th)
-              moreover have "cntCS (e # s) th = cntCS s th"
-                apply (unfold cntCS_def holdents_test eq_e step_RAG_p[OF vtp])
-                by (auto simp:False)
-              moreover note eq_cnp eq_cnv ih[of th]
-              moreover from is_runing have "th \<in> readys s"
-                by (simp add:runing_def eq_th)
-              ultimately show ?thesis by auto
-            qed
-          qed
-        } ultimately show ?thesis by blast
-      qed
-    next
-      case (thread_V thread cs)
-      from assms vt stp ih thread_V have vtv: "vt (V thread cs # s)" by auto
-      then interpret vt_v: valid_trace "(V thread cs # s)" by (unfold_locales, simp)
-      assume eq_e: "e = V thread cs"
-        and is_runing: "thread \<in> runing s"
-        and hold: "holding s thread cs"
-      from hold obtain rest 
-        where eq_wq: "wq s cs = thread # rest"
-        by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
-      have eq_threads: "threads (e#s) = threads s" by (simp add: eq_e)
-      have eq_set: "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-      proof(rule someI2)
-        from vt_v.wq_distinct[of cs] and eq_wq
-        show "distinct rest \<and> set rest = set rest"
-          by (metis distinct.simps(2) vt_s.wq_distinct)
-      next
-        show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
-          by auto
-      qed
-      show ?thesis
-      proof -
-        { assume eq_th: "th = thread"
-          from eq_th have eq_cnp: "cntP (e # s) th = cntP s th"
-            by (unfold eq_e, simp add:cntP_def count_def)
-          moreover from eq_th have eq_cnv: "cntV (e#s) th = 1 + cntV s th"
-            by (unfold eq_e, simp add:cntV_def count_def)
-          moreover from cntCS_v_dec [OF vtv] 
-          have "cntCS (e # s) thread + 1 = cntCS s thread"
-            by (simp add:eq_e)
-          moreover from is_runing have rd_before: "thread \<in> readys s"
-            by (unfold runing_def, simp)
-          moreover have "thread \<in> readys (e # s)"
-          proof -
-            from is_runing
-            have "thread \<in> threads (e#s)" 
-              by (unfold eq_e, auto simp:runing_def readys_def)
-            moreover have "\<forall> cs1. \<not> waiting (e#s) thread cs1"
-            proof
-              fix cs1
-              { assume eq_cs: "cs1 = cs" 
-                have "\<not> waiting (e # s) thread cs1"
-                proof -
-                  from eq_wq
-                  have "thread \<notin> set (wq (e#s) cs1)"
-                    apply(unfold eq_e wq_def eq_cs s_holding_def)
-                    apply (auto simp:Let_def)
-                  proof -
-                    assume "thread \<in> set (SOME q. distinct q \<and> set q = set rest)"
-                    with eq_set have "thread \<in> set rest" by simp
-                    with vt_v.wq_distinct[of cs]
-                    and eq_wq show False
-                        by (metis distinct.simps(2) vt_s.wq_distinct)
-                  qed
-                  thus ?thesis by (simp add:wq_def s_waiting_def)
-                qed
-              } moreover {
-                assume neq_cs: "cs1 \<noteq> cs"
-                  have "\<not> waiting (e # s) thread cs1" 
-                  proof -
-                    from wq_v_neq [OF neq_cs[symmetric]]
-                    have "wq (V thread cs # s) cs1 = wq s cs1" .
-                    moreover have "\<not> waiting s thread cs1" 
-                    proof -
-                      from runing_ready and is_runing
-                      have "thread \<in> readys s" by auto
-                      thus ?thesis by (simp add:readys_def)
-                    qed
-                    ultimately show ?thesis 
-                      by (auto simp:wq_def s_waiting_def eq_e)
-                  qed
-              } ultimately show "\<not> waiting (e # s) thread cs1" by blast
-            qed
-            ultimately show ?thesis by (simp add:readys_def)
-          qed
-          moreover note eq_th ih
-          ultimately have ?thesis by auto
-        } moreover {
-          assume neq_th: "th \<noteq> thread"
-          from neq_th eq_e have eq_cnp: "cntP (e # s) th = cntP s th" 
-            by (simp add:cntP_def count_def)
-          from neq_th eq_e have eq_cnv: "cntV (e # s) th = cntV s th" 
-            by (simp add:cntV_def count_def)
-          have ?thesis
-          proof(cases "th \<in> set rest")
-            case False
-            have "(th \<in> readys (e # s)) = (th \<in> readys s)"
-              apply (insert step_back_vt[OF vtv])
-              by (simp add: False eq_e eq_wq neq_th vt_s.readys_v_eq)
-            moreover have "cntCS (e#s) th = cntCS s th"
-              apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
-              proof -
-                have "{csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
-                      {cs. (Cs cs, Th th) \<in> RAG s}"
-                proof -
-                  from False eq_wq
-                  have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> RAG s"
-                    apply (unfold next_th_def, auto)
-                  proof -
-                    assume ne: "rest \<noteq> []"
-                      and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-                      and eq_wq: "wq s cs = thread # rest"
-                    from eq_set ni have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
-                                  set (SOME q. distinct q \<and> set q = set rest)
-                                  " by simp
-                    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-                    proof(rule someI2)
-                      from vt_s.wq_distinct[ of cs] and eq_wq
-                      show "distinct rest \<and> set rest = set rest" by auto
-                    next
-                      fix x assume "distinct x \<and> set x = set rest"
-                      with ne show "x \<noteq> []" by auto
-                    qed
-                    ultimately show 
-                      "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
-                      by auto
-                  qed    
-                  thus ?thesis by auto
-                qed
-                thus "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
-                             card {cs. (Cs cs, Th th) \<in> RAG s}" by simp 
-              qed
-            moreover note ih eq_cnp eq_cnv eq_threads
-            ultimately show ?thesis by auto
-          next
-            case True
-            assume th_in: "th \<in> set rest"
-            show ?thesis
-            proof(cases "next_th s thread cs th")
-              case False
-              with eq_wq and th_in have 
-                neq_hd: "th \<noteq> hd (SOME q. distinct q \<and> set q = set rest)" (is "th \<noteq> hd ?rest")
-                by (auto simp:next_th_def)
-              have "(th \<in> readys (e # s)) = (th \<in> readys s)"
-              proof -
-                from eq_wq and th_in
-                have "\<not> th \<in> readys s"
-                  apply (auto simp:readys_def s_waiting_def)
-                  apply (rule_tac x = cs in exI, auto)
-                  by (insert vt_s.wq_distinct[of cs], auto simp add: wq_def)
-                moreover 
-                from eq_wq and th_in and neq_hd
-                have "\<not> (th \<in> readys (e # s))"
-                  apply (auto simp:readys_def s_waiting_def eq_e wq_def Let_def split:list.splits)
-                  by (rule_tac x = cs in exI, auto simp:eq_set)
-                ultimately show ?thesis by auto
-              qed
-              moreover have "cntCS (e#s) th = cntCS s th" 
-              proof -
-                from eq_wq and  th_in and neq_hd
-                have "(holdents (e # s) th) = (holdents s th)"
-                  apply (unfold eq_e step_RAG_v[OF vtv], 
-                         auto simp:next_th_def eq_set s_RAG_def holdents_test wq_def
-                                   Let_def cs_holding_def)
-                  by (insert vt_s.wq_distinct[of cs], auto simp:wq_def)
-                thus ?thesis by (simp add:cntCS_def)
-              qed
-              moreover note ih eq_cnp eq_cnv eq_threads
-              ultimately show ?thesis by auto
-            next
-              case True
-              let ?rest = " (SOME q. distinct q \<and> set q = set rest)"
-              let ?t = "hd ?rest"
-              from True eq_wq th_in neq_th
-              have "th \<in> readys (e # s)"
-                apply (auto simp:eq_e readys_def s_waiting_def wq_def
-                        Let_def next_th_def)
-              proof -
-                assume eq_wq: "wq_fun (schs s) cs = thread # rest"
-                  and t_in: "?t \<in> set rest"
-                show "?t \<in> threads s"
-                proof(rule vt_s.wq_threads)
-                  from eq_wq and t_in
-                  show "?t \<in> set (wq s cs)" by (auto simp:wq_def)
-                qed
-              next
-                fix csa
-                assume eq_wq: "wq_fun (schs s) cs = thread # rest"
-                  and t_in: "?t \<in> set rest"
-                  and neq_cs: "csa \<noteq> cs"
-                  and t_in': "?t \<in>  set (wq_fun (schs s) csa)"
-                show "?t = hd (wq_fun (schs s) csa)"
-                proof -
-                  { assume neq_hd': "?t \<noteq> hd (wq_fun (schs s) csa)"
-                    from vt_s.wq_distinct[of cs] and 
-                    eq_wq[folded wq_def] and t_in eq_wq
-                    have "?t \<noteq> thread" by auto
-                    with eq_wq and t_in
-                    have w1: "waiting s ?t cs"
-                      by (auto simp:s_waiting_def wq_def)
-                    from t_in' neq_hd'
-                    have w2: "waiting s ?t csa"
-                      by (auto simp:s_waiting_def wq_def)
-                    from vt_s.waiting_unique[OF w1 w2]
-                    and neq_cs have "False" by auto
-                  } thus ?thesis by auto
-                qed
-              qed
-              moreover have "cntP s th = cntV s th + cntCS s th + 1"
-              proof -
-                have "th \<notin> readys s" 
-                proof -
-                  from True eq_wq neq_th th_in
-                  show ?thesis
-                    apply (unfold readys_def s_waiting_def, auto)
-                    by (rule_tac x = cs in exI, auto simp add: wq_def)
-                qed
-                moreover have "th \<in> threads s"
-                proof -
-                  from th_in eq_wq
-                  have "th \<in> set (wq s cs)" by simp
-                  from vt_s.wq_threads [OF this] 
-                  show ?thesis .
-                qed
-                ultimately show ?thesis using ih by auto
-              qed
-              moreover from True neq_th have "cntCS (e # s) th = 1 + cntCS s th"
-                apply (unfold cntCS_def holdents_test eq_e step_RAG_v[OF vtv], auto)
-              proof -
-                show "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs} =
-                               Suc (card {cs. (Cs cs, Th th) \<in> RAG s})"
-                  (is "card ?A = Suc (card ?B)")
-                proof -
-                  have "?A = insert cs ?B" by auto
-                  hence "card ?A = card (insert cs ?B)" by simp
-                  also have "\<dots> = Suc (card ?B)"
-                  proof(rule card_insert_disjoint)
-                    have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` RAG s)" 
-                      apply (auto simp:image_def)
-                      by (rule_tac x = "(Cs x, Th th)" in bexI, auto)
-                    with vt_s.finite_RAG
-                    show "finite {cs. (Cs cs, Th th) \<in> RAG s}" by (auto intro:finite_subset)
-                  next
-                    show "cs \<notin> {cs. (Cs cs, Th th) \<in> RAG s}"
-                    proof
-                      assume "cs \<in> {cs. (Cs cs, Th th) \<in> RAG s}"
-                      hence "(Cs cs, Th th) \<in> RAG s" by simp
-                      with True neq_th eq_wq show False
-                        by (auto simp:next_th_def s_RAG_def cs_holding_def)
-                    qed
-                  qed
-                  finally show ?thesis .
-                qed
-              qed
-              moreover note eq_cnp eq_cnv
-              ultimately show ?thesis by simp
-            qed
-          qed
-        } ultimately show ?thesis by blast
-      qed
-    next
-      case (thread_set thread prio)
-      assume eq_e: "e = Set thread prio"
-        and is_runing: "thread \<in> runing s"
-      show ?thesis
-      proof -
-        from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
-        from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
-        have eq_cncs: "cntCS (e#s) th = cntCS s th"
-          unfolding cntCS_def holdents_test
-          by (simp add:RAG_set_unchanged eq_e)
-        from eq_e have eq_readys: "readys (e#s) = readys s" 
-          by (simp add:readys_def cs_waiting_def s_waiting_def wq_def,
-                  auto simp:Let_def)
-        { assume "th \<noteq> thread"
-          with eq_readys eq_e
-          have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
-                      (th \<in> readys (s) \<or> th \<notin> threads (s))" 
-            by (simp add:threads.simps)
-          with eq_cnp eq_cnv eq_cncs ih is_runing
-          have ?thesis by simp
-        } moreover {
-          assume eq_th: "th = thread"
-          with is_runing ih have " cntP s th  = cntV s th + cntCS s th" 
-            by (unfold runing_def, auto)
-          moreover from eq_th and eq_readys is_runing have "th \<in> readys (e#s)"
-            by (simp add:runing_def)
-          moreover note eq_cnp eq_cnv eq_cncs
-          ultimately have ?thesis by auto
-        } ultimately show ?thesis by blast
-      qed   
-    qed
-  next
-    case vt_nil
-    show ?case 
-      by (unfold cntP_def cntV_def cntCS_def, 
-        auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
-  qed
-qed
-
-lemma not_thread_cncs:
-  assumes not_in: "th \<notin> threads s" 
-  shows "cntCS s th = 0"
-proof -
-  from vt not_in show ?thesis
-  proof(induct arbitrary:th)
-    case (vt_cons s e th)
-    interpret vt_s: valid_trace s using vt_cons(1)
-       by (unfold_locales, simp)
-    assume vt: "vt s"
-      and ih: "\<And>th. th \<notin> threads s \<Longrightarrow> cntCS s th = 0"
-      and stp: "step s e"
-      and not_in: "th \<notin> threads (e # s)"
-    from stp show ?case
-    proof(cases)
-      case (thread_create thread prio)
-      assume eq_e: "e = Create thread prio"
-        and not_in': "thread \<notin> threads s"
-      have "cntCS (e # s) th = cntCS s th"
-        apply (unfold eq_e cntCS_def holdents_test)
-        by (simp add:RAG_create_unchanged)
-      moreover have "th \<notin> threads s" 
-      proof -
-        from not_in eq_e show ?thesis by simp
-      qed
-      moreover note ih ultimately show ?thesis by auto
-    next
-      case (thread_exit thread)
-      assume eq_e: "e = Exit thread"
-      and nh: "holdents s thread = {}"
-      have eq_cns: "cntCS (e # s) th = cntCS s th"
-        apply (unfold eq_e cntCS_def holdents_test)
-        by (simp add:RAG_exit_unchanged)
-      show ?thesis
-      proof(cases "th = thread")
-        case True
-        have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True)
-        with eq_cns show ?thesis by simp
-      next
-        case False
-        with not_in and eq_e
-        have "th \<notin> threads s" by simp
-        from ih[OF this] and eq_cns show ?thesis by simp
-      qed
-    next
-      case (thread_P thread cs)
-      assume eq_e: "e = P thread cs"
-      and is_runing: "thread \<in> runing s"
-      from assms thread_P ih vt stp thread_P have vtp: "vt (P thread cs#s)" by auto
-      have neq_th: "th \<noteq> thread" 
-      proof -
-        from not_in eq_e have "th \<notin> threads s" by simp
-        moreover from is_runing have "thread \<in> threads s"
-          by (simp add:runing_def readys_def)
-        ultimately show ?thesis by auto
-      qed
-      hence "cntCS (e # s) th  = cntCS s th "
-        apply (unfold cntCS_def holdents_test eq_e)
-        by (unfold step_RAG_p[OF vtp], auto)
-      moreover have "cntCS s th = 0"
-      proof(rule ih)
-        from not_in eq_e show "th \<notin> threads s" by simp
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (thread_V thread cs)
-      assume eq_e: "e = V thread cs"
-        and is_runing: "thread \<in> runing s"
-        and hold: "holding s thread cs"
-      have neq_th: "th \<noteq> thread" 
-      proof -
-        from not_in eq_e have "th \<notin> threads s" by simp
-        moreover from is_runing have "thread \<in> threads s"
-          by (simp add:runing_def readys_def)
-        ultimately show ?thesis by auto
-      qed
-      from assms thread_V vt stp ih 
-      have vtv: "vt (V thread cs#s)" by auto
-      then interpret vt_v: valid_trace "(V thread cs#s)"
-        by (unfold_locales, simp)
-      from hold obtain rest 
-        where eq_wq: "wq s cs = thread # rest"
-        by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
-      from not_in eq_e eq_wq
-      have "\<not> next_th s thread cs th"
-        apply (auto simp:next_th_def)
-      proof -
-        assume ne: "rest \<noteq> []"
-          and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> threads s" (is "?t \<notin> threads s")
-        have "?t \<in> set rest"
-        proof(rule someI2)
-          from vt_v.wq_distinct[of cs] and eq_wq
-          show "distinct rest \<and> set rest = set rest"
-            by (metis distinct.simps(2) vt_s.wq_distinct) 
-        next
-          fix x assume "distinct x \<and> set x = set rest" with ne
-          show "hd x \<in> set rest" by (cases x, auto)
-        qed
-        with eq_wq have "?t \<in> set (wq s cs)" by simp
-        from vt_s.wq_threads[OF this] and ni
-        show False
-          using `hd (SOME q. distinct q \<and> set q = set rest) \<in> set (wq s cs)` 
-            ni vt_s.wq_threads by blast 
-      qed
-      moreover note neq_th eq_wq
-      ultimately have "cntCS (e # s) th  = cntCS s th"
-        by (unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
-      moreover have "cntCS s th = 0"
-      proof(rule ih)
-        from not_in eq_e show "th \<notin> threads s" by simp
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (thread_set thread prio)
-      print_facts
-      assume eq_e: "e = Set thread prio"
-        and is_runing: "thread \<in> runing s"
-      from not_in and eq_e have "th \<notin> threads s" by auto
-      from ih [OF this] and eq_e
-      show ?thesis 
-        apply (unfold eq_e cntCS_def holdents_test)
-        by (simp add:RAG_set_unchanged)
-    qed
-    next
-      case vt_nil
-      show ?case
-      by (unfold cntCS_def, 
-        auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
-  qed
-qed
-
-end
-
-
-context valid_trace
-begin
-
-lemma dm_RAG_threads:
-  assumes in_dom: "(Th th) \<in> Domain (RAG s)"
-  shows "th \<in> threads s"
-proof -
-  from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
-  moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
-  ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
-  hence "th \<in> set (wq s cs)"
-    by (unfold s_RAG_def, auto simp:cs_waiting_def)
-  from wq_threads [OF this] show ?thesis .
-qed
-
-end
-
-lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
-unfolding cp_def wq_def
-apply(induct s rule: schs.induct)
-thm cpreced_initial
-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
-
-context valid_trace
-begin
-
-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
-    apply(simp)
-    done
-  hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1)) =
-                 Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependants (wq s) th2))"
-    (is "Max (?f ` ?A) = Max (?f ` ?B)")
-    unfolding cp_eq_cpreced 
-    unfolding cpreced_def .
-  obtain th1' where th1_in: "th1' \<in> ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)"
-  proof -
-    have h1: "finite (?f ` ?A)"
-    proof -
-      have "finite ?A" 
-      proof -
-        have "finite (dependants (wq s) th1)"
-        proof-
-          have "finite {th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+}"
-          proof -
-            let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
-              apply (auto simp:image_def)
-              by (rule_tac x = "(Th x, Th th1)" in bexI, auto)
-            moreover have "finite \<dots>"
-            proof -
-              from finite_RAG have "finite (RAG s)" .
-              hence "finite ((RAG (wq s))\<^sup>+)"
-                apply (unfold finite_trancl)
-                by (auto simp: s_RAG_def cs_RAG_def wq_def)
-              thus ?thesis by auto
-            qed
-            ultimately show ?thesis by (auto intro:finite_subset)
-          qed
-          thus ?thesis by (simp add:cs_dependants_def)
-        qed
-        thus ?thesis by simp
-      qed
-      thus ?thesis by auto
-    qed
-    moreover have h2: "(?f ` ?A) \<noteq> {}"
-    proof -
-      have "?A \<noteq> {}" by simp
-      thus ?thesis by simp
-    qed
-    from Max_in [OF h1 h2]
-    have "Max (?f ` ?A) \<in> (?f ` ?A)" .
-    thus ?thesis 
-      thm cpreced_def
-      unfolding cpreced_def[symmetric] 
-      unfolding cp_eq_cpreced[symmetric] 
-      unfolding cpreced_def 
-      using that[intro] by (auto)
-  qed
-  obtain th2' where th2_in: "th2' \<in> ?B" and eq_f_th2: "?f th2' = Max (?f ` ?B)"
-  proof -
-    have h1: "finite (?f ` ?B)"
-    proof -
-      have "finite ?B" 
-      proof -
-        have "finite (dependants (wq s) th2)"
-        proof-
-          have "finite {th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+}"
-          proof -
-            let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
-              apply (auto simp:image_def)
-              by (rule_tac x = "(Th x, Th th2)" in bexI, auto)
-            moreover have "finite \<dots>"
-            proof -
-              from finite_RAG have "finite (RAG s)" .
-              hence "finite ((RAG (wq s))\<^sup>+)"
-                apply (unfold finite_trancl)
-                by (auto simp: s_RAG_def cs_RAG_def wq_def)
-              thus ?thesis by auto
-            qed
-            ultimately show ?thesis by (auto intro:finite_subset)
-          qed
-          thus ?thesis by (simp add:cs_dependants_def)
-        qed
-        thus ?thesis by simp
-      qed
-      thus ?thesis by auto
-    qed
-    moreover have h2: "(?f ` ?B) \<noteq> {}"
-    proof -
-      have "?B \<noteq> {}" by simp
-      thus ?thesis by simp
-    qed
-    from Max_in [OF h1 h2]
-    have "Max (?f ` ?B) \<in> (?f ` ?B)" .
-    thus ?thesis by (auto intro:that)
-  qed
-  from eq_f_th1 eq_f_th2 eq_max 
-  have eq_preced: "preced th1' s = preced th2' s" by auto
-  hence eq_th12: "th1' = th2'"
-  proof (rule preced_unique)
-    from th1_in have "th1' = th1 \<or> (th1' \<in> dependants (wq s) th1)" by simp
-    thus "th1' \<in> threads s"
-    proof
-      assume "th1' \<in> dependants (wq s) th1"
-      hence "(Th th1') \<in> Domain ((RAG s)^+)"
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "(Th th1') \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      from dm_RAG_threads[OF this] show ?thesis .
-    next
-      assume "th1' = th1"
-      with runing_1 show ?thesis
-        by (unfold runing_def readys_def, auto)
-    qed
-  next
-    from th2_in have "th2' = th2 \<or> (th2' \<in> dependants (wq s) th2)" by simp
-    thus "th2' \<in> threads s"
-    proof
-      assume "th2' \<in> dependants (wq s) th2"
-      hence "(Th th2') \<in> Domain ((RAG s)^+)"
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "(Th th2') \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      from dm_RAG_threads[OF this] show ?thesis .
-    next
-      assume "th2' = th2"
-      with runing_2 show ?thesis
-        by (unfold runing_def readys_def, auto)
-    qed
-  qed
-  from th1_in have "th1' = th1 \<or> th1' \<in> dependants (wq s) th1" by simp
-  thus ?thesis
-  proof
-    assume eq_th': "th1' = th1"
-    from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
-    thus ?thesis
-    proof
-      assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
-    next
-      assume "th2' \<in> dependants (wq s) th2"
-      with eq_th12 eq_th' have "th1 \<in> dependants (wq s) th2" by simp
-      hence "(Th th1, Th th2) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      hence "Th th1 \<in> Domain ((RAG s)^+)" 
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "Th th1 \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      then obtain n where d: "(Th th1, n) \<in> RAG s" by (auto simp:Domain_def)
-      from RAG_target_th [OF this]
-      obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th1, Cs cs') \<in> RAG s" by simp
-      with runing_1 have "False"
-        apply (unfold runing_def readys_def s_RAG_def)
-        by (auto simp:waiting_eq)
-      thus ?thesis by simp
-    qed
-  next
-    assume th1'_in: "th1' \<in> dependants (wq s) th1"
-    from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
-    thus ?thesis 
-    proof
-      assume "th2' = th2"
-      with th1'_in eq_th12 have "th2 \<in> dependants (wq s) th1" by simp
-      hence "(Th th2, Th th1) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      hence "Th th2 \<in> Domain ((RAG s)^+)" 
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "Th th2 \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      then obtain n where d: "(Th th2, n) \<in> RAG s" by (auto simp:Domain_def)
-      from RAG_target_th [OF this]
-      obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th2, Cs cs') \<in> RAG s" by simp
-      with runing_2 have "False"
-        apply (unfold runing_def readys_def s_RAG_def)
-        by (auto simp:waiting_eq)
-      thus ?thesis by simp
-    next
-      assume "th2' \<in> dependants (wq s) th2"
-      with eq_th12 have "th1' \<in> dependants (wq s) th2" by simp
-      hence h1: "(Th th1', Th th2) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      from th1'_in have h2: "(Th th1', Th th1) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      show ?thesis
-      proof(rule dchain_unique[OF h1 _ h2, symmetric])
-        from runing_1 show "th1 \<in> readys s" by (simp add:runing_def)
-        from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) 
-      qed
-    qed
-  qed
-qed
-
-
-lemma "card (runing s) \<le> 1"
-apply(subgoal_tac "finite (runing s)")
-prefer 2
-apply (metis finite_nat_set_iff_bounded lessI runing_unique)
-apply(rule ccontr)
-apply(simp)
-apply(case_tac "Suc (Suc 0) \<le> card (runing s)")
-apply(subst (asm) card_le_Suc_iff)
-apply(simp)
-apply(auto)[1]
-apply (metis insertCI runing_unique)
-apply(auto) 
-done
-
-end
-
-
-lemma create_pre:
-  assumes stp: "step s e"
-  and not_in: "th \<notin> threads s"
-  and is_in: "th \<in> threads (e#s)"
-  obtains prio where "e = Create th prio"
-proof -
-  from assms  
-  show ?thesis
-  proof(cases)
-    case (thread_create thread prio)
-    with is_in not_in have "e = Create th prio" by simp
-    from that[OF this] show ?thesis .
-  next
-    case (thread_exit thread)
-    with assms show ?thesis by (auto intro!:that)
-  next
-    case (thread_P thread)
-    with assms show ?thesis by (auto intro!:that)
-  next
-    case (thread_V thread)
-    with assms show ?thesis by (auto intro!:that)
-  next 
-    case (thread_set thread)
-    with assms show ?thesis by (auto intro!:that)
-  qed
-qed
-
-context valid_trace
-begin
-
-lemma cnp_cnv_eq:
-  assumes "th \<notin> threads s"
-  shows "cntP s th = cntV s th"
-  using assms
-  using cnp_cnv_cncs not_thread_cncs by auto
-
-end
-
-
-lemma eq_RAG: 
-  "RAG (wq s) = RAG s"
-by (unfold cs_RAG_def s_RAG_def, auto)
-
-context valid_trace
-begin
-
-lemma count_eq_dependants:
-  assumes eq_pv: "cntP s th = cntV s th"
-  shows "dependants (wq s) th = {}"
-proof -
-  from cnp_cnv_cncs and eq_pv
-  have "cntCS s th = 0" 
-    by (auto split:if_splits)
-  moreover have "finite {cs. (Cs cs, Th th) \<in> RAG s}"
-  proof -
-    from finite_holding[of th] show ?thesis
-      by (simp add:holdents_test)
-  qed
-  ultimately have h: "{cs. (Cs cs, Th th) \<in> RAG s} = {}"
-    by (unfold cntCS_def holdents_test cs_dependants_def, auto)
-  show ?thesis
-  proof(unfold cs_dependants_def)
-    { assume "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}"
-      then obtain th' where "(Th th', Th th) \<in> (RAG (wq s))\<^sup>+" by auto
-      hence "False"
-      proof(cases)
-        assume "(Th th', Th th) \<in> RAG (wq s)"
-        thus "False" by (auto simp:cs_RAG_def)
-      next
-        fix c
-        assume "(c, Th th) \<in> RAG (wq s)"
-        with h and eq_RAG show "False"
-          by (cases c, auto simp:cs_RAG_def)
-      qed
-    } thus "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} = {}" by auto
-  qed
-qed
-
-lemma dependants_threads:
-  shows "dependants (wq s) th \<subseteq> threads s"
-proof
-  { fix th th'
-    assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (RAG (wq s))\<^sup>+}"
-    have "Th th \<in> Domain (RAG s)"
-    proof -
-      from h obtain th' where "(Th th, Th th') \<in> (RAG (wq s))\<^sup>+" by auto
-      hence "(Th th) \<in> Domain ( (RAG (wq s))\<^sup>+)" by (auto simp:Domain_def)
-      with trancl_domain have "(Th th) \<in> Domain (RAG (wq s))" by simp
-      thus ?thesis using eq_RAG by simp
-    qed
-    from dm_RAG_threads[OF this]
-    have "th \<in> threads s" .
-  } note hh = this
-  fix th1 
-  assume "th1 \<in> dependants (wq s) th"
-  hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (RAG (wq s))\<^sup>+}"
-    by (unfold cs_dependants_def, simp)
-  from hh [OF this] show "th1 \<in> threads s" .
-qed
-
-lemma finite_threads:
-  shows "finite (threads s)"
-using vt by (induct) (auto elim: step.cases)
-
-end
-
-lemma Max_f_mono:
-  assumes seq: "A \<subseteq> B"
-  and np: "A \<noteq> {}"
-  and fnt: "finite B"
-  shows "Max (f ` A) \<le> Max (f ` B)"
-proof(rule Max_mono)
-  from seq show "f ` A \<subseteq> f ` B" by auto
-next
-  from np show "f ` A \<noteq> {}" by auto
-next
-  from fnt and seq show "finite (f ` B)" by auto
-qed
-
-context valid_trace
-begin
-
-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 le_cp:
-  shows "preced th s \<le> cp s th"
-proof(unfold cp_eq_cpreced preced_def cpreced_def, simp)
-  show "Prc (priority th s) (last_set th s)
-    \<le> Max (insert (Prc (priority th s) (last_set th s))
-            ((\<lambda>th. Prc (priority th s) (last_set th s)) ` dependants (wq s) th))"
-    (is "?l \<le> Max (insert ?l ?A)")
-  proof(cases "?A = {}")
-    case False
-    have "finite ?A" (is "finite (?f ` ?B)")
-    proof -
-      have "finite ?B" 
-      proof-
-        have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}"
-        proof -
-          let ?F = "\<lambda> (x, y). the_th x"
-          have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
-            apply (auto simp:image_def)
-            by (rule_tac x = "(Th x, Th th)" in bexI, auto)
-          moreover have "finite \<dots>"
-          proof -
-            from finite_RAG have "finite (RAG s)" .
-            hence "finite ((RAG (wq s))\<^sup>+)"
-              apply (unfold finite_trancl)
-              by (auto simp: s_RAG_def cs_RAG_def wq_def)
-            thus ?thesis by auto
-          qed
-          ultimately show ?thesis by (auto intro:finite_subset)
-        qed
-        thus ?thesis by (simp add:cs_dependants_def)
-      qed
-      thus ?thesis by simp
-    qed
-    from Max_insert [OF this False, of ?l] show ?thesis by auto
-  next
-    case True
-    thus ?thesis by auto
-  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_readys_threads_pre:
-  assumes np: "threads s \<noteq> {}"
-  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
-proof(unfold max_cp_eq)
-  show "Max (cp s ` readys s) = Max ((\<lambda>th. preced th s) ` threads s)"
-  proof -
-    let ?p = "Max ((\<lambda>th. preced th s) ` threads s)" 
-    let ?f = "(\<lambda>th. preced th s)"
-    have "?p \<in> ((\<lambda>th. preced th s) ` threads s)"
-    proof(rule Max_in)
-      from finite_threads show "finite (?f ` threads s)" by simp
-    next
-      from np show "?f ` threads s \<noteq> {}" by simp
-    qed
-    then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
-      by (auto simp:Image_def)
-    from th_chain_to_ready [OF tm_in]
-    have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" .
-    thus ?thesis
-    proof
-      assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ "
-      then obtain th' where th'_in: "th' \<in> readys s" 
-        and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto
-      have "cp s th' = ?f tm"
-      proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
-        from dependants_threads finite_threads
-        show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" 
-          by (auto intro:finite_subset)
-      next
-        fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
-        from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" .
-        moreover have "p \<le> \<dots>"
-        proof(rule Max_ge)
-          from finite_threads
-          show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        next
-          from p_in and th'_in and dependants_threads[of th']
-          show "p \<in> (\<lambda>th. preced th s) ` threads s"
-            by (auto simp:readys_def)
-        qed
-        ultimately show "p \<le> preced tm s" by auto
-      next
-        show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
-        proof -
-          from tm_chain
-          have "tm \<in> dependants (wq s) th'"
-            by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto)
-          thus ?thesis by auto
-        qed
-      qed
-      with tm_max
-      have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
-      show ?thesis
-      proof (fold h, rule Max_eqI)
-        fix q 
-        assume "q \<in> cp s ` readys s"
-        then obtain th1 where th1_in: "th1 \<in> readys s"
-          and eq_q: "q = cp s th1" by auto
-        show "q \<le> cp s th'"
-          apply (unfold h eq_q)
-          apply (unfold cp_eq_cpreced cpreced_def)
-          apply (rule Max_mono)
-        proof -
-          from dependants_threads [of th1] th1_in
-          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<subseteq> 
-                 (\<lambda>th. preced th s) ` threads s"
-            by (auto simp:readys_def)
-        next
-          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}" by simp
-        next
-          from finite_threads 
-          show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        qed
-      next
-        from finite_threads
-        show "finite (cp s ` readys s)" by (auto simp:readys_def)
-      next
-        from th'_in
-        show "cp s th' \<in> cp s ` readys s" by simp
-      qed
-    next
-      assume tm_ready: "tm \<in> readys s"
-      show ?thesis
-      proof(fold tm_max)
-        have cp_eq_p: "cp s tm = preced tm s"
-        proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
-          fix y 
-          assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
-          show "y \<le> preced tm s"
-          proof -
-            { fix y'
-              assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependants (wq s) tm)"
-              have "y' \<le> preced tm s"
-              proof(unfold tm_max, rule Max_ge)
-                from hy' dependants_threads[of tm]
-                show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
-              next
-                from finite_threads
-                show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-              qed
-            } with hy show ?thesis by auto
-          qed
-        next
-          from dependants_threads[of tm] finite_threads
-          show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm))"
-            by (auto intro:finite_subset)
-        next
-          show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
-            by simp
-        qed 
-        moreover have "Max (cp s ` readys s) = cp s tm"
-        proof(rule Max_eqI)
-          from tm_ready show "cp s tm \<in> cp s ` readys s" by simp
-        next
-          from finite_threads
-          show "finite (cp s ` readys s)" by (auto simp:readys_def)
-        next
-          fix y assume "y \<in> cp s ` readys s"
-          then obtain th1 where th1_readys: "th1 \<in> readys s"
-            and h: "y = cp s th1" by auto
-          show "y \<le> cp s tm"
-            apply(unfold cp_eq_p h)
-            apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono)
-          proof -
-            from finite_threads
-            show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-          next
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}"
-              by simp
-          next
-            from dependants_threads[of th1] th1_readys
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) 
-                    \<subseteq> (\<lambda>th. preced th s) ` threads s"
-              by (auto simp:readys_def)
-          qed
-        qed
-        ultimately show " Max (cp s ` readys s) = preced tm s" by simp
-      qed 
-    qed
-  qed
-qed
-
-text {* (* ccc *) \noindent
-  Since the current precedence of the threads in ready queue will always be boosted,
-  there must be one inside it has the maximum precedence of the whole system. 
-*}
-lemma max_cp_readys_threads:
-  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
-proof(cases "threads s = {}")
-  case True
-  thus ?thesis 
-    by (auto simp:readys_def)
-next
-  case False
-  show ?thesis by (rule max_cp_readys_threads_pre[OF False])
-qed
-
-end
-
-lemma eq_holding: "holding (wq s) th cs = holding s th cs"
-  apply (unfold s_holding_def cs_holding_def wq_def, simp)
-  done
-
-lemma f_image_eq:
-  assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
-  shows "f ` A = g ` A"
-proof
-  show "f ` A \<subseteq> g ` A"
-    by(rule image_subsetI, auto intro:h)
-next
-  show "g ` A \<subseteq> f ` A"
-   by (rule image_subsetI, auto intro:h[symmetric])
-qed
-
-
-definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
-  where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
-
-lemma detached_test:
-  shows "detached s th = (Th th \<notin> Field (RAG s))"
-apply(simp add: detached_def Field_def)
-apply(simp add: s_RAG_def)
-apply(simp add: s_holding_abv s_waiting_abv)
-apply(simp add: Domain_iff Range_iff)
-apply(simp add: wq_def)
-apply(auto)
-done
-
-context valid_trace
-begin
-
-lemma detached_intro:
-  assumes eq_pv: "cntP s th = cntV s th"
-  shows "detached s th"
-proof -
- from cnp_cnv_cncs
-  have eq_cnt: "cntP s th =
-    cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
-  hence cncs_zero: "cntCS s th = 0"
-    by (auto simp:eq_pv split:if_splits)
-  with eq_cnt
-  have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:eq_pv)
-  thus ?thesis
-  proof
-    assume "th \<notin> threads s"
-    with range_in dm_RAG_threads
-    show ?thesis
-      by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def Domain_iff Range_iff)
-  next
-    assume "th \<in> readys s"
-    moreover have "Th th \<notin> Range (RAG s)"
-    proof -
-      from card_0_eq [OF finite_holding] and cncs_zero
-      have "holdents s th = {}"
-        by (simp add:cntCS_def)
-      thus ?thesis
-        apply(auto simp:holdents_test)
-        apply(case_tac a)
-        apply(auto simp:holdents_test s_RAG_def)
-        done
-    qed
-    ultimately show ?thesis
-      by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def readys_def)
-  qed
-qed
-
-lemma detached_elim:
-  assumes dtc: "detached s th"
-  shows "cntP s th = cntV s th"
-proof -
-  from cnp_cnv_cncs
-  have eq_pv: " cntP s th =
-    cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
-  have cncs_z: "cntCS s th = 0"
-  proof -
-    from dtc have "holdents s th = {}"
-      unfolding detached_def holdents_test s_RAG_def
-      by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
-    thus ?thesis by (auto simp:cntCS_def)
-  qed
-  show ?thesis
-  proof(cases "th \<in> threads s")
-    case True
-    with dtc 
-    have "th \<in> readys s"
-      by (unfold readys_def detached_def Field_def Domain_def Range_def, 
-           auto simp:waiting_eq s_RAG_def)
-    with cncs_z and eq_pv show ?thesis by simp
-  next
-    case False
-    with cncs_z and eq_pv show ?thesis by simp
-  qed
-qed
-
-lemma detached_eq:
-  shows "(detached s th) = (cntP s th = cntV s th)"
-  by (insert vt, auto intro:detached_intro detached_elim)
-
-end
-
-text {* 
-  The lemmas in this .thy file are all obvious lemmas, however, they still needs to be derived
-  from the concise and miniature model of PIP given in PrioGDef.thy.
-*}
-
-lemma eq_dependants: "dependants (wq s) = dependants s"
-  by (simp add: s_dependants_abv wq_def)
-
-lemma next_th_unique: 
-  assumes nt1: "next_th s th cs th1"
-  and nt2: "next_th s th cs th2"
-  shows "th1 = th2"
-using assms by (unfold next_th_def, auto)
-
-lemma birth_time_lt:  "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
-  apply (induct s, simp)
-proof -
-  fix a s
-  assume ih: "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
-    and eq_as: "a # s \<noteq> []"
-  show "last_set th (a # s) < length (a # s)"
-  proof(cases "s \<noteq> []")
-    case False
-    from False show ?thesis
-      by (cases a, auto simp:last_set.simps)
-  next
-    case True
-    from ih [OF True] show ?thesis
-      by (cases a, auto simp:last_set.simps)
-  qed
-qed
-
-lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
-  by (induct s, auto simp:threads.simps)
-
-lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
-  apply (drule_tac th_in_ne)
-  by (unfold preced_def, auto intro: birth_time_lt)
-
-lemma inj_the_preced: 
-  "inj_on (the_preced s) (threads s)"
-  by (metis inj_onI preced_unique the_preced_def)
-
-lemma tRAG_alt_def: 
-  "tRAG s = {(Th th1, Th th2) | th1 th2. 
-                  \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
- by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
-
-lemma tRAG_Field:
-  "Field (tRAG s) \<subseteq> Field (RAG s)"
-  by (unfold tRAG_alt_def Field_def, auto)
-
-lemma tRAG_ancestorsE:
-  assumes "x \<in> ancestors (tRAG s) u"
-  obtains th where "x = Th th"
-proof -
-  from assms have "(u, x) \<in> (tRAG s)^+" 
-      by (unfold ancestors_def, auto)
-  from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
-  then obtain th where "x = Th th"
-    by (unfold tRAG_alt_def, auto)
-  from that[OF this] show ?thesis .
-qed
-
-lemma tRAG_mono:
-  assumes "RAG s' \<subseteq> RAG s"
-  shows "tRAG s' \<subseteq> tRAG s"
-  using assms 
-  by (unfold tRAG_alt_def, auto)
-
-lemma holding_next_thI:
-  assumes "holding s th cs"
-  and "length (wq s cs) > 1"
-  obtains th' where "next_th s th cs th'"
-proof -
-  from assms(1)[folded eq_holding, unfolded cs_holding_def]
-  have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" .
-  then obtain rest where h1: "wq s cs = th#rest" 
-    by (cases "wq s cs", auto)
-  with assms(2) have h2: "rest \<noteq> []" by auto
-  let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
-  have "next_th s th cs ?th'" using  h1(1) h2 
-    by (unfold next_th_def, auto)
-  from that[OF this] show ?thesis .
-qed
-
-lemma RAG_tRAG_transfer:
-  assumes "vt s'"
-  assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
-  and "(Cs cs, Th th'') \<in> RAG s'"
-  shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
-proof -
-  interpret vt_s': valid_trace "s'" using assms(1)
-    by (unfold_locales, simp)
-  interpret rtree: rtree "RAG s'"
-  proof
-  show "single_valued (RAG s')"
-  apply (intro_locales)
-    by (unfold single_valued_def, 
-        auto intro:vt_s'.unique_RAG)
-
-  show "acyclic (RAG s')"
-     by (rule vt_s'.acyclic_RAG)
-  qed
-  { fix n1 n2
-    assume "(n1, n2) \<in> ?L"
-    from this[unfolded tRAG_alt_def]
-    obtain th1 th2 cs' where 
-      h: "n1 = Th th1" "n2 = Th th2" 
-         "(Th th1, Cs cs') \<in> RAG s"
-         "(Cs cs', Th th2) \<in> RAG s" by auto
-    from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto
-    from h(3) and assms(2) 
-    have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
-          (Th th1, Cs cs') \<in> RAG s'" by auto
-    hence "(n1, n2) \<in> ?R"
-    proof
-      assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
-      hence eq_th1: "th1 = th" by simp
-      moreover have "th2 = th''"
-      proof -
-        from h1 have "cs' = cs" by simp
-        from assms(3) cs_in[unfolded this] rtree.sgv
-        show ?thesis
-          by (unfold single_valued_def, auto)
-      qed
-      ultimately show ?thesis using h(1,2) by auto
-    next
-      assume "(Th th1, Cs cs') \<in> RAG s'"
-      with cs_in have "(Th th1, Th th2) \<in> tRAG s'"
-        by (unfold tRAG_alt_def, auto)
-      from this[folded h(1, 2)] show ?thesis by auto
-    qed
-  } moreover {
-    fix n1 n2
-    assume "(n1, n2) \<in> ?R"
-    hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto
-    hence "(n1, n2) \<in> ?L" 
-    proof
-      assume "(n1, n2) \<in> tRAG s'"
-      moreover have "... \<subseteq> ?L"
-      proof(rule tRAG_mono)
-        show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto)
-      qed
-      ultimately show ?thesis by auto
-    next
-      assume eq_n: "(n1, n2) = (Th th, Th th'')"
-      from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto
-      moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto
-      ultimately show ?thesis 
-        by (unfold eq_n tRAG_alt_def, auto)
-    qed
-  } ultimately show ?thesis by auto
-qed
-
-context valid_trace
-begin
-
-lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt]
-
-end
-
-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
-
-lemma cp_gen_alt_def:
-  "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
-    by (auto simp:cp_gen_def)
-
-lemma tRAG_nodeE:
-  assumes "(n1, n2) \<in> tRAG s"
-  obtains th1 th2 where "n1 = Th th1" "n2 = Th th2"
-  using assms
-  by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def)
-
-lemma subtree_nodeE:
-  assumes "n \<in> subtree (tRAG s) (Th th)"
-  obtains th1 where "n = Th th1"
-proof -
-  show ?thesis
-  proof(rule subtreeE[OF assms])
-    assume "n = Th th"
-    from that[OF this] show ?thesis .
-  next
-    assume "Th th \<in> ancestors (tRAG s) n"
-    hence "(n, Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def)
-    hence "\<exists> th1. n = Th th1"
-    proof(induct)
-      case (base y)
-      from tRAG_nodeE[OF this] show ?case by metis
-    next
-      case (step y z)
-      thus ?case by auto
-    qed
-    with that show ?thesis by auto
-  qed
-qed
-
-lemma tRAG_star_RAG: "(tRAG s)^* \<subseteq> (RAG s)^*"
-proof -
-  have "(wRAG s O hRAG s)^* \<subseteq> (RAG s O RAG s)^*" 
-    by (rule rtrancl_mono, auto simp:RAG_split)
-  also have "... \<subseteq> ((RAG s)^*)^*"
-    by (rule rtrancl_mono, auto)
-  also have "... = (RAG s)^*" by simp
-  finally show ?thesis by (unfold tRAG_def, simp)
-qed
-
-lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x"
-proof -
-  { fix a
-    assume "a \<in> subtree (tRAG s) x"
-    hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def)
-    with tRAG_star_RAG[of s]
-    have "(a, x) \<in> (RAG s)^*" by auto
-    hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def)
-  } thus ?thesis by auto
-qed
-
-lemma tRAG_trancl_eq:
-   "{th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
-    {th'. (Th th', Th th)  \<in> (RAG s)^+}"
-   (is "?L = ?R")
-proof -
-  { fix th'
-    assume "th' \<in> ?L"
-    hence "(Th th', Th th) \<in> (tRAG s)^+" by auto
-    from tranclD[OF this]
-    obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto
-    from tRAG_subtree_RAG[of s] and this(2)
-    have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) 
-    moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto 
-    ultimately have "th' \<in> ?R"  by auto 
-  } moreover 
-  { fix th'
-    assume "th' \<in> ?R"
-    hence "(Th th', Th th) \<in> (RAG s)^+" by (auto)
-    from plus_rpath[OF this]
-    obtain xs where rp: "rpath (RAG s) (Th th') xs (Th th)" "xs \<noteq> []" by auto
-    hence "(Th th', Th th) \<in> (tRAG s)^+"
-    proof(induct xs arbitrary:th' th rule:length_induct)
-      case (1 xs th' th)
-      then obtain x1 xs1 where Cons1: "xs = x1#xs1" by (cases xs, auto)
-      show ?case
-      proof(cases "xs1")
-        case Nil
-        from 1(2)[unfolded Cons1 Nil]
-        have rp: "rpath (RAG s) (Th th') [x1] (Th th)" .
-        hence "(Th th', x1) \<in> (RAG s)" by (cases, simp)
-        then obtain cs where "x1 = Cs cs" 
-              by (unfold s_RAG_def, auto)
-        from rpath_nnl_lastE[OF rp[unfolded this]]
-        show ?thesis by auto
-      next
-        case (Cons x2 xs2)
-        from 1(2)[unfolded Cons1[unfolded this]]
-        have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" .
-        from rpath_edges_on[OF this]
-        have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" .
-        have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)"
-            by (simp add: edges_on_unfold)
-        with eds have rg1: "(Th th', x1) \<in> RAG s" by auto
-        then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto)
-        have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)"
-            by (simp add: edges_on_unfold)
-        from this eds
-        have rg2: "(x1, x2) \<in> RAG s" by auto
-        from this[unfolded eq_x1] 
-        obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto)
-        from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2]
-        have rt1: "(Th th', Th th1) \<in> tRAG s" by (unfold tRAG_alt_def, auto)
-        from rp have "rpath (RAG s) x2 xs2 (Th th)"
-           by  (elim rpath_ConsE, simp)
-        from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" .
-        show ?thesis
-        proof(cases "xs2 = []")
-          case True
-          from rpath_nilE[OF rp'[unfolded this]]
-          have "th1 = th" by auto
-          from rt1[unfolded this] show ?thesis by auto
-        next
-          case False
-          from 1(1)[rule_format, OF _ rp' this, unfolded Cons1 Cons]
-          have "(Th th1, Th th) \<in> (tRAG s)\<^sup>+" by simp
-          with rt1 show ?thesis by auto
-        qed
-      qed
-    qed
-    hence "th' \<in> ?L" by auto
-  } ultimately show ?thesis by blast
-qed
-
-lemma tRAG_trancl_eq_Th:
-   "{Th th' | th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
-    {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}"
-    using tRAG_trancl_eq by auto
-
-lemma dependants_alt_def:
-  "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
-  by (metis eq_RAG s_dependants_def tRAG_trancl_eq)
-  
-context valid_trace
-begin
-
-lemma count_eq_tRAG_plus:
-  assumes "cntP s th = cntV s th"
-  shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
-  using assms count_eq_dependants dependants_alt_def eq_dependants by auto 
-
-lemma count_eq_RAG_plus:
-  assumes "cntP s th = cntV s th"
-  shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
-  using assms count_eq_dependants cs_dependants_def eq_RAG by auto
-
-lemma count_eq_RAG_plus_Th:
-  assumes "cntP s th = cntV s th"
-  shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
-  using count_eq_RAG_plus[OF assms] by auto
-
-lemma count_eq_tRAG_plus_Th:
-  assumes "cntP s th = cntV s th"
-  shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
-   using count_eq_tRAG_plus[OF assms] by auto
-
-end
-
-lemma tRAG_subtree_eq: 
-   "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
-   (is "?L = ?R")
-proof -
-  { fix n 
-    assume h: "n \<in> ?L"
-    hence "n \<in> ?R"
-    by (smt mem_Collect_eq subsetCE subtree_def subtree_nodeE tRAG_subtree_RAG) 
-  } moreover {
-    fix n
-    assume "n \<in> ?R"
-    then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*"
-      by (auto simp:subtree_def)
-    from rtranclD[OF this(2)]
-    have "n \<in> ?L"
-    proof
-      assume "Th th' \<noteq> Th th \<and> (Th th', Th th) \<in> (RAG s)\<^sup>+"
-      with h have "n \<in> {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}" by auto
-      thus ?thesis using subtree_def tRAG_trancl_eq by fastforce
-    qed (insert h, auto simp:subtree_def)
-  } ultimately show ?thesis by auto
-qed
-
-lemma threads_set_eq: 
-   "the_thread ` (subtree (tRAG s) (Th th)) = 
-                  {th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R")
-   by (auto intro:rev_image_eqI simp:tRAG_subtree_eq)
-
-lemma cp_alt_def1: 
-  "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))"
-proof -
-  have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) =
-       ((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))"
-       by auto
-  thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto)
-qed
-
-lemma cp_gen_def_cond: 
-  assumes "x = Th th"
-  shows "cp s th = cp_gen s (Th th)"
-by (unfold cp_alt_def1 cp_gen_def, simp)
-
-lemma cp_gen_over_set:
-  assumes "\<forall> x \<in> A. \<exists> th. x = Th th"
-  shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A"
-proof(rule f_image_eq)
-  fix a
-  assume "a \<in> A"
-  from assms[rule_format, OF this]
-  obtain th where eq_a: "a = Th th" by auto
-  show "cp_gen s a = (cp s \<circ> the_thread) a"
-    by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
-qed
-
-
-context valid_trace
-begin
-
-lemma RAG_threads:
-  assumes "(Th th) \<in> Field (RAG s)"
-  shows "th \<in> threads s"
-  using assms
-  by (metis Field_def UnE dm_RAG_threads range_in vt)
-
-lemma subtree_tRAG_thread:
-  assumes "th \<in> threads s"
-  shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R")
-proof -
-  have "?L = {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
-    by (unfold tRAG_subtree_eq, simp)
-  also have "... \<subseteq> ?R"
-  proof
-    fix x
-    assume "x \<in> {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
-    then obtain th' where h: "x = Th th'" "Th th' \<in> subtree (RAG s) (Th th)" by auto
-    from this(2)
-    show "x \<in> ?R"
-    proof(cases rule:subtreeE)
-      case 1
-      thus ?thesis by (simp add: assms h(1)) 
-    next
-      case 2
-      thus ?thesis by (metis ancestors_Field dm_RAG_threads h(1) image_eqI) 
-    qed
-  qed
-  finally show ?thesis .
-qed
-
-lemma readys_root:
-  assumes "th \<in> readys s"
-  shows "root (RAG s) (Th th)"
-proof -
-  { fix x
-    assume "x \<in> ancestors (RAG s) (Th th)"
-    hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def)
-    from tranclD[OF this]
-    obtain z where "(Th th, z) \<in> RAG s" by auto
-    with assms(1) have False
-         apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def)
-         by (fold wq_def, blast)
-  } thus ?thesis by (unfold root_def, auto)
-qed
-
-lemma readys_in_no_subtree:
-  assumes "th \<in> readys s"
-  and "th' \<noteq> th"
-  shows "Th th \<notin> subtree (RAG s) (Th th')" 
-proof
-   assume "Th th \<in> subtree (RAG s) (Th th')"
-   thus False
-   proof(cases rule:subtreeE)
-      case 1
-      with assms show ?thesis by auto
-   next
-      case 2
-      with readys_root[OF assms(1)]
-      show ?thesis by (auto simp:root_def)
-   qed
-qed
-
-lemma not_in_thread_isolated:
-  assumes "th \<notin> threads s"
-  shows "(Th th) \<notin> Field (RAG s)"
-proof
-  assume "(Th th) \<in> Field (RAG s)"
-  with dm_RAG_threads and range_in assms
-  show False by (unfold Field_def, blast)
-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 holding_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 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
-
-sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
-proof -
-  have "fsubtree (tRAG s)"
-  proof -
-    have "fbranch (tRAG s)"
-    proof(unfold tRAG_def, rule fbranch_compose)
-        show "fbranch (wRAG s)"
-        proof(rule finite_fbranchI)
-           from finite_RAG show "finite (wRAG s)"
-           by (unfold RAG_split, auto)
-        qed
-    next
-        show "fbranch (hRAG s)"
-        proof(rule finite_fbranchI)
-           from finite_RAG 
-           show "finite (hRAG s)" by (unfold RAG_split, auto)
-        qed
-    qed
-    moreover have "wf (tRAG s)"
-    proof(rule wf_subset)
-      show "wf (RAG s O RAG s)" using wf_RAG
-        by (fold wf_comp_self, simp)
-    next
-      show "tRAG s \<subseteq> (RAG s O RAG s)"
-        by (unfold tRAG_alt_def, auto)
-    qed
-    ultimately show ?thesis
-      by (unfold fsubtree_def fsubtree_axioms_def,auto)
-  qed
-  from this[folded tRAG_def] show "fsubtree (tRAG s)" .
-qed
-
-lemma Max_UNION: 
-  assumes "finite A"
-  and "A \<noteq> {}"
-  and "\<forall> M \<in> f ` A. finite M"
-  and "\<forall> M \<in> f ` A. M \<noteq> {}"
-  shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R")
-  using assms[simp]
-proof -
-  have "?L = Max (\<Union>(f ` A))"
-    by (fold Union_image_eq, simp)
-  also have "... = ?R"
-    by (subst Max_Union, simp+)
-  finally show ?thesis .
-qed
-
-lemma max_Max_eq:
-  assumes "finite A"
-    and "A \<noteq> {}"
-    and "x = y"
-  shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R")
-proof -
-  have "?R = Max (insert y A)" by simp
-  also from assms have "... = ?L"
-      by (subst Max.insert, simp+)
-  finally show ?thesis by simp
-qed
-
-context valid_trace
-begin
-
-(* ddd *)
-lemma cp_gen_rec:
-  assumes "x = Th th"
-  shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
-proof(cases "children (tRAG s) x = {}")
-  case True
-  show ?thesis
-    by (unfold True cp_gen_def subtree_children, simp add:assms)
-next
-  case False
-  hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
-  note fsbttRAGs.finite_subtree[simp]
-  have [simp]: "finite (children (tRAG s) x)"
-     by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], 
-            rule children_subtree)
-  { fix r x
-    have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
-  } note this[simp]
-  have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
-  proof -
-    from False obtain q where "q \<in> children (tRAG s) x" by blast
-    moreover have "subtree (tRAG s) q \<noteq> {}" by simp
-    ultimately show ?thesis by blast
-  qed
-  have h: "Max ((the_preced s \<circ> the_thread) `
-                ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
-        Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
-                     (is "?L = ?R")
-  proof -
-    let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L
-    let "Max (_ \<union> (?h ` ?B))" = ?R
-    let ?L1 = "?f ` \<Union>(?g ` ?B)"
-    have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)"
-    proof -
-      have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp
-      also have "... =  (\<Union> x \<in> ?B. ?f ` (?g x))" by auto
-      finally have "Max ?L1 = Max ..." by simp
-      also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)"
-        by (subst Max_UNION, simp+)
-      also have "... = Max (cp_gen s ` children (tRAG s) x)"
-          by (unfold image_comp cp_gen_alt_def, simp)
-      finally show ?thesis .
-    qed
-    show ?thesis
-    proof -
-      have "?L = Max (?f ` ?A \<union> ?L1)" by simp
-      also have "... = max (the_preced s (the_thread x)) (Max ?L1)"
-            by (subst Max_Un, simp+)
-      also have "... = max (?f x) (Max (?h ` ?B))"
-        by (unfold eq_Max_L1, simp)
-      also have "... =?R"
-        by (rule max_Max_eq, (simp)+, unfold assms, simp)
-      finally show ?thesis .
-    qed
-  qed  thus ?thesis 
-          by (fold h subtree_children, unfold cp_gen_def, simp) 
-qed
-
-lemma cp_rec:
-  "cp s th = Max ({the_preced s th} \<union> 
-                     (cp s o the_thread) ` children (tRAG s) (Th th))"
-proof -
-  have "Th th = Th th" by simp
-  note h =  cp_gen_def_cond[OF this] cp_gen_rec[OF this]
-  show ?thesis 
-  proof -
-    have "cp_gen s ` children (tRAG s) (Th th) = 
-                (cp s \<circ> the_thread) ` children (tRAG s) (Th th)"
-    proof(rule cp_gen_over_set)
-      show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th"
-        by (unfold tRAG_alt_def, auto simp:children_def)
-    qed
-    thus ?thesis by (subst (1) h(1), unfold h(2), simp)
-  qed
-qed
-
-end
-
-(* keep *)
-lemma next_th_holding:
-  assumes vt: "vt s"
-  and nxt: "next_th s th cs th'"
-  shows "holding (wq s) th cs"
-proof -
-  from nxt[unfolded next_th_def]
-  obtain rest where h: "wq s cs = th # rest"
-                       "rest \<noteq> []" 
-                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
-  thus ?thesis
-    by (unfold cs_holding_def, auto)
-qed
-
-context valid_trace
-begin
-
-lemma next_th_waiting:
-  assumes nxt: "next_th s th cs th'"
-  shows "waiting (wq s) th' cs"
-proof -
-  from nxt[unfolded next_th_def]
-  obtain rest where h: "wq s cs = th # rest"
-                       "rest \<noteq> []" 
-                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
-  from wq_distinct[of cs, unfolded h]
-  have dst: "distinct (th # rest)" .
-  have in_rest: "th' \<in> set rest"
-  proof(unfold h, rule someI2)
-    show "distinct rest \<and> set rest = set rest" using dst by auto
-  next
-    fix x assume "distinct x \<and> set x = set rest"
-    with h(2)
-    show "hd x \<in> set (rest)" by (cases x, auto)
-  qed
-  hence "th' \<in> set (wq s cs)" by (unfold h(1), auto)
-  moreover have "th' \<noteq> hd (wq s cs)"
-    by (unfold h(1), insert in_rest dst, auto)
-  ultimately show ?thesis by (auto simp:cs_waiting_def)
-qed
-
-lemma next_th_RAG:
-  assumes nxt: "next_th (s::event list) th cs th'"
-  shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s"
-  using vt assms next_th_holding next_th_waiting
-  by (unfold s_RAG_def, simp)
-
-end
-
--- {* A useless definition *}
-definition cps:: "state \<Rightarrow> (thread \<times> precedence) set"
-where "cps s = {(th, cp s th) | th . th \<in> threads s}"
-
-lemma "wq (V th cs # s) cs1 = ttt"
-  apply (unfold wq_def, auto simp:Let_def)
-
-end
-
--- a/CpsG.thy~	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,4656 +0,0 @@
-theory CpsG
-imports PIPDefs
-begin
-
-lemma f_image_eq:
-  assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
-  shows "f ` A = g ` A"
-proof
-  show "f ` A \<subseteq> g ` A"
-    by(rule image_subsetI, auto intro:h)
-next
-  show "g ` A \<subseteq> f ` A"
-   by (rule image_subsetI, auto intro:h[symmetric])
-qed
-
-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> {}"
-  and fnt: "finite B"
-  shows "Max (f ` A) \<le> Max (f ` B)"
-proof(rule Max_mono)
-  from seq show "f ` A \<subseteq> f ` B" by auto
-next
-  from np show "f ` A \<noteq> {}" by auto
-next
-  from fnt and seq show "finite (f ` B)" by auto
-qed
-
-lemma Max_UNION: 
-  assumes "finite A"
-  and "A \<noteq> {}"
-  and "\<forall> M \<in> f ` A. finite M"
-  and "\<forall> M \<in> f ` A. M \<noteq> {}"
-  shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R")
-  using assms[simp]
-proof -
-  have "?L = Max (\<Union>(f ` A))"
-    by (fold Union_image_eq, simp)
-  also have "... = ?R"
-    by (subst Max_Union, simp+)
-  finally show ?thesis .
-qed
-
-lemma max_Max_eq:
-  assumes "finite A"
-    and "A \<noteq> {}"
-    and "x = y"
-  shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R")
-proof -
-  have "?R = Max (insert y A)" by simp
-  also from assms have "... = ?L"
-      by (subst Max.insert, simp+)
-  finally show ?thesis by simp
-qed
-
-lemma birth_time_lt:  
-  assumes "s \<noteq> []"
-  shows "last_set th s < length s"
-  using assms
-proof(induct s)
-  case (Cons a s)
-  show ?case
-  proof(cases "s \<noteq> []")
-    case False
-    thus ?thesis
-      by (cases a, auto)
-  next
-    case True
-    show ?thesis using Cons(1)[OF True]
-      by (cases a, auto)
-  qed
-qed simp
-
-lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
-  by (induct s, auto)
-
-lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
-  by (drule_tac th_in_ne, unfold preced_def, auto intro: birth_time_lt)
-
-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
-  assumes vt : "vt s"
-
-locale valid_trace_e = valid_trace +
-  fixes e
-  assumes vt_e: "vt (e#s)"
-begin
-
-lemma pip_e: "PIP s e"
-  using vt_e by (cases, simp)  
-
-end
-
-locale valid_trace_create = valid_trace_e + 
-  fixes th prio
-  assumes is_create: "e = Create th prio"
-
-locale valid_trace_exit = valid_trace_e + 
-  fixes th
-  assumes is_exit: "e = Exit th"
-
-locale valid_trace_p = valid_trace_e + 
-  fixes th cs
-  assumes is_p: "e = P th cs"
-
-locale valid_trace_v = valid_trace_e + 
-  fixes th cs
-  assumes is_v: "e = V th cs"
-begin
-  definition "rest = tl (wq s cs)"
-  definition "wq' = (SOME q. distinct q \<and> set q = set rest)"
-end
-
-locale valid_trace_v_n = valid_trace_v +
-  assumes rest_nnl: "rest \<noteq> []"
-
-locale valid_trace_v_e = valid_trace_v +
-  assumes rest_nil: "rest = []"
-
-locale valid_trace_set= valid_trace_e + 
-  fixes th prio
-  assumes is_set: "e = Set th prio"
-
-context valid_trace
-begin
-
-lemma ind [consumes 0, case_names Nil Cons, induct type]:
-  assumes "PP []"
-     and "(\<And>s e. valid_trace_e s e \<Longrightarrow>
-                   PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))"
-     shows "PP s"
-proof(induct rule:vt.induct[OF vt, case_names Init Step])
-  case Init
-  from assms(1) show ?case .
-next
-  case (Step s e)
-  show ?case
-  proof(rule assms(2))
-    show "valid_trace_e s e" using Step by (unfold_locales, auto)
-  next
-    show "PP s" using Step by simp
-  next
-    show "PIP s e" using Step by simp
-  qed
-qed
-
-lemma  vt_moment: "\<And> t. vt (moment t s)"
-proof(induct rule:ind)
-  case Nil
-  thus ?case by (simp add:vt_nil)
-next
-  case (Cons s e t)
-  show ?case
-  proof(cases "t \<ge> length (e#s)")
-    case True
-    from True have "moment t (e#s) = e#s" by simp
-    thus ?thesis using Cons
-      by (simp add:valid_trace_def valid_trace_e_def, auto)
-  next
-    case False
-    from Cons have "vt (moment t s)" by simp
-    moreover have "moment t (e#s) = moment t s"
-    proof -
-      from False have "t \<le> length s" by simp
-      from moment_app [OF this, of "[e]"] 
-      show ?thesis by simp
-    qed
-    ultimately show ?thesis by simp
-  qed
-qed
-
-lemma finite_threads:
-  shows "finite (threads s)"
-using vt by (induct) (auto elim: step.cases)
-
-end
-
-lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
-  by (unfold s_RAG_def, auto)
-
-locale valid_moment = valid_trace + 
-  fixes i :: nat
-
-sublocale valid_moment < vat_moment: valid_trace "(moment i s)"
-  by (unfold_locales, insert vt_moment, auto)
-
-lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
-  by  (unfold s_waiting_def cs_waiting_def wq_def, auto)
-
-lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
-  by (unfold s_holding_def wq_def cs_holding_def, simp)
-
-lemma runing_ready: 
-  shows "runing s \<subseteq> readys s"
-  unfolding runing_def readys_def
-  by auto 
-
-lemma readys_threads:
-  shows "readys s \<subseteq> threads s"
-  unfolding readys_def
-  by auto
-
-lemma wq_v_neq [simp]:
-   "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
-  by (auto simp:wq_def Let_def cp_def split:list.splits)
-
-lemma runing_head:
-  assumes "th \<in> runing s"
-  and "th \<in> set (wq_fun (schs s) cs)"
-  shows "th = hd (wq_fun (schs s) cs)"
-  using assms
-  by (simp add:runing_def readys_def s_waiting_def wq_def)
-
-context valid_trace
-begin
-
-lemma runing_wqE:
-  assumes "th \<in> runing s"
-  and "th \<in> set (wq s cs)"
-  obtains rest where "wq s cs = th#rest"
-proof -
-  from assms(2) obtain th' rest where eq_wq: "wq s cs = th'#rest"
-    by (meson list.set_cases)
-  have "th' = th"
-  proof(rule ccontr)
-    assume "th' \<noteq> th"
-    hence "th \<noteq> hd (wq s cs)" using eq_wq by auto 
-    with assms(2)
-    have "waiting s th cs" 
-      by (unfold s_waiting_def, fold wq_def, auto)
-    with assms show False 
-      by (unfold runing_def readys_def, auto)
-  qed
-  with eq_wq that show ?thesis by metis
-qed
-
-end
-
-context valid_trace_create
-begin
-
-lemma wq_neq_simp [simp]:
-  shows "wq (e#s) cs' = wq s cs'"
-    using assms unfolding is_create wq_def
-  by (auto simp:Let_def)
-
-lemma wq_distinct_kept:
-  assumes "distinct (wq s cs')"
-  shows "distinct (wq (e#s) cs')"
-  using assms by simp
-end
-
-context valid_trace_exit
-begin
-
-lemma wq_neq_simp [simp]:
-  shows "wq (e#s) cs' = wq s cs'"
-    using assms unfolding is_exit wq_def
-  by (auto simp:Let_def)
-
-lemma wq_distinct_kept:
-  assumes "distinct (wq s cs')"
-  shows "distinct (wq (e#s) cs')"
-  using assms by simp
-end
-
-context valid_trace_p
-begin
-
-lemma wq_neq_simp [simp]:
-  assumes "cs' \<noteq> cs"
-  shows "wq (e#s) cs' = wq s cs'"
-    using assms unfolding is_p wq_def
-  by (auto simp:Let_def)
-
-lemma runing_th_s:
-  shows "th \<in> runing s"
-proof -
-  from pip_e[unfolded is_p]
-  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 -
-  have "th \<in> readys s"
-    using runing_ready runing_th_s by blast 
-  thus ?thesis
-    by (unfold readys_def, auto)
-qed
-
-lemma waiting_neq_th: 
-  assumes "waiting s t c"
-  shows "t \<noteq> th"
-  using assms using th_not_waiting by blast 
-
-lemma th_not_in_wq: 
-  shows "th \<notin> set (wq s cs)"
-proof
-  assume otherwise: "th \<in> set (wq s cs)"
-  from runing_wqE[OF runing_th_s this]
-  obtain rest where eq_wq: "wq s cs = th#rest" by blast
-  with otherwise
-  have "holding s th cs"
-    by (unfold s_holding_def, fold wq_def, simp)
-  hence cs_th_RAG: "(Cs cs, Th th) \<in> RAG s"
-    by (unfold s_RAG_def, fold holding_eq, auto)
-  from pip_e[unfolded is_p]
-  show False
-  proof(cases)
-    case (thread_P)
-    with cs_th_RAG show ?thesis by auto
-  qed
-qed
-
-lemma wq_es_cs: 
-  "wq (e#s) cs =  wq s cs @ [th]"
-  by (unfold is_p wq_def, auto simp:Let_def)
-
-lemma wq_distinct_kept:
-  assumes "distinct (wq s cs')"
-  shows "distinct (wq (e#s) cs')"
-proof(cases "cs' = cs")
-  case True
-  show ?thesis using True assms th_not_in_wq
-    by (unfold True wq_es_cs, auto)
-qed (insert assms, simp)
-
-end
-
-context valid_trace_v
-begin
-
-lemma wq_neq_simp [simp]:
-  assumes "cs' \<noteq> cs"
-  shows "wq (e#s) cs' = wq s cs'"
-    using assms unfolding is_v wq_def
-  by (auto simp:Let_def)
-
-lemma runing_th_s:
-  shows "th \<in> runing s"
-proof -
-  from pip_e[unfolded is_v]
-  show ?thesis by (cases, simp)
-qed
-
-lemma th_not_waiting: 
-  "\<not> waiting s th c"
-proof -
-  have "th \<in> readys s"
-    using runing_ready runing_th_s by blast 
-  thus ?thesis
-    by (unfold readys_def, auto)
-qed
-
-lemma waiting_neq_th: 
-  assumes "waiting s t c"
-  shows "t \<noteq> th"
-  using assms using th_not_waiting by blast 
-
-lemma wq_s_cs:
-  "wq s cs = th#rest"
-proof -
-  from pip_e[unfolded is_v]
-  show ?thesis
-  proof(cases)
-    case (thread_V)
-    from this(2) show ?thesis
-      by (unfold rest_def s_holding_def, fold wq_def,
-                 metis empty_iff list.collapse list.set(1))
-  qed
-qed
-
-lemma wq_es_cs:
-  "wq (e#s) cs = wq'"
- using wq_s_cs[unfolded wq_def]
- by (auto simp:Let_def wq_def rest_def wq'_def is_v, simp) 
-
-lemma wq_distinct_kept:
-  assumes "distinct (wq s cs')"
-  shows "distinct (wq (e#s) cs')"
-proof(cases "cs' = cs")
-  case True
-  show ?thesis
-  proof(unfold True wq_es_cs wq'_def, rule someI2)
-    show "distinct rest \<and> set rest = set rest"
-        using assms[unfolded True wq_s_cs] by auto
-  qed simp
-qed (insert assms, simp)
-
-end
-
-context valid_trace_set
-begin
-
-lemma wq_neq_simp [simp]:
-  shows "wq (e#s) cs' = wq s cs'"
-    using assms unfolding is_set wq_def
-  by (auto simp:Let_def)
-
-lemma wq_distinct_kept:
-  assumes "distinct (wq s cs')"
-  shows "distinct (wq (e#s) cs')"
-  using assms by simp
-end
-
-context valid_trace
-begin
-
-lemma actor_inv: 
-  assumes "PIP s e"
-  and "\<not> isCreate e"
-  shows "actor e \<in> runing s"
-  using assms
-  by (induct, auto)
-
-lemma isP_E:
-  assumes "isP e"
-  obtains cs where "e = P (actor e) cs"
-  using assms by (cases e, auto)
-
-lemma isV_E:
-  assumes "isV e"
-  obtains cs where "e = V (actor e) cs"
-  using assms by (cases e, auto) 
-
-lemma wq_distinct: "distinct (wq s cs)"
-proof(induct rule:ind)
-  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.wq_distinct_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.wq_distinct_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.wq_distinct_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.wq_distinct_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.wq_distinct_kept) 
-  qed
-qed (unfold wq_def Let_def, simp)
-
-end
-
-context valid_trace_e
-begin
-
-text {*
-  The following lemma shows that only the @{text "P"}
-  operation can add new thread into waiting queues. 
-  Such kind of lemmas are very obvious, but need to be checked formally.
-  This is a kind of confirmation that our modelling is correct.
-*}
-
-lemma wq_in_inv: 
-  assumes s_ni: "thread \<notin> set (wq s cs)"
-  and s_i: "thread \<in> set (wq (e#s) cs)"
-  shows "e = P thread cs"
-proof(cases e)
-  -- {* This is the only non-trivial case: *}
-  case (V th cs1)
-  have False
-  proof(cases "cs1 = cs")
-    case True
-    show ?thesis
-    proof(cases "(wq s cs1)")
-      case (Cons w_hd w_tl)
-      have "set (wq (e#s) cs) \<subseteq> set (wq s cs)"
-      proof -
-        have "(wq (e#s) cs) = (SOME q. distinct q \<and> set q = set w_tl)"
-          using  Cons V by (auto simp:wq_def Let_def True split:if_splits)
-        moreover have "set ... \<subseteq> set (wq s cs)"
-        proof(rule someI2)
-          show "distinct w_tl \<and> set w_tl = set w_tl"
-            by (metis distinct.simps(2) local.Cons wq_distinct)
-        qed (insert Cons True, auto)
-        ultimately show ?thesis by simp
-      qed
-      with assms show ?thesis by auto
-    qed (insert assms V True, auto simp:wq_def Let_def split:if_splits)
-  qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
-  thus ?thesis by auto
-qed (insert assms, auto simp:wq_def Let_def split:if_splits)
-
-lemma wq_out_inv: 
-  assumes s_in: "thread \<in> set (wq s cs)"
-  and s_hd: "thread = hd (wq s cs)"
-  and s_i: "thread \<noteq> hd (wq (e#s) cs)"
-  shows "e = V thread cs"
-proof(cases e)
--- {* There are only two non-trivial cases: *}
-  case (V th cs1)
-  show ?thesis
-  proof(cases "cs1 = cs")
-    case True
-    have "PIP s (V th cs)" using pip_e[unfolded V[unfolded True]] .
-    thus ?thesis
-    proof(cases)
-      case (thread_V)
-      moreover have "th = thread" using thread_V(2) s_hd
-          by (unfold s_holding_def wq_def, simp)
-      ultimately show ?thesis using V True by simp
-    qed
-  qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
-next
-  case (P th cs1)
-  show ?thesis
-  proof(cases "cs1 = cs")
-    case True
-    with P have "wq (e#s) cs = wq_fun (schs s) cs @ [th]"
-      by (auto simp:wq_def Let_def split:if_splits)
-    with s_i s_hd s_in have False
-      by (metis empty_iff hd_append2 list.set(1) wq_def) 
-    thus ?thesis by simp
-  qed (insert assms P, auto simp:wq_def Let_def split:if_splits)
-qed (insert assms, auto simp:wq_def Let_def split:if_splits)
-
-end
-
-
-context valid_trace
-begin
-
-
-text {* (* ddd *)
-  The nature of the work is like this: since it starts from a very simple and basic 
-  model, even intuitively very `basic` and `obvious` properties need to derived from scratch.
-  For instance, the fact 
-  that one thread can not be blocked by two critical resources at the same time
-  is obvious, because only running threads can make new requests, if one is waiting for 
-  a critical resource and get blocked, it can not make another resource request and get 
-  blocked the second time (because it is not running). 
-
-  To derive this fact, one needs to prove by contraction and 
-  reason about time (or @{text "moement"}). The reasoning is based on a generic theorem
-  named @{text "p_split"}, which is about status changing along the time axis. It says if 
-  a condition @{text "Q"} is @{text "True"} at a state @{text "s"},
-  but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} 
-  in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history 
-  of events leading to it), such that @{text "Q"} switched 
-  from being @{text "False"} to @{text "True"} and kept being @{text "True"}
-  till the last moment of @{text "s"}.
-
-  Suppose a thread @{text "th"} is blocked
-  on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, 
-  since no thread is blocked at the very beginning, by applying 
-  @{text "p_split"} to these two blocking facts, there exist 
-  two moments @{text "t1"} and @{text "t2"}  in @{text "s"}, such that 
-  @{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} 
-  and kept on blocked on them respectively ever since.
- 
-  Without lost of generality, we assume @{text "t1"} is earlier than @{text "t2"}.
-  However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
-  in blocked state at moment @{text "t2"} and could not
-  make any request and get blocked the second time: Contradiction.
-*}
-
-lemma waiting_unique_pre: (* ddd *)
-  assumes h11: "thread \<in> set (wq s cs1)"
-  and h12: "thread \<noteq> hd (wq s cs1)"
-  assumes h21: "thread \<in> set (wq s cs2)"
-  and h22: "thread \<noteq> hd (wq s cs2)"
-  and neq12: "cs1 \<noteq> cs2"
-  shows "False"
-proof -
-  let "?Q" = "\<lambda> cs s. thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
-  from h11 and h12 have q1: "?Q cs1 s" by simp
-  from h21 and h22 have q2: "?Q cs2 s" by simp
-  have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
-  have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
-  from p_split [of "?Q cs1", OF q1 nq1]
-  obtain t1 where lt1: "t1 < length s"
-    and np1: "\<not> ?Q cs1 (moment t1 s)"
-    and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))" by auto
-  from p_split [of "?Q cs2", OF q2 nq2]
-  obtain t2 where lt2: "t2 < length s"
-    and np2: "\<not> ?Q cs2 (moment t2 s)"
-    and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))" by auto
-  { fix s cs
-    assume q: "?Q cs s"
-    have "thread \<notin> runing s"
-    proof
-      assume "thread \<in> runing s"
-      hence " \<forall>cs. \<not> (thread \<in> set (wq_fun (schs s) cs) \<and> 
-                 thread \<noteq> hd (wq_fun (schs s) cs))"
-        by (unfold runing_def s_waiting_def readys_def, auto)
-      from this[rule_format, of cs] q 
-      show False by (simp add: wq_def) 
-    qed
-  } note q_not_runing = this
-  { fix t1 t2 cs1 cs2
-    assume  lt1: "t1 < length s"
-    and np1: "\<not> ?Q cs1 (moment t1 s)"
-    and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))"
-    and lt2: "t2 < length s"
-    and np2: "\<not> ?Q cs2 (moment t2 s)"
-    and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))"
-    and lt12: "t1 < t2"
-    let ?t3 = "Suc t2"
-    from lt2 have le_t3: "?t3 \<le> length s" by auto
-    from moment_plus [OF this] 
-    obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
-    have "t2 < ?t3" by simp
-    from nn2 [rule_format, OF this] and eq_m
-    have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
-         h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
-    have "vt (e#moment t2 s)"
-    proof -
-      from vt_moment 
-      have "vt (moment ?t3 s)" .
-      with eq_m show ?thesis by simp
-    qed
-    then interpret vt_e: valid_trace_e "moment t2 s" "e"
-        by (unfold_locales, auto, cases, simp)
-    have ?thesis
-    proof -
-      have "thread \<in> runing (moment t2 s)"
-      proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
-        case True
-        have "e = V thread cs2"
-        proof -
-          have eq_th: "thread = hd (wq (moment t2 s) cs2)" 
-              using True and np2  by auto 
-          from vt_e.wq_out_inv[OF True this h2]
-          show ?thesis .
-        qed
-        thus ?thesis using vt_e.actor_inv[OF vt_e.pip_e] by auto
-      next
-        case False
-        have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] .
-        with vt_e.actor_inv[OF vt_e.pip_e]
-        show ?thesis by auto
-      qed
-      moreover have "thread \<notin> runing (moment t2 s)"
-        by (rule q_not_runing[OF nn1[rule_format, OF lt12]])
-      ultimately show ?thesis by simp
-    qed
-  } note lt_case = this
-  show ?thesis
-  proof -
-    { assume "t1 < t2"
-      from lt_case[OF lt1 np1 nn1 lt2 np2 nn2 this]
-      have ?thesis .
-    } moreover {
-      assume "t2 < t1"
-      from lt_case[OF lt2 np2 nn2 lt1 np1 nn1 this]
-      have ?thesis .
-    } moreover {
-      assume eq_12: "t1 = t2"
-      let ?t3 = "Suc t2"
-      from lt2 have le_t3: "?t3 \<le> length s" by auto
-      from moment_plus [OF this] 
-      obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
-      have lt_2: "t2 < ?t3" by simp
-      from nn2 [rule_format, OF this] and eq_m
-      have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
-           h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
-      from nn1[rule_format, OF lt_2[folded eq_12]] eq_m[folded eq_12]
-      have g1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
-           g2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
-      have "vt (e#moment t2 s)"
-      proof -
-        from vt_moment 
-        have "vt (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      then interpret vt_e: valid_trace_e "moment t2 s" "e"
-          by (unfold_locales, auto, cases, simp)
-      have "e = V thread cs2 \<or> e = P thread cs2"
-      proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
-        case True
-        have "e = V thread cs2"
-        proof -
-          have eq_th: "thread = hd (wq (moment t2 s) cs2)" 
-              using True and np2  by auto 
-          from vt_e.wq_out_inv[OF True this h2]
-          show ?thesis .
-        qed
-        thus ?thesis by auto
-      next
-        case False
-        have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] .
-        thus ?thesis by auto
-      qed
-      moreover have "e = V thread cs1 \<or> e = P thread cs1"
-      proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
-        case True
-        have eq_th: "thread = hd (wq (moment t1 s) cs1)" 
-              using True and np1  by auto 
-        from vt_e.wq_out_inv[folded eq_12, OF True this g2]
-        have "e = V thread cs1" .
-        thus ?thesis by auto
-      next
-        case False
-        have "e = P thread cs1" using vt_e.wq_in_inv[folded eq_12, OF False g1] .
-        thus ?thesis by auto
-      qed
-      ultimately have ?thesis using neq12 by auto
-    } ultimately show ?thesis using nat_neq_iff by blast 
-  qed
-qed
-
-text {*
-  This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
-*}
-
-lemma waiting_unique:
-  assumes "waiting s th cs1"
-  and "waiting s th cs2"
-  shows "cs1 = cs2"
-  using waiting_unique_pre assms
-  unfolding wq_def s_waiting_def
-  by auto
-
-end
-
-(* not used *)
-text {*
-  Every thread can only be blocked on one critical resource, 
-  symmetrically, every critical resource can only be held by one thread. 
-  This fact is much more easier according to our definition. 
-*}
-lemma held_unique:
-  assumes "holding (s::event list) th1 cs"
-  and "holding s th2 cs"
-  shows "th1 = th2"
- by (insert assms, unfold s_holding_def, auto)
-
-lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s"
-  apply (induct s, auto)
-  by (case_tac a, auto split:if_splits)
-
-lemma last_set_unique: 
-  "\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk>
-          \<Longrightarrow> th1 = th2"
-  apply (induct s, auto)
-  by (case_tac a, auto split:if_splits dest:last_set_lt)
-
-lemma preced_unique : 
-  assumes pcd_eq: "preced th1 s = preced th2 s"
-  and th_in1: "th1 \<in> threads s"
-  and th_in2: " th2 \<in> threads s"
-  shows "th1 = th2"
-proof -
-  from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def)
-  from last_set_unique [OF this th_in1 th_in2]
-  show ?thesis .
-qed
-                      
-lemma preced_linorder: 
-  assumes neq_12: "th1 \<noteq> th2"
-  and th_in1: "th1 \<in> threads s"
-  and th_in2: " th2 \<in> threads s"
-  shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
-proof -
-  from preced_unique [OF _ th_in1 th_in2] and neq_12 
-  have "preced th1 s \<noteq> preced th2 s" by auto
-  thus ?thesis by auto
-qed
-
-text {*
-  The following three lemmas show that @{text "RAG"} does not change
-  by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
-  events, respectively.
-*}
-
-lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
-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
-
-lemma distinct_rest: "distinct rest"
-  by (simp add: distinct_tl rest_def wq_distinct)
-
-lemma holding_cs_eq_th:
-  assumes "holding s t cs"
-  shows "t = th"
-proof -
-  from pip_e[unfolded is_v]
-  show ?thesis
-  proof(cases)
-    case (thread_V)
-    from held_unique[OF this(2) assms]
-    show ?thesis by simp
-  qed
-qed
-
-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 set_wq' by simp
-
-lemma neq_t_th: 
-  assumes "waiting (e#s) t c"
-  shows "t \<noteq> th"
-proof
-  assume otherwise: "t = th"
-  show False
-  proof(cases "c = cs")
-    case True
-    have "t \<in> set wq'" 
-     using assms[unfolded True s_waiting_def, folded wq_def, unfolded wq_es_cs]
-     by simp 
-    from th'_in_inv[OF this] have "t \<in> set rest" .
-    with wq_s_cs[folded otherwise] wq_distinct[of cs]
-    show ?thesis by simp
-  next
-    case False
-    have "wq (e#s) c = wq s c" using False
-        by (unfold is_v, simp)
-    hence "waiting s t c" using assms 
-        by (simp add: cs_waiting_def waiting_eq)
-    hence "t \<notin> readys s" by (unfold readys_def, auto)
-    hence "t \<notin> runing s" using runing_ready by auto 
-    with runing_th_s[folded otherwise] show ?thesis by auto
-  qed
-qed
-
-lemma waiting_esI1:
-  assumes "waiting s t c"
-      and "c \<noteq> cs" 
-  shows "waiting (e#s) t c" 
-proof -
-  have "wq (e#s) c = wq s c" 
-    using assms(2) is_v by auto
-  with assms(1) show ?thesis 
-    using cs_waiting_def waiting_eq by auto 
-qed
-
-lemma holding_esI2:
-  assumes "c \<noteq> cs" 
-  and "holding s t c"
-  shows "holding (e#s) t c"
-proof -
-  from assms(1) have "wq (e#s) c = wq s c" using is_v by auto
-  from assms(2)[unfolded s_holding_def, folded wq_def, 
-                folded this, unfolded wq_def, folded s_holding_def]
-  show ?thesis .
-qed
-
-lemma holding_esI1:
-  assumes "holding s t c"
-  and "t \<noteq> th"
-  shows "holding (e#s) t c"
-proof -
-  have "c \<noteq> cs" using assms using holding_cs_eq_th by blast 
-  from holding_esI2[OF this assms(1)]
-  show ?thesis .
-qed
-
-end
-
-context valid_trace_v_n
-begin
-
-lemma neq_wq': "wq' \<noteq> []" 
-proof (unfold wq'_def, rule someI2)
-  show "distinct rest \<and> set rest = set rest"
-    by (simp add: distinct_rest) 
-next
-  fix x
-  assume " distinct x \<and> set x = set rest" 
-  thus "x \<noteq> []" using rest_nnl by auto
-qed 
-
-definition "taker = hd wq'"
-
-definition "rest' = tl wq'"
-
-lemma eq_wq': "wq' = taker # rest'"
-  by (simp add: neq_wq' rest'_def taker_def)
-
-lemma next_th_taker: 
-  shows "next_th s th cs taker"
-  using rest_nnl taker_def wq'_def wq_s_cs 
-  by (auto simp:next_th_def)
-
-lemma taker_unique: 
-  assumes "next_th s th cs taker'"
-  shows "taker' = taker"
-proof -
-  from assms
-  obtain rest' where 
-    h: "wq s cs = th # rest'" 
-       "taker' = hd (SOME q. distinct q \<and> set q = set rest')"
-          by (unfold next_th_def, auto)
-  with wq_s_cs have "rest' = rest" by auto
-  thus ?thesis using h(2) taker_def wq'_def by auto 
-qed
-
-lemma waiting_set_eq:
-  "{(Th th', Cs cs) |th'. next_th s th cs th'} = {(Th taker, Cs cs)}"
-  by (smt all_not_in_conv bot.extremum insertI1 insert_subset 
-      mem_Collect_eq next_th_taker subsetI subset_antisym taker_def taker_unique)
-
-lemma holding_set_eq:
-  "{(Cs cs, Th th') |th'.  next_th s th cs th'} = {(Cs cs, Th taker)}"
-  using next_th_taker taker_def waiting_set_eq 
-  by fastforce
-   
-lemma holding_taker:
-  shows "holding (e#s) taker cs"
-    by (unfold s_holding_def, fold wq_def, unfold wq_es_cs, 
-        auto simp:neq_wq' taker_def)
-
-lemma waiting_esI2:
-  assumes "waiting s t cs"
-      and "t \<noteq> taker"
-  shows "waiting (e#s) t cs" 
-proof -
-  have "t \<in> set wq'" 
-  proof(unfold wq'_def, rule someI2)
-    show "distinct rest \<and> set rest = set rest"
-          by (simp add: distinct_rest)
-  next
-    fix x
-    assume "distinct x \<and> set x = set rest"
-    moreover have "t \<in> set rest"
-        using assms(1) cs_waiting_def waiting_eq wq_s_cs by auto 
-    ultimately show "t \<in> set x" by simp
-  qed
-  moreover have "t \<noteq> hd wq'"
-    using assms(2) taker_def by auto 
-  ultimately show ?thesis
-    by (unfold s_waiting_def, fold wq_def, unfold wq_es_cs, simp)
-qed
-
-lemma waiting_esE:
-  assumes "waiting (e#s) t c" 
-  obtains "c \<noteq> cs" "waiting s t c"
-     |    "c = cs" "t \<noteq> taker" "waiting s t cs" "t \<in> set rest'"
-proof(cases "c = cs")
-  case False
-  hence "wq (e#s) c = wq s c" using is_v by auto
-  with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto 
-  from that(1)[OF False this] show ?thesis .
-next
-  case True
-  from assms[unfolded s_waiting_def True, folded wq_def, unfolded wq_es_cs]
-  have "t \<noteq> hd wq'" "t \<in> set wq'" by auto
-  hence "t \<noteq> taker" by (simp add: taker_def) 
-  moreover hence "t \<noteq> th" using assms neq_t_th by blast 
-  moreover have "t \<in> set rest" by (simp add: `t \<in> set wq'` th'_in_inv) 
-  ultimately have "waiting s t cs"
-    by (metis cs_waiting_def list.distinct(2) list.sel(1) 
-                list.set_sel(2) rest_def waiting_eq wq_s_cs)  
-  show ?thesis using that(2)
-  using True `t \<in> set wq'` `t \<noteq> taker` `waiting s t cs` eq_wq' by auto   
-qed
-
-lemma holding_esI1:
-  assumes "c = cs"
-  and "t = taker"
-  shows "holding (e#s) t c"
-  by (unfold assms, simp add: holding_taker)
-
-lemma holding_esE:
-  assumes "holding (e#s) t c" 
-  obtains "c = cs" "t = taker"
-      | "c \<noteq> cs" "holding s t c"
-proof(cases "c = cs")
-  case True
-  from assms[unfolded True, unfolded s_holding_def, 
-             folded wq_def, unfolded wq_es_cs]
-  have "t = taker" by (simp add: taker_def) 
-  from that(1)[OF True this] show ?thesis .
-next
-  case False
-  hence "wq (e#s) c = wq s c" using is_v by auto
-  from assms[unfolded s_holding_def, folded wq_def, 
-             unfolded this, unfolded wq_def, folded s_holding_def]
-  have "holding s t c"  .
-  from that(2)[OF False this] show ?thesis .
-qed
-
-end 
-
-
-context valid_trace_v_e
-begin
-
-lemma nil_wq': "wq' = []" 
-proof (unfold wq'_def, rule someI2)
-  show "distinct rest \<and> set rest = set rest"
-    by (simp add: distinct_rest) 
-next
-  fix x
-  assume " distinct x \<and> set x = set rest" 
-  thus "x = []" using rest_nil by auto
-qed 
-
-lemma no_taker: 
-  assumes "next_th s th cs taker"
-  shows "False"
-proof -
-  from assms[unfolded next_th_def]
-  obtain rest' where "wq s cs = th # rest'" "rest' \<noteq> []"
-    by auto
-  thus ?thesis using rest_def rest_nil by auto 
-qed
-
-lemma waiting_set_eq:
-  "{(Th th', Cs cs) |th'. next_th s th cs th'} = {}"
-  using no_taker by auto
-
-lemma holding_set_eq:
-  "{(Cs cs, Th th') |th'.  next_th s th cs th'} = {}"
-  using no_taker by auto
-   
-lemma no_holding:
-  assumes "holding (e#s) taker cs"
-  shows False
-proof -
-  from wq_es_cs[unfolded nil_wq']
-  have " wq (e # s) cs = []" .
-  from assms[unfolded s_holding_def, folded wq_def, unfolded this]
-  show ?thesis by auto
-qed
-
-lemma no_waiting:
-  assumes "waiting (e#s) t cs"
-  shows False
-proof -
-  from wq_es_cs[unfolded nil_wq']
-  have " wq (e # s) cs = []" .
-  from assms[unfolded s_waiting_def, folded wq_def, unfolded this]
-  show ?thesis by auto
-qed
-
-lemma waiting_esI2:
-  assumes "waiting s t c"
-  shows "waiting (e#s) t c"
-proof -
-  have "c \<noteq> cs" using assms
-    using cs_waiting_def rest_nil waiting_eq wq_s_cs by auto 
-  from waiting_esI1[OF assms this]
-  show ?thesis .
-qed
-
-lemma waiting_esE:
-  assumes "waiting (e#s) t c" 
-  obtains "c \<noteq> cs" "waiting s t c"
-proof(cases "c = cs")
-  case False
-  hence "wq (e#s) c = wq s c" using is_v by auto
-  with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto 
-  from that(1)[OF False this] show ?thesis .
-next
-  case True
-  from no_waiting[OF assms[unfolded True]]
-  show ?thesis by auto
-qed
-
-lemma holding_esE:
-  assumes "holding (e#s) t c" 
-  obtains "c \<noteq> cs" "holding s t c"
-proof(cases "c = cs")
-  case True
-  from no_holding[OF assms[unfolded True]] 
-  show ?thesis by auto
-next
-  case False
-  hence "wq (e#s) c = wq s c" using is_v by auto
-  from assms[unfolded s_holding_def, folded wq_def, 
-             unfolded this, unfolded wq_def, folded s_holding_def]
-  have "holding s t c"  .
-  from that[OF False this] show ?thesis .
-qed
-
-end 
-
-lemma rel_eqI:
-  assumes "\<And> x y. (x,y) \<in> A \<Longrightarrow> (x,y) \<in> B"
-  and "\<And> x y. (x,y) \<in> B \<Longrightarrow> (x, y) \<in> A"
-  shows "A = B"
-  using assms by auto
-
-lemma in_RAG_E:
-  assumes "(n1, n2) \<in> RAG (s::state)"
-  obtains (waiting) th cs where "n1 = Th th" "n2 = Cs cs" "waiting s th cs"
-      | (holding) th cs where "n1 = Cs cs" "n2 = Th th" "holding s th cs"
-  using assms[unfolded s_RAG_def, folded waiting_eq holding_eq]
-  by auto
-  
-context valid_trace_v
-begin
-
-lemma RAG_es:
-  "RAG (e # s) =
-   RAG s - {(Cs cs, Th th)} -
-     {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-     {(Cs cs, Th th') |th'.  next_th s th cs th'}" (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')
-    show ?thesis
-    proof(cases "rest = []")
-      case False
-      interpret h_n: valid_trace_v_n s e th cs
-        by (unfold_locales, insert False, simp)
-      from waiting(3)
-      show ?thesis
-      proof(cases rule:h_n.waiting_esE)
-        case 1
-        with waiting(1,2)
-        show ?thesis
-        by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
-             fold waiting_eq, auto)
-      next
-        case 2
-        with waiting(1,2)
-        show ?thesis
-         by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
-             fold waiting_eq, auto)
-      qed
-    next
-      case True
-      interpret h_e: valid_trace_v_e s e th cs
-        by (unfold_locales, insert True, simp)
-      from waiting(3)
-      show ?thesis
-      proof(cases rule:h_e.waiting_esE)
-        case 1
-        with waiting(1,2)
-        show ?thesis
-        by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, 
-             fold waiting_eq, auto)
-      qed
-    qed
-  next
-    case (holding th' cs')
-    show ?thesis
-    proof(cases "rest = []")
-      case False
-      interpret h_n: valid_trace_v_n s e th cs
-        by (unfold_locales, insert False, simp)
-      from holding(3)
-      show ?thesis
-      proof(cases rule:h_n.holding_esE)
-        case 1
-        with holding(1,2)
-        show ?thesis
-        by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
-             fold waiting_eq, auto)
-      next
-        case 2
-        with holding(1,2)
-        show ?thesis
-         by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
-             fold holding_eq, auto)
-      qed
-    next
-      case True
-      interpret h_e: valid_trace_v_e s e th cs
-        by (unfold_locales, insert True, simp)
-      from holding(3)
-      show ?thesis
-      proof(cases rule:h_e.holding_esE)
-        case 1
-        with holding(1,2)
-        show ?thesis
-        by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, 
-             fold holding_eq, auto)
-      qed
-    qed
-  qed
-next
-  fix n1 n2
-  assume h: "(n1, n2) \<in> ?R"
-  show "(n1, n2) \<in> ?L"
-  proof(cases "rest = []")
-    case False
-    interpret h_n: valid_trace_v_n s e th cs
-        by (unfold_locales, insert False, simp)
-    from h[unfolded h_n.waiting_set_eq h_n.holding_set_eq]
-    have "((n1, n2) \<in> RAG s \<and> (n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th)
-                            \<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)) \<or> 
-          (n2 = Th h_n.taker \<and> n1 = Cs cs)" 
-      by auto
-   thus ?thesis
-   proof
-      assume "n2 = Th h_n.taker \<and> n1 = Cs cs"
-      with h_n.holding_taker
-      show ?thesis 
-        by (unfold s_RAG_def, fold holding_eq, auto)
-   next
-    assume h: "(n1, n2) \<in> RAG s \<and>
-        (n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th) \<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)"
-    hence "(n1, n2) \<in> RAG s" by simp
-    thus ?thesis
-    proof(cases rule:in_RAG_E)
-      case (waiting th' cs')
-      from h and this(1,2)
-      have "th' \<noteq> h_n.taker \<or> cs' \<noteq> cs" by auto
-      hence "waiting (e#s) th' cs'" 
-      proof
-        assume "cs' \<noteq> cs"
-        from waiting_esI1[OF waiting(3) this] 
-        show ?thesis .
-      next
-        assume neq_th': "th' \<noteq> h_n.taker"
-        show ?thesis
-        proof(cases "cs' = cs")
-          case False
-          from waiting_esI1[OF waiting(3) this] 
-          show ?thesis .
-        next
-          case True
-          from h_n.waiting_esI2[OF waiting(3)[unfolded True] neq_th', folded True]
-          show ?thesis .
-        qed
-      qed
-      thus ?thesis using waiting(1,2)
-        by (unfold s_RAG_def, fold waiting_eq, auto)
-    next
-      case (holding th' cs')
-      from h this(1,2)
-      have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto
-      hence "holding (e#s) th' cs'"
-      proof
-        assume "cs' \<noteq> cs"
-        from holding_esI2[OF this holding(3)] 
-        show ?thesis .
-      next
-        assume "th' \<noteq> th"
-        from holding_esI1[OF holding(3) this]
-        show ?thesis .
-      qed
-      thus ?thesis using holding(1,2)
-        by (unfold s_RAG_def, fold holding_eq, auto)
-    qed
-   qed
- next
-   case True
-   interpret h_e: valid_trace_v_e s e th cs
-        by (unfold_locales, insert True, simp)
-   from h[unfolded h_e.waiting_set_eq h_e.holding_set_eq]
-   have h_s: "(n1, n2) \<in> RAG s" "(n1, n2) \<noteq> (Cs cs, Th th)" 
-      by auto
-   from h_s(1)
-   show ?thesis
-   proof(cases rule:in_RAG_E)
-    case (waiting th' cs')
-    from h_e.waiting_esI2[OF this(3)]
-    show ?thesis using waiting(1,2)
-      by (unfold s_RAG_def, fold waiting_eq, auto)
-   next
-    case (holding th' cs')
-    with h_s(2)
-    have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto
-    thus ?thesis
-    proof
-      assume neq_cs: "cs' \<noteq> cs"
-      from holding_esI2[OF this holding(3)]
-      show ?thesis using holding(1,2)
-        by (unfold s_RAG_def, fold holding_eq, auto)
-    next
-      assume "th' \<noteq> th"
-      from holding_esI1[OF holding(3) this]
-      show ?thesis using holding(1,2)
-        by (unfold s_RAG_def, fold holding_eq, auto)
-    qed
-   qed
- qed
-qed
-
-end
-
-lemma step_RAG_v: 
-assumes vt:
-  "vt (V th cs#s)"
-shows "
-  RAG (V th cs # s) =
-  RAG s - {(Cs cs, Th th)} -
-  {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-  {(Cs cs, Th th') |th'.  next_th s th cs th'}" (is "?L = ?R")
-proof -
-  interpret vt_v: valid_trace_v s "V th cs"
-    using assms step_back_vt by (unfold_locales, auto) 
-  show ?thesis using vt_v.RAG_es .
-qed
-
-lemma (in valid_trace_create)
-  th_not_in_threads: "th \<notin> threads s"
-proof -
-  from pip_e[unfolded is_create]
-  show ?thesis by (cases, simp)
-qed
-
-lemma (in valid_trace_create)
-  threads_es [simp]: "threads (e#s) = threads s \<union> {th}"
-  by (unfold is_create, simp)
-
-lemma (in valid_trace_exit)
-  threads_es [simp]: "threads (e#s) = threads s - {th}"
-  by (unfold is_exit, simp)
-
-lemma (in valid_trace_p)
-  threads_es [simp]: "threads (e#s) = threads s"
-  by (unfold is_p, simp)
-
-lemma (in valid_trace_v)
-  threads_es [simp]: "threads (e#s) = threads s"
-  by (unfold is_v, simp)
-
-lemma (in valid_trace_v)
-  th_not_in_rest[simp]: "th \<notin> set rest"
-proof
-  assume otherwise: "th \<in> set rest"
-  have "distinct (wq s cs)" by (simp add: wq_distinct)
-  from this[unfolded wq_s_cs] and otherwise
-  show False by auto
-qed
-
-lemma (in valid_trace_v)
-  set_wq_es_cs [simp]: "set (wq (e#s) cs) = set (wq s cs) - {th}"
-proof(unfold wq_es_cs wq'_def, rule someI2)
-  show "distinct rest \<and> set rest = set rest"
-    by (simp add: distinct_rest)
-next
-  fix x
-  assume "distinct x \<and> set x = set rest"
-  thus "set x = set (wq s cs) - {th}" 
-      by (unfold wq_s_cs, simp)
-qed
-
-lemma (in valid_trace_exit)
-  th_not_in_wq: "th \<notin> set (wq s cs)"
-proof -
-  from pip_e[unfolded is_exit]
-  show ?thesis
-  by (cases, unfold holdents_def s_holding_def, fold wq_def, 
-             auto elim!:runing_wqE)
-qed
-
-lemma (in valid_trace) wq_threads: 
-  assumes "th \<in> set (wq s cs)"
-  shows "th \<in> threads s"
-  using assms
-proof(induct rule:ind)
-  case (Nil)
-  thus ?case by (auto simp:wq_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.hyps(2) Cons.prems by auto
-  next
-    case (Exit th')
-    interpret vt: valid_trace_exit s e th'
-      using Exit by (unfold_locales, simp)
-    show ?thesis
-      using Cons.hyps(2) Cons.prems vt.th_not_in_wq by auto 
-  next
-    case (P th' cs')
-    interpret vt: valid_trace_p s e th' cs'
-      using P by (unfold_locales, simp)
-    show ?thesis
-      using Cons.hyps(2) Cons.prems readys_threads 
-        runing_ready vt.is_p vt.runing_th_s vt_e.wq_in_inv 
-        by fastforce 
-  next
-    case (V th' cs')
-    interpret vt: valid_trace_v s e th' cs'
-      using V by (unfold_locales, simp)
-    show ?thesis using Cons
-      using vt.is_v vt.threads_es vt_e.wq_in_inv by blast
-  next
-    case (Set th' prio)
-    interpret vt: valid_trace_set s e th' prio
-      using Set by (unfold_locales, simp)
-    show ?thesis using Cons.hyps(2) Cons.prems vt.is_set 
-        by (auto simp:wq_def Let_def)
-  qed
-qed 
-
-context valid_trace
-begin
-
-lemma  dm_RAG_threads:
-  assumes in_dom: "(Th th) \<in> Domain (RAG s)"
-  shows "th \<in> threads s"
-proof -
-  from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
-  moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
-  ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
-  hence "th \<in> set (wq s cs)"
-    by (unfold s_RAG_def, auto simp:cs_waiting_def)
-  from wq_threads [OF this] show ?thesis .
-qed
-
-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)
-
-lemma RAG_threads:
-  assumes "(Th th) \<in> Field (RAG s)"
-  shows "th \<in> threads s"
-  using assms
-  by (metis Field_def UnE dm_RAG_threads rg_RAG_threads)
-
-end
-
-lemma (in valid_trace_v)
-  preced_es [simp]: "preced th (e#s) = preced th s"
-  by (unfold is_v preced_def, simp)
-
-lemma the_preced_v[simp]: "the_preced (V th cs#s) = the_preced s"
-proof
-  fix th'
-  show "the_preced (V th cs # s) th' = the_preced s th'"
-    by (unfold the_preced_def preced_def, simp)
-qed
-
-lemma (in valid_trace_v)
-  the_preced_es: "the_preced (e#s) = the_preced s"
-  by (unfold is_v preced_def, simp)
-
-context valid_trace_p
-begin
-
-lemma not_holding_s_th_cs: "\<not> holding s th cs"
-proof
-  assume otherwise: "holding s th cs"
-  from pip_e[unfolded is_p]
-  show False
-  proof(cases)
-    case (thread_P)
-    moreover have "(Cs cs, Th th) \<in> RAG s"
-      using otherwise cs_holding_def 
-            holding_eq th_not_in_wq by auto
-    ultimately show ?thesis by auto
-  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 wne: "wq s cs \<noteq> []"
-begin
-
-definition "holder = hd (wq s cs)"
-definition "waiters = tl (wq s cs)"
-definition "waiters' = waiters @ [th]"
-
-lemma wq_s_cs: "wq s cs = holder#waiters"
-    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)
-
-lemma waiting_es_th_cs: "waiting (e#s) th cs"
-  using cs_waiting_def th_not_in_wq waiting_eq wq_es_cs' wq_s_cs by auto
-
-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
-begin
-
-lemma wq_es_cs': "wq (e#s) cs = [th]"
-  using wq_es_cs[unfolded we] by simp
-
-lemma holding_es_th_cs: 
-  shows "holding (e#s) th cs"
-proof -
-  from wq_es_cs'
-  have "th \<in> set (wq (e#s) cs)" "th = hd (wq (e#s) cs)" by auto
-  thus ?thesis using cs_holding_def holding_eq by blast 
-qed
-
-lemma RAG_edge: "(Cs cs, Th th) \<in> RAG (e#s)"
-  by (unfold s_RAG_def, fold holding_eq, insert holding_es_th_cs, auto)
-
-lemma waiting_esE:
-  assumes "waiting (e#s) th' cs'"
-  obtains "waiting s th' cs'"
-  using assms
-  by (metis cs_waiting_def event.distinct(15) is_p list.sel(1) 
-        set_ConsD waiting_eq we wq_es_cs' wq_neq_simp wq_out_inv)
-  
-lemma holding_esE:
-  assumes "holding (e#s) th' cs'"
-  obtains "cs' \<noteq> cs" "holding s th' cs'"
-    | "cs' = cs" "th' = th"
-proof(cases "cs' = cs")
-  case True
-  from held_unique[OF holding_es_th_cs assms[unfolded True]]
-  have "th' = th" by simp
-  from that(2)[OF True this] show ?thesis .
-next
-  case False
-  have "holding s th' cs'" using assms
-    using False cs_holding_def holding_eq by auto
-  from that(1)[OF False this] show ?thesis .
-qed
-
-lemma RAG_es: "RAG (e # s) =  RAG s \<union> {(Cs cs, Th th)}" (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)
-    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)
-    next
-      case 2
-      with holding(1,2) show ?thesis by auto
-    qed
-  qed
-next
-  fix n1 n2
-  assume "(n1, n2) \<in> ?R"
-  hence "(n1, n2) \<in> RAG s \<or> (n1 = Cs cs \<and> n2 = Th th)" 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 = 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
-
-lemma tRAG_alt_def: 
-  "tRAG s = {(Th th1, Th th2) | th1 th2. 
-                  \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
- by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
-
-sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
-proof -
-  have "fsubtree (tRAG s)"
-  proof -
-    have "fbranch (tRAG s)"
-    proof(unfold tRAG_def, rule fbranch_compose)
-        show "fbranch (wRAG s)"
-        proof(rule finite_fbranchI)
-           from finite_RAG show "finite (wRAG s)"
-           by (unfold RAG_split, auto)
-        qed
-    next
-        show "fbranch (hRAG s)"
-        proof(rule finite_fbranchI)
-           from finite_RAG 
-           show "finite (hRAG s)" by (unfold RAG_split, auto)
-        qed
-    qed
-    moreover have "wf (tRAG s)"
-    proof(rule wf_subset)
-      show "wf (RAG s O RAG s)" using wf_RAG
-        by (fold wf_comp_self, simp)
-    next
-      show "tRAG s \<subseteq> (RAG s O RAG s)"
-        by (unfold tRAG_alt_def, auto)
-    qed
-    ultimately show ?thesis
-      by (unfold fsubtree_def fsubtree_axioms_def,auto)
-  qed
-  from this[folded tRAG_def] show "fsubtree (tRAG s)" .
-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: fsbtRAGs.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 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)
-
-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 fsbtRAGs.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: (* ddd *)
-  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
-
-
-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
-
-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
-  from rp1 rp2
-  show ?thesis
-  proof(cases)
-    case (less_1 xs')
-    moreover have "xs' = []"
-    proof(rule ccontr)
-      assume otherwise: "xs' \<noteq> []"
-      from rpath_plus[OF less_1(3) this]
-      have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+" .
-      from tranclD[OF this]
-      obtain cs where "waiting s th1 cs"
-        by (unfold s_RAG_def, fold waiting_eq, auto)
-      with runing_1 show False
-        by (unfold runing_def readys_def, auto)
-    qed
-    ultimately have "xs2 = xs1" by simp
-    from rpath_dest_eq[OF rp1 rp2[unfolded this]]
-    show ?thesis by simp
-  next
-    case (less_2 xs')
-    moreover have "xs' = []"
-    proof(rule ccontr)
-      assume otherwise: "xs' \<noteq> []"
-      from rpath_plus[OF less_2(3) this]
-      have "(Th th2, Th th1) \<in> (RAG s)\<^sup>+" .
-      from tranclD[OF this]
-      obtain cs where "waiting s th2 cs"
-        by (unfold s_RAG_def, fold waiting_eq, auto)
-      with runing_2 show False
-        by (unfold runing_def readys_def, auto)
-    qed
-    ultimately have "xs2 = xs1" by simp
-    from rpath_dest_eq[OF rp1 rp2[unfolded this]]
-    show ?thesis by simp
-  qed
-qed
-
-lemma card_runing: "card (runing s) \<le> 1"
-proof(cases "runing s = {}")
-  case True
-  thus ?thesis by auto
-next
-  case False
-  then obtain th where [simp]: "th \<in> runing s" by auto
-  from runing_unique[OF this]
-  have "runing s = {th}" by auto
-  thus ?thesis by auto
-qed
-
-lemma create_pre:
-  assumes stp: "step s e"
-  and not_in: "th \<notin> threads s"
-  and is_in: "th \<in> threads (e#s)"
-  obtains prio where "e = Create th prio"
-proof -
-  from assms  
-  show ?thesis
-  proof(cases)
-    case (thread_create thread prio)
-    with is_in not_in have "e = Create th prio" by simp
-    from that[OF this] show ?thesis .
-  next
-    case (thread_exit thread)
-    with assms show ?thesis by (auto intro!:that)
-  next
-    case (thread_P thread)
-    with assms show ?thesis by (auto intro!:that)
-  next
-    case (thread_V thread)
-    with assms show ?thesis by (auto intro!:that)
-  next 
-    case (thread_set thread)
-    with assms show ?thesis by (auto intro!:that)
-  qed
-qed
-
-lemma eq_pv_children:
-  assumes eq_pv: "cntP s th = cntV s th"
-  shows "children (RAG s) (Th th) = {}"
-proof -
-    from cnp_cnv_cncs and eq_pv
-    have "cntCS s th = 0" 
-      by (auto split:if_splits)
-    from this[unfolded cntCS_def holdents_alt_def]
-    have card_0: "card (the_cs ` children (RAG s) (Th th)) = 0" .
-    have "finite (the_cs ` children (RAG s) (Th th))"
-      by (simp add: fsbtRAGs.finite_children)
-    from card_0[unfolded card_0_eq[OF this]]
-    show ?thesis by auto
-qed
-
-lemma eq_pv_holdents:
-  assumes eq_pv: "cntP s th = cntV s th"
-  shows "holdents s th = {}"
-  by (unfold holdents_alt_def eq_pv_children[OF assms], simp)
-
-lemma eq_pv_subtree:
-  assumes eq_pv: "cntP s th = cntV s th"
-  shows "subtree (RAG s) (Th th) = {Th th}"
-  using eq_pv_children[OF assms]
-    by (unfold subtree_children, simp)
-
-end
-
-lemma cp_gen_alt_def:
-  "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
-    by (auto simp:cp_gen_def)
-
-lemma tRAG_nodeE:
-  assumes "(n1, n2) \<in> tRAG s"
-  obtains th1 th2 where "n1 = Th th1" "n2 = Th th2"
-  using assms
-  by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def)
-
-lemma subtree_nodeE:
-  assumes "n \<in> subtree (tRAG s) (Th th)"
-  obtains th1 where "n = Th th1"
-proof -
-  show ?thesis
-  proof(rule subtreeE[OF assms])
-    assume "n = Th th"
-    from that[OF this] show ?thesis .
-  next
-    assume "Th th \<in> ancestors (tRAG s) n"
-    hence "(n, Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def)
-    hence "\<exists> th1. n = Th th1"
-    proof(induct)
-      case (base y)
-      from tRAG_nodeE[OF this] show ?case by metis
-    next
-      case (step y z)
-      thus ?case by auto
-    qed
-    with that show ?thesis by auto
-  qed
-qed
-
-lemma tRAG_star_RAG: "(tRAG s)^* \<subseteq> (RAG s)^*"
-proof -
-  have "(wRAG s O hRAG s)^* \<subseteq> (RAG s O RAG s)^*" 
-    by (rule rtrancl_mono, auto simp:RAG_split)
-  also have "... \<subseteq> ((RAG s)^*)^*"
-    by (rule rtrancl_mono, auto)
-  also have "... = (RAG s)^*" by simp
-  finally show ?thesis by (unfold tRAG_def, simp)
-qed
-
-lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x"
-proof -
-  { fix a
-    assume "a \<in> subtree (tRAG s) x"
-    hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def)
-    with tRAG_star_RAG
-    have "(a, x) \<in> (RAG s)^*" by auto
-    hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def)
-  } thus ?thesis by auto
-qed
-
-lemma tRAG_trancl_eq:
-   "{th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
-    {th'. (Th th', Th th)  \<in> (RAG s)^+}"
-   (is "?L = ?R")
-proof -
-  { fix th'
-    assume "th' \<in> ?L"
-    hence "(Th th', Th th) \<in> (tRAG s)^+" by auto
-    from tranclD[OF this]
-    obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto
-    from tRAG_subtree_RAG and this(2)
-    have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) 
-    moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto 
-    ultimately have "th' \<in> ?R"  by auto 
-  } moreover 
-  { fix th'
-    assume "th' \<in> ?R"
-    hence "(Th th', Th th) \<in> (RAG s)^+" by (auto)
-    from plus_rpath[OF this]
-    obtain xs where rp: "rpath (RAG s) (Th th') xs (Th th)" "xs \<noteq> []" by auto
-    hence "(Th th', Th th) \<in> (tRAG s)^+"
-    proof(induct xs arbitrary:th' th rule:length_induct)
-      case (1 xs th' th)
-      then obtain x1 xs1 where Cons1: "xs = x1#xs1" by (cases xs, auto)
-      show ?case
-      proof(cases "xs1")
-        case Nil
-        from 1(2)[unfolded Cons1 Nil]
-        have rp: "rpath (RAG s) (Th th') [x1] (Th th)" .
-        hence "(Th th', x1) \<in> (RAG s)" 
-          by (cases, auto)
-        then obtain cs where "x1 = Cs cs" 
-              by (unfold s_RAG_def, auto)
-        from rpath_nnl_lastE[OF rp[unfolded this]]
-        show ?thesis by auto
-      next
-        case (Cons x2 xs2)
-        from 1(2)[unfolded Cons1[unfolded this]]
-        have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" .
-        from rpath_edges_on[OF this]
-        have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" .
-        have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)"
-            by (simp add: edges_on_unfold)
-        with eds have rg1: "(Th th', x1) \<in> RAG s" by auto
-        then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto)
-        have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)"
-            by (simp add: edges_on_unfold)
-        from this eds
-        have rg2: "(x1, x2) \<in> RAG s" by auto
-        from this[unfolded eq_x1] 
-        obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto)
-        from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2]
-        have rt1: "(Th th', Th th1) \<in> tRAG s" by (unfold tRAG_alt_def, auto)
-        from rp have "rpath (RAG s) x2 xs2 (Th th)"
-           by  (elim rpath_ConsE, simp)
-        from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" .
-        show ?thesis
-        proof(cases "xs2 = []")
-          case True
-          from rpath_nilE[OF rp'[unfolded this]]
-          have "th1 = th" by auto
-          from rt1[unfolded this] show ?thesis by auto
-        next
-          case False
-          from 1(1)[rule_format, OF _ rp' this, unfolded Cons1 Cons]
-          have "(Th th1, Th th) \<in> (tRAG s)\<^sup>+" by simp
-          with rt1 show ?thesis by auto
-        qed
-      qed
-    qed
-    hence "th' \<in> ?L" by auto
-  } ultimately show ?thesis by blast
-qed
-
-lemma tRAG_trancl_eq_Th:
-   "{Th th' | th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
-    {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}"
-    using tRAG_trancl_eq by auto
-
-lemma dependants_alt_def:
-  "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
-  by (metis eq_RAG s_dependants_def tRAG_trancl_eq)
-
-lemma dependants_alt_def1:
-  "dependants (s::state) th = {th'. (Th th', Th th) \<in> (RAG s)^+}"
-  using dependants_alt_def tRAG_trancl_eq by auto
-
-context valid_trace
-begin
-lemma count_eq_RAG_plus:
-  assumes "cntP s th = cntV s th"
-  shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
-proof(rule ccontr)
-    assume otherwise: "{th'. (Th th', Th th) \<in> (RAG s)\<^sup>+} \<noteq> {}"
-    then obtain th' where "(Th th', Th th) \<in> (RAG s)^+" by auto
-    from tranclD2[OF this]
-    obtain z where "z \<in> children (RAG s) (Th th)" 
-      by (auto simp:children_def)
-    with eq_pv_children[OF assms]
-    show False by simp
-qed
-
-lemma eq_pv_dependants:
-  assumes eq_pv: "cntP s th = cntV s th"
-  shows "dependants s th = {}"
-proof -
-  from count_eq_RAG_plus[OF assms, folded dependants_alt_def1]
-  show ?thesis .
-qed
-
-end
-
-lemma eq_dependants: "dependants (wq s) = dependants s"
-  by (simp add: s_dependants_abv wq_def)
-
-context valid_trace
-begin
-
-lemma count_eq_tRAG_plus:
-  assumes "cntP s th = cntV s th"
-  shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
-  using assms eq_pv_dependants dependants_alt_def eq_dependants by auto 
-
-lemma count_eq_RAG_plus_Th:
-  assumes "cntP s th = cntV s th"
-  shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
-  using count_eq_RAG_plus[OF assms] by auto
-
-lemma count_eq_tRAG_plus_Th:
-  assumes "cntP s th = cntV s th"
-  shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
-   using count_eq_tRAG_plus[OF assms] by auto
-end
-
-lemma inj_the_preced: 
-  "inj_on (the_preced s) (threads s)"
-  by (metis inj_onI preced_unique the_preced_def)
-
-lemma tRAG_Field:
-  "Field (tRAG s) \<subseteq> Field (RAG s)"
-  by (unfold tRAG_alt_def Field_def, auto)
-
-lemma tRAG_ancestorsE:
-  assumes "x \<in> ancestors (tRAG s) u"
-  obtains th where "x = Th th"
-proof -
-  from assms have "(u, x) \<in> (tRAG s)^+" 
-      by (unfold ancestors_def, auto)
-  from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
-  then obtain th where "x = Th th"
-    by (unfold tRAG_alt_def, auto)
-  from that[OF this] show ?thesis .
-qed
-
-lemma tRAG_mono:
-  assumes "RAG s' \<subseteq> RAG s"
-  shows "tRAG s' \<subseteq> tRAG s"
-  using assms 
-  by (unfold tRAG_alt_def, auto)
-
-lemma holding_next_thI:
-  assumes "holding s th cs"
-  and "length (wq s cs) > 1"
-  obtains th' where "next_th s th cs th'"
-proof -
-  from assms(1)[folded holding_eq, unfolded cs_holding_def]
-  have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" 
-    by (unfold s_holding_def, fold wq_def, auto)
-  then obtain rest where h1: "wq s cs = th#rest" 
-    by (cases "wq s cs", auto)
-  with assms(2) have h2: "rest \<noteq> []" by auto
-  let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
-  have "next_th s th cs ?th'" using  h1(1) h2 
-    by (unfold next_th_def, auto)
-  from that[OF this] show ?thesis .
-qed
-
-lemma RAG_tRAG_transfer:
-  assumes "vt s'"
-  assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
-  and "(Cs cs, Th th'') \<in> RAG s'"
-  shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
-proof -
-  interpret vt_s': valid_trace "s'" using assms(1)
-    by (unfold_locales, simp)
-  { fix n1 n2
-    assume "(n1, n2) \<in> ?L"
-    from this[unfolded tRAG_alt_def]
-    obtain th1 th2 cs' where 
-      h: "n1 = Th th1" "n2 = Th th2" 
-         "(Th th1, Cs cs') \<in> RAG s"
-         "(Cs cs', Th th2) \<in> RAG s" by auto
-    from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto
-    from h(3) and assms(2) 
-    have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
-          (Th th1, Cs cs') \<in> RAG s'" by auto
-    hence "(n1, n2) \<in> ?R"
-    proof
-      assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
-      hence eq_th1: "th1 = th" by simp
-      moreover have "th2 = th''"
-      proof -
-        from h1 have "cs' = cs" by simp
-        from assms(3) cs_in[unfolded this]
-        show ?thesis using vt_s'.unique_RAG by auto 
-      qed
-      ultimately show ?thesis using h(1,2) by auto
-    next
-      assume "(Th th1, Cs cs') \<in> RAG s'"
-      with cs_in have "(Th th1, Th th2) \<in> tRAG s'"
-        by (unfold tRAG_alt_def, auto)
-      from this[folded h(1, 2)] show ?thesis by auto
-    qed
-  } moreover {
-    fix n1 n2
-    assume "(n1, n2) \<in> ?R"
-    hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto
-    hence "(n1, n2) \<in> ?L" 
-    proof
-      assume "(n1, n2) \<in> tRAG s'"
-      moreover have "... \<subseteq> ?L"
-      proof(rule tRAG_mono)
-        show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto)
-      qed
-      ultimately show ?thesis by auto
-    next
-      assume eq_n: "(n1, n2) = (Th th, Th th'')"
-      from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto
-      moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto
-      ultimately show ?thesis 
-        by (unfold eq_n tRAG_alt_def, auto)
-    qed
-  } ultimately show ?thesis by auto
-qed
-
-context valid_trace
-begin
-
-lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt]
-
-end
-
-lemma tRAG_subtree_eq: 
-   "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
-   (is "?L = ?R")
-proof -
-  { fix n 
-    assume h: "n \<in> ?L"
-    hence "n \<in> ?R"
-    by (smt mem_Collect_eq subsetCE subtree_def subtree_nodeE tRAG_subtree_RAG) 
-  } moreover {
-    fix n
-    assume "n \<in> ?R"
-    then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*"
-      by (auto simp:subtree_def)
-    from rtranclD[OF this(2)]
-    have "n \<in> ?L"
-    proof
-      assume "Th th' \<noteq> Th th \<and> (Th th', Th th) \<in> (RAG s)\<^sup>+"
-      with h have "n \<in> {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}" by auto
-      thus ?thesis using subtree_def tRAG_trancl_eq by fastforce
-    qed (insert h, auto simp:subtree_def)
-  } ultimately show ?thesis by auto
-qed
-
-lemma threads_set_eq: 
-   "the_thread ` (subtree (tRAG s) (Th th)) = 
-                  {th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R")
-   by (auto intro:rev_image_eqI simp:tRAG_subtree_eq)
-
-lemma cp_alt_def1: 
-  "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))"
-proof -
-  have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) =
-       ((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))"
-       by auto
-  thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto)
-qed
-
-lemma cp_gen_def_cond: 
-  assumes "x = Th th"
-  shows "cp s th = cp_gen s (Th th)"
-by (unfold cp_alt_def1 cp_gen_def, simp)
-
-lemma cp_gen_over_set:
-  assumes "\<forall> x \<in> A. \<exists> th. x = Th th"
-  shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A"
-proof(rule f_image_eq)
-  fix a
-  assume "a \<in> A"
-  from assms[rule_format, OF this]
-  obtain th where eq_a: "a = Th th" by auto
-  show "cp_gen s a = (cp s \<circ> the_thread) a"
-    by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
-qed
-
-context valid_trace
-begin
-
-lemma subtree_tRAG_thread:
-  assumes "th \<in> threads s"
-  shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R")
-proof -
-  have "?L = {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
-    by (unfold tRAG_subtree_eq, simp)
-  also have "... \<subseteq> ?R"
-  proof
-    fix x
-    assume "x \<in> {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
-    then obtain th' where h: "x = Th th'" "Th th' \<in> subtree (RAG s) (Th th)" by auto
-    from this(2)
-    show "x \<in> ?R"
-    proof(cases rule:subtreeE)
-      case 1
-      thus ?thesis by (simp add: assms h(1)) 
-    next
-      case 2
-      thus ?thesis by (metis ancestors_Field dm_RAG_threads h(1) image_eqI) 
-    qed
-  qed
-  finally show ?thesis .
-qed
-
-lemma readys_root:
-  assumes "th \<in> readys s"
-  shows "root (RAG s) (Th th)"
-proof -
-  { fix x
-    assume "x \<in> ancestors (RAG s) (Th th)"
-    hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def)
-    from tranclD[OF this]
-    obtain z where "(Th th, z) \<in> RAG s" by auto
-    with assms(1) have False
-         apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def)
-         by (fold wq_def, blast)
-  } thus ?thesis by (unfold root_def, auto)
-qed
-
-lemma readys_in_no_subtree:
-  assumes "th \<in> readys s"
-  and "th' \<noteq> th"
-  shows "Th th \<notin> subtree (RAG s) (Th th')" 
-proof
-   assume "Th th \<in> subtree (RAG s) (Th th')"
-   thus False
-   proof(cases rule:subtreeE)
-      case 1
-      with assms show ?thesis by auto
-   next
-      case 2
-      with readys_root[OF assms(1)]
-      show ?thesis by (auto simp:root_def)
-   qed
-qed
-
-lemma not_in_thread_isolated:
-  assumes "th \<notin> threads s"
-  shows "(Th th) \<notin> Field (RAG s)"
-proof
-  assume "(Th th) \<in> Field (RAG s)"
-  with dm_RAG_threads and rg_RAG_threads assms
-  show False by (unfold Field_def, blast)
-qed
-
-end
-
-definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
-  where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
-
-
-lemma detached_test:
-  shows "detached s th = (Th th \<notin> Field (RAG s))"
-apply(simp add: detached_def Field_def)
-apply(simp add: s_RAG_def)
-apply(simp add: s_holding_abv s_waiting_abv)
-apply(simp add: Domain_iff Range_iff)
-apply(simp add: wq_def)
-apply(auto)
-done
-
-context valid_trace
-begin
-
-lemma detached_intro:
-  assumes eq_pv: "cntP s th = cntV s th"
-  shows "detached s th"
-proof -
-  from eq_pv cnp_cnv_cncs
-  have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:pvD_def)
-  thus ?thesis
-  proof
-    assume "th \<notin> threads s"
-    with rg_RAG_threads dm_RAG_threads
-    show ?thesis
-      by (auto simp add: detached_def s_RAG_def s_waiting_abv 
-              s_holding_abv wq_def Domain_iff Range_iff)
-  next
-    assume "th \<in> readys s"
-    moreover have "Th th \<notin> Range (RAG s)"
-    proof -
-      from eq_pv_children[OF assms]
-      have "children (RAG s) (Th th) = {}" .
-      thus ?thesis
-      by (unfold children_def, auto)
-    qed
-    ultimately show ?thesis
-      by (auto simp add: detached_def s_RAG_def s_waiting_abv 
-              s_holding_abv wq_def readys_def)
-  qed
-qed
-
-lemma detached_elim:
-  assumes dtc: "detached s th"
-  shows "cntP s th = cntV s th"
-proof -
-  have cncs_z: "cntCS s th = 0"
-  proof -
-    from dtc have "holdents s th = {}"
-      unfolding detached_def holdents_test s_RAG_def
-      by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
-    thus ?thesis by (auto simp:cntCS_def)
-  qed
-  show ?thesis
-  proof(cases "th \<in> threads s")
-    case True
-    with dtc 
-    have "th \<in> readys s"
-      by (unfold readys_def detached_def Field_def Domain_def Range_def, 
-           auto simp:waiting_eq s_RAG_def)
-    with cncs_z  show ?thesis using cnp_cnv_cncs by (simp add:pvD_def)
-  next
-    case False
-    with cncs_z and cnp_cnv_cncs show ?thesis by (simp add:pvD_def)
-  qed
-qed
-
-lemma detached_eq:
-  shows "(detached s th) = (cntP s th = cntV s th)"
-  by (insert vt, auto intro:detached_intro detached_elim)
-
-end
-
-context valid_trace
-begin
-(* ddd *)
-lemma cp_gen_rec:
-  assumes "x = Th th"
-  shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
-proof(cases "children (tRAG s) x = {}")
-  case True
-  show ?thesis
-    by (unfold True cp_gen_def subtree_children, simp add:assms)
-next
-  case False
-  hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
-  note fsbttRAGs.finite_subtree[simp]
-  have [simp]: "finite (children (tRAG s) x)"
-     by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], 
-            rule children_subtree)
-  { fix r x
-    have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
-  } note this[simp]
-  have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
-  proof -
-    from False obtain q where "q \<in> children (tRAG s) x" by blast
-    moreover have "subtree (tRAG s) q \<noteq> {}" by simp
-    ultimately show ?thesis by blast
-  qed
-  have h: "Max ((the_preced s \<circ> the_thread) `
-                ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
-        Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
-                     (is "?L = ?R")
-  proof -
-    let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L
-    let "Max (_ \<union> (?h ` ?B))" = ?R
-    let ?L1 = "?f ` \<Union>(?g ` ?B)"
-    have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)"
-    proof -
-      have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp
-      also have "... =  (\<Union> x \<in> ?B. ?f ` (?g x))" by auto
-      finally have "Max ?L1 = Max ..." by simp
-      also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)"
-        by (subst Max_UNION, simp+)
-      also have "... = Max (cp_gen s ` children (tRAG s) x)"
-          by (unfold image_comp cp_gen_alt_def, simp)
-      finally show ?thesis .
-    qed
-    show ?thesis
-    proof -
-      have "?L = Max (?f ` ?A \<union> ?L1)" by simp
-      also have "... = max (the_preced s (the_thread x)) (Max ?L1)"
-            by (subst Max_Un, simp+)
-      also have "... = max (?f x) (Max (?h ` ?B))"
-        by (unfold eq_Max_L1, simp)
-      also have "... =?R"
-        by (rule max_Max_eq, (simp)+, unfold assms, simp)
-      finally show ?thesis .
-    qed
-  qed  thus ?thesis 
-          by (fold h subtree_children, unfold cp_gen_def, simp) 
-qed
-
-lemma cp_rec:
-  "cp s th = Max ({the_preced s th} \<union> 
-                     (cp s o the_thread) ` children (tRAG s) (Th th))"
-proof -
-  have "Th th = Th th" by simp
-  note h =  cp_gen_def_cond[OF this] cp_gen_rec[OF this]
-  show ?thesis 
-  proof -
-    have "cp_gen s ` children (tRAG s) (Th th) = 
-                (cp s \<circ> the_thread) ` children (tRAG s) (Th th)"
-    proof(rule cp_gen_over_set)
-      show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th"
-        by (unfold tRAG_alt_def, auto simp:children_def)
-    qed
-    thus ?thesis by (subst (1) h(1), unfold h(2), simp)
-  qed
-qed
-
-lemma next_th_holding:
-  assumes nxt: "next_th s th cs th'"
-  shows "holding (wq s) th cs"
-proof -
-  from nxt[unfolded next_th_def]
-  obtain rest where h: "wq s cs = th # rest"
-                       "rest \<noteq> []" 
-                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
-  thus ?thesis
-    by (unfold cs_holding_def, auto)
-qed
-
-lemma next_th_waiting:
-  assumes nxt: "next_th s th cs th'"
-  shows "waiting (wq s) th' cs"
-proof -
-  from nxt[unfolded next_th_def]
-  obtain rest where h: "wq s cs = th # rest"
-                       "rest \<noteq> []" 
-                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
-  from wq_distinct[of cs, unfolded h]
-  have dst: "distinct (th # rest)" .
-  have in_rest: "th' \<in> set rest"
-  proof(unfold h, rule someI2)
-    show "distinct rest \<and> set rest = set rest" using dst by auto
-  next
-    fix x assume "distinct x \<and> set x = set rest"
-    with h(2)
-    show "hd x \<in> set (rest)" by (cases x, auto)
-  qed
-  hence "th' \<in> set (wq s cs)" by (unfold h(1), auto)
-  moreover have "th' \<noteq> hd (wq s cs)"
-    by (unfold h(1), insert in_rest dst, auto)
-  ultimately show ?thesis by (auto simp:cs_waiting_def)
-qed
-
-lemma next_th_RAG:
-  assumes nxt: "next_th (s::event list) th cs th'"
-  shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s"
-  using vt assms next_th_holding next_th_waiting
-  by (unfold s_RAG_def, simp)
-
-end
-
-lemma next_th_unique: 
-  assumes nt1: "next_th s th cs th1"
-  and nt2: "next_th s th cs th2"
-  shows "th1 = th2"
-using assms by (unfold next_th_def, auto)
-
-context valid_trace
-begin
-
-thm th_chain_to_ready
-
-find_theorems subtree Th RAG
-
-lemma "(threads s) = (\<Union> th \<in> readys s. {th'. Th th' \<in> subtree (RAG s) (Th th)})"
-    (is "?L = ?R")
-proof -
-  { fix th1
-    assume "th1 \<in> ?L"
-    from th_chain_to_ready[OF this]
-    have "th1 \<in> readys s \<or> (\<exists>th'a. th'a \<in> readys s \<and> (Th th1, Th th'a) \<in> (RAG s)\<^sup>+)" .
-    hence "th1 \<in> ?R"
-    proof
-      assume "th1 \<in> readys s"
-      thus ?thesis by (auto simp:subtree_def)
-    next
-      assume "\<exists>th'a. th'a \<in> readys s \<and> (Th th1, Th th'a) \<in> (RAG s)\<^sup>+"
-      thus ?thesis 
-    qed
-  } moreover 
-  { fix th'
-    assume "th' \<in> ?R"
-    have "th' \<in> ?L" sorry
-  } ultimately show ?thesis by auto
-qed
-
-lemma max_cp_readys_threads_pre: (* ccc *)
-  assumes np: "threads s \<noteq> {}"
-  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
-proof(unfold max_cp_eq)
-  show "Max (cp s ` readys s) = Max (the_preced s ` threads s)"
-  proof -
-    let ?p = "Max (the_preced s ` threads s)" 
-    let ?f = "the_preced s"
-    have "?p \<in> (?f ` threads s)"
-    proof(rule Max_in)
-      from finite_threads show "finite (?f ` threads s)" by simp
-    next
-      from np show "?f ` threads s \<noteq> {}" by simp
-    qed
-    then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
-      by (auto simp:Image_def)
-    from th_chain_to_ready [OF tm_in]
-    have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" .
-    thus ?thesis
-    proof
-      assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ "
-      then obtain th' where th'_in: "th' \<in> readys s" 
-        and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto
-      have "cp s th' = ?f tm"
-      proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
-        from dependants_threads finite_threads
-        show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" 
-          by (auto intro:finite_subset)
-      next
-        fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
-        from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" .
-        moreover have "p \<le> \<dots>"
-        proof(rule Max_ge)
-          from finite_threads
-          show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        next
-          from p_in and th'_in and dependants_threads[of th']
-          show "p \<in> (\<lambda>th. preced th s) ` threads s"
-            by (auto simp:readys_def)
-        qed
-        ultimately show "p \<le> preced tm s" by auto
-      next
-        show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
-        proof -
-          from tm_chain
-          have "tm \<in> dependants (wq s) th'"
-            by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto)
-          thus ?thesis by auto
-        qed
-      qed
-      with tm_max
-      have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
-      show ?thesis
-      proof (fold h, rule Max_eqI)
-        fix q 
-        assume "q \<in> cp s ` readys s"
-        then obtain th1 where th1_in: "th1 \<in> readys s"
-          and eq_q: "q = cp s th1" by auto
-        show "q \<le> cp s th'"
-          apply (unfold h eq_q)
-          apply (unfold cp_eq_cpreced cpreced_def)
-          apply (rule Max_mono)
-        proof -
-          from dependants_threads [of th1] th1_in
-          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<subseteq> 
-                 (\<lambda>th. preced th s) ` threads s"
-            by (auto simp:readys_def)
-        next
-          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}" by simp
-        next
-          from finite_threads 
-          show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        qed
-      next
-        from finite_threads
-        show "finite (cp s ` readys s)" by (auto simp:readys_def)
-      next
-        from th'_in
-        show "cp s th' \<in> cp s ` readys s" by simp
-      qed
-    next
-      assume tm_ready: "tm \<in> readys s"
-      show ?thesis
-      proof(fold tm_max)
-        have cp_eq_p: "cp s tm = preced tm s"
-        proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
-          fix y 
-          assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
-          show "y \<le> preced tm s"
-          proof -
-            { fix y'
-              assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependants (wq s) tm)"
-              have "y' \<le> preced tm s"
-              proof(unfold tm_max, rule Max_ge)
-                from hy' dependants_threads[of tm]
-                show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
-              next
-                from finite_threads
-                show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-              qed
-            } with hy show ?thesis by auto
-          qed
-        next
-          from dependants_threads[of tm] finite_threads
-          show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm))"
-            by (auto intro:finite_subset)
-        next
-          show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
-            by simp
-        qed 
-        moreover have "Max (cp s ` readys s) = cp s tm"
-        proof(rule Max_eqI)
-          from tm_ready show "cp s tm \<in> cp s ` readys s" by simp
-        next
-          from finite_threads
-          show "finite (cp s ` readys s)" by (auto simp:readys_def)
-        next
-          fix y assume "y \<in> cp s ` readys s"
-          then obtain th1 where th1_readys: "th1 \<in> readys s"
-            and h: "y = cp s th1" by auto
-          show "y \<le> cp s tm"
-            apply(unfold cp_eq_p h)
-            apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono)
-          proof -
-            from finite_threads
-            show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-          next
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}"
-              by simp
-          next
-            from dependants_threads[of th1] th1_readys
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) 
-                    \<subseteq> (\<lambda>th. preced th s) ` threads s"
-              by (auto simp:readys_def)
-          qed
-        qed
-        ultimately show " Max (cp s ` readys s) = preced tm s" by simp
-      qed 
-    qed
-  qed
-qed
-
-text {* (* ccc *) \noindent
-  Since the current precedence of the threads in ready queue will always be boosted,
-  there must be one inside it has the maximum precedence of the whole system. 
-*}
-lemma max_cp_readys_threads:
-  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
-proof(cases "threads s = {}")
-  case True
-  thus ?thesis 
-    by (auto simp:readys_def)
-next
-  case False
-  show ?thesis by (rule max_cp_readys_threads_pre[OF False])
-qed
-
-end
-
-end
-
--- a/ExtGG.thy~	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,920 +0,0 @@
-section {*
-  This file contains lemmas used to guide the recalculation of current precedence 
-  after every system call (or system operation)
-*}
-theory Implementation
-imports PIPBasics
-begin
-
-text {* (* ddd *)
-  One beauty of our modelling is that we follow the definitional extension tradition of HOL.
-  The benefit of such a concise and miniature model is that  large number of intuitively 
-  obvious facts are derived as lemmas, rather than asserted as axioms.
-*}
-
-text {*
-  However, the lemmas in the forthcoming several locales are no longer 
-  obvious. These lemmas show how the current precedences should be recalculated 
-  after every execution step (in our model, every step is represented by an event, 
-  which in turn, represents a system call, or operation). Each operation is 
-  treated in a separate locale.
-
-  The complication of current precedence recalculation comes 
-  because the changing of RAG needs to be taken into account, 
-  in addition to the changing of precedence. 
-
-  The reason RAG changing affects current precedence is that,
-  according to the definition, current precedence 
-  of a thread is the maximum of the precedences of every threads in its subtree, 
-  where the notion of sub-tree in RAG is defined in RTree.thy.
-
-  Therefore, for each operation, lemmas about the change of precedences 
-  and RAG are derived first, on which lemmas about current precedence 
-  recalculation are based on.
-*}
-
-section {* The @{term Set} operation *}
-
-text {* (* ddd *)
-  The following locale @{text "step_set_cps"} investigates the recalculation 
-  after the @{text "Set"} operation.
-*}
-locale step_set_cps =
-  fixes s' th prio s 
-  -- {* @{text "s'"} is the system state before the operation *}
-  -- {* @{text "s"} is the system state after the operation *}
-  defines s_def : "s \<equiv> (Set th prio#s')" 
-  -- {* @{text "s"} is assumed to be a legitimate state, from which
-         the legitimacy of @{text "s"} can be derived. *}
-  assumes vt_s: "vt s"
-
-sublocale step_set_cps < vat_s : valid_trace "s"
-proof
-  from vt_s show "vt s" .
-qed
-
-sublocale step_set_cps < vat_s' : valid_trace "s'"
-proof
-  from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" .
-qed
-
-context step_set_cps 
-begin
-
-text {* (* ddd *)
-  The following two lemmas confirm that @{text "Set"}-operation
-  only changes the precedence of the initiating thread (or actor)
-  of the operation (or event).
-*}
-
-lemma eq_preced:
-  assumes "th' \<noteq> th"
-  shows "preced th' s = preced th' s'"
-proof -
-  from assms show ?thesis 
-    by (unfold s_def, auto simp:preced_def)
-qed
-
-lemma eq_the_preced: 
-  assumes "th' \<noteq> th"
-  shows "the_preced s th' = the_preced s' th'"
-  using assms
-  by (unfold the_preced_def, intro eq_preced, simp)
-
-text {*
-  The following lemma assures that the resetting of priority does not change the RAG. 
-*}
-
-lemma eq_dep: "RAG s = RAG s'"
-  by (unfold s_def RAG_set_unchanged, auto)
-
-text {* (* ddd *)
-  Th following lemma @{text "eq_cp_pre"} says that the priority change of @{text "th"}
-  only affects those threads, which as @{text "Th th"} in their sub-trees.
-  
-  The proof of this lemma is simplified by using the alternative definition 
-  of @{text "cp"}. 
-*}
-
-lemma eq_cp_pre:
-  assumes nd: "Th th \<notin> subtree (RAG s') (Th th')"
-  shows "cp s th' = cp s' th'"
-proof -
-  -- {* After unfolding using the alternative definition, elements 
-        affecting the @{term "cp"}-value of threads become explicit. 
-        We only need to prove the following: *}
-  have "Max (the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')}) =
-        Max (the_preced s' ` {th'a. Th th'a \<in> subtree (RAG s') (Th th')})"
-        (is "Max (?f ` ?S1) = Max (?g ` ?S2)")
-  proof -
-    -- {* The base sets are equal. *}
-    have "?S1 = ?S2" using eq_dep by simp
-    -- {* The function values on the base set are equal as well. *}
-    moreover have "\<forall> e \<in> ?S2. ?f e = ?g e"
-    proof
-      fix th1
-      assume "th1 \<in> ?S2"
-      with nd have "th1 \<noteq> th" by (auto)
-      from eq_the_preced[OF this]
-      show "the_preced s th1 = the_preced s' th1" .
-    qed
-    -- {* Therefore, the image of the functions are equal. *}
-    ultimately have "(?f ` ?S1) = (?g ` ?S2)" by (auto intro!:f_image_eq)
-    thus ?thesis by simp
-  qed
-  thus ?thesis by (simp add:cp_alt_def)
-qed
-
-text {*
-  The following lemma shows that @{term "th"} is not in the 
-  sub-tree of any other thread. 
-*}
-lemma th_in_no_subtree:
-  assumes "th' \<noteq> th"
-  shows "Th th \<notin> subtree (RAG s') (Th th')"
-proof -
-  have "th \<in> readys s'"
-  proof -
-    from step_back_step [OF vt_s[unfolded s_def]]
-    have "step s' (Set th prio)" .
-    hence "th \<in> runing s'" by (cases, simp)
-    thus ?thesis by (simp add:readys_def runing_def)
-  qed
-  from vat_s'.readys_in_no_subtree[OF this assms(1)]
-  show ?thesis by blast
-qed
-
-text {* 
-  By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"}, 
-  it is obvious that the change of priority only affects the @{text "cp"}-value 
-  of the initiating thread @{text "th"}.
-*}
-lemma eq_cp:
-  assumes "th' \<noteq> th"
-  shows "cp s th' = cp s' th'"
-  by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]])
-
-end
-
-section {* The @{term V} operation *}
-
-text {*
-  The following @{text "step_v_cps"} is the locale for @{text "V"}-operation.
-*}
-
-locale step_v_cps =
-  -- {* @{text "th"} is the initiating thread *}
-  -- {* @{text "cs"} is the critical resource release by the @{text "V"}-operation *}
-  fixes s' th cs s    -- {* @{text "s'"} is the state before operation*}
-  defines s_def : "s \<equiv> (V th cs#s')" -- {* @{text "s"} is the state after operation*}
-  -- {* @{text "s"} is assumed to be valid, which implies the validity of @{text "s'"} *}
-  assumes vt_s: "vt s"
-
-sublocale step_v_cps < vat_s : valid_trace "s"
-proof
-  from vt_s show "vt s" .
-qed
-
-sublocale step_v_cps < vat_s' : valid_trace "s'"
-proof
-  from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" .
-qed
-
-context step_v_cps
-begin
-
-lemma ready_th_s': "th \<in> readys s'"
-  using step_back_step[OF vt_s[unfolded s_def]]
-  by (cases, simp add:runing_def)
-
-lemma ancestors_th: "ancestors (RAG s') (Th th) = {}"
-proof -
-  from vat_s'.readys_root[OF ready_th_s']
-  show ?thesis
-  by (unfold root_def, simp)
-qed
-
-lemma holding_th: "holding s' th cs"
-proof -
-  from vt_s[unfolded s_def]
-  have " PIP s' (V th cs)" by (cases, simp)
-  thus ?thesis by (cases, auto)
-qed
-
-lemma edge_of_th:
-    "(Cs cs, Th th) \<in> RAG s'" 
-proof -
- from holding_th
- show ?thesis 
-    by (unfold s_RAG_def holding_eq, auto)
-qed
-
-lemma ancestors_cs: 
-  "ancestors (RAG s') (Cs cs) = {Th th}"
-proof -
-  have "ancestors (RAG s') (Cs cs) = ancestors (RAG s') (Th th)  \<union>  {Th th}"
-  proof(rule vat_s'.rtree_RAG.ancestors_accum)
-    from vt_s[unfolded s_def]
-    have " PIP s' (V th cs)" by (cases, simp)
-    thus "(Cs cs, Th th) \<in> RAG s'" 
-    proof(cases)
-      assume "holding s' th cs"
-      from this[unfolded holding_eq]
-      show ?thesis by (unfold s_RAG_def, auto)
-    qed
-  qed
-  from this[unfolded ancestors_th] show ?thesis by simp
-qed
-
-lemma preced_kept: "the_preced s = the_preced s'"
-  by (auto simp: s_def the_preced_def preced_def)
-
-end
-
-text {*
-  The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation, 
-  which represents the case when there is another thread @{text "th'"}
-  to take over the critical resource released by the initiating thread @{text "th"}.
-*}
-locale step_v_cps_nt = step_v_cps +
-  fixes th'
-  -- {* @{text "th'"} is assumed to take over @{text "cs"} *}
-  assumes nt: "next_th s' th cs th'" 
-
-context step_v_cps_nt
-begin
-
-text {*
-  Lemma @{text "RAG_s"} confirms the change of RAG:
-  two edges removed and one added, as shown by the following diagram.
-*}
-
-(*
-  RAG before the V-operation
-    th1 ----|
-            |
-    th' ----|
-            |----> cs -----|
-    th2 ----|              |
-            |              |
-    th3 ----|              |
-                           |------> th
-    th4 ----|              |
-            |              |
-    th5 ----|              |
-            |----> cs'-----|
-    th6 ----|
-            |
-    th7 ----|
-
- RAG after the V-operation
-    th1 ----|
-            |
-            |----> cs ----> th'
-    th2 ----|              
-            |              
-    th3 ----|              
-                           
-    th4 ----|              
-            |              
-    th5 ----|              
-            |----> cs'----> th
-    th6 ----|
-            |
-    th7 ----|
-*)
-
-lemma sub_RAGs': "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s'"
-                using next_th_RAG[OF nt]  .
-
-lemma ancestors_th': 
-  "ancestors (RAG s') (Th th') = {Th th, Cs cs}" 
-proof -
-  have "ancestors (RAG s') (Th th') = ancestors (RAG s') (Cs cs) \<union> {Cs cs}"
-  proof(rule  vat_s'.rtree_RAG.ancestors_accum)
-    from sub_RAGs' show "(Th th', Cs cs) \<in> RAG s'" by auto
-  qed
-  thus ?thesis using ancestors_th ancestors_cs by auto
-qed
-
-lemma RAG_s:
-  "RAG s = (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) \<union>
-                                         {(Cs cs, Th th')}"
-proof -
-  from step_RAG_v[OF vt_s[unfolded s_def], folded s_def]
-    and nt show ?thesis  by (auto intro:next_th_unique)
-qed
-
-lemma subtree_kept: (* ddd *)
-  assumes "th1 \<notin> {th, th'}"
-  shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" (is "_ = ?R")
-proof -
-  let ?RAG' = "(RAG s' - {(Cs cs, Th th), (Th th', Cs cs)})"
-  let ?RAG'' = "?RAG' \<union> {(Cs cs, Th th')}"
-  have "subtree ?RAG' (Th th1) = ?R" 
-  proof(rule subset_del_subtree_outside)
-    show "Range {(Cs cs, Th th), (Th th', Cs cs)} \<inter> subtree (RAG s') (Th th1) = {}"
-    proof -
-      have "(Th th) \<notin> subtree (RAG s') (Th th1)"
-      proof(rule subtree_refute)
-        show "Th th1 \<notin> ancestors (RAG s') (Th th)"
-          by (unfold ancestors_th, simp)
-      next
-        from assms show "Th th1 \<noteq> Th th" by simp
-      qed
-      moreover have "(Cs cs) \<notin>  subtree (RAG s') (Th th1)"
-      proof(rule subtree_refute)
-        show "Th th1 \<notin> ancestors (RAG s') (Cs cs)"
-          by (unfold ancestors_cs, insert assms, auto)
-      qed simp
-      ultimately have "{Th th, Cs cs} \<inter> subtree (RAG s') (Th th1) = {}" by auto
-      thus ?thesis by simp
-     qed
-  qed
-  moreover have "subtree ?RAG'' (Th th1) =  subtree ?RAG' (Th th1)"
-  proof(rule subtree_insert_next)
-    show "Th th' \<notin> subtree (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th1)"
-    proof(rule subtree_refute)
-      show "Th th1 \<notin> ancestors (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th')"
-            (is "_ \<notin> ?R")
-      proof -
-          have "?R \<subseteq> ancestors (RAG s') (Th th')" by (rule ancestors_mono, auto)
-          moreover have "Th th1 \<notin> ..." using ancestors_th' assms by simp
-          ultimately show ?thesis by auto
-      qed
-    next
-      from assms show "Th th1 \<noteq> Th th'" by simp
-    qed
-  qed
-  ultimately show ?thesis by (unfold RAG_s, simp)
-qed
-
-lemma cp_kept:
-  assumes "th1 \<notin> {th, th'}"
-  shows "cp s th1 = cp s' th1"
-    by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp)
-
-end
-
-locale step_v_cps_nnt = step_v_cps +
-  assumes nnt: "\<And> th'. (\<not> next_th s' th cs th')"
-
-context step_v_cps_nnt
-begin
-
-lemma RAG_s: "RAG s = RAG s' - {(Cs cs, Th th)}"
-proof -
-  from nnt and  step_RAG_v[OF vt_s[unfolded s_def], folded s_def]
-  show ?thesis by auto
-qed
-
-lemma subtree_kept:
-  assumes "th1 \<noteq> th"
-  shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)"
-proof(unfold RAG_s, rule subset_del_subtree_outside)
-  show "Range {(Cs cs, Th th)} \<inter> subtree (RAG s') (Th th1) = {}"
-  proof -
-    have "(Th th) \<notin> subtree (RAG s') (Th th1)"
-    proof(rule subtree_refute)
-      show "Th th1 \<notin> ancestors (RAG s') (Th th)"
-          by (unfold ancestors_th, simp)
-    next
-      from assms show "Th th1 \<noteq> Th th" by simp
-    qed
-    thus ?thesis by auto
-  qed
-qed
-
-lemma cp_kept_1:
-  assumes "th1 \<noteq> th"
-  shows "cp s th1 = cp s' th1"
-    by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp)
-
-lemma subtree_cs: "subtree (RAG s') (Cs cs) = {Cs cs}"
-proof -
-  { fix n
-    have "(Cs cs) \<notin> ancestors (RAG s') n"
-    proof
-      assume "Cs cs \<in> ancestors (RAG s') n"
-      hence "(n, Cs cs) \<in> (RAG s')^+" by (auto simp:ancestors_def)
-      from tranclE[OF this] obtain nn where h: "(nn, Cs cs) \<in> RAG s'" by auto
-      then obtain th' where "nn = Th th'"
-        by (unfold s_RAG_def, auto)
-      from h[unfolded this] have "(Th th', Cs cs) \<in> RAG s'" .
-      from this[unfolded s_RAG_def]
-      have "waiting (wq s') th' cs" by auto
-      from this[unfolded cs_waiting_def]
-      have "1 < length (wq s' cs)"
-          by (cases "wq s' cs", auto)
-      from holding_next_thI[OF holding_th this]
-      obtain th' where "next_th s' th cs th'" by auto
-      with nnt show False by auto
-    qed
-  } note h = this
-  {  fix n
-     assume "n \<in> subtree (RAG s') (Cs cs)"
-     hence "n = (Cs cs)"
-     by (elim subtreeE, insert h, auto)
-  } moreover have "(Cs cs) \<in> subtree (RAG s') (Cs cs)"
-      by (auto simp:subtree_def)
-  ultimately show ?thesis by auto 
-qed
-
-lemma subtree_th: 
-  "subtree (RAG s) (Th th) = subtree (RAG s') (Th th) - {Cs cs}"
-proof(unfold RAG_s, fold subtree_cs, rule vat_s'.rtree_RAG.subtree_del_inside)
-  from edge_of_th
-  show "(Cs cs, Th th) \<in> edges_in (RAG s') (Th th)"
-    by (unfold edges_in_def, auto simp:subtree_def)
-qed
-
-lemma cp_kept_2: 
-  shows "cp s th = cp s' th" 
- by (unfold cp_alt_def subtree_th preced_kept, auto)
-
-lemma eq_cp:
-  shows "cp s th' = cp s' th'"
-  using cp_kept_1 cp_kept_2
-  by (cases "th' = th", auto)
-end
-
-
-locale step_P_cps =
-  fixes s' th cs s 
-  defines s_def : "s \<equiv> (P th cs#s')"
-  assumes vt_s: "vt s"
-
-sublocale step_P_cps < vat_s : valid_trace "s"
-proof
-  from vt_s show "vt s" .
-qed
-
-section {* The @{term P} operation *}
-
-sublocale step_P_cps < vat_s' : valid_trace "s'"
-proof
-  from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" .
-qed
-
-context step_P_cps
-begin
-
-lemma readys_th: "th \<in> readys s'"
-proof -
-    from step_back_step [OF vt_s[unfolded s_def]]
-    have "PIP s' (P th cs)" .
-    hence "th \<in> runing s'" by (cases, simp)
-    thus ?thesis by (simp add:readys_def runing_def)
-qed
-
-lemma root_th: "root (RAG s') (Th th)"
-  using readys_root[OF readys_th] .
-
-lemma in_no_others_subtree:
-  assumes "th' \<noteq> th"
-  shows "Th th \<notin> subtree (RAG s') (Th th')"
-proof
-  assume "Th th \<in> subtree (RAG s') (Th th')"
-  thus False
-  proof(cases rule:subtreeE)
-    case 1
-    with assms show ?thesis by auto
-  next
-    case 2
-    with root_th show ?thesis by (auto simp:root_def)
-  qed
-qed
-
-lemma preced_kept: "the_preced s = the_preced s'"
-  by (auto simp: s_def the_preced_def preced_def)
-
-end
-
-locale step_P_cps_ne =step_P_cps +
-  fixes th'
-  assumes ne: "wq s' cs \<noteq> []"
-  defines th'_def: "th' \<equiv> hd (wq s' cs)"
-
-locale step_P_cps_e =step_P_cps +
-  assumes ee: "wq s' cs = []"
-
-context step_P_cps_e
-begin
-
-lemma RAG_s: "RAG s = RAG s' \<union> {(Cs cs, Th th)}"
-proof -
-  from ee and  step_RAG_p[OF vt_s[unfolded s_def], folded s_def]
-  show ?thesis by auto
-qed
-
-lemma subtree_kept:
-  assumes "th' \<noteq> th"
-  shows "subtree (RAG s) (Th th') = subtree (RAG s') (Th th')"
-proof(unfold RAG_s, rule subtree_insert_next)
-  from in_no_others_subtree[OF assms] 
-  show "Th th \<notin> subtree (RAG s') (Th th')" .
-qed
-
-lemma cp_kept: 
-  assumes "th' \<noteq> th"
-  shows "cp s th' = cp s' th'"
-proof -
-  have "(the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')}) =
-        (the_preced s' ` {th'a. Th th'a \<in> subtree (RAG s') (Th th')})"
-        by (unfold preced_kept subtree_kept[OF assms], simp)
-  thus ?thesis by (unfold cp_alt_def, simp)
-qed
-
-end
-
-context step_P_cps_ne 
-begin
-
-lemma RAG_s: "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
-proof -
-  from step_RAG_p[OF vt_s[unfolded s_def]] and ne
-  show ?thesis by (simp add:s_def)
-qed
-
-lemma cs_held: "(Cs cs, Th th') \<in> RAG s'"
-proof -
-  have "(Cs cs, Th th') \<in> hRAG s'"
-  proof -
-    from ne
-    have " holding s' th' cs"
-      by (unfold th'_def holding_eq cs_holding_def, auto)
-    thus ?thesis 
-      by (unfold hRAG_def, auto)
-  qed
-  thus ?thesis by (unfold RAG_split, auto)
-qed
-
-lemma tRAG_s: 
-  "tRAG s = tRAG s' \<union> {(Th th, Th th')}"
-  using RAG_tRAG_transfer[OF RAG_s cs_held] .
-
-lemma cp_kept:
-  assumes "Th th'' \<notin> ancestors (tRAG s) (Th th)"
-  shows "cp s th'' = cp s' th''"
-proof -
-  have h: "subtree (tRAG s) (Th th'') = subtree (tRAG s') (Th th'')"
-  proof -
-    have "Th th' \<notin> subtree (tRAG s') (Th th'')"
-    proof
-      assume "Th th' \<in> subtree (tRAG s') (Th th'')"
-      thus False
-      proof(rule subtreeE)
-         assume "Th th' = Th th''"
-         from assms[unfolded tRAG_s ancestors_def, folded this]
-         show ?thesis by auto
-      next
-         assume "Th th'' \<in> ancestors (tRAG s') (Th th')"
-         moreover have "... \<subseteq> ancestors (tRAG s) (Th th')"
-         proof(rule ancestors_mono)
-            show "tRAG s' \<subseteq> tRAG s" by (unfold tRAG_s, auto)
-         qed 
-         ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th')" by auto
-         moreover have "Th th' \<in> ancestors (tRAG s) (Th th)"
-           by (unfold tRAG_s, auto simp:ancestors_def)
-         ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th)"
-                       by (auto simp:ancestors_def)
-         with assms show ?thesis by auto
-      qed
-    qed
-    from subtree_insert_next[OF this]
-    have "subtree (tRAG s' \<union> {(Th th, Th th')}) (Th th'') = subtree (tRAG s') (Th th'')" .
-    from this[folded tRAG_s] show ?thesis .
-  qed
-  show ?thesis by (unfold cp_alt_def1 h preced_kept, simp)
-qed
-
-lemma cp_gen_update_stop: (* ddd *)
-  assumes "u \<in> ancestors (tRAG s) (Th th)"
-  and "cp_gen s u = cp_gen s' u"
-  and "y \<in> ancestors (tRAG s) u"
-  shows "cp_gen s y = cp_gen s' y"
-  using assms(3)
-proof(induct rule:wf_induct[OF vat_s.fsbttRAGs.wf])
-  case (1 x)
-  show ?case (is "?L = ?R")
-  proof -
-    from tRAG_ancestorsE[OF 1(2)]
-    obtain th2 where eq_x: "x = Th th2" by blast
-    from vat_s.cp_gen_rec[OF this]
-    have "?L = 
-          Max ({the_preced s th2} \<union> cp_gen s ` RTree.children (tRAG s) x)" .
-    also have "... = 
-          Max ({the_preced s' th2} \<union> cp_gen s' ` RTree.children (tRAG s') x)"
-  
-    proof -
-      from preced_kept have "the_preced s th2 = the_preced s' th2" by simp
-      moreover have "cp_gen s ` RTree.children (tRAG s) x =
-                     cp_gen s' ` RTree.children (tRAG s') x"
-      proof -
-        have "RTree.children (tRAG s) x =  RTree.children (tRAG s') x"
-        proof(unfold tRAG_s, rule children_union_kept)
-          have start: "(Th th, Th th') \<in> tRAG s"
-            by (unfold tRAG_s, auto)
-          note x_u = 1(2)
-          show "x \<notin> Range {(Th th, Th th')}"
-          proof
-            assume "x \<in> Range {(Th th, Th th')}"
-            hence eq_x: "x = Th th'" using RangeE by auto
-            show False
-            proof(cases rule:vat_s.rtree_s.ancestors_headE[OF assms(1) start])
-              case 1
-              from x_u[folded this, unfolded eq_x] vat_s.acyclic_tRAG
-              show ?thesis by (auto simp:ancestors_def acyclic_def)
-            next
-              case 2
-              with x_u[unfolded eq_x]
-              have "(Th th', Th th') \<in> (tRAG s)^+" by (auto simp:ancestors_def)
-              with vat_s.acyclic_tRAG show ?thesis by (auto simp:acyclic_def)
-            qed
-          qed
-        qed
-        moreover have "cp_gen s ` RTree.children (tRAG s) x =
-                       cp_gen s' ` RTree.children (tRAG s) x" (is "?f ` ?A = ?g ` ?A")
-        proof(rule f_image_eq)
-          fix a
-          assume a_in: "a \<in> ?A"
-          from 1(2)
-          show "?f a = ?g a"
-          proof(cases rule:vat_s.rtree_s.ancestors_childrenE[case_names in_ch out_ch])
-             case in_ch
-             show ?thesis
-             proof(cases "a = u")
-                case True
-                from assms(2)[folded this] show ?thesis .
-             next
-                case False
-                have a_not_in: "a \<notin> ancestors (tRAG s) (Th th)"
-                proof
-                  assume a_in': "a \<in> ancestors (tRAG s) (Th th)"
-                  have "a = u"
-                  proof(rule vat_s.rtree_s.ancestors_children_unique)
-                    from a_in' a_in show "a \<in> ancestors (tRAG s) (Th th) \<inter> 
-                                          RTree.children (tRAG s) x" by auto
-                  next 
-                    from assms(1) in_ch show "u \<in> ancestors (tRAG s) (Th th) \<inter> 
-                                      RTree.children (tRAG s) x" by auto
-                  qed
-                  with False show False by simp
-                qed
-                from a_in obtain th_a where eq_a: "a = Th th_a" 
-                    by (unfold RTree.children_def tRAG_alt_def, auto)
-                from cp_kept[OF a_not_in[unfolded eq_a]]
-                have "cp s th_a = cp s' th_a" .
-                from this [unfolded cp_gen_def_cond[OF eq_a], folded eq_a]
-                show ?thesis .
-             qed
-          next
-            case (out_ch z)
-            hence h: "z \<in> ancestors (tRAG s) u" "z \<in> RTree.children (tRAG s) x" by auto
-            show ?thesis
-            proof(cases "a = z")
-              case True
-              from h(2) have zx_in: "(z, x) \<in> (tRAG s)" by (auto simp:RTree.children_def)
-              from 1(1)[rule_format, OF this h(1)]
-              have eq_cp_gen: "cp_gen s z = cp_gen s' z" .
-              with True show ?thesis by metis
-            next
-              case False
-              from a_in obtain th_a where eq_a: "a = Th th_a"
-                by (auto simp:RTree.children_def tRAG_alt_def)
-              have "a \<notin> ancestors (tRAG s) (Th th)"
-              proof
-                assume a_in': "a \<in> ancestors (tRAG s) (Th th)"
-                have "a = z"
-                proof(rule vat_s.rtree_s.ancestors_children_unique)
-                  from assms(1) h(1) have "z \<in> ancestors (tRAG s) (Th th)"
-                      by (auto simp:ancestors_def)
-                  with h(2) show " z \<in> ancestors (tRAG s) (Th th) \<inter> 
-                                       RTree.children (tRAG s) x" by auto
-                next
-                  from a_in a_in'
-                  show "a \<in> ancestors (tRAG s) (Th th) \<inter> RTree.children (tRAG s) x"
-                    by auto
-                qed
-                with False show False by auto
-              qed
-              from cp_kept[OF this[unfolded eq_a]]
-              have "cp s th_a = cp s' th_a" .
-              from this[unfolded cp_gen_def_cond[OF eq_a], folded eq_a]
-              show ?thesis .
-            qed
-          qed
-        qed
-        ultimately show ?thesis by metis
-      qed
-      ultimately show ?thesis by simp
-    qed
-    also have "... = ?R"
-      by (fold vat_s'.cp_gen_rec[OF eq_x], simp)
-    finally show ?thesis .
-  qed
-qed
-
-lemma cp_up:
-  assumes "(Th th') \<in> ancestors (tRAG s) (Th th)"
-  and "cp s th' = cp s' th'"
-  and "(Th th'') \<in> ancestors (tRAG s) (Th th')"
-  shows "cp s th'' = cp s' th''"
-proof -
-  have "cp_gen s (Th th'') = cp_gen s' (Th th'')"
-  proof(rule cp_gen_update_stop[OF assms(1) _ assms(3)])
-    from assms(2) cp_gen_def_cond[OF refl[of "Th th'"]]
-    show "cp_gen s (Th th') = cp_gen s' (Th th')" by metis
-  qed
-  with cp_gen_def_cond[OF refl[of "Th th''"]]
-  show ?thesis by metis
-qed
-
-end
-
-section {* The @{term Create} operation *}
-
-locale step_create_cps =
-  fixes s' th prio s 
-  defines s_def : "s \<equiv> (Create th prio#s')"
-  assumes vt_s: "vt s"
-
-sublocale step_create_cps < vat_s: valid_trace "s"
-  by (unfold_locales, insert vt_s, simp)
-
-sublocale step_create_cps < vat_s': valid_trace "s'"
-  by (unfold_locales, insert step_back_vt[OF vt_s[unfolded s_def]], simp)
-
-context step_create_cps
-begin
-
-lemma RAG_kept: "RAG s = RAG s'"
-  by (unfold s_def RAG_create_unchanged, auto)
-
-lemma tRAG_kept: "tRAG s = tRAG s'"
-  by (unfold tRAG_alt_def RAG_kept, auto)
-
-lemma preced_kept:
-  assumes "th' \<noteq> th"
-  shows "the_preced s th' = the_preced s' th'"
-  by (unfold s_def the_preced_def preced_def, insert assms, auto)
-
-lemma th_not_in: "Th th \<notin> Field (tRAG s')"
-proof -
-  from vt_s[unfolded s_def]
-  have "PIP s' (Create th prio)" by (cases, simp)
-  hence "th \<notin> threads s'" by(cases, simp)
-  from vat_s'.not_in_thread_isolated[OF this]
-  have "Th th \<notin> Field (RAG s')" .
-  with tRAG_Field show ?thesis by auto
-qed
-
-lemma eq_cp:
-  assumes neq_th: "th' \<noteq> th"
-  shows "cp s th' = cp s' th'"
-proof -
-  have "(the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th') =
-        (the_preced s' \<circ> the_thread) ` subtree (tRAG s') (Th th')"
-  proof(unfold tRAG_kept, rule f_image_eq)
-    fix a
-    assume a_in: "a \<in> subtree (tRAG s') (Th th')"
-    then obtain th_a where eq_a: "a = Th th_a" 
-    proof(cases rule:subtreeE)
-      case 2
-      from ancestors_Field[OF 2(2)]
-      and that show ?thesis by (unfold tRAG_alt_def, auto)
-    qed auto
-    have neq_th_a: "th_a \<noteq> th"
-    proof -
-      have "(Th th) \<notin> subtree (tRAG s') (Th th')"
-      proof
-        assume "Th th \<in> subtree (tRAG s') (Th th')"
-        thus False
-        proof(cases rule:subtreeE)
-          case 2
-          from ancestors_Field[OF this(2)]
-          and th_not_in[unfolded Field_def]
-          show ?thesis by auto
-        qed (insert assms, auto)
-      qed
-      with a_in[unfolded eq_a] show ?thesis by auto
-    qed
-    from preced_kept[OF this]
-    show "(the_preced s \<circ> the_thread) a = (the_preced s' \<circ> the_thread) a"
-      by (unfold eq_a, simp)
-  qed
-  thus ?thesis by (unfold cp_alt_def1, simp)
-qed
-
-lemma children_of_th: "RTree.children (tRAG s) (Th th) = {}"
-proof -
-  { fix a
-    assume "a \<in> RTree.children (tRAG s) (Th th)"
-    hence "(a, Th th) \<in> tRAG s" by (auto simp:RTree.children_def)
-    with th_not_in have False 
-     by (unfold Field_def tRAG_kept, auto)
-  } thus ?thesis by auto
-qed
-
-lemma eq_cp_th: "cp s th = preced th s"
- by (unfold vat_s.cp_rec children_of_th, simp add:the_preced_def)
-
-end
-
-locale step_exit_cps =
-  fixes s' th prio s 
-  defines s_def : "s \<equiv> Exit th # s'"
-  assumes vt_s: "vt s"
-
-sublocale step_exit_cps < vat_s: valid_trace "s"
-  by (unfold_locales, insert vt_s, simp)
-
-sublocale step_exit_cps < vat_s': valid_trace "s'"
-  by (unfold_locales, insert step_back_vt[OF vt_s[unfolded s_def]], simp)
-
-context step_exit_cps
-begin
-
-lemma preced_kept:
-  assumes "th' \<noteq> th"
-  shows "the_preced s th' = the_preced s' th'"
-  by (unfold s_def the_preced_def preced_def, insert assms, auto)
-
-lemma RAG_kept: "RAG s = RAG s'"
-  by (unfold s_def RAG_exit_unchanged, auto)
-
-lemma tRAG_kept: "tRAG s = tRAG s'"
-  by (unfold tRAG_alt_def RAG_kept, auto)
-
-lemma th_ready: "th \<in> readys s'"
-proof -
-  from vt_s[unfolded s_def]
-  have "PIP s' (Exit th)" by (cases, simp)
-  hence h: "th \<in> runing s' \<and> holdents s' th = {}" by (cases, metis)
-  thus ?thesis by (unfold runing_def, auto)
-qed
-
-lemma th_holdents: "holdents s' th = {}"
-proof -
- from vt_s[unfolded s_def]
-  have "PIP s' (Exit th)" by (cases, simp)
-  thus ?thesis by (cases, metis)
-qed
-
-lemma th_RAG: "Th th \<notin> Field (RAG s')"
-proof -
-  have "Th th \<notin> Range (RAG s')"
-  proof
-    assume "Th th \<in> Range (RAG s')"
-    then obtain cs where "holding (wq s') th cs"
-      by (unfold Range_iff s_RAG_def, auto)
-    with th_holdents[unfolded holdents_def]
-    show False by (unfold eq_holding, auto)
-  qed
-  moreover have "Th th \<notin> Domain (RAG s')"
-  proof
-    assume "Th th \<in> Domain (RAG s')"
-    then obtain cs where "waiting (wq s') th cs"
-      by (unfold Domain_iff s_RAG_def, auto)
-    with th_ready show False by (unfold readys_def eq_waiting, auto)
-  qed
-  ultimately show ?thesis by (auto simp:Field_def)
-qed
-
-lemma th_tRAG: "(Th th) \<notin> Field (tRAG s')"
-  using th_RAG tRAG_Field[of s'] by auto
-
-lemma eq_cp:
-  assumes neq_th: "th' \<noteq> th"
-  shows "cp s th' = cp s' th'"
-proof -
-  have "(the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th') =
-        (the_preced s' \<circ> the_thread) ` subtree (tRAG s') (Th th')"
-  proof(unfold tRAG_kept, rule f_image_eq)
-    fix a
-    assume a_in: "a \<in> subtree (tRAG s') (Th th')"
-    then obtain th_a where eq_a: "a = Th th_a" 
-    proof(cases rule:subtreeE)
-      case 2
-      from ancestors_Field[OF 2(2)]
-      and that show ?thesis by (unfold tRAG_alt_def, auto)
-    qed auto
-    have neq_th_a: "th_a \<noteq> th"
-    proof -
-      from vat_s'.readys_in_no_subtree[OF th_ready assms]
-      have "(Th th) \<notin> subtree (RAG s') (Th th')" .
-      with tRAG_subtree_RAG[of s' "Th th'"]
-      have "(Th th) \<notin> subtree (tRAG s') (Th th')" by auto
-      with a_in[unfolded eq_a] show ?thesis by auto
-    qed
-    from preced_kept[OF this]
-    show "(the_preced s \<circ> the_thread) a = (the_preced s' \<circ> the_thread) a"
-      by (unfold eq_a, simp)
-  qed
-  thus ?thesis by (unfold cp_alt_def1, simp)
-qed
-
-end
-
-end
-
--- a/Implementation.thy~	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,913 +0,0 @@
-section {*
-  This file contains lemmas used to guide the recalculation of current precedence 
-  after every system call (or system operation)
-*}
-theory Implementation
-imports PIPBasics
-begin
-
-text {* (* ddd *)
-  One beauty of our modelling is that we follow the definitional extension tradition of HOL.
-  The benefit of such a concise and miniature model is that  large number of intuitively 
-  obvious facts are derived as lemmas, rather than asserted as axioms.
-*}
-
-text {*
-  However, the lemmas in the forthcoming several locales are no longer 
-  obvious. These lemmas show how the current precedences should be recalculated 
-  after every execution step (in our model, every step is represented by an event, 
-  which in turn, represents a system call, or operation). Each operation is 
-  treated in a separate locale.
-
-  The complication of current precedence recalculation comes 
-  because the changing of RAG needs to be taken into account, 
-  in addition to the changing of precedence. 
-  The reason RAG changing affects current precedence is that,
-  according to the definition, current precedence 
-  of a thread is the maximum of the precedences of its dependants, 
-  where the dependants are defined in terms of RAG.
-
-  Therefore, each operation, lemmas concerning the change of the precedences 
-  and RAG are derived first, so that the lemmas about
-  current precedence recalculation can be based on.
-*}
-
-text {* (* ddd *)
-  The following locale @{text "step_set_cps"} investigates the recalculation 
-  after the @{text "Set"} operation.
-*}
-locale step_set_cps =
-  fixes s' th prio s 
-  -- {* @{text "s'"} is the system state before the operation *}
-  -- {* @{text "s"} is the system state after the operation *}
-  defines s_def : "s \<equiv> (Set th prio#s')" 
-  -- {* @{text "s"} is assumed to be a legitimate state, from which
-         the legitimacy of @{text "s"} can be derived. *}
-  assumes vt_s: "vt s"
-
-sublocale step_set_cps < vat_s : valid_trace "s"
-proof
-  from vt_s show "vt s" .
-qed
-
-sublocale step_set_cps < vat_s' : valid_trace "s'"
-proof
-  from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" .
-qed
-
-context step_set_cps 
-begin
-
-text {* (* ddd *)
-  The following two lemmas confirm that @{text "Set"}-operating only changes the precedence 
-  of the initiating thread.
-*}
-
-lemma eq_preced:
-  assumes "th' \<noteq> th"
-  shows "preced th' s = preced th' s'"
-proof -
-  from assms show ?thesis 
-    by (unfold s_def, auto simp:preced_def)
-qed
-
-lemma eq_the_preced: 
-  fixes th'
-  assumes "th' \<noteq> th"
-  shows "the_preced s th' = the_preced s' th'"
-  using assms
-  by (unfold the_preced_def, intro eq_preced, simp)
-
-text {*
-  The following lemma assures that the resetting of priority does not change the RAG. 
-*}
-
-lemma eq_dep: "RAG s = RAG s'"
-  by (unfold s_def RAG_set_unchanged, auto)
-
-text {* (* ddd *)
-  Th following lemma @{text "eq_cp_pre"} says the priority change of @{text "th"}
-  only affects those threads, which as @{text "Th th"} in their sub-trees.
-  
-  The proof of this lemma is simplified by using the alternative definition of @{text "cp"}. 
-*}
-
-lemma eq_cp_pre:
-  fixes th' 
-  assumes nd: "Th th \<notin> subtree (RAG s') (Th th')"
-  shows "cp s th' = cp s' th'"
-proof -
-  -- {* After unfolding using the alternative definition, elements 
-        affecting the @{term "cp"}-value of threads become explicit. 
-        We only need to prove the following: *}
-  have "Max (the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')}) =
-        Max (the_preced s' ` {th'a. Th th'a \<in> subtree (RAG s') (Th th')})"
-        (is "Max (?f ` ?S1) = Max (?g ` ?S2)")
-  proof -
-    -- {* The base sets are equal. *}
-    have "?S1 = ?S2" using eq_dep by simp
-    -- {* The function values on the base set are equal as well. *}
-    moreover have "\<forall> e \<in> ?S2. ?f e = ?g e"
-    proof
-      fix th1
-      assume "th1 \<in> ?S2"
-      with nd have "th1 \<noteq> th" by (auto)
-      from eq_the_preced[OF this]
-      show "the_preced s th1 = the_preced s' th1" .
-    qed
-    -- {* Therefore, the image of the functions are equal. *}
-    ultimately have "(?f ` ?S1) = (?g ` ?S2)" by (auto intro!:f_image_eq)
-    thus ?thesis by simp
-  qed
-  thus ?thesis by (simp add:cp_alt_def)
-qed
-
-text {*
-  The following lemma shows that @{term "th"} is not in the 
-  sub-tree of any other thread. 
-*}
-lemma th_in_no_subtree:
-  assumes "th' \<noteq> th"
-  shows "Th th \<notin> subtree (RAG s') (Th th')"
-proof -
-  have "th \<in> readys s'"
-  proof -
-    from step_back_step [OF vt_s[unfolded s_def]]
-    have "step s' (Set th prio)" .
-    hence "th \<in> runing s'" by (cases, simp)
-    thus ?thesis by (simp add:readys_def runing_def)
-  qed
-  from vat_s'.readys_in_no_subtree[OF this assms(1)]
-  show ?thesis by blast
-qed
-
-text {* 
-  By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"}, 
-  it is obvious that the change of priority only affects the @{text "cp"}-value 
-  of the initiating thread @{text "th"}.
-*}
-lemma eq_cp:
-  fixes th' 
-  assumes "th' \<noteq> th"
-  shows "cp s th' = cp s' th'"
-  by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]])
-
-end
-
-text {*
-  The following @{text "step_v_cps"} is the locale for @{text "V"}-operation.
-*}
-
-locale step_v_cps =
-  -- {* @{text "th"} is the initiating thread *}
-  -- {* @{text "cs"} is the critical resource release by the @{text "V"}-operation *}
-  fixes s' th cs s    -- {* @{text "s'"} is the state before operation*}
-  defines s_def : "s \<equiv> (V th cs#s')" -- {* @{text "s"} is the state after operation*}
-  -- {* @{text "s"} is assumed to be valid, which implies the validity of @{text "s'"} *}
-  assumes vt_s: "vt s"
-
-sublocale step_v_cps < vat_s : valid_trace "s"
-proof
-  from vt_s show "vt s" .
-qed
-
-sublocale step_v_cps < vat_s' : valid_trace "s'"
-proof
-  from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" .
-qed
-
-context step_v_cps
-begin
-
-lemma ready_th_s': "th \<in> readys s'"
-  using step_back_step[OF vt_s[unfolded s_def]]
-  by (cases, simp add:runing_def)
-
-lemma ancestors_th: "ancestors (RAG s') (Th th) = {}"
-proof -
-  from vat_s'.readys_root[OF ready_th_s']
-  show ?thesis
-  by (unfold root_def, simp)
-qed
-
-lemma holding_th: "holding s' th cs"
-proof -
-  from vt_s[unfolded s_def]
-  have " PIP s' (V th cs)" by (cases, simp)
-  thus ?thesis by (cases, auto)
-qed
-
-lemma edge_of_th:
-    "(Cs cs, Th th) \<in> RAG s'" 
-proof -
- from holding_th
- show ?thesis 
-    by (unfold s_RAG_def holding_eq, auto)
-qed
-
-lemma ancestors_cs: 
-  "ancestors (RAG s') (Cs cs) = {Th th}"
-proof -
-  have "ancestors (RAG s') (Cs cs) = ancestors (RAG s') (Th th)  \<union>  {Th th}"
-  proof(rule vat_s'.rtree_RAG.ancestors_accum)
-    from vt_s[unfolded s_def]
-    have " PIP s' (V th cs)" by (cases, simp)
-    thus "(Cs cs, Th th) \<in> RAG s'" 
-    proof(cases)
-      assume "holding s' th cs"
-      from this[unfolded holding_eq]
-      show ?thesis by (unfold s_RAG_def, auto)
-    qed
-  qed
-  from this[unfolded ancestors_th] show ?thesis by simp
-qed
-
-lemma preced_kept: "the_preced s = the_preced s'"
-  by (auto simp: s_def the_preced_def preced_def)
-
-end
-
-text {*
-  The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation, 
-  which represents the case when there is another thread @{text "th'"}
-  to take over the critical resource released by the initiating thread @{text "th"}.
-*}
-locale step_v_cps_nt = step_v_cps +
-  fixes th'
-  -- {* @{text "th'"} is assumed to take over @{text "cs"} *}
-  assumes nt: "next_th s' th cs th'" 
-
-context step_v_cps_nt
-begin
-
-text {*
-  Lemma @{text "RAG_s"} confirms the change of RAG:
-  two edges removed and one added, as shown by the following diagram.
-*}
-
-(*
-  RAG before the V-operation
-    th1 ----|
-            |
-    th' ----|
-            |----> cs -----|
-    th2 ----|              |
-            |              |
-    th3 ----|              |
-                           |------> th
-    th4 ----|              |
-            |              |
-    th5 ----|              |
-            |----> cs'-----|
-    th6 ----|
-            |
-    th7 ----|
-
- RAG after the V-operation
-    th1 ----|
-            |
-            |----> cs ----> th'
-    th2 ----|              
-            |              
-    th3 ----|              
-                           
-    th4 ----|              
-            |              
-    th5 ----|              
-            |----> cs'----> th
-    th6 ----|
-            |
-    th7 ----|
-*)
-
-lemma sub_RAGs': "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s'"
-                using next_th_RAG[OF nt]  .
-
-lemma ancestors_th': 
-  "ancestors (RAG s') (Th th') = {Th th, Cs cs}" 
-proof -
-  have "ancestors (RAG s') (Th th') = ancestors (RAG s') (Cs cs) \<union> {Cs cs}"
-  proof(rule  vat_s'.rtree_RAG.ancestors_accum)
-    from sub_RAGs' show "(Th th', Cs cs) \<in> RAG s'" by auto
-  qed
-  thus ?thesis using ancestors_th ancestors_cs by auto
-qed
-
-lemma RAG_s:
-  "RAG s = (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) \<union>
-                                         {(Cs cs, Th th')}"
-proof -
-  from step_RAG_v[OF vt_s[unfolded s_def], folded s_def]
-    and nt show ?thesis  by (auto intro:next_th_unique)
-qed
-
-lemma subtree_kept:
-  assumes "th1 \<notin> {th, th'}"
-  shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" (is "_ = ?R")
-proof -
-  let ?RAG' = "(RAG s' - {(Cs cs, Th th), (Th th', Cs cs)})"
-  let ?RAG'' = "?RAG' \<union> {(Cs cs, Th th')}"
-  have "subtree ?RAG' (Th th1) = ?R" 
-  proof(rule subset_del_subtree_outside)
-    show "Range {(Cs cs, Th th), (Th th', Cs cs)} \<inter> subtree (RAG s') (Th th1) = {}"
-    proof -
-      have "(Th th) \<notin> subtree (RAG s') (Th th1)"
-      proof(rule subtree_refute)
-        show "Th th1 \<notin> ancestors (RAG s') (Th th)"
-          by (unfold ancestors_th, simp)
-      next
-        from assms show "Th th1 \<noteq> Th th" by simp
-      qed
-      moreover have "(Cs cs) \<notin>  subtree (RAG s') (Th th1)"
-      proof(rule subtree_refute)
-        show "Th th1 \<notin> ancestors (RAG s') (Cs cs)"
-          by (unfold ancestors_cs, insert assms, auto)
-      qed simp
-      ultimately have "{Th th, Cs cs} \<inter> subtree (RAG s') (Th th1) = {}" by auto
-      thus ?thesis by simp
-     qed
-  qed
-  moreover have "subtree ?RAG'' (Th th1) =  subtree ?RAG' (Th th1)"
-  proof(rule subtree_insert_next)
-    show "Th th' \<notin> subtree (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th1)"
-    proof(rule subtree_refute)
-      show "Th th1 \<notin> ancestors (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th')"
-            (is "_ \<notin> ?R")
-      proof -
-          have "?R \<subseteq> ancestors (RAG s') (Th th')" by (rule ancestors_mono, auto)
-          moreover have "Th th1 \<notin> ..." using ancestors_th' assms by simp
-          ultimately show ?thesis by auto
-      qed
-    next
-      from assms show "Th th1 \<noteq> Th th'" by simp
-    qed
-  qed
-  ultimately show ?thesis by (unfold RAG_s, simp)
-qed
-
-lemma cp_kept:
-  assumes "th1 \<notin> {th, th'}"
-  shows "cp s th1 = cp s' th1"
-    by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp)
-
-end
-
-locale step_v_cps_nnt = step_v_cps +
-  assumes nnt: "\<And> th'. (\<not> next_th s' th cs th')"
-
-context step_v_cps_nnt
-begin
-
-lemma RAG_s: "RAG s = RAG s' - {(Cs cs, Th th)}"
-proof -
-  from nnt and  step_RAG_v[OF vt_s[unfolded s_def], folded s_def]
-  show ?thesis by auto
-qed
-
-lemma subtree_kept:
-  assumes "th1 \<noteq> th"
-  shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)"
-proof(unfold RAG_s, rule subset_del_subtree_outside)
-  show "Range {(Cs cs, Th th)} \<inter> subtree (RAG s') (Th th1) = {}"
-  proof -
-    have "(Th th) \<notin> subtree (RAG s') (Th th1)"
-    proof(rule subtree_refute)
-      show "Th th1 \<notin> ancestors (RAG s') (Th th)"
-          by (unfold ancestors_th, simp)
-    next
-      from assms show "Th th1 \<noteq> Th th" by simp
-    qed
-    thus ?thesis by auto
-  qed
-qed
-
-lemma cp_kept_1:
-  assumes "th1 \<noteq> th"
-  shows "cp s th1 = cp s' th1"
-    by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp)
-
-lemma subtree_cs: "subtree (RAG s') (Cs cs) = {Cs cs}"
-proof -
-  { fix n
-    have "(Cs cs) \<notin> ancestors (RAG s') n"
-    proof
-      assume "Cs cs \<in> ancestors (RAG s') n"
-      hence "(n, Cs cs) \<in> (RAG s')^+" by (auto simp:ancestors_def)
-      from tranclE[OF this] obtain nn where h: "(nn, Cs cs) \<in> RAG s'" by auto
-      then obtain th' where "nn = Th th'"
-        by (unfold s_RAG_def, auto)
-      from h[unfolded this] have "(Th th', Cs cs) \<in> RAG s'" .
-      from this[unfolded s_RAG_def]
-      have "waiting (wq s') th' cs" by auto
-      from this[unfolded cs_waiting_def]
-      have "1 < length (wq s' cs)"
-          by (cases "wq s' cs", auto)
-      from holding_next_thI[OF holding_th this]
-      obtain th' where "next_th s' th cs th'" by auto
-      with nnt show False by auto
-    qed
-  } note h = this
-  {  fix n
-     assume "n \<in> subtree (RAG s') (Cs cs)"
-     hence "n = (Cs cs)"
-     by (elim subtreeE, insert h, auto)
-  } moreover have "(Cs cs) \<in> subtree (RAG s') (Cs cs)"
-      by (auto simp:subtree_def)
-  ultimately show ?thesis by auto 
-qed
-
-lemma subtree_th: 
-  "subtree (RAG s) (Th th) = subtree (RAG s') (Th th) - {Cs cs}"
-proof(unfold RAG_s, fold subtree_cs, rule vat_s'.rtree_RAG.subtree_del_inside)
-  from edge_of_th
-  show "(Cs cs, Th th) \<in> edges_in (RAG s') (Th th)"
-    by (unfold edges_in_def, auto simp:subtree_def)
-qed
-
-lemma cp_kept_2: 
-  shows "cp s th = cp s' th" 
- by (unfold cp_alt_def subtree_th preced_kept, auto)
-
-lemma eq_cp:
-  fixes th' 
-  shows "cp s th' = cp s' th'"
-  using cp_kept_1 cp_kept_2
-  by (cases "th' = th", auto)
-end
-
-
-locale step_P_cps =
-  fixes s' th cs s 
-  defines s_def : "s \<equiv> (P th cs#s')"
-  assumes vt_s: "vt s"
-
-sublocale step_P_cps < vat_s : valid_trace "s"
-proof
-  from vt_s show "vt s" .
-qed
-
-sublocale step_P_cps < vat_s' : valid_trace "s'"
-proof
-  from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" .
-qed
-
-context step_P_cps
-begin
-
-lemma readys_th: "th \<in> readys s'"
-proof -
-    from step_back_step [OF vt_s[unfolded s_def]]
-    have "PIP s' (P th cs)" .
-    hence "th \<in> runing s'" by (cases, simp)
-    thus ?thesis by (simp add:readys_def runing_def)
-qed
-
-lemma root_th: "root (RAG s') (Th th)"
-  using readys_root[OF readys_th] .
-
-lemma in_no_others_subtree:
-  assumes "th' \<noteq> th"
-  shows "Th th \<notin> subtree (RAG s') (Th th')"
-proof
-  assume "Th th \<in> subtree (RAG s') (Th th')"
-  thus False
-  proof(cases rule:subtreeE)
-    case 1
-    with assms show ?thesis by auto
-  next
-    case 2
-    with root_th show ?thesis by (auto simp:root_def)
-  qed
-qed
-
-lemma preced_kept: "the_preced s = the_preced s'"
-  by (auto simp: s_def the_preced_def preced_def)
-
-end
-
-locale step_P_cps_ne =step_P_cps +
-  fixes th'
-  assumes ne: "wq s' cs \<noteq> []"
-  defines th'_def: "th' \<equiv> hd (wq s' cs)"
-
-locale step_P_cps_e =step_P_cps +
-  assumes ee: "wq s' cs = []"
-
-context step_P_cps_e
-begin
-
-lemma RAG_s: "RAG s = RAG s' \<union> {(Cs cs, Th th)}"
-proof -
-  from ee and  step_RAG_p[OF vt_s[unfolded s_def], folded s_def]
-  show ?thesis by auto
-qed
-
-lemma subtree_kept:
-  assumes "th' \<noteq> th"
-  shows "subtree (RAG s) (Th th') = subtree (RAG s') (Th th')"
-proof(unfold RAG_s, rule subtree_insert_next)
-  from in_no_others_subtree[OF assms] 
-  show "Th th \<notin> subtree (RAG s') (Th th')" .
-qed
-
-lemma cp_kept: 
-  assumes "th' \<noteq> th"
-  shows "cp s th' = cp s' th'"
-proof -
-  have "(the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')}) =
-        (the_preced s' ` {th'a. Th th'a \<in> subtree (RAG s') (Th th')})"
-        by (unfold preced_kept subtree_kept[OF assms], simp)
-  thus ?thesis by (unfold cp_alt_def, simp)
-qed
-
-end
-
-context step_P_cps_ne 
-begin
-
-lemma RAG_s: "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
-proof -
-  from step_RAG_p[OF vt_s[unfolded s_def]] and ne
-  show ?thesis by (simp add:s_def)
-qed
-
-lemma cs_held: "(Cs cs, Th th') \<in> RAG s'"
-proof -
-  have "(Cs cs, Th th') \<in> hRAG s'"
-  proof -
-    from ne
-    have " holding s' th' cs"
-      by (unfold th'_def holding_eq cs_holding_def, auto)
-    thus ?thesis 
-      by (unfold hRAG_def, auto)
-  qed
-  thus ?thesis by (unfold RAG_split, auto)
-qed
-
-lemma tRAG_s: 
-  "tRAG s = tRAG s' \<union> {(Th th, Th th')}"
-  using RAG_tRAG_transfer[OF RAG_s cs_held] .
-
-lemma cp_kept:
-  assumes "Th th'' \<notin> ancestors (tRAG s) (Th th)"
-  shows "cp s th'' = cp s' th''"
-proof -
-  have h: "subtree (tRAG s) (Th th'') = subtree (tRAG s') (Th th'')"
-  proof -
-    have "Th th' \<notin> subtree (tRAG s') (Th th'')"
-    proof
-      assume "Th th' \<in> subtree (tRAG s') (Th th'')"
-      thus False
-      proof(rule subtreeE)
-         assume "Th th' = Th th''"
-         from assms[unfolded tRAG_s ancestors_def, folded this]
-         show ?thesis by auto
-      next
-         assume "Th th'' \<in> ancestors (tRAG s') (Th th')"
-         moreover have "... \<subseteq> ancestors (tRAG s) (Th th')"
-         proof(rule ancestors_mono)
-            show "tRAG s' \<subseteq> tRAG s" by (unfold tRAG_s, auto)
-         qed 
-         ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th')" by auto
-         moreover have "Th th' \<in> ancestors (tRAG s) (Th th)"
-           by (unfold tRAG_s, auto simp:ancestors_def)
-         ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th)"
-                       by (auto simp:ancestors_def)
-         with assms show ?thesis by auto
-      qed
-    qed
-    from subtree_insert_next[OF this]
-    have "subtree (tRAG s' \<union> {(Th th, Th th')}) (Th th'') = subtree (tRAG s') (Th th'')" .
-    from this[folded tRAG_s] show ?thesis .
-  qed
-  show ?thesis by (unfold cp_alt_def1 h preced_kept, simp)
-qed
-
-lemma cp_gen_update_stop: (* ddd *)
-  assumes "u \<in> ancestors (tRAG s) (Th th)"
-  and "cp_gen s u = cp_gen s' u"
-  and "y \<in> ancestors (tRAG s) u"
-  shows "cp_gen s y = cp_gen s' y"
-  using assms(3)
-proof(induct rule:wf_induct[OF vat_s.fsbttRAGs.wf])
-  case (1 x)
-  show ?case (is "?L = ?R")
-  proof -
-    from tRAG_ancestorsE[OF 1(2)]
-    obtain th2 where eq_x: "x = Th th2" by blast
-    from vat_s.cp_gen_rec[OF this]
-    have "?L = 
-          Max ({the_preced s th2} \<union> cp_gen s ` RTree.children (tRAG s) x)" .
-    also have "... = 
-          Max ({the_preced s' th2} \<union> cp_gen s' ` RTree.children (tRAG s') x)"
-  
-    proof -
-      from preced_kept have "the_preced s th2 = the_preced s' th2" by simp
-      moreover have "cp_gen s ` RTree.children (tRAG s) x =
-                     cp_gen s' ` RTree.children (tRAG s') x"
-      proof -
-        have "RTree.children (tRAG s) x =  RTree.children (tRAG s') x"
-        proof(unfold tRAG_s, rule children_union_kept)
-          have start: "(Th th, Th th') \<in> tRAG s"
-            by (unfold tRAG_s, auto)
-          note x_u = 1(2)
-          show "x \<notin> Range {(Th th, Th th')}"
-          proof
-            assume "x \<in> Range {(Th th, Th th')}"
-            hence eq_x: "x = Th th'" using RangeE by auto
-            show False
-            proof(cases rule:vat_s.rtree_s.ancestors_headE[OF assms(1) start])
-              case 1
-              from x_u[folded this, unfolded eq_x] vat_s.acyclic_tRAG
-              show ?thesis by (auto simp:ancestors_def acyclic_def)
-            next
-              case 2
-              with x_u[unfolded eq_x]
-              have "(Th th', Th th') \<in> (tRAG s)^+" by (auto simp:ancestors_def)
-              with vat_s.acyclic_tRAG show ?thesis by (auto simp:acyclic_def)
-            qed
-          qed
-        qed
-        moreover have "cp_gen s ` RTree.children (tRAG s) x =
-                       cp_gen s' ` RTree.children (tRAG s) x" (is "?f ` ?A = ?g ` ?A")
-        proof(rule f_image_eq)
-          fix a
-          assume a_in: "a \<in> ?A"
-          from 1(2)
-          show "?f a = ?g a"
-          proof(cases rule:vat_s.rtree_s.ancestors_childrenE[case_names in_ch out_ch])
-             case in_ch
-             show ?thesis
-             proof(cases "a = u")
-                case True
-                from assms(2)[folded this] show ?thesis .
-             next
-                case False
-                have a_not_in: "a \<notin> ancestors (tRAG s) (Th th)"
-                proof
-                  assume a_in': "a \<in> ancestors (tRAG s) (Th th)"
-                  have "a = u"
-                  proof(rule vat_s.rtree_s.ancestors_children_unique)
-                    from a_in' a_in show "a \<in> ancestors (tRAG s) (Th th) \<inter> 
-                                          RTree.children (tRAG s) x" by auto
-                  next 
-                    from assms(1) in_ch show "u \<in> ancestors (tRAG s) (Th th) \<inter> 
-                                      RTree.children (tRAG s) x" by auto
-                  qed
-                  with False show False by simp
-                qed
-                from a_in obtain th_a where eq_a: "a = Th th_a" 
-                    by (unfold RTree.children_def tRAG_alt_def, auto)
-                from cp_kept[OF a_not_in[unfolded eq_a]]
-                have "cp s th_a = cp s' th_a" .
-                from this [unfolded cp_gen_def_cond[OF eq_a], folded eq_a]
-                show ?thesis .
-             qed
-          next
-            case (out_ch z)
-            hence h: "z \<in> ancestors (tRAG s) u" "z \<in> RTree.children (tRAG s) x" by auto
-            show ?thesis
-            proof(cases "a = z")
-              case True
-              from h(2) have zx_in: "(z, x) \<in> (tRAG s)" by (auto simp:RTree.children_def)
-              from 1(1)[rule_format, OF this h(1)]
-              have eq_cp_gen: "cp_gen s z = cp_gen s' z" .
-              with True show ?thesis by metis
-            next
-              case False
-              from a_in obtain th_a where eq_a: "a = Th th_a"
-                by (auto simp:RTree.children_def tRAG_alt_def)
-              have "a \<notin> ancestors (tRAG s) (Th th)"
-              proof
-                assume a_in': "a \<in> ancestors (tRAG s) (Th th)"
-                have "a = z"
-                proof(rule vat_s.rtree_s.ancestors_children_unique)
-                  from assms(1) h(1) have "z \<in> ancestors (tRAG s) (Th th)"
-                      by (auto simp:ancestors_def)
-                  with h(2) show " z \<in> ancestors (tRAG s) (Th th) \<inter> 
-                                       RTree.children (tRAG s) x" by auto
-                next
-                  from a_in a_in'
-                  show "a \<in> ancestors (tRAG s) (Th th) \<inter> RTree.children (tRAG s) x"
-                    by auto
-                qed
-                with False show False by auto
-              qed
-              from cp_kept[OF this[unfolded eq_a]]
-              have "cp s th_a = cp s' th_a" .
-              from this[unfolded cp_gen_def_cond[OF eq_a], folded eq_a]
-              show ?thesis .
-            qed
-          qed
-        qed
-        ultimately show ?thesis by metis
-      qed
-      ultimately show ?thesis by simp
-    qed
-    also have "... = ?R"
-      by (fold vat_s'.cp_gen_rec[OF eq_x], simp)
-    finally show ?thesis .
-  qed
-qed
-
-lemma cp_up:
-  assumes "(Th th') \<in> ancestors (tRAG s) (Th th)"
-  and "cp s th' = cp s' th'"
-  and "(Th th'') \<in> ancestors (tRAG s) (Th th')"
-  shows "cp s th'' = cp s' th''"
-proof -
-  have "cp_gen s (Th th'') = cp_gen s' (Th th'')"
-  proof(rule cp_gen_update_stop[OF assms(1) _ assms(3)])
-    from assms(2) cp_gen_def_cond[OF refl[of "Th th'"]]
-    show "cp_gen s (Th th') = cp_gen s' (Th th')" by metis
-  qed
-  with cp_gen_def_cond[OF refl[of "Th th''"]]
-  show ?thesis by metis
-qed
-
-end
-
-locale step_create_cps =
-  fixes s' th prio s 
-  defines s_def : "s \<equiv> (Create th prio#s')"
-  assumes vt_s: "vt s"
-
-sublocale step_create_cps < vat_s: valid_trace "s"
-  by (unfold_locales, insert vt_s, simp)
-
-sublocale step_create_cps < vat_s': valid_trace "s'"
-  by (unfold_locales, insert step_back_vt[OF vt_s[unfolded s_def]], simp)
-
-context step_create_cps
-begin
-
-lemma RAG_kept: "RAG s = RAG s'"
-  by (unfold s_def RAG_create_unchanged, auto)
-
-lemma tRAG_kept: "tRAG s = tRAG s'"
-  by (unfold tRAG_alt_def RAG_kept, auto)
-
-lemma preced_kept:
-  assumes "th' \<noteq> th"
-  shows "the_preced s th' = the_preced s' th'"
-  by (unfold s_def the_preced_def preced_def, insert assms, auto)
-
-lemma th_not_in: "Th th \<notin> Field (tRAG s')"
-proof -
-  from vt_s[unfolded s_def]
-  have "PIP s' (Create th prio)" by (cases, simp)
-  hence "th \<notin> threads s'" by(cases, simp)
-  from vat_s'.not_in_thread_isolated[OF this]
-  have "Th th \<notin> Field (RAG s')" .
-  with tRAG_Field show ?thesis by auto
-qed
-
-lemma eq_cp:
-  assumes neq_th: "th' \<noteq> th"
-  shows "cp s th' = cp s' th'"
-proof -
-  have "(the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th') =
-        (the_preced s' \<circ> the_thread) ` subtree (tRAG s') (Th th')"
-  proof(unfold tRAG_kept, rule f_image_eq)
-    fix a
-    assume a_in: "a \<in> subtree (tRAG s') (Th th')"
-    then obtain th_a where eq_a: "a = Th th_a" 
-    proof(cases rule:subtreeE)
-      case 2
-      from ancestors_Field[OF 2(2)]
-      and that show ?thesis by (unfold tRAG_alt_def, auto)
-    qed auto
-    have neq_th_a: "th_a \<noteq> th"
-    proof -
-      have "(Th th) \<notin> subtree (tRAG s') (Th th')"
-      proof
-        assume "Th th \<in> subtree (tRAG s') (Th th')"
-        thus False
-        proof(cases rule:subtreeE)
-          case 2
-          from ancestors_Field[OF this(2)]
-          and th_not_in[unfolded Field_def]
-          show ?thesis by auto
-        qed (insert assms, auto)
-      qed
-      with a_in[unfolded eq_a] show ?thesis by auto
-    qed
-    from preced_kept[OF this]
-    show "(the_preced s \<circ> the_thread) a = (the_preced s' \<circ> the_thread) a"
-      by (unfold eq_a, simp)
-  qed
-  thus ?thesis by (unfold cp_alt_def1, simp)
-qed
-
-lemma children_of_th: "RTree.children (tRAG s) (Th th) = {}"
-proof -
-  { fix a
-    assume "a \<in> RTree.children (tRAG s) (Th th)"
-    hence "(a, Th th) \<in> tRAG s" by (auto simp:RTree.children_def)
-    with th_not_in have False 
-     by (unfold Field_def tRAG_kept, auto)
-  } thus ?thesis by auto
-qed
-
-lemma eq_cp_th: "cp s th = preced th s"
- by (unfold vat_s.cp_rec children_of_th, simp add:the_preced_def)
-
-end
-
-locale step_exit_cps =
-  fixes s' th prio s 
-  defines s_def : "s \<equiv> Exit th # s'"
-  assumes vt_s: "vt s"
-
-sublocale step_exit_cps < vat_s: valid_trace "s"
-  by (unfold_locales, insert vt_s, simp)
-
-sublocale step_exit_cps < vat_s': valid_trace "s'"
-  by (unfold_locales, insert step_back_vt[OF vt_s[unfolded s_def]], simp)
-
-context step_exit_cps
-begin
-
-lemma preced_kept:
-  assumes "th' \<noteq> th"
-  shows "the_preced s th' = the_preced s' th'"
-  by (unfold s_def the_preced_def preced_def, insert assms, auto)
-
-lemma RAG_kept: "RAG s = RAG s'"
-  by (unfold s_def RAG_exit_unchanged, auto)
-
-lemma tRAG_kept: "tRAG s = tRAG s'"
-  by (unfold tRAG_alt_def RAG_kept, auto)
-
-lemma th_ready: "th \<in> readys s'"
-proof -
-  from vt_s[unfolded s_def]
-  have "PIP s' (Exit th)" by (cases, simp)
-  hence h: "th \<in> runing s' \<and> holdents s' th = {}" by (cases, metis)
-  thus ?thesis by (unfold runing_def, auto)
-qed
-
-lemma th_holdents: "holdents s' th = {}"
-proof -
- from vt_s[unfolded s_def]
-  have "PIP s' (Exit th)" by (cases, simp)
-  thus ?thesis by (cases, metis)
-qed
-
-lemma th_RAG: "Th th \<notin> Field (RAG s')"
-proof -
-  have "Th th \<notin> Range (RAG s')"
-  proof
-    assume "Th th \<in> Range (RAG s')"
-    then obtain cs where "holding (wq s') th cs"
-      by (unfold Range_iff s_RAG_def, auto)
-    with th_holdents[unfolded holdents_def]
-    show False by (unfold eq_holding, auto)
-  qed
-  moreover have "Th th \<notin> Domain (RAG s')"
-  proof
-    assume "Th th \<in> Domain (RAG s')"
-    then obtain cs where "waiting (wq s') th cs"
-      by (unfold Domain_iff s_RAG_def, auto)
-    with th_ready show False by (unfold readys_def eq_waiting, auto)
-  qed
-  ultimately show ?thesis by (auto simp:Field_def)
-qed
-
-lemma th_tRAG: "(Th th) \<notin> Field (tRAG s')"
-  using th_RAG tRAG_Field[of s'] by auto
-
-lemma eq_cp:
-  assumes neq_th: "th' \<noteq> th"
-  shows "cp s th' = cp s' th'"
-proof -
-  have "(the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th') =
-        (the_preced s' \<circ> the_thread) ` subtree (tRAG s') (Th th')"
-  proof(unfold tRAG_kept, rule f_image_eq)
-    fix a
-    assume a_in: "a \<in> subtree (tRAG s') (Th th')"
-    then obtain th_a where eq_a: "a = Th th_a" 
-    proof(cases rule:subtreeE)
-      case 2
-      from ancestors_Field[OF 2(2)]
-      and that show ?thesis by (unfold tRAG_alt_def, auto)
-    qed auto
-    have neq_th_a: "th_a \<noteq> th"
-    proof -
-      from vat_s'.readys_in_no_subtree[OF th_ready assms]
-      have "(Th th) \<notin> subtree (RAG s') (Th th')" .
-      with tRAG_subtree_RAG[of s' "Th th'"]
-      have "(Th th) \<notin> subtree (tRAG s') (Th th')" by auto
-      with a_in[unfolded eq_a] show ?thesis by auto
-    qed
-    from preced_kept[OF this]
-    show "(the_preced s \<circ> the_thread) a = (the_preced s' \<circ> the_thread) a"
-      by (unfold eq_a, simp)
-  qed
-  thus ?thesis by (unfold cp_alt_def1, simp)
-qed
-
-end
-
-end
-
--- a/Moment.thy~	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,225 +0,0 @@
-theory Moment
-imports Main
-begin
-
-definition moment :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where "moment n s = rev (take n (rev s))"
-
-value "moment 3 [0, 1, 2, 3, 4, 5, 6, 7, 8, 9::int]"
-value "moment 2 [5, 4, 3, 2, 1, 0::int]"
-
-(*
-lemma length_moment_le:
-  assumes le_k: "k \<le> length s"
-  shows "length (moment k s) = k"
-using le_k unfolding moment_def by auto
-*)
-
-(*
-lemma length_moment_ge:
-  assumes le_k: "length s \<le> k"
-  shows "length (moment k s) = (length s)"
-using assms unfolding moment_def by simp
-*)
-
-lemma moment_app [simp]:
-  assumes ile: "i \<le> length s"
-  shows "moment i (s' @ s) = moment i s"
-using assms unfolding moment_def by simp
-
-lemma moment_eq [simp]: "moment (length s) (s' @ s) = s"
-  unfolding moment_def by simp
-
-lemma moment_ge [simp]: "length s \<le> n \<Longrightarrow> moment n s = s"
-  by (unfold moment_def, simp)
-
-lemma moment_zero [simp]: "moment 0 s = []"
-  by (simp add:moment_def)
-
-lemma p_split_gen: 
-  "\<lbrakk>Q s; \<not> Q (moment k s)\<rbrakk> \<Longrightarrow>
-  (\<exists> i. i < length s \<and> k \<le> i \<and> \<not> Q (moment i s) \<and> (\<forall> i' > i. Q (moment i' s)))"
-proof (induct s, simp)
-  fix a s
-  assume ih: "\<lbrakk>Q s; \<not> Q (moment k s)\<rbrakk>
-           \<Longrightarrow> \<exists>i<length s. k \<le> i \<and> \<not> Q (moment i s) \<and> (\<forall>i'>i. Q (moment i' s))"
-    and nq: "\<not> Q (moment k (a # s))" and qa: "Q (a # s)"
-  have le_k: "k \<le> length s"
-  proof -
-    { assume "length s < k"
-      hence "length (a#s) \<le> k" by simp
-      from moment_ge [OF this] and nq and qa
-      have "False" by auto
-    } thus ?thesis by arith
-  qed
-  have nq_k: "\<not> Q (moment k s)"
-  proof -
-    have "moment k (a#s) = moment k s"
-    proof -
-      from moment_app [OF le_k, of "[a]"] show ?thesis by simp
-    qed
-    with nq show ?thesis by simp
-  qed
-  show "\<exists>i<length (a # s). k \<le> i \<and> \<not> Q (moment i (a # s)) \<and> (\<forall>i'>i. Q (moment i' (a # s)))"
-  proof -
-    { assume "Q s"
-      from ih [OF this nq_k]
-      obtain i where lti: "i < length s" 
-        and nq: "\<not> Q (moment i s)" 
-        and rst: "\<forall>i'>i. Q (moment i' s)" 
-        and lki: "k \<le> i" by auto
-      have ?thesis 
-      proof -
-        from lti have "i < length (a # s)" by auto
-        moreover have " \<not> Q (moment i (a # s))"
-        proof -
-          from lti have "i \<le> (length s)" by simp
-          from moment_app [OF this, of "[a]"]
-          have "moment i (a # s) = moment i s" by simp
-          with nq show ?thesis by auto
-        qed
-        moreover have " (\<forall>i'>i. Q (moment i' (a # s)))"
-        proof -
-          {
-            fix i'
-            assume lti': "i < i'"
-            have "Q (moment i' (a # s))"
-            proof(cases "length (a#s) \<le> i'")
-              case True
-              from True have "moment i' (a#s) = a#s" by simp
-              with qa show ?thesis by simp
-            next
-              case False
-              from False have "i' \<le> length s" by simp
-              from moment_app [OF this, of "[a]"]
-              have "moment i' (a#s) = moment i' s" by simp
-              with rst lti' show ?thesis by auto
-            qed
-          } thus ?thesis by auto
-        qed
-        moreover note lki
-        ultimately show ?thesis by auto
-      qed
-    } moreover {
-      assume ns: "\<not> Q s"
-      have ?thesis
-      proof -
-        let ?i = "length s"
-        have "\<not> Q (moment ?i (a#s))"
-        proof -
-          have "?i \<le> length s" by simp
-          from moment_app [OF this, of "[a]"]
-          have "moment ?i (a#s) = moment ?i s" by simp
-          moreover have "\<dots> = s" by simp
-          ultimately show ?thesis using ns by auto
-        qed
-        moreover have "\<forall> i' > ?i. Q (moment i' (a#s))" 
-        proof -
-          { fix i'
-            assume "i' > ?i"
-            hence "length (a#s) \<le> i'" by simp
-            from moment_ge [OF this] 
-            have " moment i' (a # s) = a # s" .
-            with qa have "Q (moment i' (a#s))" by simp
-          } thus ?thesis by auto
-        qed
-        moreover have "?i < length (a#s)" by simp
-        moreover note le_k
-        ultimately show ?thesis by auto
-      qed
-    } ultimately show ?thesis by auto
-  qed
-qed
-
-lemma p_split: 
-  "\<lbrakk>Q s; \<not> Q []\<rbrakk> \<Longrightarrow> 
-       (\<exists> i. i < length s \<and> \<not> Q (moment i s) \<and> (\<forall> i' > i. Q (moment i' s)))"
-proof -
-  fix s Q
-  assume qs: "Q s" and nq: "\<not> Q []"
-  from nq have "\<not> Q (moment 0 s)" by simp
-  from p_split_gen [of Q s 0, OF qs this]
-  show "(\<exists> i. i < length s \<and> \<not> Q (moment i s) \<and> (\<forall> i' > i. Q (moment i' s)))"
-    by auto
-qed
-
-lemma moment_Suc_tl:
-  assumes "Suc i \<le> length s"
-  shows "tl (moment (Suc i) s) = moment i s"
-  using assms unfolding moment_def rev_take
-by (simp, metis Suc_diff_le diff_Suc_Suc drop_Suc tl_drop)
-
-lemma moment_plus:
-  assumes "Suc i \<le> length s"
-  shows "(moment (Suc i) s) = (hd (moment (Suc i) s)) # (moment i s)"
-proof -
-  have "(moment (Suc i) s) \<noteq> []"
-  using assms by (auto simp add: moment_def)
-  hence "(moment (Suc i) s) = (hd (moment (Suc i) s)) #  tl (moment (Suc i) s)"
-    by auto
-<<<<<<< local
-  have "moment (i + length s) (t @ s) = moment i t @ moment (length s) (t @ s)" 
-    by (simp add: moment_def)
-  with moment_app show ?thesis by auto
-qed
-
-lemma moment_Suc_tl:
-  assumes "Suc i \<le> length s"
-  shows "tl (moment (Suc i) s) = moment i s"
-  using assms unfolding moment_def rev_take
-  by (simp, metis Suc_diff_le diff_Suc_Suc drop_Suc tl_drop)
-  
-lemma moment_plus':
-  assumes "Suc i \<le> length s"
-  shows "(moment (Suc i) s) = (hd (moment (Suc i) s)) # (moment i s)"
-proof -
-  have "(moment (Suc i) s) \<noteq> []"
-  using assms length_moment_le by fastforce 
-  hence "(moment (Suc i) s) = (hd (moment (Suc i) s)) #  tl (moment (Suc i) s)"
-    by auto
-  with moment_Suc_tl[OF assms]
-  show ?thesis by metis
-qed
-
-lemma moment_plus: 
-  "Suc i \<le> length s \<Longrightarrow> moment (Suc i) s = (hd (moment (Suc i) s)) # (moment i s)"
-proof(induct s, simp+)
-  fix a s
-  assume ih: "Suc i \<le> length s \<Longrightarrow> moment (Suc i) s = hd (moment (Suc i) s) # moment i s"
-    and le_i: "i \<le> length s"
-  show "moment (Suc i) (a # s) = hd (moment (Suc i) (a # s)) # moment i (a # s)"
-  proof(cases "i= length s")
-    case True
-    hence "Suc i = length (a#s)" by simp
-    with moment_eq have "moment (Suc i) (a#s) = a#s" by auto
-    moreover have "moment i (a#s) = s"
-    proof -
-      from moment_app [OF le_i, of "[a]"]
-      and True show ?thesis by simp
-    qed
-    ultimately show ?thesis by auto
-  next
-    case False
-    from False and le_i have lti: "i < length s" by arith
-    hence les_i: "Suc i \<le> length s" by arith
-    show ?thesis 
-    proof -
-      from moment_app [OF les_i, of "[a]"]
-      have "moment (Suc i) (a # s) = moment (Suc i) s" by simp
-      moreover have "moment i (a#s) = moment i s" 
-      proof -
-        from lti have "i \<le> length s" by simp
-        from moment_app [OF this, of "[a]"] show ?thesis by simp
-      qed
-      moreover note ih [OF les_i]
-      ultimately show ?thesis by auto
-    qed
-  qed
-=======
-  with moment_Suc_tl[OF assms]
-  show ?thesis by metis
->>>>>>> other
-qed
-
-end
-
--- a/PIPBasics.thy~	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3788 +0,0 @@
-theory PIPBasics
-imports PIPDefs 
-begin
-
-locale valid_trace = 
-  fixes s
-  assumes vt : "vt s"
-
-locale valid_trace_e = valid_trace +
-  fixes e
-  assumes vt_e: "vt (e#s)"
-begin
-
-lemma pip_e: "PIP s e"
-  using vt_e by (cases, simp)  
-
-end
-
-lemma runing_ready: 
-  shows "runing s \<subseteq> readys s"
-  unfolding runing_def readys_def
-  by auto 
-
-lemma readys_threads:
-  shows "readys s \<subseteq> threads s"
-  unfolding readys_def
-  by auto
-
-lemma wq_v_neq:
-   "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
-  by (auto simp:wq_def Let_def cp_def split:list.splits)
-
-lemma runing_head:
-  assumes "th \<in> runing s"
-  and "th \<in> set (wq_fun (schs s) cs)"
-  shows "th = hd (wq_fun (schs s) cs)"
-  using assms
-  by (simp add:runing_def readys_def s_waiting_def wq_def)
-
-context valid_trace
-begin
-
-lemma actor_inv: 
-  assumes "PIP s e"
-  and "\<not> isCreate e"
-  shows "actor e \<in> runing s"
-  using assms
-  by (induct, auto)
-
-lemma ind [consumes 0, case_names Nil Cons, induct type]:
-  assumes "PP []"
-     and "(\<And>s e. valid_trace s \<Longrightarrow> valid_trace (e#s) \<Longrightarrow>
-                   PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))"
-     shows "PP s"
-proof(rule vt.induct[OF vt])
-  from assms(1) show "PP []" .
-next
-  fix s e
-  assume h: "vt s" "PP s" "PIP s e"
-  show "PP (e # s)"
-  proof(cases rule:assms(2))
-    from h(1) show v1: "valid_trace s" by (unfold_locales, simp)
-  next
-    from h(1,3) have "vt (e#s)" by auto
-    thus "valid_trace (e # s)" by (unfold_locales, simp)
-  qed (insert h, auto)
-qed
-
-lemma wq_distinct: "distinct (wq s cs)"
-proof(induct rule:ind)
-  case (Cons s e)
-  from Cons(4,3)
-  show ?case 
-  proof(induct)
-    case (thread_P th s cs1)
-    show ?case 
-    proof(cases "cs = cs1")
-      case True
-      thus ?thesis (is "distinct ?L")
-      proof - 
-        have "?L = wq_fun (schs s) cs1 @ [th]" using True
-          by (simp add:wq_def wf_def Let_def split:list.splits)
-        moreover have "distinct ..."
-        proof -
-          have "th \<notin> set (wq_fun (schs s) cs1)"
-          proof
-            assume otherwise: "th \<in> set (wq_fun (schs s) cs1)"
-            from runing_head[OF thread_P(1) this]
-            have "th = hd (wq_fun (schs s) cs1)" .
-            hence "(Cs cs1, Th th) \<in> (RAG s)" using otherwise
-              by (simp add:s_RAG_def s_holding_def wq_def cs_holding_def)
-            with thread_P(2) show False by auto
-          qed
-          moreover have "distinct (wq_fun (schs s) cs1)"
-              using True thread_P wq_def by auto 
-          ultimately show ?thesis by auto
-        qed
-        ultimately show ?thesis by simp
-      qed
-    next
-      case False
-      with thread_P(3)
-      show ?thesis
-        by (auto simp:wq_def wf_def Let_def split:list.splits)
-    qed
-  next
-    case (thread_V th s cs1)
-    thus ?case
-    proof(cases "cs = cs1")
-      case True
-      show ?thesis (is "distinct ?L")
-      proof(cases "(wq s cs)")
-        case Nil
-        thus ?thesis
-          by (auto simp:wq_def wf_def Let_def split:list.splits)
-      next
-        case (Cons w_hd w_tl)
-        moreover have "distinct (SOME q. distinct q \<and> set q = set w_tl)"
-        proof(rule someI2)
-          from thread_V(3)[unfolded Cons]
-          show  "distinct w_tl \<and> set w_tl = set w_tl" by auto
-        qed auto
-        ultimately show ?thesis
-          by (auto simp:wq_def wf_def Let_def True split:list.splits)
-      qed 
-    next
-      case False
-      with thread_V(3)
-      show ?thesis
-        by (auto simp:wq_def wf_def Let_def split:list.splits)
-    qed
-  qed (insert Cons, auto simp: wq_def Let_def split:list.splits)
-qed (unfold wq_def Let_def, simp)
-
-end
-
-
-context valid_trace_e
-begin
-
-text {*
-  The following lemma shows that only the @{text "P"}
-  operation can add new thread into waiting queues. 
-  Such kind of lemmas are very obvious, but need to be checked formally.
-  This is a kind of confirmation that our modelling is correct.
-*}
-
-lemma block_pre: 
-  assumes s_ni: "thread \<notin> set (wq s cs)"
-  and s_i: "thread \<in> set (wq (e#s) cs)"
-  shows "e = P thread cs"
-proof(cases e)
-  -- {* This is the only non-trivial case: *}
-  case (V th cs1)
-  have False
-  proof(cases "cs1 = cs")
-    case True
-    show ?thesis
-    proof(cases "(wq s cs1)")
-      case (Cons w_hd w_tl)
-      have "set (wq (e#s) cs) \<subseteq> set (wq s cs)"
-      proof -
-        have "(wq (e#s) cs) = (SOME q. distinct q \<and> set q = set w_tl)"
-          using  Cons V by (auto simp:wq_def Let_def True split:if_splits)
-        moreover have "set ... \<subseteq> set (wq s cs)"
-        proof(rule someI2)
-          show "distinct w_tl \<and> set w_tl = set w_tl"
-            by (metis distinct.simps(2) local.Cons wq_distinct)
-        qed (insert Cons True, auto)
-        ultimately show ?thesis by simp
-      qed
-      with assms show ?thesis by auto
-    qed (insert assms V True, auto simp:wq_def Let_def split:if_splits)
-  qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
-  thus ?thesis by auto
-qed (insert assms, auto simp:wq_def Let_def split:if_splits)
-
-end
-
-text {*
-  The following lemmas is also obvious and shallow. It says
-  that only running thread can request for a critical resource 
-  and that the requested resource must be one which is
-  not current held by the thread.
-*}
-
-lemma p_pre: "\<lbrakk>vt ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
-  thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (RAG s)^+"
-apply (ind_cases "vt ((P thread cs)#s)")
-apply (ind_cases "step s (P thread cs)")
-by auto
-
-lemma abs1:
-  assumes ein: "e \<in> set es"
-  and neq: "hd es \<noteq> hd (es @ [x])"
-  shows "False"
-proof -
-  from ein have "es \<noteq> []" by auto
-  then obtain e ess where "es = e # ess" by (cases es, auto)
-  with neq show ?thesis by auto
-qed
-
-lemma q_head: "Q (hd es) \<Longrightarrow> hd es = hd [th\<leftarrow>es . Q th]"
-  by (cases es, auto)
-
-inductive_cases evt_cons: "vt (a#s)"
-
-context valid_trace_e
-begin
-
-lemma abs2:
-  assumes inq: "thread \<in> set (wq s cs)"
-  and nh: "thread = hd (wq s cs)"
-  and qt: "thread \<noteq> hd (wq (e#s) cs)"
-  and inq': "thread \<in> set (wq (e#s) cs)"
-  shows "False"
-proof -
-  from vt_e assms show "False"
-    apply (cases e)
-    apply ((simp split:if_splits add:Let_def wq_def)[1])+
-    apply (insert abs1, fast)[1]
-    apply (auto simp:wq_def simp:Let_def split:if_splits list.splits)
-  proof -
-    fix th qs
-    assume vt: "vt (V th cs # s)"
-      and th_in: "thread \<in> set (SOME q. distinct q \<and> set q = set qs)"
-      and eq_wq: "wq_fun (schs s) cs = thread # qs"
-    show "False"
-    proof -
-      from wq_distinct[of cs]
-        and eq_wq[folded wq_def] have "distinct (thread#qs)" by simp
-      moreover have "thread \<in> set qs"
-      proof -
-        have "set (SOME q. distinct q \<and> set q = set qs) = set qs"
-        proof(rule someI2)
-          from wq_distinct [of cs]
-          and eq_wq [folded wq_def]
-          show "distinct qs \<and> set qs = set qs" by auto
-        next
-          fix x assume "distinct x \<and> set x = set qs"
-          thus "set x = set qs" by auto
-        qed
-        with th_in show ?thesis by auto
-      qed
-      ultimately show ?thesis by auto
-    qed
-  qed
-qed
-
-end
-
-
-context valid_trace
-begin
-lemma  vt_moment: "\<And> t. vt (moment t s)"
-proof(induct rule:ind)
-  case Nil
-  thus ?case by (simp add:vt_nil)
-next
-  case (Cons s e t)
-  show ?case
-  proof(cases "t \<ge> length (e#s)")
-    case True
-    from True have "moment t (e#s) = e#s" by simp
-    thus ?thesis using Cons
-      by (simp add:valid_trace_def)
-  next
-    case False
-    from Cons have "vt (moment t s)" by simp
-    moreover have "moment t (e#s) = moment t s"
-    proof -
-      from False have "t \<le> length s" by simp
-      from moment_app [OF this, of "[e]"] 
-      show ?thesis by simp
-    qed
-    ultimately show ?thesis by simp
-  qed
-qed
-end
-
-locale valid_moment = valid_trace + 
-  fixes i :: nat
-
-sublocale valid_moment < vat_moment: valid_trace "(moment i s)"
-  by (unfold_locales, insert vt_moment, auto)
-
-context valid_trace
-begin
-
-
-text {* (* ddd *)
-  The nature of the work is like this: since it starts from a very simple and basic 
-  model, even intuitively very `basic` and `obvious` properties need to derived from scratch.
-  For instance, the fact 
-  that one thread can not be blocked by two critical resources at the same time
-  is obvious, because only running threads can make new requests, if one is waiting for 
-  a critical resource and get blocked, it can not make another resource request and get 
-  blocked the second time (because it is not running). 
-
-  To derive this fact, one needs to prove by contraction and 
-  reason about time (or @{text "moement"}). The reasoning is based on a generic theorem
-  named @{text "p_split"}, which is about status changing along the time axis. It says if 
-  a condition @{text "Q"} is @{text "True"} at a state @{text "s"},
-  but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} 
-  in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history 
-  of events leading to it), such that @{text "Q"} switched 
-  from being @{text "False"} to @{text "True"} and kept being @{text "True"}
-  till the last moment of @{text "s"}.
-
-  Suppose a thread @{text "th"} is blocked
-  on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, 
-  since no thread is blocked at the very beginning, by applying 
-  @{text "p_split"} to these two blocking facts, there exist 
-  two moments @{text "t1"} and @{text "t2"}  in @{text "s"}, such that 
-  @{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} 
-  and kept on blocked on them respectively ever since.
- 
-  Without lost of generality, we assume @{text "t1"} is earlier than @{text "t2"}.
-  However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
-  in blocked state at moment @{text "t2"} and could not
-  make any request and get blocked the second time: Contradiction.
-*}
-
-lemma waiting_unique_pre: (* ccc *)
-  assumes h11: "thread \<in> set (wq s cs1)"
-  and h12: "thread \<noteq> hd (wq s cs1)"
-  assumes h21: "thread \<in> set (wq s cs2)"
-  and h22: "thread \<noteq> hd (wq s cs2)"
-  and neq12: "cs1 \<noteq> cs2"
-  shows "False"
-proof -
-  let "?Q cs s" = "thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
-  from h11 and h12 have q1: "?Q cs1 s" by simp
-  from h21 and h22 have q2: "?Q cs2 s" by simp
-  have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
-  have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
-  from p_split [of "?Q cs1", OF q1 nq1]
-  obtain t1 where lt1: "t1 < length s"
-    and np1: "\<not>(thread \<in> set (wq (moment t1 s) cs1) \<and>
-        thread \<noteq> hd (wq (moment t1 s) cs1))"
-    and nn1: "(\<forall>i'>t1. thread \<in> set (wq (moment i' s) cs1) \<and>
-             thread \<noteq> hd (wq (moment i' s) cs1))" by auto
-  from p_split [of "?Q cs2", OF q2 nq2]
-  obtain t2 where lt2: "t2 < length s"
-    and np2: "\<not>(thread \<in> set (wq (moment t2 s) cs2) \<and>
-        thread \<noteq> hd (wq (moment t2 s) cs2))"
-    and nn2: "(\<forall>i'>t2. thread \<in> set (wq (moment i' s) cs2) \<and>
-             thread \<noteq> hd (wq (moment i' s) cs2))" by auto
-  show ?thesis
-  proof -
-    { 
-      assume lt12: "t1 < t2"
-      let ?t3 = "Suc t2"
-      from lt2 have le_t3: "?t3 \<le> length s" by auto
-      from moment_plus [OF this] 
-      obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
-      have "t2 < ?t3" by simp
-      from nn2 [rule_format, OF this] and eq_m
-      have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
-        h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
-      have "vt (e#moment t2 s)"
-      proof -
-        from vt_moment 
-        have "vt (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      then interpret vt_e: valid_trace_e "moment t2 s" "e"
-        by (unfold_locales, auto, cases, simp)
-      have ?thesis
-      proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
-        case True
-        from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
-          by auto 
-        from vt_e.abs2 [OF True eq_th h2 h1]
-        show ?thesis by auto
-      next
-        case False
-        from vt_e.block_pre[OF False h1]
-        have "e = P thread cs2" .
-        with vt_e.vt_e have "vt ((P thread cs2)# moment t2 s)" by simp
-        from p_pre [OF this] have "thread \<in> runing (moment t2 s)" by simp
-        with runing_ready have "thread \<in> readys (moment t2 s)" by auto
-        with nn1 [rule_format, OF lt12]
-        show ?thesis  by (simp add:readys_def wq_def s_waiting_def, auto)
-      qed
-    } moreover {
-      assume lt12: "t2 < t1"
-      let ?t3 = "Suc t1"
-      from lt1 have le_t3: "?t3 \<le> length s" by auto
-      from moment_plus [OF this] 
-      obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
-      have lt_t3: "t1 < ?t3" by simp
-      from nn1 [rule_format, OF this] and eq_m
-      have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
-        h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
-      have "vt  (e#moment t1 s)"
-      proof -
-        from vt_moment
-        have "vt (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      then interpret vt_e: valid_trace_e "moment t1 s" e
-        by (unfold_locales, auto, cases, auto)
-      have ?thesis
-      proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
-        case True
-        from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
-          by auto
-        from vt_e.abs2 True eq_th h2 h1
-        show ?thesis by auto
-      next
-        case False
-        from vt_e.block_pre [OF False h1]
-        have "e = P thread cs1" .
-        with vt_e.vt_e have "vt ((P thread cs1)# moment t1 s)" by simp
-        from p_pre [OF this] have "thread \<in> runing (moment t1 s)" by simp
-        with runing_ready have "thread \<in> readys (moment t1 s)" by auto
-        with nn2 [rule_format, OF lt12]
-        show ?thesis  by (simp add:readys_def wq_def s_waiting_def, auto)
-      qed
-    } moreover {
-      assume eqt12: "t1 = t2"
-      let ?t3 = "Suc t1"
-      from lt1 have le_t3: "?t3 \<le> length s" by auto
-      from moment_plus [OF this] 
-      obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
-      have lt_t3: "t1 < ?t3" by simp
-      from nn1 [rule_format, OF this] and eq_m
-      have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
-        h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
-      have vt_e: "vt (e#moment t1 s)"
-      proof -
-        from vt_moment
-        have "vt (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      then interpret vt_e: valid_trace_e "moment t1 s" e
-        by (unfold_locales, auto, cases, auto)
-      have ?thesis
-      proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
-        case True
-        from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
-          by auto
-        from vt_e.abs2 [OF True eq_th h2 h1]
-        show ?thesis by auto
-      next
-        case False
-        from vt_e.block_pre [OF False h1]
-        have eq_e1: "e = P thread cs1" .
-        have lt_t3: "t1 < ?t3" by simp
-        with eqt12 have "t2 < ?t3" by simp
-        from nn2 [rule_format, OF this] and eq_m and eqt12
-        have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
-          h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
-        show ?thesis
-        proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
-          case True
-          from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
-            by auto
-          from vt_e and eqt12 have "vt (e#moment t2 s)" by simp 
-          then interpret vt_e2: valid_trace_e "moment t2 s" e
-            by (unfold_locales, auto, cases, auto)
-          from vt_e2.abs2 [OF True eq_th h2 h1]
-          show ?thesis .
-        next
-          case False
-          have "vt (e#moment t2 s)"
-          proof -
-            from vt_moment eqt12
-            have "vt (moment (Suc t2) s)" by auto
-            with eq_m eqt12 show ?thesis by simp
-          qed
-          then interpret vt_e2: valid_trace_e "moment t2 s" e
-            by (unfold_locales, auto, cases, auto)
-          from vt_e2.block_pre [OF False h1]
-          have "e = P thread cs2" .
-          with eq_e1 neq12 show ?thesis by auto
-        qed
-      qed
-    } ultimately show ?thesis by arith
-  qed
-qed
-
-text {*
-  This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
-*}
-
-lemma waiting_unique:
-  assumes "waiting s th cs1"
-  and "waiting s th cs2"
-  shows "cs1 = cs2"
-using waiting_unique_pre assms
-unfolding wq_def s_waiting_def
-by auto
-
-end
-
-(* not used *)
-text {*
-  Every thread can only be blocked on one critical resource, 
-  symmetrically, every critical resource can only be held by one thread. 
-  This fact is much more easier according to our definition. 
-*}
-lemma held_unique:
-  assumes "holding (s::event list) th1 cs"
-  and "holding s th2 cs"
-  shows "th1 = th2"
- by (insert assms, unfold s_holding_def, auto)
-
-
-lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s"
-  apply (induct s, auto)
-  by (case_tac a, auto split:if_splits)
-
-lemma last_set_unique: 
-  "\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk>
-          \<Longrightarrow> th1 = th2"
-  apply (induct s, auto)
-  by (case_tac a, auto split:if_splits dest:last_set_lt)
-
-lemma preced_unique : 
-  assumes pcd_eq: "preced th1 s = preced th2 s"
-  and th_in1: "th1 \<in> threads s"
-  and th_in2: " th2 \<in> threads s"
-  shows "th1 = th2"
-proof -
-  from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def)
-  from last_set_unique [OF this th_in1 th_in2]
-  show ?thesis .
-qed
-
-lemma preced_linorder: 
-  assumes neq_12: "th1 \<noteq> th2"
-  and th_in1: "th1 \<in> threads s"
-  and th_in2: " th2 \<in> threads s"
-  shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
-proof -
-  from preced_unique [OF _ th_in1 th_in2] and neq_12 
-  have "preced th1 s \<noteq> preced th2 s" by auto
-  thus ?thesis by auto
-qed
-
-(* An aux lemma used later *)
-lemma unique_minus:
-  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
-  and xy: "(x, y) \<in> r"
-  and xz: "(x, z) \<in> r^+"
-  and neq: "y \<noteq> z"
-  shows "(y, z) \<in> r^+"
-proof -
- from xz and neq show ?thesis
- proof(induct)
-   case (base ya)
-   have "(x, ya) \<in> r" by fact
-   from unique [OF xy this] have "y = ya" .
-   with base show ?case by auto
- next
-   case (step ya z)
-   show ?case
-   proof(cases "y = ya")
-     case True
-     from step True show ?thesis by simp
-   next
-     case False
-     from step False
-     show ?thesis by auto
-   qed
- qed
-qed
-
-lemma unique_base:
-  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
-  and xy: "(x, y) \<in> r"
-  and xz: "(x, z) \<in> r^+"
-  and neq_yz: "y \<noteq> z"
-  shows "(y, z) \<in> r^+"
-proof -
-  from xz neq_yz show ?thesis
-  proof(induct)
-    case (base ya)
-    from xy unique base show ?case by auto
-  next
-    case (step ya z)
-    show ?case
-    proof(cases "y = ya")
-      case True
-      from True step show ?thesis by auto
-    next
-      case False
-      from False step 
-      have "(y, ya) \<in> r\<^sup>+" by auto
-      with step show ?thesis by auto
-    qed
-  qed
-qed
-
-lemma unique_chain:
-  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
-  and xy: "(x, y) \<in> r^+"
-  and xz: "(x, z) \<in> r^+"
-  and neq_yz: "y \<noteq> z"
-  shows "(y, z) \<in> r^+ \<or> (z, y) \<in> r^+"
-proof -
-  from xy xz neq_yz show ?thesis
-  proof(induct)
-    case (base y)
-    have h1: "(x, y) \<in> r" and h2: "(x, z) \<in> r\<^sup>+" and h3: "y \<noteq> z" using base by auto
-    from unique_base [OF _ h1 h2 h3] and unique show ?case by auto
-  next
-    case (step y za)
-    show ?case
-    proof(cases "y = z")
-      case True
-      from True step show ?thesis by auto
-    next
-      case False
-      from False step have "(y, z) \<in> r\<^sup>+ \<or> (z, y) \<in> r\<^sup>+" by auto
-      thus ?thesis
-      proof
-        assume "(z, y) \<in> r\<^sup>+"
-        with step have "(z, za) \<in> r\<^sup>+" by auto
-        thus ?thesis by auto
-      next
-        assume h: "(y, z) \<in> r\<^sup>+"
-        from step have yza: "(y, za) \<in> r" by simp
-        from step have "za \<noteq> z" by simp
-        from unique_minus [OF _ yza h this] and unique
-        have "(za, z) \<in> r\<^sup>+" by auto
-        thus ?thesis by auto
-      qed
-    qed
-  qed
-qed
-
-text {*
-  The following three lemmas show that @{text "RAG"} does not change
-  by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
-  events, respectively.
-*}
-
-lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
-apply (unfold s_RAG_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-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 RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
-apply (unfold s_RAG_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-
-text {* 
-  The following lemmas are used in the proof of 
-  lemma @{text "step_RAG_v"}, which characterizes how the @{text "RAG"} is changed
-  by @{text "V"}-events. 
-  However, since our model is very concise, such  seemingly obvious lemmas need to be derived from scratch,
-  starting from the model definitions.
-*}
-lemma step_v_hold_inv[elim_format]:
-  "\<And>c t. \<lbrakk>vt (V th cs # s); 
-          \<not> holding (wq s) t c; holding (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> 
-            next_th s th cs t \<and> c = cs"
-proof -
-  fix c t
-  assume vt: "vt (V th cs # s)"
-    and nhd: "\<not> holding (wq s) t c"
-    and hd: "holding (wq (V th cs # s)) t c"
-  show "next_th s th cs t \<and> c = cs"
-  proof(cases "c = cs")
-    case False
-    with nhd hd show ?thesis
-      by (unfold cs_holding_def wq_def, auto simp:Let_def)
-  next
-    case True
-    with step_back_step [OF vt] 
-    have "step s (V th c)" by simp
-    hence "next_th s th cs t"
-    proof(cases)
-      assume "holding s th c"
-      with nhd hd show ?thesis
-        apply (unfold s_holding_def cs_holding_def wq_def next_th_def,
-               auto simp:Let_def split:list.splits if_splits)
-        proof -
-          assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
-          moreover have "\<dots> = set []"
-          proof(rule someI2)
-            show "distinct [] \<and> [] = []" by auto
-          next
-            fix x assume "distinct x \<and> x = []"
-            thus "set x = set []" by auto
-          qed
-          ultimately show False by auto
-        next
-          assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
-          moreover have "\<dots> = set []"
-          proof(rule someI2)
-            show "distinct [] \<and> [] = []" by auto
-          next
-            fix x assume "distinct x \<and> x = []"
-            thus "set x = set []" by auto
-          qed
-          ultimately show False by auto
-        qed
-    qed
-    with True show ?thesis by auto
-  qed
-qed
-
-text {* 
-  The following @{text "step_v_wait_inv"} is also an obvious lemma, which, however, needs to be
-  derived from scratch, which confirms the correctness of the definition of @{text "next_th"}.
-*}
-lemma step_v_wait_inv[elim_format]:
-    "\<And>t c. \<lbrakk>vt (V th cs # s); \<not> waiting (wq (V th cs # s)) t c; waiting (wq s) t c
-           \<rbrakk>
-          \<Longrightarrow> (next_th s th cs t \<and> cs = c)"
-proof -
-  fix t c 
-  assume vt: "vt (V th cs # s)"
-    and nw: "\<not> waiting (wq (V th cs # s)) t c"
-    and wt: "waiting (wq s) t c"
-  from vt interpret vt_v: valid_trace_e s "V th cs" 
-    by  (cases, unfold_locales, simp)
-  show "next_th s th cs t \<and> cs = c"
-  proof(cases "cs = c")
-    case False
-    with nw wt show ?thesis
-      by (auto simp:cs_waiting_def wq_def Let_def)
-  next
-    case True
-    from nw[folded True] wt[folded True]
-    have "next_th s th cs t"
-      apply (unfold next_th_def, auto simp:cs_waiting_def wq_def Let_def split:list.splits)
-    proof -
-      fix a list
-      assume t_in: "t \<in> set list"
-        and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
-        and eq_wq: "wq_fun (schs s) cs = a # list"
-      have " set (SOME q. distinct q \<and> set q = set list) = set list"
-      proof(rule someI2)
-        from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def]
-        show "distinct list \<and> set list = set list" by auto
-      next
-        show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
-          by auto
-      qed
-      with t_ni and t_in show "a = th" by auto
-    next
-      fix a list
-      assume t_in: "t \<in> set list"
-        and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
-        and eq_wq: "wq_fun (schs s) cs = a # list"
-      have " set (SOME q. distinct q \<and> set q = set list) = set list"
-      proof(rule someI2)
-        from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def]
-        show "distinct list \<and> set list = set list" by auto
-      next
-        show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
-          by auto
-      qed
-      with t_ni and t_in show "t = hd (SOME q. distinct q \<and> set q = set list)" by auto
-    next
-      fix a list
-      assume eq_wq: "wq_fun (schs s) cs = a # list"
-      from step_back_step[OF vt]
-      show "a = th"
-      proof(cases)
-        assume "holding s th cs"
-        with eq_wq show ?thesis
-          by (unfold s_holding_def wq_def, auto)
-      qed
-    qed
-    with True show ?thesis by simp
-  qed
-qed
-
-lemma step_v_not_wait[consumes 3]:
-  "\<lbrakk>vt (V th cs # s); next_th s th cs t; waiting (wq (V th cs # s)) t cs\<rbrakk> \<Longrightarrow> False"
-  by (unfold next_th_def cs_waiting_def wq_def, auto simp:Let_def)
-
-lemma step_v_release:
-  "\<lbrakk>vt (V th cs # s); holding (wq (V th cs # s)) th cs\<rbrakk> \<Longrightarrow> False"
-proof -
-  assume vt: "vt (V th cs # s)"
-    and hd: "holding (wq (V th cs # s)) th cs"
-  from vt interpret vt_v: valid_trace_e s "V th cs"
-    by (cases, unfold_locales, simp+)
-  from step_back_step [OF vt] and hd
-  show "False"
-  proof(cases)
-    assume "holding (wq (V th cs # s)) th cs" and "holding s th cs"
-    thus ?thesis
-      apply (unfold s_holding_def wq_def cs_holding_def)
-      apply (auto simp:Let_def split:list.splits)
-    proof -
-      fix list
-      assume eq_wq[folded wq_def]: 
-        "wq_fun (schs s) cs = hd (SOME q. distinct q \<and> set q = set list) # list"
-      and hd_in: "hd (SOME q. distinct q \<and> set q = set list)
-            \<in> set (SOME q. distinct q \<and> set q = set list)"
-      have "set (SOME q. distinct q \<and> set q = set list) = set list"
-      proof(rule someI2)
-        from vt_v.wq_distinct[of cs] and eq_wq
-        show "distinct list \<and> set list = set list" by auto
-      next
-        show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
-          by auto
-      qed
-      moreover have "distinct  (hd (SOME q. distinct q \<and> set q = set list) # list)"
-      proof -
-        from vt_v.wq_distinct[of cs] and eq_wq
-        show ?thesis by auto
-      qed
-      moreover note eq_wq and hd_in
-      ultimately show "False" by auto
-    qed
-  qed
-qed
-
-lemma step_v_get_hold:
-  "\<And>th'. \<lbrakk>vt (V th cs # s); \<not> holding (wq (V th cs # s)) th' cs; next_th s th cs th'\<rbrakk> \<Longrightarrow> False"
-  apply (unfold cs_holding_def next_th_def wq_def,
-         auto simp:Let_def)
-proof -
-  fix rest
-  assume vt: "vt (V th cs # s)"
-    and eq_wq[folded wq_def]: " wq_fun (schs s) cs = th # rest"
-    and nrest: "rest \<noteq> []"
-    and ni: "hd (SOME q. distinct q \<and> set q = set rest)
-            \<notin> set (SOME q. distinct q \<and> set q = set rest)"
-  from vt interpret vt_v: valid_trace_e s "V th cs"
-    by (cases, unfold_locales, simp+)
-  have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-  proof(rule someI2)
-    from vt_v.wq_distinct[of cs] and eq_wq
-    show "distinct rest \<and> set rest = set rest" by auto
-  next
-    fix x assume "distinct x \<and> set x = set rest"
-    hence "set x = set rest" by auto
-    with nrest
-    show "x \<noteq> []" by (case_tac x, auto)
-  qed
-  with ni show "False" by auto
-qed
-
-lemma step_v_release_inv[elim_format]:
-"\<And>c t. \<lbrakk>vt (V th cs # s); \<not> holding (wq (V th cs # s)) t c; holding (wq s) t c\<rbrakk> \<Longrightarrow> 
-  c = cs \<and> t = th"
-  apply (unfold cs_holding_def wq_def, auto simp:Let_def split:if_splits list.splits)
-  proof -
-    fix a list
-    assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list"
-    from step_back_step [OF vt] show "a = th"
-    proof(cases)
-      assume "holding s th cs" with eq_wq
-      show ?thesis
-        by (unfold s_holding_def wq_def, auto)
-    qed
-  next
-    fix a list
-    assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list"
-    from step_back_step [OF vt] show "a = th"
-    proof(cases)
-      assume "holding s th cs" with eq_wq
-      show ?thesis
-        by (unfold s_holding_def wq_def, auto)
-    qed
-  qed
-
-lemma step_v_waiting_mono:
-  "\<And>t c. \<lbrakk>vt (V th cs # s); waiting (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> waiting (wq s) t c"
-proof -
-  fix t c
-  let ?s' = "(V th cs # s)"
-  assume vt: "vt ?s'" 
-    and wt: "waiting (wq ?s') t c"
-  from vt interpret vt_v: valid_trace_e s "V th cs"
-    by (cases, unfold_locales, simp+)
-  show "waiting (wq s) t c"
-  proof(cases "c = cs")
-    case False
-    assume neq_cs: "c \<noteq> cs"
-    hence "waiting (wq ?s') t c = waiting (wq s) t c"
-      by (unfold cs_waiting_def wq_def, auto simp:Let_def)
-    with wt show ?thesis by simp
-  next
-    case True
-    with wt show ?thesis
-      apply (unfold cs_waiting_def wq_def, auto simp:Let_def split:list.splits)
-    proof -
-      fix a list
-      assume not_in: "t \<notin> set list"
-        and is_in: "t \<in> set (SOME q. distinct q \<and> set q = set list)"
-        and eq_wq: "wq_fun (schs s) cs = a # list"
-      have "set (SOME q. distinct q \<and> set q = set list) = set list"
-      proof(rule someI2)
-        from vt_v.wq_distinct [of cs]
-        and eq_wq[folded wq_def]
-        show "distinct list \<and> set list = set list" by auto
-      next
-        fix x assume "distinct x \<and> set x = set list"
-        thus "set x = set list" by auto
-      qed
-      with not_in is_in show "t = a" by auto
-    next
-      fix list
-      assume is_waiting: "waiting (wq (V th cs # s)) t cs"
-      and eq_wq: "wq_fun (schs s) cs = t # list"
-      hence "t \<in> set list"
-        apply (unfold wq_def, auto simp:Let_def cs_waiting_def)
-      proof -
-        assume " t \<in> set (SOME q. distinct q \<and> set q = set list)"
-        moreover have "\<dots> = set list" 
-        proof(rule someI2)
-          from vt_v.wq_distinct [of cs]
-            and eq_wq[folded wq_def]
-          show "distinct list \<and> set list = set list" by auto
-        next
-          fix x assume "distinct x \<and> set x = set list" 
-          thus "set x = set list" by auto
-        qed
-        ultimately show "t \<in> set list" by simp
-      qed
-      with eq_wq and vt_v.wq_distinct [of cs, unfolded wq_def]
-      show False by auto
-    qed
-  qed
-qed
-
-text {* (* ddd *) 
-  The following @{text "step_RAG_v"} lemma charaterizes how @{text "RAG"} is changed
-  with the happening of @{text "V"}-events:
-*}
-lemma step_RAG_v:
-assumes vt:
-  "vt (V th cs#s)"
-shows "
-  RAG (V th cs # s) =
-  RAG s - {(Cs cs, Th th)} -
-  {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-  {(Cs cs, Th th') |th'.  next_th s th cs th'}"
-  apply (insert vt, unfold s_RAG_def) 
-  apply (auto split:if_splits list.splits simp:Let_def)
-  apply (auto elim: step_v_waiting_mono step_v_hold_inv 
-              step_v_release step_v_wait_inv
-              step_v_get_hold step_v_release_inv)
-  apply (erule_tac step_v_not_wait, auto)
-  done
-
-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
-
-
-lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
-  by (unfold s_RAG_def, auto)
-
-context valid_trace
-begin
-
-text {*
-  The following lemma shows that @{text "RAG"} is acyclic.
-  The overall structure is by induction on the formation of @{text "vt s"}
-  and then case analysis on event @{text "e"}, where the non-trivial cases 
-  for those for @{text "V"} and @{text "P"} events.
-*}
-lemma acyclic_RAG:
-  shows "acyclic (RAG s)"
-using vt
-proof(induct)
-  case (vt_cons s e)
-  interpret vt_s: valid_trace s using vt_cons(1)
-    by (unfold_locales, simp)
-  assume ih: "acyclic (RAG s)"
-    and stp: "step s e"
-    and vt: "vt s"
-  show ?case
-  proof(cases e)
-    case (Create th prio)
-    with ih
-    show ?thesis by (simp add:RAG_create_unchanged)
-  next
-    case (Exit th)
-    with ih show ?thesis by (simp add:RAG_exit_unchanged)
-  next
-    case (V th cs)
-    from V vt stp have vtt: "vt (V th cs#s)" by auto
-    from step_RAG_v [OF this]
-    have eq_de: 
-      "RAG (e # s) = 
-      RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-      {(Cs cs, Th th') |th'. next_th s th cs th'}"
-      (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
-    from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
-    from step_back_step [OF vtt]
-    have "step s (V th cs)" .
-    thus ?thesis
-    proof(cases)
-      assume "holding s th cs"
-      hence th_in: "th \<in> set (wq s cs)" and
-        eq_hd: "th = hd (wq s cs)" unfolding s_holding_def wq_def by auto
-      then obtain rest where
-        eq_wq: "wq s cs = th#rest"
-        by (cases "wq s cs", auto)
-      show ?thesis
-      proof(cases "rest = []")
-        case False
-        let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
-        from eq_wq False have eq_D: "?D = {(Cs cs, Th ?th')}"
-          by (unfold next_th_def, auto)
-        let ?E = "(?A - ?B - ?C)"
-        have "(Th ?th', Cs cs) \<notin> ?E\<^sup>*"
-        proof
-          assume "(Th ?th', Cs cs) \<in> ?E\<^sup>*"
-          hence " (Th ?th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
-          from tranclD [OF this]
-          obtain x where th'_e: "(Th ?th', x) \<in> ?E" by blast
-          hence th_d: "(Th ?th', x) \<in> ?A" by simp
-          from RAG_target_th [OF this]
-          obtain cs' where eq_x: "x = Cs cs'" by auto
-          with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp
-          hence wt_th': "waiting s ?th' cs'"
-            unfolding s_RAG_def s_waiting_def cs_waiting_def wq_def by simp
-          hence "cs' = cs"
-          proof(rule vt_s.waiting_unique)
-            from eq_wq vt_s.wq_distinct[of cs]
-            show "waiting s ?th' cs" 
-              apply (unfold s_waiting_def wq_def, auto)
-            proof -
-              assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-                and eq_wq: "wq_fun (schs s) cs = th # rest"
-              have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
-              next
-                fix x assume "distinct x \<and> set x = set rest"
-                with False show "x \<noteq> []" by auto
-              qed
-              hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
-                set (SOME q. distinct q \<and> set q = set rest)" by auto
-              moreover have "\<dots> = set rest" 
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
-              next
-                show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-              qed
-              moreover note hd_in
-              ultimately show "hd (SOME q. distinct q \<and> set q = set rest) = th" by auto
-            next
-              assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-                and eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest"
-              have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" by auto
-              next
-                fix x assume "distinct x \<and> set x = set rest"
-                with False show "x \<noteq> []" by auto
-              qed
-              hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
-                set (SOME q. distinct q \<and> set q = set rest)" by auto
-              moreover have "\<dots> = set rest" 
-              proof(rule someI2)
-                from vt_s.wq_distinct[of cs] and eq_wq
-                show "distinct rest \<and> set rest = set rest" by auto
-              next
-                show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-              qed
-              moreover note hd_in
-              ultimately show False by auto
-            qed
-          qed
-          with th'_e eq_x have "(Th ?th', Cs cs) \<in> ?E" by simp
-          with False
-          show "False" by (auto simp: next_th_def eq_wq)
-        qed
-        with acyclic_insert[symmetric] and ac
-          and eq_de eq_D show ?thesis by auto
-      next
-        case True
-        with eq_wq
-        have eq_D: "?D = {}"
-          by (unfold next_th_def, auto)
-        with eq_de ac
-        show ?thesis by auto
-      qed 
-    qed
-  next
-    case (P th cs)
-    from P vt stp have vtt: "vt (P th cs#s)" by auto
-    from step_RAG_p [OF this] P
-    have "RAG (e # s) = 
-      (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
-      RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
-      by simp
-    moreover have "acyclic ?R"
-    proof(cases "wq s cs = []")
-      case True
-      hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
-      have "(Th th, Cs cs) \<notin> (RAG s)\<^sup>*"
-      proof
-        assume "(Th th, Cs cs) \<in> (RAG s)\<^sup>*"
-        hence "(Th th, Cs cs) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
-        from tranclD2 [OF this]
-        obtain x where "(x, Cs cs) \<in> RAG s" by auto
-        with True show False by (auto simp:s_RAG_def cs_waiting_def)
-      qed
-      with acyclic_insert ih eq_r show ?thesis by auto
-    next
-      case False
-      hence eq_r: "?R =  RAG s \<union> {(Th th, Cs cs)}" by simp
-      have "(Cs cs, Th th) \<notin> (RAG s)\<^sup>*"
-      proof
-        assume "(Cs cs, Th th) \<in> (RAG s)\<^sup>*"
-        hence "(Cs cs, Th th) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
-        moreover from step_back_step [OF vtt] have "step s (P th cs)" .
-        ultimately show False
-        proof -
-          show " \<lbrakk>(Cs cs, Th th) \<in> (RAG s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
-            by (ind_cases "step s (P th cs)", simp)
-        qed
-      qed
-      with acyclic_insert ih eq_r show ?thesis by auto
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (Set thread prio)
-      with ih
-      thm RAG_set_unchanged
-      show ?thesis by (simp add:RAG_set_unchanged)
-    qed
-  next
-    case vt_nil
-    show "acyclic (RAG ([]::state))"
-      by (auto simp: s_RAG_def cs_waiting_def 
-        cs_holding_def wq_def acyclic_def)
-qed
-
-
-lemma finite_RAG:
-  shows "finite (RAG s)"
-proof -
-  from vt show ?thesis
-  proof(induct)
-    case (vt_cons s e)
-    interpret vt_s: valid_trace s using vt_cons(1)
-      by (unfold_locales, simp)
-    assume ih: "finite (RAG s)"
-      and stp: "step s e"
-      and vt: "vt s"
-    show ?case
-    proof(cases e)
-      case (Create th prio)
-      with ih
-      show ?thesis by (simp add:RAG_create_unchanged)
-    next
-      case (Exit th)
-      with ih show ?thesis by (simp add:RAG_exit_unchanged)
-    next
-      case (V th cs)
-      from V vt stp have vtt: "vt (V th cs#s)" by auto
-      from step_RAG_v [OF this]
-      have eq_de: "RAG (e # s) = 
-                   RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
-                      {(Cs cs, Th th') |th'. next_th s th cs th'}
-"
-        (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
-      moreover from ih have ac: "finite (?A - ?B - ?C)" by simp
-      moreover have "finite ?D"
-      proof -
-        have "?D = {} \<or> (\<exists> a. ?D = {a})" 
-          by (unfold next_th_def, auto)
-        thus ?thesis
-        proof
-          assume h: "?D = {}"
-          show ?thesis by (unfold h, simp)
-        next
-          assume "\<exists> a. ?D = {a}"
-          thus ?thesis
-            by (metis finite.simps)
-        qed
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (P th cs)
-      from P vt stp have vtt: "vt (P th cs#s)" by auto
-      from step_RAG_p [OF this] P
-      have "RAG (e # s) = 
-              (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
-                                    RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
-        by simp
-      moreover have "finite ?R"
-      proof(cases "wq s cs = []")
-        case True
-        hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
-        with True and ih show ?thesis by auto
-      next
-        case False
-        hence "?R = RAG s \<union> {(Th th, Cs cs)}" by simp
-        with False and ih show ?thesis by auto
-      qed
-      ultimately show ?thesis by auto
-    next
-      case (Set thread prio)
-      with ih
-      show ?thesis by (simp add:RAG_set_unchanged)
-    qed
-  next
-    case vt_nil
-    show "finite (RAG ([]::state))"
-      by (auto simp: s_RAG_def cs_waiting_def 
-                   cs_holding_def wq_def acyclic_def)
-  qed
-qed
-
-text {* Several useful lemmas *}
-
-lemma wf_dep_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
-
-end
-
-lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l"
-  by (induct l, auto)
-
-lemma th_chasing: "(Th th, Cs cs) \<in> RAG (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> RAG s"
-  by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
-
-context valid_trace
-begin
-
-lemma wq_threads: 
-  assumes h: "th \<in> set (wq s cs)"
-  shows "th \<in> threads s"
-proof -
- from vt and h show ?thesis
-  proof(induct arbitrary: th cs)
-    case (vt_cons s e)
-    interpret vt_s: valid_trace s
-      using vt_cons(1) by (unfold_locales, auto)
-    assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s"
-      and stp: "step s e"
-      and vt: "vt s"
-      and h: "th \<in> set (wq (e # s) cs)"
-    show ?case
-    proof(cases e)
-      case (Create th' prio)
-      with ih h show ?thesis
-        by (auto simp:wq_def Let_def)
-    next
-      case (Exit th')
-      with stp ih h show ?thesis
-        apply (auto simp:wq_def Let_def)
-        apply (ind_cases "step s (Exit th')")
-        apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def
-               s_RAG_def s_holding_def cs_holding_def)
-        done
-    next
-      case (V th' cs')
-      show ?thesis
-      proof(cases "cs' = cs")
-        case False
-        with h
-        show ?thesis
-          apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def)
-          by (drule_tac ih, simp)
-      next
-        case True
-        from h
-        show ?thesis
-        proof(unfold V wq_def)
-          assume th_in: "th \<in> set (wq_fun (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l")
-          show "th \<in> threads (V th' cs' # s)"
-          proof(cases "cs = cs'")
-            case False
-            hence "?l = wq_fun (schs s) cs" by (simp add:Let_def)
-            with th_in have " th \<in> set (wq s cs)" 
-              by (fold wq_def, simp)
-            from ih [OF this] show ?thesis by simp
-          next
-            case True
-            show ?thesis
-            proof(cases "wq_fun (schs s) cs'")
-              case Nil
-              with h V show ?thesis
-                apply (auto simp:wq_def Let_def split:if_splits)
-                by (fold wq_def, drule_tac ih, simp)
-            next
-              case (Cons a rest)
-              assume eq_wq: "wq_fun (schs s) cs' = a # rest"
-              with h V show ?thesis
-                apply (auto simp:Let_def wq_def split:if_splits)
-              proof -
-                assume th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
-                have "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
-                proof(rule someI2)
-                  from vt_s.wq_distinct[of cs'] and eq_wq[folded wq_def]
-                  show "distinct rest \<and> set rest = set rest" by auto
-                next
-                  show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
-                    by auto
-                qed
-                with eq_wq th_in have "th \<in> set (wq_fun (schs s) cs')" by auto
-                from ih[OF this[folded wq_def]] show "th \<in> threads s" .
-              next
-                assume th_in: "th \<in> set (wq_fun (schs s) cs)"
-                from ih[OF this[folded wq_def]]
-                show "th \<in> threads s" .
-              qed
-            qed
-          qed
-        qed
-      qed
-    next
-      case (P th' cs')
-      from h stp
-      show ?thesis
-        apply (unfold P wq_def)
-        apply (auto simp:Let_def split:if_splits, fold wq_def)
-        apply (auto intro:ih)
-        apply(ind_cases "step s (P th' cs')")
-        by (unfold runing_def readys_def, auto)
-    next
-      case (Set thread prio)
-      with ih h show ?thesis
-        by (auto simp:wq_def Let_def)
-    qed
-  next
-    case vt_nil
-    thus ?case by (auto simp:wq_def)
-  qed
-qed
-
-lemma range_in: "\<lbrakk>(Th th) \<in> Range (RAG (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
-  apply(unfold s_RAG_def cs_waiting_def cs_holding_def)
-  by (auto intro:wq_threads)
-
-lemma readys_v_eq:
-  assumes neq_th: "th \<noteq> thread"
-  and eq_wq: "wq s cs = thread#rest"
-  and not_in: "th \<notin>  set rest"
-  shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)"
-proof -
-  from assms show ?thesis
-    apply (auto simp:readys_def)
-    apply(simp add:s_waiting_def[folded wq_def])
-    apply (erule_tac x = csa in allE)
-    apply (simp add:s_waiting_def wq_def Let_def split:if_splits)
-    apply (case_tac "csa = cs", simp)
-    apply (erule_tac x = cs in allE)
-    apply(auto simp add: s_waiting_def[folded wq_def] Let_def split: list.splits)
-    apply(auto simp add: wq_def)
-    apply (auto simp:s_waiting_def wq_def Let_def split:list.splits)
-    proof -
-       assume th_nin: "th \<notin> set rest"
-        and th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
-        and eq_wq: "wq_fun (schs s) cs = thread # rest"
-      have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-      proof(rule someI2)
-        from wq_distinct[of cs, unfolded wq_def] and eq_wq[unfolded wq_def]
-        show "distinct rest \<and> set rest = set rest" by auto
-      next
-        show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-      qed
-      with th_nin th_in show False by auto
-    qed
-qed
-
-text {* \noindent
-  The following lemmas shows that: starting from any node in @{text "RAG"}, 
-  by chasing out-going edges, it is always possible to reach a node representing a ready
-  thread. In this lemma, it is the @{text "th'"}.
-*}
-
-lemma chain_building:
-  shows "node \<in> Domain (RAG s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (RAG s)^+)"
-proof -
-  from wf_dep_converse
-  have h: "wf ((RAG s)\<inverse>)" .
-  show ?thesis
-  proof(induct rule:wf_induct [OF h])
-    fix x
-    assume ih [rule_format]: 
-      "\<forall>y. (y, x) \<in> (RAG s)\<inverse> \<longrightarrow> 
-           y \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (RAG s)\<^sup>+)"
-    show "x \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+)"
-    proof
-      assume x_d: "x \<in> Domain (RAG s)"
-      show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+"
-      proof(cases x)
-        case (Th th)
-        from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> RAG s" by (auto simp:s_RAG_def)
-        with Th have x_in_r: "(Cs cs, x) \<in> (RAG s)^-1" by simp
-        from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> RAG s" by blast
-        hence "Cs cs \<in> Domain (RAG s)" by auto
-        from ih [OF x_in_r this] obtain th'
-          where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (RAG s)\<^sup>+" by auto
-        have "(x, Th th') \<in> (RAG s)\<^sup>+" using Th x_in cs_in by auto
-        with th'_ready show ?thesis by auto
-      next
-        case (Cs cs)
-        from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (RAG s)^-1" by (auto simp:s_RAG_def)
-        show ?thesis
-        proof(cases "th' \<in> readys s")
-          case True
-          from True and th'_d show ?thesis by auto
-        next
-          case False
-          from th'_d and range_in  have "th' \<in> threads s" by auto
-          with False have "Th th' \<in> Domain (RAG s)" 
-            by (auto simp:readys_def wq_def s_waiting_def s_RAG_def cs_waiting_def Domain_def)
-          from ih [OF th'_d this]
-          obtain th'' where 
-            th''_r: "th'' \<in> readys s" and 
-            th''_in: "(Th th', Th th'') \<in> (RAG s)\<^sup>+" by auto
-          from th'_d and th''_in 
-          have "(x, Th th'') \<in> (RAG s)\<^sup>+" by auto
-          with th''_r show ?thesis by auto
-        qed
-      qed
-    qed
-  qed
-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 waiting_eq: "waiting s th cs = waiting (wq s) th cs"
-  by  (unfold s_waiting_def cs_waiting_def wq_def, auto)
-
-lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
-  by (unfold s_holding_def wq_def cs_holding_def, simp)
-
-lemma holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2"
-  by (unfold s_holding_def cs_holding_def, auto)
-
-context valid_trace
-begin
-
-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 holding_unique)
-
-end
-
-
-lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
-by (induct rule:trancl_induct, auto)
-
-context valid_trace
-begin
-
-lemma dchain_unique:
-  assumes th1_d: "(n, Th th1) \<in> (RAG s)^+"
-  and th1_r: "th1 \<in> readys s"
-  and th2_d: "(n, Th th2) \<in> (RAG s)^+"
-  and th2_r: "th2 \<in> readys s"
-  shows "th1 = th2"
-proof -
-  { assume neq: "th1 \<noteq> th2"
-    hence "Th th1 \<noteq> Th th2" by simp
-    from unique_chain [OF _ th1_d th2_d this] and unique_RAG 
-    have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG s)\<^sup>+" by auto
-    hence "False"
-    proof
-      assume "(Th th1, Th th2) \<in> (RAG s)\<^sup>+"
-      from trancl_split [OF this]
-      obtain n where dd: "(Th th1, n) \<in> RAG s" by auto
-      then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
-      from dd eq_n have "th1 \<notin> readys s"
-        by (auto simp:readys_def s_RAG_def wq_def s_waiting_def cs_waiting_def)
-      with th1_r show ?thesis by auto
-    next
-      assume "(Th th2, Th th1) \<in> (RAG s)\<^sup>+"
-      from trancl_split [OF this]
-      obtain n where dd: "(Th th2, n) \<in> RAG s" by auto
-      then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
-      from dd eq_n have "th2 \<notin> readys s"
-        by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def)
-      with th2_r show ?thesis by auto
-    qed
-  } thus ?thesis by auto
-qed
-
-end
-             
-
-lemma step_holdents_p_add:
-  assumes vt: "vt (P th cs#s)"
-  and "wq s cs = []"
-  shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
-proof -
-  from assms show ?thesis
-  unfolding  holdents_test step_RAG_p[OF vt] by (auto)
-qed
-
-lemma step_holdents_p_eq:
-  assumes vt: "vt (P th cs#s)"
-  and "wq s cs \<noteq> []"
-  shows "holdents (P th cs#s) th = holdents s th"
-proof -
-  from assms show ?thesis
-  unfolding  holdents_test step_RAG_p[OF vt] by auto
-qed
-
-
-lemma (in valid_trace) finite_holding :
-  shows "finite (holdents s th)"
-proof -
-  let ?F = "\<lambda> (x, y). the_cs x"
-  from finite_RAG 
-  have "finite (RAG s)" .
-  hence "finite (?F `(RAG s))" by simp
-  moreover have "{cs . (Cs cs, Th th) \<in> RAG s} \<subseteq> \<dots>" 
-  proof -
-    { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
-      fix x assume "(Cs x, Th th) \<in> RAG s"
-      hence "?F (Cs x, Th th) \<in> ?F `(RAG s)" by (rule h)
-      moreover have "?F (Cs x, Th th) = x" by simp
-      ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` RAG s" by simp 
-    } thus ?thesis by auto
-  qed
-  ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset)
-qed
-
-lemma cntCS_v_dec: 
-  assumes vtv: "vt (V thread cs#s)"
-  shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
-proof -
-  from vtv interpret vt_s: valid_trace s
-    by (cases, unfold_locales, simp)
-  from vtv interpret vt_v: valid_trace "V thread cs#s"
-     by (unfold_locales, simp)
-  from step_back_step[OF vtv]
-  have cs_in: "cs \<in> holdents s thread" 
-    apply (cases, unfold holdents_test s_RAG_def, simp)
-    by (unfold cs_holding_def s_holding_def wq_def, auto)
-  moreover have cs_not_in: 
-    "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
-    apply (insert vt_s.wq_distinct[of cs])
-    apply (unfold holdents_test, unfold step_RAG_v[OF vtv],
-            auto simp:next_th_def)
-  proof -
-    fix rest
-    assume dst: "distinct (rest::thread list)"
-      and ne: "rest \<noteq> []"
-    and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-    moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-    qed
-    ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
-                     set (SOME q. distinct q \<and> set q = set rest)" by simp
-    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      fix x assume " distinct x \<and> set x = set rest" with ne
-      show "x \<noteq> []" by auto
-    qed
-    ultimately 
-    show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
-      by auto
-  next
-    fix rest
-    assume dst: "distinct (rest::thread list)"
-      and ne: "rest \<noteq> []"
-    and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-    moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
-    qed
-    ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
-                     set (SOME q. distinct q \<and> set q = set rest)" by simp
-    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-    proof(rule someI2)
-      from dst show "distinct rest \<and> set rest = set rest" by auto
-    next
-      fix x assume " distinct x \<and> set x = set rest" with ne
-      show "x \<noteq> []" by auto
-    qed
-    ultimately show "False" by auto 
-  qed
-  ultimately 
-  have "holdents s thread = insert cs (holdents (V thread cs#s) thread)"
-    by auto
-  moreover have "card \<dots> = 
-                    Suc (card ((holdents (V thread cs#s) thread) - {cs}))"
-  proof(rule card_insert)
-    from vt_v.finite_holding
-    show " finite (holdents (V thread cs # s) thread)" .
-  qed
-  moreover from cs_not_in 
-  have "cs \<notin> (holdents (V thread cs#s) thread)" by auto
-  ultimately show ?thesis by (simp add:cntCS_def)
-qed 
-
-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_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 isP_E:
-  assumes "isP e"
-  obtains cs where "e = P (actor e) cs"
-  using assms by (cases e, auto)
-
-lemma isV_E:
-  assumes "isV e"
-  obtains cs where "e = V (actor e) cs"
-  using assms by (cases e, auto) (* ccc *)
-
-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)
-
-context valid_trace
-begin
-
-text {* (* ddd *) \noindent
-  The relationship between @{text "cntP"}, @{text "cntV"} and @{text "cntCS"} 
-  of one particular thread. 
-*} 
-
-lemma cnp_cnv_cncs:
-  shows "cntP s th = cntV s th + (if (th \<in> readys s \<or> th \<notin> threads s) 
-                                       then cntCS s th else cntCS s th + 1)"
-proof -
-  from vt show ?thesis
-  proof(induct arbitrary:th)
-    case (vt_cons s e)
-    interpret vt_s: valid_trace s using vt_cons(1) by (unfold_locales, simp)
-    assume vt: "vt s"
-    and ih: "\<And>th. cntP s th  = cntV s th +
-               (if (th \<in> readys s \<or> th \<notin> threads s) then cntCS s th else cntCS s th + 1)"
-    and stp: "step s e"
-    from stp show ?case
-    proof(cases)
-      case (thread_create thread prio)
-      assume eq_e: "e = Create thread prio"
-        and not_in: "thread \<notin> threads s"
-      show ?thesis
-      proof -
-        { fix cs 
-          assume "thread \<in> set (wq s cs)"
-          from vt_s.wq_threads [OF this] have "thread \<in> threads s" .
-          with not_in have "False" by simp
-        } with eq_e have eq_readys: "readys (e#s) = readys s \<union> {thread}"
-          by (auto simp:readys_def threads.simps s_waiting_def 
-            wq_def cs_waiting_def Let_def)
-        from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
-        from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
-        have eq_cncs: "cntCS (e#s) th = cntCS s th"
-          unfolding cntCS_def holdents_test
-          by (simp add:RAG_create_unchanged eq_e)
-        { assume "th \<noteq> thread"
-          with eq_readys eq_e
-          have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
-                      (th \<in> readys (s) \<or> th \<notin> threads (s))" 
-            by (simp add:threads.simps)
-          with eq_cnp eq_cnv eq_cncs ih not_in
-          have ?thesis by simp
-        } moreover {
-          assume eq_th: "th = thread"
-          with not_in ih have " cntP s th  = cntV s th + cntCS s th" by simp
-          moreover from eq_th and eq_readys have "th \<in> readys (e#s)" by simp
-          moreover note eq_cnp eq_cnv eq_cncs
-          ultimately have ?thesis by auto
-        } ultimately show ?thesis by blast
-      qed
-    next
-      case (thread_exit thread)
-      assume eq_e: "e = Exit thread" 
-      and is_runing: "thread \<in> runing s"
-      and no_hold: "holdents s thread = {}"
-      from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
-      from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
-      have eq_cncs: "cntCS (e#s) th = cntCS s th"
-        unfolding cntCS_def holdents_test
-        by (simp add:RAG_exit_unchanged eq_e)
-      { assume "th \<noteq> thread"
-        with eq_e
-        have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
-          (th \<in> readys (s) \<or> th \<notin> threads (s))" 
-          apply (simp add:threads.simps readys_def)
-          apply (subst s_waiting_def)
-          apply (simp add:Let_def)
-          apply (subst s_waiting_def, simp)
-          done
-        with eq_cnp eq_cnv eq_cncs ih
-        have ?thesis by simp
-      } moreover {
-        assume eq_th: "th = thread"
-        with ih is_runing have " cntP s th = cntV s th + cntCS s th" 
-          by (simp add:runing_def)
-        moreover from eq_th eq_e have "th \<notin> threads (e#s)"
-          by simp
-        moreover note eq_cnp eq_cnv eq_cncs
-        ultimately have ?thesis by auto
-      } ultimately show ?thesis by blast
-    next
-      case (thread_P thread cs)
-      assume eq_e: "e = P thread cs"
-        and is_runing: "thread \<in> runing s"
-        and no_dep: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+"
-      from thread_P vt stp ih  have vtp: "vt (P thread cs#s)" by auto
-      then interpret vt_p: valid_trace "(P thread cs#s)"
-        by (unfold_locales, simp)
-      show ?thesis 
-      proof -
-        { have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast
-          assume neq_th: "th \<noteq> thread"
-          with eq_e
-          have eq_readys: "(th \<in> readys (e#s)) = (th \<in> readys (s))"
-            apply (simp add:readys_def s_waiting_def wq_def Let_def)
-            apply (rule_tac hh)
-             apply (intro iffI allI, clarify)
-            apply (erule_tac x = csa in allE, auto)
-            apply (subgoal_tac "wq_fun (schs s) cs \<noteq> []", auto)
-            apply (erule_tac x = cs in allE, auto)
-            by (case_tac "(wq_fun (schs s) cs)", auto)
-          moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th"
-            apply (simp add:cntCS_def holdents_test)
-            by (unfold  step_RAG_p [OF vtp], auto)
-          moreover from eq_e neq_th have "cntP (e # s) th = cntP s th"
-            by (simp add:cntP_def count_def)
-          moreover from eq_e neq_th have "cntV (e#s) th = cntV s th"
-            by (simp add:cntV_def count_def)
-          moreover from eq_e neq_th have "threads (e#s) = threads s" by simp
-          moreover note ih [of th] 
-          ultimately have ?thesis by simp
-        } moreover {
-          assume eq_th: "th = thread"
-          have ?thesis
-          proof -
-            from eq_e eq_th have eq_cnp: "cntP (e # s) th  = 1 + (cntP s th)" 
-              by (simp add:cntP_def count_def)
-            from eq_e eq_th have eq_cnv: "cntV (e#s) th = cntV s th"
-              by (simp add:cntV_def count_def)
-            show ?thesis
-            proof (cases "wq s cs = []")
-              case True
-              with is_runing
-              have "th \<in> readys (e#s)"
-                apply (unfold eq_e wq_def, unfold readys_def s_RAG_def)
-                apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def)
-                by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def)
-              moreover have "cntCS (e # s) th = 1 + cntCS s th"
-              proof -
-                have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> RAG s} =
-                  Suc (card {cs. (Cs cs, Th thread) \<in> RAG s})" (is "card ?L = Suc (card ?R)")
-                proof -
-                  have "?L = insert cs ?R" by auto
-                  moreover have "card \<dots> = Suc (card (?R - {cs}))" 
-                  proof(rule card_insert)
-                    from vt_s.finite_holding [of thread]
-                    show " finite {cs. (Cs cs, Th thread) \<in> RAG s}"
-                      by (unfold holdents_test, simp)
-                  qed
-                  moreover have "?R - {cs} = ?R"
-                  proof -
-                    have "cs \<notin> ?R"
-                    proof
-                      assume "cs \<in> {cs. (Cs cs, Th thread) \<in> RAG s}"
-                      with no_dep show False by auto
-                    qed
-                    thus ?thesis by auto
-                  qed
-                  ultimately show ?thesis by auto
-                qed
-                thus ?thesis
-                  apply (unfold eq_e eq_th cntCS_def)
-                  apply (simp add: holdents_test)
-                  by (unfold step_RAG_p [OF vtp], auto simp:True)
-              qed
-              moreover from is_runing have "th \<in> readys s"
-                by (simp add:runing_def eq_th)
-              moreover note eq_cnp eq_cnv ih [of th]
-              ultimately show ?thesis by auto
-            next
-              case False
-              have eq_wq: "wq (e#s) cs = wq s cs @ [th]"
-                    by (unfold eq_th eq_e wq_def, auto simp:Let_def)
-              have "th \<notin> readys (e#s)"
-              proof
-                assume "th \<in> readys (e#s)"
-                hence "\<forall>cs. \<not> waiting (e # s) th cs" by (simp add:readys_def)
-                from this[rule_format, of cs] have " \<not> waiting (e # s) th cs" .
-                hence "th \<in> set (wq (e#s) cs) \<Longrightarrow> th = hd (wq (e#s) cs)" 
-                  by (simp add:s_waiting_def wq_def)
-                moreover from eq_wq have "th \<in> set (wq (e#s) cs)" by auto
-                ultimately have "th = hd (wq (e#s) cs)" by blast
-                with eq_wq have "th = hd (wq s cs @ [th])" by simp
-                hence "th = hd (wq s cs)" using False by auto
-                with False eq_wq vt_p.wq_distinct [of cs]
-                show False by (fold eq_e, auto)
-              qed
-              moreover from is_runing have "th \<in> threads (e#s)" 
-                by (unfold eq_e, auto simp:runing_def readys_def eq_th)
-              moreover have "cntCS (e # s) th = cntCS s th"
-                apply (unfold cntCS_def holdents_test eq_e step_RAG_p[OF vtp])
-                by (auto simp:False)
-              moreover note eq_cnp eq_cnv ih[of th]
-              moreover from is_runing have "th \<in> readys s"
-                by (simp add:runing_def eq_th)
-              ultimately show ?thesis by auto
-            qed
-          qed
-        } ultimately show ?thesis by blast
-      qed
-    next
-      case (thread_V thread cs)
-      from assms vt stp ih thread_V have vtv: "vt (V thread cs # s)" by auto
-      then interpret vt_v: valid_trace "(V thread cs # s)" by (unfold_locales, simp)
-      assume eq_e: "e = V thread cs"
-        and is_runing: "thread \<in> runing s"
-        and hold: "holding s thread cs"
-      from hold obtain rest 
-        where eq_wq: "wq s cs = thread # rest"
-        by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
-      have eq_threads: "threads (e#s) = threads s" by (simp add: eq_e)
-      have eq_set: "set (SOME q. distinct q \<and> set q = set rest) = set rest"
-      proof(rule someI2)
-        from vt_v.wq_distinct[of cs] and eq_wq
-        show "distinct rest \<and> set rest = set rest"
-          by (metis distinct.simps(2) vt_s.wq_distinct)
-      next
-        show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
-          by auto
-      qed
-      show ?thesis
-      proof -
-        { assume eq_th: "th = thread"
-          from eq_th have eq_cnp: "cntP (e # s) th = cntP s th"
-            by (unfold eq_e, simp add:cntP_def count_def)
-          moreover from eq_th have eq_cnv: "cntV (e#s) th = 1 + cntV s th"
-            by (unfold eq_e, simp add:cntV_def count_def)
-          moreover from cntCS_v_dec [OF vtv] 
-          have "cntCS (e # s) thread + 1 = cntCS s thread"
-            by (simp add:eq_e)
-          moreover from is_runing have rd_before: "thread \<in> readys s"
-            by (unfold runing_def, simp)
-          moreover have "thread \<in> readys (e # s)"
-          proof -
-            from is_runing
-            have "thread \<in> threads (e#s)" 
-              by (unfold eq_e, auto simp:runing_def readys_def)
-            moreover have "\<forall> cs1. \<not> waiting (e#s) thread cs1"
-            proof
-              fix cs1
-              { assume eq_cs: "cs1 = cs" 
-                have "\<not> waiting (e # s) thread cs1"
-                proof -
-                  from eq_wq
-                  have "thread \<notin> set (wq (e#s) cs1)"
-                    apply(unfold eq_e wq_def eq_cs s_holding_def)
-                    apply (auto simp:Let_def)
-                  proof -
-                    assume "thread \<in> set (SOME q. distinct q \<and> set q = set rest)"
-                    with eq_set have "thread \<in> set rest" by simp
-                    with vt_v.wq_distinct[of cs]
-                    and eq_wq show False
-                        by (metis distinct.simps(2) vt_s.wq_distinct)
-                  qed
-                  thus ?thesis by (simp add:wq_def s_waiting_def)
-                qed
-              } moreover {
-                assume neq_cs: "cs1 \<noteq> cs"
-                  have "\<not> waiting (e # s) thread cs1" 
-                  proof -
-                    from wq_v_neq [OF neq_cs[symmetric]]
-                    have "wq (V thread cs # s) cs1 = wq s cs1" .
-                    moreover have "\<not> waiting s thread cs1" 
-                    proof -
-                      from runing_ready and is_runing
-                      have "thread \<in> readys s" by auto
-                      thus ?thesis by (simp add:readys_def)
-                    qed
-                    ultimately show ?thesis 
-                      by (auto simp:wq_def s_waiting_def eq_e)
-                  qed
-              } ultimately show "\<not> waiting (e # s) thread cs1" by blast
-            qed
-            ultimately show ?thesis by (simp add:readys_def)
-          qed
-          moreover note eq_th ih
-          ultimately have ?thesis by auto
-        } moreover {
-          assume neq_th: "th \<noteq> thread"
-          from neq_th eq_e have eq_cnp: "cntP (e # s) th = cntP s th" 
-            by (simp add:cntP_def count_def)
-          from neq_th eq_e have eq_cnv: "cntV (e # s) th = cntV s th" 
-            by (simp add:cntV_def count_def)
-          have ?thesis
-          proof(cases "th \<in> set rest")
-            case False
-            have "(th \<in> readys (e # s)) = (th \<in> readys s)"
-              apply (insert step_back_vt[OF vtv])
-              by (simp add: False eq_e eq_wq neq_th vt_s.readys_v_eq)
-            moreover have "cntCS (e#s) th = cntCS s th"
-              apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
-              proof -
-                have "{csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
-                      {cs. (Cs cs, Th th) \<in> RAG s}"
-                proof -
-                  from False eq_wq
-                  have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> RAG s"
-                    apply (unfold next_th_def, auto)
-                  proof -
-                    assume ne: "rest \<noteq> []"
-                      and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-                      and eq_wq: "wq s cs = thread # rest"
-                    from eq_set ni have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
-                                  set (SOME q. distinct q \<and> set q = set rest)
-                                  " by simp
-                    moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-                    proof(rule someI2)
-                      from vt_s.wq_distinct[ of cs] and eq_wq
-                      show "distinct rest \<and> set rest = set rest" by auto
-                    next
-                      fix x assume "distinct x \<and> set x = set rest"
-                      with ne show "x \<noteq> []" by auto
-                    qed
-                    ultimately show 
-                      "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
-                      by auto
-                  qed    
-                  thus ?thesis by auto
-                qed
-                thus "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
-                             card {cs. (Cs cs, Th th) \<in> RAG s}" by simp 
-              qed
-            moreover note ih eq_cnp eq_cnv eq_threads
-            ultimately show ?thesis by auto
-          next
-            case True
-            assume th_in: "th \<in> set rest"
-            show ?thesis
-            proof(cases "next_th s thread cs th")
-              case False
-              with eq_wq and th_in have 
-                neq_hd: "th \<noteq> hd (SOME q. distinct q \<and> set q = set rest)" (is "th \<noteq> hd ?rest")
-                by (auto simp:next_th_def)
-              have "(th \<in> readys (e # s)) = (th \<in> readys s)"
-              proof -
-                from eq_wq and th_in
-                have "\<not> th \<in> readys s"
-                  apply (auto simp:readys_def s_waiting_def)
-                  apply (rule_tac x = cs in exI, auto)
-                  by (insert vt_s.wq_distinct[of cs], auto simp add: wq_def)
-                moreover 
-                from eq_wq and th_in and neq_hd
-                have "\<not> (th \<in> readys (e # s))"
-                  apply (auto simp:readys_def s_waiting_def eq_e wq_def Let_def split:list.splits)
-                  by (rule_tac x = cs in exI, auto simp:eq_set)
-                ultimately show ?thesis by auto
-              qed
-              moreover have "cntCS (e#s) th = cntCS s th" 
-              proof -
-                from eq_wq and  th_in and neq_hd
-                have "(holdents (e # s) th) = (holdents s th)"
-                  apply (unfold eq_e step_RAG_v[OF vtv], 
-                         auto simp:next_th_def eq_set s_RAG_def holdents_test wq_def
-                                   Let_def cs_holding_def)
-                  by (insert vt_s.wq_distinct[of cs], auto simp:wq_def)
-                thus ?thesis by (simp add:cntCS_def)
-              qed
-              moreover note ih eq_cnp eq_cnv eq_threads
-              ultimately show ?thesis by auto
-            next
-              case True
-              let ?rest = " (SOME q. distinct q \<and> set q = set rest)"
-              let ?t = "hd ?rest"
-              from True eq_wq th_in neq_th
-              have "th \<in> readys (e # s)"
-                apply (auto simp:eq_e readys_def s_waiting_def wq_def
-                        Let_def next_th_def)
-              proof -
-                assume eq_wq: "wq_fun (schs s) cs = thread # rest"
-                  and t_in: "?t \<in> set rest"
-                show "?t \<in> threads s"
-                proof(rule vt_s.wq_threads)
-                  from eq_wq and t_in
-                  show "?t \<in> set (wq s cs)" by (auto simp:wq_def)
-                qed
-              next
-                fix csa
-                assume eq_wq: "wq_fun (schs s) cs = thread # rest"
-                  and t_in: "?t \<in> set rest"
-                  and neq_cs: "csa \<noteq> cs"
-                  and t_in': "?t \<in>  set (wq_fun (schs s) csa)"
-                show "?t = hd (wq_fun (schs s) csa)"
-                proof -
-                  { assume neq_hd': "?t \<noteq> hd (wq_fun (schs s) csa)"
-                    from vt_s.wq_distinct[of cs] and 
-                    eq_wq[folded wq_def] and t_in eq_wq
-                    have "?t \<noteq> thread" by auto
-                    with eq_wq and t_in
-                    have w1: "waiting s ?t cs"
-                      by (auto simp:s_waiting_def wq_def)
-                    from t_in' neq_hd'
-                    have w2: "waiting s ?t csa"
-                      by (auto simp:s_waiting_def wq_def)
-                    from vt_s.waiting_unique[OF w1 w2]
-                    and neq_cs have "False" by auto
-                  } thus ?thesis by auto
-                qed
-              qed
-              moreover have "cntP s th = cntV s th + cntCS s th + 1"
-              proof -
-                have "th \<notin> readys s" 
-                proof -
-                  from True eq_wq neq_th th_in
-                  show ?thesis
-                    apply (unfold readys_def s_waiting_def, auto)
-                    by (rule_tac x = cs in exI, auto simp add: wq_def)
-                qed
-                moreover have "th \<in> threads s"
-                proof -
-                  from th_in eq_wq
-                  have "th \<in> set (wq s cs)" by simp
-                  from vt_s.wq_threads [OF this] 
-                  show ?thesis .
-                qed
-                ultimately show ?thesis using ih by auto
-              qed
-              moreover from True neq_th have "cntCS (e # s) th = 1 + cntCS s th"
-                apply (unfold cntCS_def holdents_test eq_e step_RAG_v[OF vtv], auto)
-              proof -
-                show "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs} =
-                               Suc (card {cs. (Cs cs, Th th) \<in> RAG s})"
-                  (is "card ?A = Suc (card ?B)")
-                proof -
-                  have "?A = insert cs ?B" by auto
-                  hence "card ?A = card (insert cs ?B)" by simp
-                  also have "\<dots> = Suc (card ?B)"
-                  proof(rule card_insert_disjoint)
-                    have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` RAG s)" 
-                      apply (auto simp:image_def)
-                      by (rule_tac x = "(Cs x, Th th)" in bexI, auto)
-                    with vt_s.finite_RAG
-                    show "finite {cs. (Cs cs, Th th) \<in> RAG s}" by (auto intro:finite_subset)
-                  next
-                    show "cs \<notin> {cs. (Cs cs, Th th) \<in> RAG s}"
-                    proof
-                      assume "cs \<in> {cs. (Cs cs, Th th) \<in> RAG s}"
-                      hence "(Cs cs, Th th) \<in> RAG s" by simp
-                      with True neq_th eq_wq show False
-                        by (auto simp:next_th_def s_RAG_def cs_holding_def)
-                    qed
-                  qed
-                  finally show ?thesis .
-                qed
-              qed
-              moreover note eq_cnp eq_cnv
-              ultimately show ?thesis by simp
-            qed
-          qed
-        } ultimately show ?thesis by blast
-      qed
-    next
-      case (thread_set thread prio)
-      assume eq_e: "e = Set thread prio"
-        and is_runing: "thread \<in> runing s"
-      show ?thesis
-      proof -
-        from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
-        from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
-        have eq_cncs: "cntCS (e#s) th = cntCS s th"
-          unfolding cntCS_def holdents_test
-          by (simp add:RAG_set_unchanged eq_e)
-        from eq_e have eq_readys: "readys (e#s) = readys s" 
-          by (simp add:readys_def cs_waiting_def s_waiting_def wq_def,
-                  auto simp:Let_def)
-        { assume "th \<noteq> thread"
-          with eq_readys eq_e
-          have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
-                      (th \<in> readys (s) \<or> th \<notin> threads (s))" 
-            by (simp add:threads.simps)
-          with eq_cnp eq_cnv eq_cncs ih is_runing
-          have ?thesis by simp
-        } moreover {
-          assume eq_th: "th = thread"
-          with is_runing ih have " cntP s th  = cntV s th + cntCS s th" 
-            by (unfold runing_def, auto)
-          moreover from eq_th and eq_readys is_runing have "th \<in> readys (e#s)"
-            by (simp add:runing_def)
-          moreover note eq_cnp eq_cnv eq_cncs
-          ultimately have ?thesis by auto
-        } ultimately show ?thesis by blast
-      qed   
-    qed
-  next
-    case vt_nil
-    show ?case 
-      by (unfold cntP_def cntV_def cntCS_def, 
-        auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
-  qed
-qed
-
-lemma not_thread_cncs:
-  assumes not_in: "th \<notin> threads s" 
-  shows "cntCS s th = 0"
-proof -
-  from vt not_in show ?thesis
-  proof(induct arbitrary:th)
-    case (vt_cons s e th)
-    interpret vt_s: valid_trace s using vt_cons(1)
-       by (unfold_locales, simp)
-    assume vt: "vt s"
-      and ih: "\<And>th. th \<notin> threads s \<Longrightarrow> cntCS s th = 0"
-      and stp: "step s e"
-      and not_in: "th \<notin> threads (e # s)"
-    from stp show ?case
-    proof(cases)
-      case (thread_create thread prio)
-      assume eq_e: "e = Create thread prio"
-        and not_in': "thread \<notin> threads s"
-      have "cntCS (e # s) th = cntCS s th"
-        apply (unfold eq_e cntCS_def holdents_test)
-        by (simp add:RAG_create_unchanged)
-      moreover have "th \<notin> threads s" 
-      proof -
-        from not_in eq_e show ?thesis by simp
-      qed
-      moreover note ih ultimately show ?thesis by auto
-    next
-      case (thread_exit thread)
-      assume eq_e: "e = Exit thread"
-      and nh: "holdents s thread = {}"
-      have eq_cns: "cntCS (e # s) th = cntCS s th"
-        apply (unfold eq_e cntCS_def holdents_test)
-        by (simp add:RAG_exit_unchanged)
-      show ?thesis
-      proof(cases "th = thread")
-        case True
-        have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True)
-        with eq_cns show ?thesis by simp
-      next
-        case False
-        with not_in and eq_e
-        have "th \<notin> threads s" by simp
-        from ih[OF this] and eq_cns show ?thesis by simp
-      qed
-    next
-      case (thread_P thread cs)
-      assume eq_e: "e = P thread cs"
-      and is_runing: "thread \<in> runing s"
-      from assms thread_P ih vt stp thread_P have vtp: "vt (P thread cs#s)" by auto
-      have neq_th: "th \<noteq> thread" 
-      proof -
-        from not_in eq_e have "th \<notin> threads s" by simp
-        moreover from is_runing have "thread \<in> threads s"
-          by (simp add:runing_def readys_def)
-        ultimately show ?thesis by auto
-      qed
-      hence "cntCS (e # s) th  = cntCS s th "
-        apply (unfold cntCS_def holdents_test eq_e)
-        by (unfold step_RAG_p[OF vtp], auto)
-      moreover have "cntCS s th = 0"
-      proof(rule ih)
-        from not_in eq_e show "th \<notin> threads s" by simp
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (thread_V thread cs)
-      assume eq_e: "e = V thread cs"
-        and is_runing: "thread \<in> runing s"
-        and hold: "holding s thread cs"
-      have neq_th: "th \<noteq> thread" 
-      proof -
-        from not_in eq_e have "th \<notin> threads s" by simp
-        moreover from is_runing have "thread \<in> threads s"
-          by (simp add:runing_def readys_def)
-        ultimately show ?thesis by auto
-      qed
-      from assms thread_V vt stp ih 
-      have vtv: "vt (V thread cs#s)" by auto
-      then interpret vt_v: valid_trace "(V thread cs#s)"
-        by (unfold_locales, simp)
-      from hold obtain rest 
-        where eq_wq: "wq s cs = thread # rest"
-        by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
-      from not_in eq_e eq_wq
-      have "\<not> next_th s thread cs th"
-        apply (auto simp:next_th_def)
-      proof -
-        assume ne: "rest \<noteq> []"
-          and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> threads s" (is "?t \<notin> threads s")
-        have "?t \<in> set rest"
-        proof(rule someI2)
-          from vt_v.wq_distinct[of cs] and eq_wq
-          show "distinct rest \<and> set rest = set rest"
-            by (metis distinct.simps(2) vt_s.wq_distinct) 
-        next
-          fix x assume "distinct x \<and> set x = set rest" with ne
-          show "hd x \<in> set rest" by (cases x, auto)
-        qed
-        with eq_wq have "?t \<in> set (wq s cs)" by simp
-        from vt_s.wq_threads[OF this] and ni
-        show False
-          using `hd (SOME q. distinct q \<and> set q = set rest) \<in> set (wq s cs)` 
-            ni vt_s.wq_threads by blast 
-      qed
-      moreover note neq_th eq_wq
-      ultimately have "cntCS (e # s) th  = cntCS s th"
-        by (unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
-      moreover have "cntCS s th = 0"
-      proof(rule ih)
-        from not_in eq_e show "th \<notin> threads s" by simp
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (thread_set thread prio)
-      print_facts
-      assume eq_e: "e = Set thread prio"
-        and is_runing: "thread \<in> runing s"
-      from not_in and eq_e have "th \<notin> threads s" by auto
-      from ih [OF this] and eq_e
-      show ?thesis 
-        apply (unfold eq_e cntCS_def holdents_test)
-        by (simp add:RAG_set_unchanged)
-    qed
-    next
-      case vt_nil
-      show ?case
-      by (unfold cntCS_def, 
-        auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
-  qed
-qed
-
-end
-
-lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs"
-  by (auto simp:s_waiting_def cs_waiting_def wq_def)
-
-context valid_trace
-begin
-
-lemma dm_RAG_threads:
-  assumes in_dom: "(Th th) \<in> Domain (RAG s)"
-  shows "th \<in> threads s"
-proof -
-  from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
-  moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
-  ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
-  hence "th \<in> set (wq s cs)"
-    by (unfold s_RAG_def, auto simp:cs_waiting_def)
-  from wq_threads [OF this] show ?thesis .
-qed
-
-end
-
-lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
-unfolding cp_def wq_def
-apply(induct s rule: schs.induct)
-thm cpreced_initial
-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
-
-context valid_trace
-begin
-
-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
-    apply(simp)
-    done
-  hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1)) =
-                 Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependants (wq s) th2))"
-    (is "Max (?f ` ?A) = Max (?f ` ?B)")
-    unfolding cp_eq_cpreced 
-    unfolding cpreced_def .
-  obtain th1' where th1_in: "th1' \<in> ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)"
-  proof -
-    have h1: "finite (?f ` ?A)"
-    proof -
-      have "finite ?A" 
-      proof -
-        have "finite (dependants (wq s) th1)"
-        proof-
-          have "finite {th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+}"
-          proof -
-            let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
-              apply (auto simp:image_def)
-              by (rule_tac x = "(Th x, Th th1)" in bexI, auto)
-            moreover have "finite \<dots>"
-            proof -
-              from finite_RAG have "finite (RAG s)" .
-              hence "finite ((RAG (wq s))\<^sup>+)"
-                apply (unfold finite_trancl)
-                by (auto simp: s_RAG_def cs_RAG_def wq_def)
-              thus ?thesis by auto
-            qed
-            ultimately show ?thesis by (auto intro:finite_subset)
-          qed
-          thus ?thesis by (simp add:cs_dependants_def)
-        qed
-        thus ?thesis by simp
-      qed
-      thus ?thesis by auto
-    qed
-    moreover have h2: "(?f ` ?A) \<noteq> {}"
-    proof -
-      have "?A \<noteq> {}" by simp
-      thus ?thesis by simp
-    qed
-    from Max_in [OF h1 h2]
-    have "Max (?f ` ?A) \<in> (?f ` ?A)" .
-    thus ?thesis 
-      thm cpreced_def
-      unfolding cpreced_def[symmetric] 
-      unfolding cp_eq_cpreced[symmetric] 
-      unfolding cpreced_def 
-      using that[intro] by (auto)
-  qed
-  obtain th2' where th2_in: "th2' \<in> ?B" and eq_f_th2: "?f th2' = Max (?f ` ?B)"
-  proof -
-    have h1: "finite (?f ` ?B)"
-    proof -
-      have "finite ?B" 
-      proof -
-        have "finite (dependants (wq s) th2)"
-        proof-
-          have "finite {th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+}"
-          proof -
-            let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
-              apply (auto simp:image_def)
-              by (rule_tac x = "(Th x, Th th2)" in bexI, auto)
-            moreover have "finite \<dots>"
-            proof -
-              from finite_RAG have "finite (RAG s)" .
-              hence "finite ((RAG (wq s))\<^sup>+)"
-                apply (unfold finite_trancl)
-                by (auto simp: s_RAG_def cs_RAG_def wq_def)
-              thus ?thesis by auto
-            qed
-            ultimately show ?thesis by (auto intro:finite_subset)
-          qed
-          thus ?thesis by (simp add:cs_dependants_def)
-        qed
-        thus ?thesis by simp
-      qed
-      thus ?thesis by auto
-    qed
-    moreover have h2: "(?f ` ?B) \<noteq> {}"
-    proof -
-      have "?B \<noteq> {}" by simp
-      thus ?thesis by simp
-    qed
-    from Max_in [OF h1 h2]
-    have "Max (?f ` ?B) \<in> (?f ` ?B)" .
-    thus ?thesis by (auto intro:that)
-  qed
-  from eq_f_th1 eq_f_th2 eq_max 
-  have eq_preced: "preced th1' s = preced th2' s" by auto
-  hence eq_th12: "th1' = th2'"
-  proof (rule preced_unique)
-    from th1_in have "th1' = th1 \<or> (th1' \<in> dependants (wq s) th1)" by simp
-    thus "th1' \<in> threads s"
-    proof
-      assume "th1' \<in> dependants (wq s) th1"
-      hence "(Th th1') \<in> Domain ((RAG s)^+)"
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "(Th th1') \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      from dm_RAG_threads[OF this] show ?thesis .
-    next
-      assume "th1' = th1"
-      with runing_1 show ?thesis
-        by (unfold runing_def readys_def, auto)
-    qed
-  next
-    from th2_in have "th2' = th2 \<or> (th2' \<in> dependants (wq s) th2)" by simp
-    thus "th2' \<in> threads s"
-    proof
-      assume "th2' \<in> dependants (wq s) th2"
-      hence "(Th th2') \<in> Domain ((RAG s)^+)"
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "(Th th2') \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      from dm_RAG_threads[OF this] show ?thesis .
-    next
-      assume "th2' = th2"
-      with runing_2 show ?thesis
-        by (unfold runing_def readys_def, auto)
-    qed
-  qed
-  from th1_in have "th1' = th1 \<or> th1' \<in> dependants (wq s) th1" by simp
-  thus ?thesis
-  proof
-    assume eq_th': "th1' = th1"
-    from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
-    thus ?thesis
-    proof
-      assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
-    next
-      assume "th2' \<in> dependants (wq s) th2"
-      with eq_th12 eq_th' have "th1 \<in> dependants (wq s) th2" by simp
-      hence "(Th th1, Th th2) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      hence "Th th1 \<in> Domain ((RAG s)^+)" 
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "Th th1 \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      then obtain n where d: "(Th th1, n) \<in> RAG s" by (auto simp:Domain_def)
-      from RAG_target_th [OF this]
-      obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th1, Cs cs') \<in> RAG s" by simp
-      with runing_1 have "False"
-        apply (unfold runing_def readys_def s_RAG_def)
-        by (auto simp:eq_waiting)
-      thus ?thesis by simp
-    qed
-  next
-    assume th1'_in: "th1' \<in> dependants (wq s) th1"
-    from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
-    thus ?thesis 
-    proof
-      assume "th2' = th2"
-      with th1'_in eq_th12 have "th2 \<in> dependants (wq s) th1" by simp
-      hence "(Th th2, Th th1) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      hence "Th th2 \<in> Domain ((RAG s)^+)" 
-        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
-        by (auto simp:Domain_def)
-      hence "Th th2 \<in> Domain (RAG s)" by (simp add:trancl_domain)
-      then obtain n where d: "(Th th2, n) \<in> RAG s" by (auto simp:Domain_def)
-      from RAG_target_th [OF this]
-      obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th2, Cs cs') \<in> RAG s" by simp
-      with runing_2 have "False"
-        apply (unfold runing_def readys_def s_RAG_def)
-        by (auto simp:eq_waiting)
-      thus ?thesis by simp
-    next
-      assume "th2' \<in> dependants (wq s) th2"
-      with eq_th12 have "th1' \<in> dependants (wq s) th2" by simp
-      hence h1: "(Th th1', Th th2) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      from th1'_in have h2: "(Th th1', Th th1) \<in> (RAG s)^+"
-        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
-      show ?thesis
-      proof(rule dchain_unique[OF h1 _ h2, symmetric])
-        from runing_1 show "th1 \<in> readys s" by (simp add:runing_def)
-        from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) 
-      qed
-    qed
-  qed
-qed
-
-
-lemma "card (runing s) \<le> 1"
-apply(subgoal_tac "finite (runing s)")
-prefer 2
-apply (metis finite_nat_set_iff_bounded lessI runing_unique)
-apply(rule ccontr)
-apply(simp)
-apply(case_tac "Suc (Suc 0) \<le> card (runing s)")
-apply(subst (asm) card_le_Suc_iff)
-apply(simp)
-apply(auto)[1]
-apply (metis insertCI runing_unique)
-apply(auto) 
-done
-
-end
-
-
-lemma create_pre:
-  assumes stp: "step s e"
-  and not_in: "th \<notin> threads s"
-  and is_in: "th \<in> threads (e#s)"
-  obtains prio where "e = Create th prio"
-proof -
-  from assms  
-  show ?thesis
-  proof(cases)
-    case (thread_create thread prio)
-    with is_in not_in have "e = Create th prio" by simp
-    from that[OF this] show ?thesis .
-  next
-    case (thread_exit thread)
-    with assms show ?thesis by (auto intro!:that)
-  next
-    case (thread_P thread)
-    with assms show ?thesis by (auto intro!:that)
-  next
-    case (thread_V thread)
-    with assms show ?thesis by (auto intro!:that)
-  next 
-    case (thread_set thread)
-    with assms show ?thesis by (auto intro!:that)
-  qed
-qed
-
-
-context valid_trace
-begin
-
-lemma cnp_cnv_eq:
-  assumes "th \<notin> threads s"
-  shows "cntP s th = cntV s th"
-  using assms
-  using cnp_cnv_cncs not_thread_cncs by auto
-
-end
-
-
-lemma eq_RAG: 
-  "RAG (wq s) = RAG s"
-by (unfold cs_RAG_def s_RAG_def, auto)
-
-context valid_trace
-begin
-
-lemma count_eq_dependants:
-  assumes eq_pv: "cntP s th = cntV s th"
-  shows "dependants (wq s) th = {}"
-proof -
-  from cnp_cnv_cncs and eq_pv
-  have "cntCS s th = 0" 
-    by (auto split:if_splits)
-  moreover have "finite {cs. (Cs cs, Th th) \<in> RAG s}"
-  proof -
-    from finite_holding[of th] show ?thesis
-      by (simp add:holdents_test)
-  qed
-  ultimately have h: "{cs. (Cs cs, Th th) \<in> RAG s} = {}"
-    by (unfold cntCS_def holdents_test cs_dependants_def, auto)
-  show ?thesis
-  proof(unfold cs_dependants_def)
-    { assume "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}"
-      then obtain th' where "(Th th', Th th) \<in> (RAG (wq s))\<^sup>+" by auto
-      hence "False"
-      proof(cases)
-        assume "(Th th', Th th) \<in> RAG (wq s)"
-        thus "False" by (auto simp:cs_RAG_def)
-      next
-        fix c
-        assume "(c, Th th) \<in> RAG (wq s)"
-        with h and eq_RAG show "False"
-          by (cases c, auto simp:cs_RAG_def)
-      qed
-    } thus "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} = {}" by auto
-  qed
-qed
-
-lemma dependants_threads:
-  shows "dependants (wq s) th \<subseteq> threads s"
-proof
-  { fix th th'
-    assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (RAG (wq s))\<^sup>+}"
-    have "Th th \<in> Domain (RAG s)"
-    proof -
-      from h obtain th' where "(Th th, Th th') \<in> (RAG (wq s))\<^sup>+" by auto
-      hence "(Th th) \<in> Domain ( (RAG (wq s))\<^sup>+)" by (auto simp:Domain_def)
-      with trancl_domain have "(Th th) \<in> Domain (RAG (wq s))" by simp
-      thus ?thesis using eq_RAG by simp
-    qed
-    from dm_RAG_threads[OF this]
-    have "th \<in> threads s" .
-  } note hh = this
-  fix th1 
-  assume "th1 \<in> dependants (wq s) th"
-  hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (RAG (wq s))\<^sup>+}"
-    by (unfold cs_dependants_def, simp)
-  from hh [OF this] show "th1 \<in> threads s" .
-qed
-
-lemma finite_threads:
-  shows "finite (threads s)"
-using vt by (induct) (auto elim: step.cases)
-
-end
-
-lemma Max_f_mono:
-  assumes seq: "A \<subseteq> B"
-  and np: "A \<noteq> {}"
-  and fnt: "finite B"
-  shows "Max (f ` A) \<le> Max (f ` B)"
-proof(rule Max_mono)
-  from seq show "f ` A \<subseteq> f ` B" by auto
-next
-  from np show "f ` A \<noteq> {}" by auto
-next
-  from fnt and seq show "finite (f ` B)" by auto
-qed
-
-context valid_trace
-begin
-
-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 le_cp:
-  shows "preced th s \<le> cp s th"
-proof(unfold cp_eq_cpreced preced_def cpreced_def, simp)
-  show "Prc (priority th s) (last_set th s)
-    \<le> Max (insert (Prc (priority th s) (last_set th s))
-            ((\<lambda>th. Prc (priority th s) (last_set th s)) ` dependants (wq s) th))"
-    (is "?l \<le> Max (insert ?l ?A)")
-  proof(cases "?A = {}")
-    case False
-    have "finite ?A" (is "finite (?f ` ?B)")
-    proof -
-      have "finite ?B" 
-      proof-
-        have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}"
-        proof -
-          let ?F = "\<lambda> (x, y). the_th x"
-          have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
-            apply (auto simp:image_def)
-            by (rule_tac x = "(Th x, Th th)" in bexI, auto)
-          moreover have "finite \<dots>"
-          proof -
-            from finite_RAG have "finite (RAG s)" .
-            hence "finite ((RAG (wq s))\<^sup>+)"
-              apply (unfold finite_trancl)
-              by (auto simp: s_RAG_def cs_RAG_def wq_def)
-            thus ?thesis by auto
-          qed
-          ultimately show ?thesis by (auto intro:finite_subset)
-        qed
-        thus ?thesis by (simp add:cs_dependants_def)
-      qed
-      thus ?thesis by simp
-    qed
-    from Max_insert [OF this False, of ?l] show ?thesis by auto
-  next
-    case True
-    thus ?thesis by auto
-  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_readys_threads_pre:
-  assumes np: "threads s \<noteq> {}"
-  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
-proof(unfold max_cp_eq)
-  show "Max (cp s ` readys s) = Max ((\<lambda>th. preced th s) ` threads s)"
-  proof -
-    let ?p = "Max ((\<lambda>th. preced th s) ` threads s)" 
-    let ?f = "(\<lambda>th. preced th s)"
-    have "?p \<in> ((\<lambda>th. preced th s) ` threads s)"
-    proof(rule Max_in)
-      from finite_threads show "finite (?f ` threads s)" by simp
-    next
-      from np show "?f ` threads s \<noteq> {}" by simp
-    qed
-    then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
-      by (auto simp:Image_def)
-    from th_chain_to_ready [OF tm_in]
-    have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" .
-    thus ?thesis
-    proof
-      assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ "
-      then obtain th' where th'_in: "th' \<in> readys s" 
-        and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto
-      have "cp s th' = ?f tm"
-      proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
-        from dependants_threads finite_threads
-        show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" 
-          by (auto intro:finite_subset)
-      next
-        fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
-        from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" .
-        moreover have "p \<le> \<dots>"
-        proof(rule Max_ge)
-          from finite_threads
-          show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        next
-          from p_in and th'_in and dependants_threads[of th']
-          show "p \<in> (\<lambda>th. preced th s) ` threads s"
-            by (auto simp:readys_def)
-        qed
-        ultimately show "p \<le> preced tm s" by auto
-      next
-        show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
-        proof -
-          from tm_chain
-          have "tm \<in> dependants (wq s) th'"
-            by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto)
-          thus ?thesis by auto
-        qed
-      qed
-      with tm_max
-      have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
-      show ?thesis
-      proof (fold h, rule Max_eqI)
-        fix q 
-        assume "q \<in> cp s ` readys s"
-        then obtain th1 where th1_in: "th1 \<in> readys s"
-          and eq_q: "q = cp s th1" by auto
-        show "q \<le> cp s th'"
-          apply (unfold h eq_q)
-          apply (unfold cp_eq_cpreced cpreced_def)
-          apply (rule Max_mono)
-        proof -
-          from dependants_threads [of th1] th1_in
-          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<subseteq> 
-                 (\<lambda>th. preced th s) ` threads s"
-            by (auto simp:readys_def)
-        next
-          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}" by simp
-        next
-          from finite_threads 
-          show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        qed
-      next
-        from finite_threads
-        show "finite (cp s ` readys s)" by (auto simp:readys_def)
-      next
-        from th'_in
-        show "cp s th' \<in> cp s ` readys s" by simp
-      qed
-    next
-      assume tm_ready: "tm \<in> readys s"
-      show ?thesis
-      proof(fold tm_max)
-        have cp_eq_p: "cp s tm = preced tm s"
-        proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
-          fix y 
-          assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
-          show "y \<le> preced tm s"
-          proof -
-            { fix y'
-              assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependants (wq s) tm)"
-              have "y' \<le> preced tm s"
-              proof(unfold tm_max, rule Max_ge)
-                from hy' dependants_threads[of tm]
-                show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
-              next
-                from finite_threads
-                show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-              qed
-            } with hy show ?thesis by auto
-          qed
-        next
-          from dependants_threads[of tm] finite_threads
-          show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm))"
-            by (auto intro:finite_subset)
-        next
-          show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
-            by simp
-        qed 
-        moreover have "Max (cp s ` readys s) = cp s tm"
-        proof(rule Max_eqI)
-          from tm_ready show "cp s tm \<in> cp s ` readys s" by simp
-        next
-          from finite_threads
-          show "finite (cp s ` readys s)" by (auto simp:readys_def)
-        next
-          fix y assume "y \<in> cp s ` readys s"
-          then obtain th1 where th1_readys: "th1 \<in> readys s"
-            and h: "y = cp s th1" by auto
-          show "y \<le> cp s tm"
-            apply(unfold cp_eq_p h)
-            apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono)
-          proof -
-            from finite_threads
-            show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-          next
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}"
-              by simp
-          next
-            from dependants_threads[of th1] th1_readys
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) 
-                    \<subseteq> (\<lambda>th. preced th s) ` threads s"
-              by (auto simp:readys_def)
-          qed
-        qed
-        ultimately show " Max (cp s ` readys s) = preced tm s" by simp
-      qed 
-    qed
-  qed
-qed
-
-text {* (* ccc *) \noindent
-  Since the current precedence of the threads in ready queue will always be boosted,
-  there must be one inside it has the maximum precedence of the whole system. 
-*}
-lemma max_cp_readys_threads:
-  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
-proof(cases "threads s = {}")
-  case True
-  thus ?thesis 
-    by (auto simp:readys_def)
-next
-  case False
-  show ?thesis by (rule max_cp_readys_threads_pre[OF False])
-qed
-
-end
-
-lemma eq_holding: "holding (wq s) th cs = holding s th cs"
-  apply (unfold s_holding_def cs_holding_def wq_def, simp)
-  done
-
-lemma f_image_eq:
-  assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
-  shows "f ` A = g ` A"
-proof
-  show "f ` A \<subseteq> g ` A"
-    by(rule image_subsetI, auto intro:h)
-next
-  show "g ` A \<subseteq> f ` A"
-   by (rule image_subsetI, auto intro:h[symmetric])
-qed
-
-
-definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
-  where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
-
-
-lemma detached_test:
-  shows "detached s th = (Th th \<notin> Field (RAG s))"
-apply(simp add: detached_def Field_def)
-apply(simp add: s_RAG_def)
-apply(simp add: s_holding_abv s_waiting_abv)
-apply(simp add: Domain_iff Range_iff)
-apply(simp add: wq_def)
-apply(auto)
-done
-
-context valid_trace
-begin
-
-lemma detached_intro:
-  assumes eq_pv: "cntP s th = cntV s th"
-  shows "detached s th"
-proof -
- from cnp_cnv_cncs
-  have eq_cnt: "cntP s th =
-    cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
-  hence cncs_zero: "cntCS s th = 0"
-    by (auto simp:eq_pv split:if_splits)
-  with eq_cnt
-  have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:eq_pv)
-  thus ?thesis
-  proof
-    assume "th \<notin> threads s"
-    with range_in dm_RAG_threads
-    show ?thesis
-      by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def Domain_iff Range_iff)
-  next
-    assume "th \<in> readys s"
-    moreover have "Th th \<notin> Range (RAG s)"
-    proof -
-      from card_0_eq [OF finite_holding] and cncs_zero
-      have "holdents s th = {}"
-        by (simp add:cntCS_def)
-      thus ?thesis
-        apply(auto simp:holdents_test)
-        apply(case_tac a)
-        apply(auto simp:holdents_test s_RAG_def)
-        done
-    qed
-    ultimately show ?thesis
-      by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def readys_def)
-  qed
-qed
-
-lemma detached_elim:
-  assumes dtc: "detached s th"
-  shows "cntP s th = cntV s th"
-proof -
-  from cnp_cnv_cncs
-  have eq_pv: " cntP s th =
-    cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
-  have cncs_z: "cntCS s th = 0"
-  proof -
-    from dtc have "holdents s th = {}"
-      unfolding detached_def holdents_test s_RAG_def
-      by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
-    thus ?thesis by (auto simp:cntCS_def)
-  qed
-  show ?thesis
-  proof(cases "th \<in> threads s")
-    case True
-    with dtc 
-    have "th \<in> readys s"
-      by (unfold readys_def detached_def Field_def Domain_def Range_def, 
-           auto simp:eq_waiting s_RAG_def)
-    with cncs_z and eq_pv show ?thesis by simp
-  next
-    case False
-    with cncs_z and eq_pv show ?thesis by simp
-  qed
-qed
-
-lemma detached_eq:
-  shows "(detached s th) = (cntP s th = cntV s th)"
-  by (insert vt, auto intro:detached_intro detached_elim)
-
-end
-
-text {* 
-  The lemmas in this .thy file are all obvious lemmas, however, they still needs to be derived
-  from the concise and miniature model of PIP given in PrioGDef.thy.
-*}
-
-lemma eq_dependants: "dependants (wq s) = dependants s"
-  by (simp add: s_dependants_abv wq_def)
-
-lemma next_th_unique: 
-  assumes nt1: "next_th s th cs th1"
-  and nt2: "next_th s th cs th2"
-  shows "th1 = th2"
-using assms by (unfold next_th_def, auto)
-
-lemma birth_time_lt:  "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
-  apply (induct s, simp)
-proof -
-  fix a s
-  assume ih: "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
-    and eq_as: "a # s \<noteq> []"
-  show "last_set th (a # s) < length (a # s)"
-  proof(cases "s \<noteq> []")
-    case False
-    from False show ?thesis
-      by (cases a, auto simp:last_set.simps)
-  next
-    case True
-    from ih [OF True] show ?thesis
-      by (cases a, auto simp:last_set.simps)
-  qed
-qed
-
-lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
-  by (induct s, auto simp:threads.simps)
-
-lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
-  apply (drule_tac th_in_ne)
-  by (unfold preced_def, auto intro: birth_time_lt)
-
-lemma inj_the_preced: 
-  "inj_on (the_preced s) (threads s)"
-  by (metis inj_onI preced_unique the_preced_def)
-
-lemma tRAG_alt_def: 
-  "tRAG s = {(Th th1, Th th2) | th1 th2. 
-                  \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
- by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
-
-lemma tRAG_Field:
-  "Field (tRAG s) \<subseteq> Field (RAG s)"
-  by (unfold tRAG_alt_def Field_def, auto)
-
-lemma tRAG_ancestorsE:
-  assumes "x \<in> ancestors (tRAG s) u"
-  obtains th where "x = Th th"
-proof -
-  from assms have "(u, x) \<in> (tRAG s)^+" 
-      by (unfold ancestors_def, auto)
-  from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
-  then obtain th where "x = Th th"
-    by (unfold tRAG_alt_def, auto)
-  from that[OF this] show ?thesis .
-qed
-
-lemma tRAG_mono:
-  assumes "RAG s' \<subseteq> RAG s"
-  shows "tRAG s' \<subseteq> tRAG s"
-  using assms 
-  by (unfold tRAG_alt_def, auto)
-
-lemma holding_next_thI:
-  assumes "holding s th cs"
-  and "length (wq s cs) > 1"
-  obtains th' where "next_th s th cs th'"
-proof -
-  from assms(1)[folded eq_holding, unfolded cs_holding_def]
-  have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" .
-  then obtain rest where h1: "wq s cs = th#rest" 
-    by (cases "wq s cs", auto)
-  with assms(2) have h2: "rest \<noteq> []" by auto
-  let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
-  have "next_th s th cs ?th'" using  h1(1) h2 
-    by (unfold next_th_def, auto)
-  from that[OF this] show ?thesis .
-qed
-
-lemma RAG_tRAG_transfer:
-  assumes "vt s'"
-  assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
-  and "(Cs cs, Th th'') \<in> RAG s'"
-  shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
-proof -
-  interpret vt_s': valid_trace "s'" using assms(1)
-    by (unfold_locales, simp)
-  interpret rtree: rtree "RAG s'"
-  proof
-  show "single_valued (RAG s')"
-  apply (intro_locales)
-    by (unfold single_valued_def, 
-        auto intro:vt_s'.unique_RAG)
-
-  show "acyclic (RAG s')"
-     by (rule vt_s'.acyclic_RAG)
-  qed
-  { fix n1 n2
-    assume "(n1, n2) \<in> ?L"
-    from this[unfolded tRAG_alt_def]
-    obtain th1 th2 cs' where 
-      h: "n1 = Th th1" "n2 = Th th2" 
-         "(Th th1, Cs cs') \<in> RAG s"
-         "(Cs cs', Th th2) \<in> RAG s" by auto
-    from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto
-    from h(3) and assms(2) 
-    have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
-          (Th th1, Cs cs') \<in> RAG s'" by auto
-    hence "(n1, n2) \<in> ?R"
-    proof
-      assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
-      hence eq_th1: "th1 = th" by simp
-      moreover have "th2 = th''"
-      proof -
-        from h1 have "cs' = cs" by simp
-        from assms(3) cs_in[unfolded this] rtree.sgv
-        show ?thesis
-          by (unfold single_valued_def, auto)
-      qed
-      ultimately show ?thesis using h(1,2) by auto
-    next
-      assume "(Th th1, Cs cs') \<in> RAG s'"
-      with cs_in have "(Th th1, Th th2) \<in> tRAG s'"
-        by (unfold tRAG_alt_def, auto)
-      from this[folded h(1, 2)] show ?thesis by auto
-    qed
-  } moreover {
-    fix n1 n2
-    assume "(n1, n2) \<in> ?R"
-    hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto
-    hence "(n1, n2) \<in> ?L" 
-    proof
-      assume "(n1, n2) \<in> tRAG s'"
-      moreover have "... \<subseteq> ?L"
-      proof(rule tRAG_mono)
-        show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto)
-      qed
-      ultimately show ?thesis by auto
-    next
-      assume eq_n: "(n1, n2) = (Th th, Th th'')"
-      from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto
-      moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto
-      ultimately show ?thesis 
-        by (unfold eq_n tRAG_alt_def, auto)
-    qed
-  } ultimately show ?thesis by auto
-qed
-
-context valid_trace
-begin
-
-lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt]
-
-end
-
-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
-
-lemma cp_gen_alt_def:
-  "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
-    by (auto simp:cp_gen_def)
-
-lemma tRAG_nodeE:
-  assumes "(n1, n2) \<in> tRAG s"
-  obtains th1 th2 where "n1 = Th th1" "n2 = Th th2"
-  using assms
-  by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def)
-
-lemma subtree_nodeE:
-  assumes "n \<in> subtree (tRAG s) (Th th)"
-  obtains th1 where "n = Th th1"
-proof -
-  show ?thesis
-  proof(rule subtreeE[OF assms])
-    assume "n = Th th"
-    from that[OF this] show ?thesis .
-  next
-    assume "Th th \<in> ancestors (tRAG s) n"
-    hence "(n, Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def)
-    hence "\<exists> th1. n = Th th1"
-    proof(induct)
-      case (base y)
-      from tRAG_nodeE[OF this] show ?case by metis
-    next
-      case (step y z)
-      thus ?case by auto
-    qed
-    with that show ?thesis by auto
-  qed
-qed
-
-lemma tRAG_star_RAG: "(tRAG s)^* \<subseteq> (RAG s)^*"
-proof -
-  have "(wRAG s O hRAG s)^* \<subseteq> (RAG s O RAG s)^*" 
-    by (rule rtrancl_mono, auto simp:RAG_split)
-  also have "... \<subseteq> ((RAG s)^*)^*"
-    by (rule rtrancl_mono, auto)
-  also have "... = (RAG s)^*" by simp
-  finally show ?thesis by (unfold tRAG_def, simp)
-qed
-
-lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x"
-proof -
-  { fix a
-    assume "a \<in> subtree (tRAG s) x"
-    hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def)
-    with tRAG_star_RAG[of s]
-    have "(a, x) \<in> (RAG s)^*" by auto
-    hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def)
-  } thus ?thesis by auto
-qed
-
-lemma tRAG_trancl_eq:
-   "{th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
-    {th'. (Th th', Th th)  \<in> (RAG s)^+}"
-   (is "?L = ?R")
-proof -
-  { fix th'
-    assume "th' \<in> ?L"
-    hence "(Th th', Th th) \<in> (tRAG s)^+" by auto
-    from tranclD[OF this]
-    obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto
-    from tRAG_subtree_RAG[of s] and this(2)
-    have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) 
-    moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto 
-    ultimately have "th' \<in> ?R"  by auto 
-  } moreover 
-  { fix th'
-    assume "th' \<in> ?R"
-    hence "(Th th', Th th) \<in> (RAG s)^+" by (auto)
-    from plus_rpath[OF this]
-    obtain xs where rp: "rpath (RAG s) (Th th') xs (Th th)" "xs \<noteq> []" by auto
-    hence "(Th th', Th th) \<in> (tRAG s)^+"
-    proof(induct xs arbitrary:th' th rule:length_induct)
-      case (1 xs th' th)
-      then obtain x1 xs1 where Cons1: "xs = x1#xs1" by (cases xs, auto)
-      show ?case
-      proof(cases "xs1")
-        case Nil
-        from 1(2)[unfolded Cons1 Nil]
-        have rp: "rpath (RAG s) (Th th') [x1] (Th th)" .
-        hence "(Th th', x1) \<in> (RAG s)" by (cases, simp)
-        then obtain cs where "x1 = Cs cs" 
-              by (unfold s_RAG_def, auto)
-        from rpath_nnl_lastE[OF rp[unfolded this]]
-        show ?thesis by auto
-      next
-        case (Cons x2 xs2)
-        from 1(2)[unfolded Cons1[unfolded this]]
-        have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" .
-        from rpath_edges_on[OF this]
-        have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" .
-        have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)"
-            by (simp add: edges_on_unfold)
-        with eds have rg1: "(Th th', x1) \<in> RAG s" by auto
-        then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto)
-        have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)"
-            by (simp add: edges_on_unfold)
-        from this eds
-        have rg2: "(x1, x2) \<in> RAG s" by auto
-        from this[unfolded eq_x1] 
-        obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto)
-        from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2]
-        have rt1: "(Th th', Th th1) \<in> tRAG s" by (unfold tRAG_alt_def, auto)
-        from rp have "rpath (RAG s) x2 xs2 (Th th)"
-           by  (elim rpath_ConsE, simp)
-        from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" .
-        show ?thesis
-        proof(cases "xs2 = []")
-          case True
-          from rpath_nilE[OF rp'[unfolded this]]
-          have "th1 = th" by auto
-          from rt1[unfolded this] show ?thesis by auto
-        next
-          case False
-          from 1(1)[rule_format, OF _ rp' this, unfolded Cons1 Cons]
-          have "(Th th1, Th th) \<in> (tRAG s)\<^sup>+" by simp
-          with rt1 show ?thesis by auto
-        qed
-      qed
-    qed
-    hence "th' \<in> ?L" by auto
-  } ultimately show ?thesis by blast
-qed
-
-lemma tRAG_trancl_eq_Th:
-   "{Th th' | th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
-    {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}"
-    using tRAG_trancl_eq by auto
-
-lemma dependants_alt_def:
-  "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
-  by (metis eq_RAG s_dependants_def tRAG_trancl_eq)
-  
-context valid_trace
-begin
-
-lemma count_eq_tRAG_plus:
-  assumes "cntP s th = cntV s th"
-  shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
-  using assms count_eq_dependants dependants_alt_def eq_dependants by auto 
-
-lemma count_eq_RAG_plus:
-  assumes "cntP s th = cntV s th"
-  shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
-  using assms count_eq_dependants cs_dependants_def eq_RAG by auto
-
-lemma count_eq_RAG_plus_Th:
-  assumes "cntP s th = cntV s th"
-  shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
-  using count_eq_RAG_plus[OF assms] by auto
-
-lemma count_eq_tRAG_plus_Th:
-  assumes "cntP s th = cntV s th"
-  shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
-   using count_eq_tRAG_plus[OF assms] by auto
-
-end
-
-lemma tRAG_subtree_eq: 
-   "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
-   (is "?L = ?R")
-proof -
-  { fix n 
-    assume h: "n \<in> ?L"
-    hence "n \<in> ?R"
-    by (smt mem_Collect_eq subsetCE subtree_def subtree_nodeE tRAG_subtree_RAG) 
-  } moreover {
-    fix n
-    assume "n \<in> ?R"
-    then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*"
-      by (auto simp:subtree_def)
-    from rtranclD[OF this(2)]
-    have "n \<in> ?L"
-    proof
-      assume "Th th' \<noteq> Th th \<and> (Th th', Th th) \<in> (RAG s)\<^sup>+"
-      with h have "n \<in> {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}" by auto
-      thus ?thesis using subtree_def tRAG_trancl_eq by fastforce
-    qed (insert h, auto simp:subtree_def)
-  } ultimately show ?thesis by auto
-qed
-
-lemma threads_set_eq: 
-   "the_thread ` (subtree (tRAG s) (Th th)) = 
-                  {th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R")
-   by (auto intro:rev_image_eqI simp:tRAG_subtree_eq)
-
-lemma cp_alt_def1: 
-  "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))"
-proof -
-  have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) =
-       ((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))"
-       by auto
-  thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto)
-qed
-
-lemma cp_gen_def_cond: 
-  assumes "x = Th th"
-  shows "cp s th = cp_gen s (Th th)"
-by (unfold cp_alt_def1 cp_gen_def, simp)
-
-lemma cp_gen_over_set:
-  assumes "\<forall> x \<in> A. \<exists> th. x = Th th"
-  shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A"
-proof(rule f_image_eq)
-  fix a
-  assume "a \<in> A"
-  from assms[rule_format, OF this]
-  obtain th where eq_a: "a = Th th" by auto
-  show "cp_gen s a = (cp s \<circ> the_thread) a"
-    by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
-qed
-
-
-context valid_trace
-begin
-
-lemma RAG_threads:
-  assumes "(Th th) \<in> Field (RAG s)"
-  shows "th \<in> threads s"
-  using assms
-  by (metis Field_def UnE dm_RAG_threads range_in vt)
-
-lemma subtree_tRAG_thread:
-  assumes "th \<in> threads s"
-  shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R")
-proof -
-  have "?L = {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
-    by (unfold tRAG_subtree_eq, simp)
-  also have "... \<subseteq> ?R"
-  proof
-    fix x
-    assume "x \<in> {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
-    then obtain th' where h: "x = Th th'" "Th th' \<in> subtree (RAG s) (Th th)" by auto
-    from this(2)
-    show "x \<in> ?R"
-    proof(cases rule:subtreeE)
-      case 1
-      thus ?thesis by (simp add: assms h(1)) 
-    next
-      case 2
-      thus ?thesis by (metis ancestors_Field dm_RAG_threads h(1) image_eqI) 
-    qed
-  qed
-  finally show ?thesis .
-qed
-
-lemma readys_root:
-  assumes "th \<in> readys s"
-  shows "root (RAG s) (Th th)"
-proof -
-  { fix x
-    assume "x \<in> ancestors (RAG s) (Th th)"
-    hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def)
-    from tranclD[OF this]
-    obtain z where "(Th th, z) \<in> RAG s" by auto
-    with assms(1) have False
-         apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def)
-         by (fold wq_def, blast)
-  } thus ?thesis by (unfold root_def, auto)
-qed
-
-lemma readys_in_no_subtree:
-  assumes "th \<in> readys s"
-  and "th' \<noteq> th"
-  shows "Th th \<notin> subtree (RAG s) (Th th')" 
-proof
-   assume "Th th \<in> subtree (RAG s) (Th th')"
-   thus False
-   proof(cases rule:subtreeE)
-      case 1
-      with assms show ?thesis by auto
-   next
-      case 2
-      with readys_root[OF assms(1)]
-      show ?thesis by (auto simp:root_def)
-   qed
-qed
-
-lemma not_in_thread_isolated:
-  assumes "th \<notin> threads s"
-  shows "(Th th) \<notin> Field (RAG s)"
-proof
-  assume "(Th th) \<in> Field (RAG s)"
-  with dm_RAG_threads and range_in assms
-  show False by (unfold Field_def, blast)
-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 holding_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 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
-
-sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
-proof -
-  have "fsubtree (tRAG s)"
-  proof -
-    have "fbranch (tRAG s)"
-    proof(unfold tRAG_def, rule fbranch_compose)
-        show "fbranch (wRAG s)"
-        proof(rule finite_fbranchI)
-           from finite_RAG show "finite (wRAG s)"
-           by (unfold RAG_split, auto)
-        qed
-    next
-        show "fbranch (hRAG s)"
-        proof(rule finite_fbranchI)
-           from finite_RAG 
-           show "finite (hRAG s)" by (unfold RAG_split, auto)
-        qed
-    qed
-    moreover have "wf (tRAG s)"
-    proof(rule wf_subset)
-      show "wf (RAG s O RAG s)" using wf_RAG
-        by (fold wf_comp_self, simp)
-    next
-      show "tRAG s \<subseteq> (RAG s O RAG s)"
-        by (unfold tRAG_alt_def, auto)
-    qed
-    ultimately show ?thesis
-      by (unfold fsubtree_def fsubtree_axioms_def,auto)
-  qed
-  from this[folded tRAG_def] show "fsubtree (tRAG s)" .
-qed
-
-lemma Max_UNION: 
-  assumes "finite A"
-  and "A \<noteq> {}"
-  and "\<forall> M \<in> f ` A. finite M"
-  and "\<forall> M \<in> f ` A. M \<noteq> {}"
-  shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R")
-  using assms[simp]
-proof -
-  have "?L = Max (\<Union>(f ` A))"
-    by (fold Union_image_eq, simp)
-  also have "... = ?R"
-    by (subst Max_Union, simp+)
-  finally show ?thesis .
-qed
-
-lemma max_Max_eq:
-  assumes "finite A"
-    and "A \<noteq> {}"
-    and "x = y"
-  shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R")
-proof -
-  have "?R = Max (insert y A)" by simp
-  also from assms have "... = ?L"
-      by (subst Max.insert, simp+)
-  finally show ?thesis by simp
-qed
-
-context valid_trace
-begin
-
-(* ddd *)
-lemma cp_gen_rec:
-  assumes "x = Th th"
-  shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
-proof(cases "children (tRAG s) x = {}")
-  case True
-  show ?thesis
-    by (unfold True cp_gen_def subtree_children, simp add:assms)
-next
-  case False
-  hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
-  note fsbttRAGs.finite_subtree[simp]
-  have [simp]: "finite (children (tRAG s) x)"
-     by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], 
-            rule children_subtree)
-  { fix r x
-    have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
-  } note this[simp]
-  have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
-  proof -
-    from False obtain q where "q \<in> children (tRAG s) x" by blast
-    moreover have "subtree (tRAG s) q \<noteq> {}" by simp
-    ultimately show ?thesis by blast
-  qed
-  have h: "Max ((the_preced s \<circ> the_thread) `
-                ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
-        Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
-                     (is "?L = ?R")
-  proof -
-    let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L
-    let "Max (_ \<union> (?h ` ?B))" = ?R
-    let ?L1 = "?f ` \<Union>(?g ` ?B)"
-    have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)"
-    proof -
-      have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp
-      also have "... =  (\<Union> x \<in> ?B. ?f ` (?g x))" by auto
-      finally have "Max ?L1 = Max ..." by simp
-      also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)"
-        by (subst Max_UNION, simp+)
-      also have "... = Max (cp_gen s ` children (tRAG s) x)"
-          by (unfold image_comp cp_gen_alt_def, simp)
-      finally show ?thesis .
-    qed
-    show ?thesis
-    proof -
-      have "?L = Max (?f ` ?A \<union> ?L1)" by simp
-      also have "... = max (the_preced s (the_thread x)) (Max ?L1)"
-            by (subst Max_Un, simp+)
-      also have "... = max (?f x) (Max (?h ` ?B))"
-        by (unfold eq_Max_L1, simp)
-      also have "... =?R"
-        by (rule max_Max_eq, (simp)+, unfold assms, simp)
-      finally show ?thesis .
-    qed
-  qed  thus ?thesis 
-          by (fold h subtree_children, unfold cp_gen_def, simp) 
-qed
-
-lemma cp_rec:
-  "cp s th = Max ({the_preced s th} \<union> 
-                     (cp s o the_thread) ` children (tRAG s) (Th th))"
-proof -
-  have "Th th = Th th" by simp
-  note h =  cp_gen_def_cond[OF this] cp_gen_rec[OF this]
-  show ?thesis 
-  proof -
-    have "cp_gen s ` children (tRAG s) (Th th) = 
-                (cp s \<circ> the_thread) ` children (tRAG s) (Th th)"
-    proof(rule cp_gen_over_set)
-      show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th"
-        by (unfold tRAG_alt_def, auto simp:children_def)
-    qed
-    thus ?thesis by (subst (1) h(1), unfold h(2), simp)
-  qed
-qed
-
-end
-
-(* keep *)
-lemma next_th_holding:
-  assumes vt: "vt s"
-  and nxt: "next_th s th cs th'"
-  shows "holding (wq s) th cs"
-proof -
-  from nxt[unfolded next_th_def]
-  obtain rest where h: "wq s cs = th # rest"
-                       "rest \<noteq> []" 
-                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
-  thus ?thesis
-    by (unfold cs_holding_def, auto)
-qed
-
-context valid_trace
-begin
-
-lemma next_th_waiting:
-  assumes nxt: "next_th s th cs th'"
-  shows "waiting (wq s) th' cs"
-proof -
-  from nxt[unfolded next_th_def]
-  obtain rest where h: "wq s cs = th # rest"
-                       "rest \<noteq> []" 
-                       "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
-  from wq_distinct[of cs, unfolded h]
-  have dst: "distinct (th # rest)" .
-  have in_rest: "th' \<in> set rest"
-  proof(unfold h, rule someI2)
-    show "distinct rest \<and> set rest = set rest" using dst by auto
-  next
-    fix x assume "distinct x \<and> set x = set rest"
-    with h(2)
-    show "hd x \<in> set (rest)" by (cases x, auto)
-  qed
-  hence "th' \<in> set (wq s cs)" by (unfold h(1), auto)
-  moreover have "th' \<noteq> hd (wq s cs)"
-    by (unfold h(1), insert in_rest dst, auto)
-  ultimately show ?thesis by (auto simp:cs_waiting_def)
-qed
-
-lemma next_th_RAG:
-  assumes nxt: "next_th (s::event list) th cs th'"
-  shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s"
-  using vt assms next_th_holding next_th_waiting
-  by (unfold s_RAG_def, simp)
-
-end
-
--- {* A useless definition *}
-definition cps:: "state \<Rightarrow> (thread \<times> precedence) set"
-where "cps s = {(th, cp s th) | th . th \<in> threads s}"
-
-
-find_theorems "waiting" holding
-context valid_trace
-begin
-
-find_theorems "waiting" holding
-
-end
-
-
-end
--- a/PIPDefs.thy~	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,665 +0,0 @@
-chapter {* Definitions *}
-(*<*)
-theory PIPDefs
-imports Precedence_ord Moment RTree Max
-begin
-(*>*)
-
-text {*
-  In this section, the formal model of  Priority Inheritance Protocol (PIP) is presented. 
-  The model is based on Paulson's inductive protocol verification method, where 
-  the state of the system is modelled as a list of events happened so far with the latest 
-  event put at the head. 
-*}
-
-text {*
-  To define events, the identifiers of {\em threads},
-  {\em priority} and {\em critical resources } (abbreviated as @{text "cs"}) 
-  need to be represented. All three are represetned using standard 
-  Isabelle/HOL type @{typ "nat"}:
-*}
-
-type_synonym thread = nat -- {* Type for thread identifiers. *}
-type_synonym priority = nat  -- {* Type for priorities. *}
-type_synonym cs = nat -- {* Type for critical sections (or critical resources). *}
-
-text {*
-  \noindent
-  The abstraction of Priority Inheritance Protocol (PIP) is set at the system call level.
-  Every system call is represented as an event. The format of events is defined 
-  defined as follows:
-  *}
-
-datatype event = 
-  Create thread priority | -- {* Thread @{text "thread"} is created with priority @{text "priority"}. *}
-  Exit thread | -- {* Thread @{text "thread"} finishing its execution. *}
-  P thread cs | -- {* Thread @{text "thread"} requesting critical resource @{text "cs"}. *}
-  V thread cs | -- {* Thread @{text "thread"}  releasing critical resource @{text "cs"}. *}
-  Set thread priority -- {* Thread @{text "thread"} resets its priority to @{text "priority"}. *}
-
-fun actor :: "event \<Rightarrow> thread" where
-  "actor (Create th pty) = th" |
-  "actor (Exit th) = th" |
-  "actor (P th cs) = th" |
-  "actor (V th cs) = th" |
-  "actor (Set th pty) = th"
-
-fun isCreate :: "event \<Rightarrow> bool" where
-  "isCreate (Create th pty) = True" |
-  "isCreate _ = False"
-
-fun isP :: "event \<Rightarrow> bool" where
-  "isP (P th cs) = True" |
-  "isP _ = False"
-
-fun isV :: "event \<Rightarrow> bool" where
-  "isV (V th cs) = True" |
-  "isV _ = False"
-
-text {* 
-  As mentioned earlier, in Paulson's inductive method, the states of system are represented as lists of events,
-  which is defined by the following type @{text "state"}:
-  *}
-type_synonym state = "event list"
-
-
-text {* 
-\noindent
-  Resource Allocation Graph (RAG for short) is used extensively in our formal analysis. 
-  The following type @{text "node"} is used to represent nodes in RAG.
-  *}
-datatype node = 
-   Th "thread" | -- {* Node for thread. *}
-   Cs "cs" -- {* Node for critical resource. *}
-
-text {*
-  \noindent
-  The following function
-  @{text "threads"} is used to calculate the set of live threads (@{text "threads s"})
-  in state @{text "s"}.
-  *}
-fun threads :: "state \<Rightarrow> thread set"
-  where 
-  -- {* At the start of the system, the set of threads is empty: *}
-  "threads [] = {}" | 
-  -- {* New thread is added to the @{text "threads"}: *}
-  "threads (Create thread prio#s) = {thread} \<union> threads s" | 
-  -- {* Finished thread is removed: *}
-  "threads (Exit thread # s) = (threads s) - {thread}" | 
-  -- {* Other kind of events does not affect the value of @{text "threads"}: *}
-  "threads (e#s) = threads s" 
-
-text {* 
-  \noindent
-  The function @{text "threads"} defined above is one of 
-  the so called {\em observation function}s which forms 
-  the very basis of Paulson's inductive protocol verification method.
-  Each observation function {\em observes} one particular aspect (or attribute)
-  of the system. For example, the attribute observed by  @{text "threads s"}
-  is the set of threads living in state @{text "s"}. 
-  The protocol being modelled 
-  The decision made the protocol being modelled is based on the {\em observation}s
-  returned by {\em observation function}s. Since {\observation function}s forms 
-  the very basis on which Paulson's inductive method is based, there will be 
-  a lot of such observation functions introduced in the following. In fact, any function 
-  which takes event list as argument is a {\em observation function}.
-  *}
-
-text {* \noindent
-  Observation @{text "priority th s"} is
-  the {\em original priority} of thread @{text "th"} in state @{text "s"}. 
-  The {\em original priority} is the priority 
-  assigned to a thread when it is created or when it is reset by system call 
-  (represented by event @{text "Set thread priority"}).
-*}
-
-fun priority :: "thread \<Rightarrow> state \<Rightarrow> priority"
-  where
-  -- {* @{text "0"} is assigned to threads which have never been created: *}
-  "priority thread [] = 0" |
-  "priority thread (Create thread' prio#s) = 
-     (if thread' = thread then prio else priority thread s)" |
-  "priority thread (Set thread' prio#s) = 
-     (if thread' = thread then prio else priority thread s)" |
-  "priority thread (e#s) = priority thread s"
-
-text {*
-  \noindent
-  Observation @{text "last_set th s"} is the last time when the priority of thread @{text "th"} is set, 
-  observed from state @{text "s"}.
-  The time in the system is measured by the number of events happened so far since the very beginning.
-*}
-fun last_set :: "thread \<Rightarrow> state \<Rightarrow> nat"
-  where
-  "last_set thread [] = 0" |
-  "last_set thread ((Create thread' prio)#s) = 
-       (if (thread = thread') then length s else last_set thread s)" |
-  "last_set thread ((Set thread' prio)#s) = 
-       (if (thread = thread') then length s else last_set thread s)" |
-  "last_set thread (_#s) = last_set thread s"
-
-text {*
-  \noindent 
-  The {\em precedence} is a notion derived from {\em priority}, where the {\em precedence} of 
-  a thread is the combination of its {\em original priority} and {\em time} the priority is set. 
-  The intention is to discriminate threads with the same priority by giving threads whose priority
-  is assigned earlier higher precedences, becasue such threads are more urgent to finish. 
-  This explains the following definition:
-  *}
-definition preced :: "thread \<Rightarrow> state \<Rightarrow> precedence"
-  where "preced thread s \<equiv> Prc (priority thread s) (last_set thread s)"
-
-
-text {*
-  \noindent
-  A number of important notions in PIP are represented as the following functions, 
-  defined in terms of the waiting queues of the system, where the waiting queues 
-  , as a whole, is represented by the @{text "wq"} argument of every notion function.
-  The @{text "wq"} argument is itself a functions which maps every critical resource 
-  @{text "cs"} to the list of threads which are holding or waiting for it. 
-  The thread at the head of this list is designated as the thread which is current 
-  holding the resrouce, which is slightly different from tradition where
-  all threads in the waiting queue are considered as waiting for the resource.
-  *}
-
-consts 
-  holding :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool" 
-  waiting :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool"
-  RAG :: "'b \<Rightarrow> (node \<times> node) set"
-  dependants :: "'b \<Rightarrow> thread \<Rightarrow> thread set"
-
-defs (overloaded) 
-  -- {* 
-  \begin{minipage}{0.9\textwidth}
-  This meaning of @{text "wq"} is reflected in the following definition of @{text "holding wq th cs"},
-  where @{text "holding wq th cs"} means thread @{text "th"} is holding the critical 
-  resource @{text "cs"}. This decision is based on @{text "wq"}.
-  \end{minipage}
-  *}
-
-  cs_holding_def: 
-  "holding wq thread cs \<equiv> (thread \<in> set (wq cs) \<and> thread = hd (wq cs))"
-  -- {* 
-  \begin{minipage}{0.9\textwidth}
-  In accordance with the definition of @{text "holding wq th cs"}, 
-  a thread @{text "th"} is considered waiting for @{text "cs"} if 
-  it is in the {\em waiting queue} of critical resource @{text "cs"}, but not at the head.
-  This is reflected in the definition of @{text "waiting wq th cs"} as follows:
-  \end{minipage}
-  *}
-  cs_waiting_def: 
-  "waiting wq thread cs \<equiv> (thread \<in> set (wq cs) \<and> thread \<noteq> hd (wq cs))"
-  -- {* 
-  \begin{minipage}{0.9\textwidth}
-  @{text "RAG wq"} generates RAG (a binary relations on @{text "node"})
-  out of waiting queues of the system (represented by the @{text "wq"} argument):
-  \end{minipage}
-  *}
-  cs_RAG_def: 
-  "RAG (wq::cs \<Rightarrow> thread list) \<equiv>
-      {(Th th, Cs cs) | th cs. waiting wq th cs} \<union> {(Cs cs, Th th) | cs th. holding wq th cs}"
-  -- {* 
-  \begin{minipage}{0.9\textwidth}
-  The following @{text "dependants wq th"} represents the set of threads which are RAGing on
-  thread @{text "th"} in Resource Allocation Graph @{text "RAG wq"}. 
-  Here, "RAGing" means waiting directly or indirectly on the critical resource. 
-  \end{minipage}
-  *}
-  cs_dependants_def: 
-  "dependants (wq::cs \<Rightarrow> thread list) th \<equiv> {th' . (Th th', Th th) \<in> (RAG wq)^+}"
-
-
-text {* \noindent 
-  The following
-  @{text "cpreced s th"} gives the {\em current precedence} of thread @{text "th"} under
-  state @{text "s"}. The definition of @{text "cpreced"} reflects the basic idea of 
-  Priority Inheritance that the {\em current precedence} of a thread is the precedence 
-  inherited from the maximum of all its dependants, i.e. the threads which are waiting 
-  directly or indirectly waiting for some resources from it. If no such thread exits, 
-  @{text "th"}'s {\em current precedence} equals its original precedence, i.e. 
-  @{text "preced th s"}.
-  *}
-
-definition cpreced :: "(cs \<Rightarrow> thread list) \<Rightarrow> state \<Rightarrow> thread \<Rightarrow> precedence"
-  where "cpreced wq s = (\<lambda>th. Max ((\<lambda>th'. preced th' s) ` ({th} \<union> dependants wq th)))"
-
-text {*
-  Notice that the current precedence (@{text "cpreced"}) of one thread @{text "th"} can be boosted 
-  (becoming larger than its own precedence) by those threads in 
-  the @{text "dependants wq th"}-set. If one thread get boosted, we say 
-  it inherits the priority (or, more precisely, the precedence) of 
-  its dependants. This is how the word "Inheritance" in 
-  Priority Inheritance Protocol comes.
-*}
-
-(*<*)
-lemma 
-  cpreced_def2:
-  "cpreced wq s th \<equiv> Max ({preced th s} \<union> {preced th' s | th'. th' \<in> dependants wq th})"
-  unfolding cpreced_def image_def
-  apply(rule eq_reflection)
-  apply(rule_tac f="Max" in arg_cong)
-  by (auto)
-(*>*)
-
-
-text {* \noindent
-  Assuming @{text "qs"} be the waiting queue of a critical resource, 
-  the following abbreviation "release qs" is the waiting queue after the thread 
-  holding the resource (which is thread at the head of @{text "qs"}) released
-  the resource:
-*}
-abbreviation
-  "release qs \<equiv> case qs of
-             [] => [] 
-          |  (_#qs') => (SOME q. distinct q \<and> set q = set qs')"
-text {* \noindent
-  It can be seen from the definition that the thread at the head of @{text "qs"} is removed
-  from the return value, and the value @{term "q"} is an reordering of @{text "qs'"}, the 
-  tail of @{text "qs"}. Through this reordering, one of the waiting threads (those in @{text "qs'"} }
-  is chosen nondeterministically to be the head of the new queue @{text "q"}. 
-  Therefore, this thread is the one who takes over the resource. This is a little better different 
-  from common sense that the thread who comes the earliest should take over.  
-  The intention of this definition is to show that the choice of which thread to take over the 
-  release resource does not affect the correctness of the PIP protocol. 
-*}
-
-text {*
-  The data structure used by the operating system for scheduling is referred to as 
-  {\em schedule state}. It is represented as a record consisting of 
-  a function assigning waiting queue to resources 
-  (to be used as the @{text "wq"} argument in @{text "holding"}, @{text "waiting"} 
-  and  @{text "RAG"}, etc) and a function assigning precedence to threads:
-  *}
-
-record schedule_state = 
-    wq_fun :: "cs \<Rightarrow> thread list" -- {* The function assigning waiting queue. *}
-    cprec_fun :: "thread \<Rightarrow> precedence" -- {* The function assigning precedence. *}
-
-text {* \noindent
-  The following two abbreviations (@{text "all_unlocked"} and @{text "initial_cprec"}) 
-  are used to set the initial values of the @{text "wq_fun"} @{text "cprec_fun"} fields 
-  respectively of the @{text "schedule_state"} record by the following function @{text "sch"},
-  which is used to calculate the system's {\em schedule state}.
-
-  Since there is no thread at the very beginning to make request, all critical resources 
-  are free (or unlocked). This status is represented by the abbreviation
-  @{text "all_unlocked"}. 
-  *}
-abbreviation
-  "all_unlocked \<equiv> \<lambda>_::cs. ([]::thread list)"
-
-
-text {* \noindent
-  The initial current precedence for a thread can be anything, because there is no thread then. 
-  We simply assume every thread has precedence @{text "Prc 0 0"}.
-  *}
-
-abbreviation 
-  "initial_cprec \<equiv> \<lambda>_::thread. Prc 0 0"
-
-
-text {* \noindent
-  The following function @{text "schs"} is used to calculate the system's schedule state @{text "schs s"}
-  out of the current system state @{text "s"}. It is the central function to model Priority Inheritance:
-  *}
-fun schs :: "state \<Rightarrow> schedule_state"
-  where 
-  -- {*
-  \begin{minipage}{0.9\textwidth}
-    Setting the initial value of the @{text "schedule_state"} record (see the explanations above).
-  \end{minipage}
-  *}
-  "schs [] = (| wq_fun = all_unlocked,  cprec_fun = initial_cprec |)" |
-
-  -- {*
-  \begin{minipage}{0.9\textwidth}
-  \begin{enumerate}
-  \item @{text "ps"} is the schedule state of last moment.
-  \item @{text "pwq"} is the waiting queue function of last moment.
-  \item @{text "pcp"} is the precedence function of last moment (NOT USED). 
-  \item @{text "nwq"} is the new waiting queue function. It is calculated using a @{text "case"} statement:
-  \begin{enumerate}
-      \item If the happening event is @{text "P thread cs"}, @{text "thread"} is added to 
-            the end of @{text "cs"}'s waiting queue.
-      \item If the happening event is @{text "V thread cs"} and @{text "s"} is a legal state,
-            @{text "th'"} must equal to @{text "thread"}, 
-            because @{text "thread"} is the one currently holding @{text "cs"}. 
-            The case @{text "[] \<Longrightarrow> []"} may never be executed in a legal state.
-            the @{text "(SOME q. distinct q \<and> set q = set qs)"} is used to choose arbitrarily one 
-            thread in waiting to take over the released resource @{text "cs"}. In our representation,
-            this amounts to rearrange elements in waiting queue, so that one of them is put at the head.
-      \item For other happening event, the schedule state just does not change.
-  \end{enumerate}
-  \item @{text "ncp"} is new precedence function, it is calculated from the newly updated waiting queue 
-        function. The RAGency of precedence function on waiting queue function is the reason to 
-        put them in the same record so that they can evolve together.
-  \end{enumerate}
-  
-
-  The calculation of @{text "cprec_fun"} depends on the value of @{text "wq_fun"}. 
-  Therefore, in the following cases, @{text "wq_fun"} is always calculated first, in 
-  the name of @{text "wq"} (if  @{text "wq_fun"} is not changed
-  by the happening event) or @{text "new_wq"} (if the value of @{text "wq_fun"} is changed).
-  \end{minipage}
-     *}
-   "schs (Create th prio # s) = 
-       (let wq = wq_fun (schs s) in
-          (|wq_fun = wq, cprec_fun = cpreced wq (Create th prio # s)|))"
-|  "schs (Exit th # s) = 
-       (let wq = wq_fun (schs s) in
-          (|wq_fun = wq, cprec_fun = cpreced wq (Exit th # s)|))"
-|  "schs (Set th prio # s) = 
-       (let wq = wq_fun (schs s) in
-          (|wq_fun = wq, cprec_fun = cpreced wq (Set th prio # s)|))"
-   -- {*
-   \begin{minipage}{0.9\textwidth}
-      Different from the forth coming cases, the @{text "wq_fun"} field of the schedule state 
-      is changed. So, the new value is calculated first, in the name of @{text "new_wq"}.
-   \end{minipage}   
-   *}
-|  "schs (P th cs # s) = 
-       (let wq = wq_fun (schs s) in
-        let new_wq = wq(cs := (wq cs @ [th])) in
-          (|wq_fun = new_wq, cprec_fun = cpreced new_wq (P th cs # s)|))"
-|  "schs (V th cs # s) = 
-       (let wq = wq_fun (schs s) in
-        let new_wq = wq(cs := release (wq cs)) in
-          (|wq_fun = new_wq, cprec_fun = cpreced new_wq (V th cs # s)|))"
-
-lemma cpreced_initial: 
-  "cpreced (\<lambda> cs. []) [] = (\<lambda>_. (Prc 0 0))"
-apply(simp add: cpreced_def)
-apply(simp add: cs_dependants_def cs_RAG_def cs_waiting_def cs_holding_def)
-apply(simp add: preced_def)
-done
-
-lemma sch_old_def:
-  "schs (e#s) = (let ps = schs s in 
-                  let pwq = wq_fun ps in 
-                  let nwq = case e of
-                             P th cs \<Rightarrow>  pwq(cs:=(pwq cs @ [th])) |
-                             V th cs \<Rightarrow> let nq = case (pwq cs) of
-                                                      [] \<Rightarrow> [] | 
-                                                      (_#qs) \<Rightarrow> (SOME q. distinct q \<and> set q = set qs)
-                                            in pwq(cs:=nq)                 |
-                              _ \<Rightarrow> pwq
-                  in let ncp = cpreced nwq (e#s) in 
-                     \<lparr>wq_fun = nwq, cprec_fun = ncp\<rparr>
-                 )"
-apply(cases e)
-apply(simp_all)
-done
-
-
-text {* 
-  \noindent
-  The following @{text "wq"} is a shorthand for @{text "wq_fun"}. 
-  *}
-definition wq :: "state \<Rightarrow> cs \<Rightarrow> thread list" 
-  where "wq s = wq_fun (schs s)"
-
-text {* \noindent 
-  The following @{text "cp"} is a shorthand for @{text "cprec_fun"}. 
-  *}
-definition cp :: "state \<Rightarrow> thread \<Rightarrow> precedence"
-  where "cp s \<equiv> cprec_fun (schs s)"
-
-text {* \noindent
-  Functions @{text "holding"}, @{text "waiting"}, @{text "RAG"} and 
-  @{text "dependants"} still have the 
-  same meaning, but redefined so that they no longer RAG on the 
-  fictitious {\em waiting queue function}
-  @{text "wq"}, but on system state @{text "s"}.
-  *}
-defs (overloaded) 
-  s_holding_abv: 
-  "holding (s::state) \<equiv> holding (wq_fun (schs s))"
-  s_waiting_abv: 
-  "waiting (s::state) \<equiv> waiting (wq_fun (schs s))"
-  s_RAG_abv: 
-  "RAG (s::state) \<equiv> RAG (wq_fun (schs s))"
-  s_dependants_abv: 
-  "dependants (s::state) \<equiv> dependants (wq_fun (schs s))"
-
-
-text {* 
-  The following lemma can be proved easily, and the meaning is obvious.
-  *}
-lemma
-  s_holding_def: 
-  "holding (s::state) th cs \<equiv> (th \<in> set (wq_fun (schs s) cs) \<and> th = hd (wq_fun (schs s) cs))"
-  by (auto simp:s_holding_abv wq_def cs_holding_def)
-
-lemma s_waiting_def: 
-  "waiting (s::state) th cs \<equiv> (th \<in> set (wq_fun (schs s) cs) \<and> th \<noteq> hd (wq_fun (schs s) cs))"
-  by (auto simp:s_waiting_abv wq_def cs_waiting_def)
-
-lemma s_RAG_def: 
-  "RAG (s::state) =
-    {(Th th, Cs cs) | th cs. waiting (wq s) th cs} \<union> {(Cs cs, Th th) | cs th. holding (wq s) th cs}"
-  by (auto simp:s_RAG_abv wq_def cs_RAG_def)
-
-lemma
-  s_dependants_def: 
-  "dependants (s::state) th \<equiv> {th' . (Th th', Th th) \<in> (RAG (wq s))^+}"
-  by (auto simp:s_dependants_abv wq_def cs_dependants_def)
-
-text {*
-  The following function @{text "readys"} calculates the set of ready threads. A thread is {\em ready} 
-  for running if it is a live thread and it is not waiting for any critical resource.
-  *}
-definition readys :: "state \<Rightarrow> thread set"
-  where "readys s \<equiv> {th . th \<in> threads s \<and> (\<forall> cs. \<not> waiting s th cs)}"
-
-text {* \noindent
-  The following function @{text "runing"} calculates the set of running thread, which is the ready 
-  thread with the highest precedence.  
-  *}
-definition runing :: "state \<Rightarrow> thread set"
-  where "runing s \<equiv> {th . th \<in> readys s \<and> cp s th = Max ((cp s) ` (readys s))}"
-
-text {* \noindent
-  Notice that the definition of @{text "running"} reflects the preemptive scheduling strategy,  
-  because, if the @{text "running"}-thread (the one in @{text "runing"} set) 
-  lowered its precedence by resetting its own priority to a lower
-  one, it will lose its status of being the max in @{text "ready"}-set and be superseded.
-*}
-
-text {* \noindent
-  The following function @{text "holdents s th"} returns the set of resources held by thread 
-  @{text "th"} in state @{text "s"}.
-  *}
-definition holdents :: "state \<Rightarrow> thread \<Rightarrow> cs set"
-  where "holdents s th \<equiv> {cs . holding s th cs}"
-
-lemma holdents_test: 
-  "holdents s th = {cs . (Cs cs, Th th) \<in> RAG s}"
-unfolding holdents_def
-unfolding s_RAG_def
-unfolding s_holding_abv
-unfolding wq_def
-by (simp)
-
-text {* \noindent
-  Observation @{text "cntCS s th"} returns the number of resources held by thread @{text "th"} in
-  state @{text "s"}:
-  *}
-definition cntCS :: "state \<Rightarrow> thread \<Rightarrow> nat"
-  where "cntCS s th = card (holdents s th)"
-
-text {* \noindent
-  According to the convention of Paulson's inductive method,
-  the decision made by a protocol that event @{text "e"} is eligible to happen next under state @{text "s"} 
-  is expressed as @{text "step s e"}. The predicate @{text "step"} is inductively defined as 
-  follows (notice how the decision is based on the {\em observation function}s 
-  defined above, and also notice how a complicated protocol is modeled by a few simple 
-  observations, and how such a kind of simplicity gives rise to improved trust on
-  faithfulness):
-  *}
-inductive step :: "state \<Rightarrow> event \<Rightarrow> bool"
-  where
-  -- {* 
-  A thread can be created if it is not a live thread:
-  *}
-  thread_create: "\<lbrakk>thread \<notin> threads s\<rbrakk> \<Longrightarrow> step s (Create thread prio)" |
-  -- {*
-  A thread can exit if it no longer hold any resource:
-  *}
-  thread_exit: "\<lbrakk>thread \<in> runing s; holdents s thread = {}\<rbrakk> \<Longrightarrow> step s (Exit thread)" |
-  -- {*
-  \begin{minipage}{0.9\textwidth}
-  A thread can request for an critical resource @{text "cs"}, if it is running and 
-  the request does not form a loop in the current RAG. The latter condition 
-  is set up to avoid deadlock. The condition also reflects our assumption all threads are 
-  carefully programmed so that deadlock can not happen:
-  \end{minipage}
-  *}
-  thread_P: "\<lbrakk>thread \<in> runing s;  (Cs cs, Th thread)  \<notin> (RAG s)^+\<rbrakk> \<Longrightarrow> 
-                                                                step s (P thread cs)" |
-  -- {*
-  \begin{minipage}{0.9\textwidth}
-  A thread can release a critical resource @{text "cs"} 
-  if it is running and holding that resource:
-  \end{minipage}
-  *}
-  thread_V: "\<lbrakk>thread \<in> runing s;  holding s thread cs\<rbrakk> \<Longrightarrow> step s (V thread cs)" |
-  -- {*
-  \begin{minipage}{0.9\textwidth}
-  A thread can adjust its own priority as long as it is current running. 
-  With the resetting of one thread's priority, its precedence may change. 
-  If this change lowered the precedence, according to the definition of @{text "running"}
-  function, 
-  \end{minipage}
-  *}  
-  thread_set: "\<lbrakk>thread \<in> runing s\<rbrakk> \<Longrightarrow> step s (Set thread prio)"
-
-text {*
-  In Paulson's inductive method, every protocol is defined by such a @{text "step"}
-  predicate. For instance, the predicate @{text "step"} given above 
-  defines the PIP protocol. So, it can also be called "PIP".
-*}
-
-abbreviation
-  "PIP \<equiv> step"
-
-
-text {* \noindent
-  For any protocol defined by a @{text "step"} predicate, 
-  the fact that @{text "s"} is a legal state in 
-  the protocol is expressed as: @{text "vt step s"}, where
-  the predicate @{text "vt"} can be defined as the following:
-  *}
-inductive vt :: "state \<Rightarrow> bool"
-  where
-  -- {* Empty list @{text "[]"} is a legal state in any protocol:*}
-  vt_nil[intro]: "vt []" |
-  -- {* 
-  \begin{minipage}{0.9\textwidth}
-  If @{text "s"} a legal state of the protocol defined by predicate @{text "step"}, 
-  and event @{text "e"} is allowed to happen under state @{text "s"} by the protocol 
-  predicate @{text "step"}, then @{text "e#s"} is a new legal state rendered by the 
-  happening of @{text "e"}:
-  \end{minipage}
-  *}
-  vt_cons[intro]: "\<lbrakk>vt s; step s e\<rbrakk> \<Longrightarrow> vt (e#s)"
-
-text {*  \noindent
-  It is easy to see that the definition of @{text "vt"} is generic. It can be applied to 
-  any specific protocol specified by a @{text "step"}-predicate to get the set of
-  legal states of that particular protocol.
-  *}
-
-text {* 
-  The following are two very basic properties of @{text "vt"}.
-*}
-
-lemma step_back_vt: "vt (e#s) \<Longrightarrow> vt s"
-  by(ind_cases "vt (e#s)", simp)
-
-lemma step_back_step: "vt (e#s) \<Longrightarrow> step s e"
-  by(ind_cases "vt (e#s)", simp)
-
-text {* \noindent
-  The following two auxiliary functions @{text "the_cs"} and @{text "the_th"} are used to extract
-  critical resource and thread respectively out of RAG nodes.
-  *}
-fun the_cs :: "node \<Rightarrow> cs"
-  where "the_cs (Cs cs) = cs"
-
-fun the_th :: "node \<Rightarrow> thread"
-  where "the_th (Th th) = th"
-
-text {* \noindent
-  The following predicate @{text "next_th"} describe the next thread to 
-  take over when a critical resource is released. In @{text "next_th s th cs t"}, 
-  @{text "th"} is the thread to release, @{text "t"} is the one to take over.
-  Notice how this definition is backed up by the @{text "release"} function and its use 
-  in the @{text "V"}-branch of @{text "schs"} function. This @{text "next_th"} function
-  is not needed for the execution of PIP. It is introduced as an auxiliary function 
-  to state lemmas. The correctness of this definition will be confirmed by 
-  lemmas @{text "step_v_hold_inv"}, @{text " step_v_wait_inv"}, 
-  @{text "step_v_get_hold"} and @{text "step_v_not_wait"}.
-  *}
-definition next_th:: "state \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> thread \<Rightarrow> bool"
-  where "next_th s th cs t = (\<exists> rest. wq s cs = th#rest \<and> rest \<noteq> [] \<and> 
-                                     t = hd (SOME q. distinct q \<and> set q = set rest))"
-
-text {* \noindent
-  The aux function @{text "count Q l"} is used to count the occurrence of situation @{text "Q"}
-  in list @{text "l"}:
-  *}
-definition count :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> nat"
-  where "count Q l = length (filter Q l)"
-
-text {* \noindent
-  The following observation @{text "cntP s"} returns the number of operation @{text "P"} happened 
-  before reaching state @{text "s"}.
-  *}
-definition cntP :: "state \<Rightarrow> thread \<Rightarrow> nat"
-  where "cntP s th = count (\<lambda> e. \<exists> cs. e = P th cs) s"
-
-text {* \noindent
-  The following observation @{text "cntV s"} returns the number of operation @{text "V"} happened 
-  before reaching state @{text "s"}.
-  *}
-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}"
-
-text {* The following @{text "hRAG"} is the holding sub-graph of @{text "RAG"}. *}
-definition "hRAG (s::state) =  {(Cs cs, Th th) | th cs. holding s th cs}"
-
-text {* 
-  The following @{text "tRAG"} is the thread-graph derived from @{term "RAG"}.
-  It characterizes the dependency between threads when calculating current
-  precedences. It is defined as the composition of the above two sub-graphs, 
-  names @{term "wRAG"} and @{term "hRAG"}.
- *}
-definition "tRAG s = wRAG s O hRAG s"
-
-text {* The following lemma splits @{term "RAG"} graph into the above two sub-graphs. *}
-lemma RAG_split: "RAG s = (wRAG s \<union> hRAG s)"
-  by (unfold s_RAG_abv wRAG_def hRAG_def s_waiting_abv 
-             s_holding_abv cs_RAG_def, auto)
-
-definition "cp_gen s x =
-                  Max ((the_preced s \<circ> the_thread) ` subtree (tRAG s) x)"
-
-(*<*)
-
-end
-(*>*)
-
--- a/Precedence_ord.thy~	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,45 +0,0 @@
-header {* Order on product types *}
-
-theory Precedence_ord
-imports Main
-begin
-
-datatype precedence = Prc nat nat
-
-instantiation precedence :: order
-begin
-
-definition
-  precedence_le_def: "x \<le> y \<longleftrightarrow> (case (x, y) of
-                                   (Prc fx sx, Prc fy sy) \<Rightarrow> 
-                                 fx < fy \<or> (fx \<le> fy \<and> sy \<le> sx))"
-
-definition
-  precedence_less_def: "x < y \<longleftrightarrow> (case (x, y) of
-                               (Prc fx sx, Prc fy sy) \<Rightarrow> 
-                                     fx < fy \<or> (fx \<le> fy \<and> sy < sx))"
-
-instance
-proof
-qed (auto simp: precedence_le_def precedence_less_def 
-              intro: order_trans split:precedence.splits)
-end
-
-instance precedence :: preorder ..
-
-instance precedence :: linorder 
-proof
-qed (auto simp: precedence_le_def precedence_less_def 
-              intro: order_trans split:precedence.splits)
-
-instantiation precedence :: zero
-begin
-
-definition Zero_precedence_def:
-  "0 = Prc 0 0"
-
-instance ..
-
-end
-
-end
--- a/PrioG.thy~	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1611 +0,0 @@
-theory PrioG
-imports CpsG
-begin
-
-
-text {* 
-  The following two auxiliary lemmas are used to reason about @{term Max}.
-*}
-lemma image_Max_eqI: 
-  assumes "finite B"
-  and "b \<in> B"
-  and "\<forall> x \<in> B. f x \<le> f b"
-  shows "Max (f ` B) = f b"
-  using assms
-  using Max_eqI by blast 
-
-lemma image_Max_subset:
-  assumes "finite A"
-  and "B \<subseteq> A"
-  and "a \<in> B"
-  and "Max (f ` A) = f a"
-  shows "Max (f ` B) = f a"
-proof(rule image_Max_eqI)
-  show "finite B"
-    using assms(1) assms(2) finite_subset by auto 
-next
-  show "a \<in> B" using assms by simp
-next
-  show "\<forall>x\<in>B. f x \<le> f a"
-    by (metis Max_ge assms(1) assms(2) assms(4) 
-            finite_imageI image_eqI subsetCE) 
-qed
-
-text {*
-  The following locale @{text "highest_gen"} sets the basic context for our
-  investigation: supposing thread @{text th} holds the highest @{term cp}-value
-  in state @{text s}, which means the task for @{text th} is the 
-  most urgent. We want to show that  
-  @{text th} is treated correctly by PIP, which means
-  @{text th} will not be blocked unreasonably by other less urgent
-  threads. 
-*}
-locale highest_gen =
-  fixes s th prio tm
-  assumes vt_s: "vt s"
-  and threads_s: "th \<in> threads s"
-  and highest: "preced th s = Max ((cp s)`threads s)"
-  -- {* The internal structure of @{term th}'s precedence is exposed:*}
-  and preced_th: "preced th s = Prc prio tm" 
-
--- {* @{term s} is a valid trace, so it will inherit all results derived for
-      a valid trace: *}
-sublocale highest_gen < vat_s: valid_trace "s"
-  by (unfold_locales, insert vt_s, simp)
-
-context highest_gen
-begin
-
-text {*
-  @{term tm} is the time when the precedence of @{term th} is set, so 
-  @{term tm} must be a valid moment index into @{term s}.
-*}
-lemma lt_tm: "tm < length s"
-  by (insert preced_tm_lt[OF threads_s preced_th], simp)
-
-text {*
-  Since @{term th} holds the highest precedence and @{text "cp"}
-  is the highest precedence of all threads in the sub-tree of 
-  @{text "th"} and @{text th} is among these threads, 
-  its @{term cp} must equal to its precedence:
-*}
-lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R")
-proof -
-  have "?L \<le> ?R"
-  by (unfold highest, rule Max_ge, 
-        auto simp:threads_s finite_threads)
-  moreover have "?R \<le> ?L"
-    by (unfold vat_s.cp_rec, rule Max_ge, 
-        auto simp:the_preced_def vat_s.fsbttRAGs.finite_children)
-  ultimately show ?thesis by auto
-qed
-
-lemma highest_cp_preced: "cp s th = Max (the_preced s ` threads s)"
-  using eq_cp_s_th highest max_cp_eq the_preced_def by presburger
-  
-
-lemma highest_preced_thread: "preced th s = Max (the_preced s ` threads s)"
-  by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
-
-lemma highest': "cp s th = Max (cp s ` threads s)"
-  by (simp add: eq_cp_s_th highest)
-
-end
-
-locale extend_highest_gen = highest_gen + 
-  fixes t 
-  assumes vt_t: "vt (t@s)"
-  and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
-  and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
-  and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"
-
-sublocale extend_highest_gen < vat_t: valid_trace "t@s"
-  by (unfold_locales, insert vt_t, simp)
-
-lemma step_back_vt_app: 
-  assumes vt_ts: "vt (t@s)" 
-  shows "vt s"
-proof -
-  from vt_ts show ?thesis
-  proof(induct t)
-    case Nil
-    from Nil show ?case by auto
-  next
-    case (Cons e t)
-    assume ih: " vt (t @ s) \<Longrightarrow> vt s"
-      and vt_et: "vt ((e # t) @ s)"
-    show ?case
-    proof(rule ih)
-      show "vt (t @ s)"
-      proof(rule step_back_vt)
-        from vt_et show "vt (e # t @ s)" by simp
-      qed
-    qed
-  qed
-qed
-
-(* locale red_extend_highest_gen = extend_highest_gen +
-   fixes i::nat
-*)
-
-(*
-sublocale red_extend_highest_gen <   red_moment: extend_highest_gen "s" "th" "prio" "tm" "(moment i t)"
-  apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
-  apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
-  by (unfold highest_gen_def, auto dest:step_back_vt_app)
-*)
-
-context extend_highest_gen
-begin
-
- lemma ind [consumes 0, case_names Nil Cons, induct type]:
-  assumes 
-    h0: "R []"
-  and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e; 
-                    extend_highest_gen s th prio tm t; 
-                    extend_highest_gen s th prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
-  shows "R t"
-proof -
-  from vt_t extend_highest_gen_axioms show ?thesis
-  proof(induct t)
-    from h0 show "R []" .
-  next
-    case (Cons e t')
-    assume ih: "\<lbrakk>vt (t' @ s); extend_highest_gen s th prio tm t'\<rbrakk> \<Longrightarrow> R t'"
-      and vt_e: "vt ((e # t') @ s)"
-      and et: "extend_highest_gen s th prio tm (e # t')"
-    from vt_e and step_back_step have stp: "step (t'@s) e" by auto
-    from vt_e and step_back_vt have vt_ts: "vt (t'@s)" by auto
-    show ?case
-    proof(rule h2 [OF vt_ts stp _ _ _ ])
-      show "R t'"
-      proof(rule ih)
-        from et show ext': "extend_highest_gen s th prio tm t'"
-          by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
-      next
-        from vt_ts show "vt (t' @ s)" .
-      qed
-    next
-      from et show "extend_highest_gen s th prio tm (e # t')" .
-    next
-      from et show ext': "extend_highest_gen s th prio tm t'"
-          by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
-    qed
-  qed
-qed
-
-
-lemma th_kept: "th \<in> threads (t @ s) \<and> 
-                 preced th (t@s) = preced th s" (is "?Q t") 
-proof -
-  show ?thesis
-  proof(induct rule:ind)
-    case Nil
-    from threads_s
-    show ?case
-      by auto
-  next
-    case (Cons e t)
-    interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
-    interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
-    show ?case
-    proof(cases e)
-      case (Create thread prio)
-      show ?thesis
-      proof -
-        from Cons and Create have "step (t@s) (Create thread prio)" by auto
-        hence "th \<noteq> thread"
-        proof(cases)
-          case thread_create
-          with Cons show ?thesis by auto
-        qed
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold Create, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:Create)
-      qed
-    next
-      case (Exit thread)
-      from h_e.exit_diff and Exit
-      have neq_th: "thread \<noteq> th" by auto
-      with Cons
-      show ?thesis
-        by (unfold Exit, auto simp:preced_def)
-    next
-      case (P thread cs)
-      with Cons
-      show ?thesis 
-        by (auto simp:P preced_def)
-    next
-      case (V thread cs)
-      with Cons
-      show ?thesis 
-        by (auto simp:V preced_def)
-    next
-      case (Set thread prio')
-      show ?thesis
-      proof -
-        from h_e.set_diff_low and Set
-        have "th \<noteq> thread" by auto
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold Set, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:Set)
-      qed
-    qed
-  qed
-qed
-
-text {*
-  According to @{thm th_kept}, thread @{text "th"} has its living status
-  and precedence kept along the way of @{text "t"}. The following lemma
-  shows that this preserved precedence of @{text "th"} remains as the highest
-  along the way of @{text "t"}.
-
-  The proof goes by induction over @{text "t"} using the specialized
-  induction rule @{thm ind}, followed by case analysis of each possible 
-  operations of PIP. All cases follow the same pattern rendered by the 
-  generalized introduction rule @{thm "image_Max_eqI"}. 
-
-  The very essence is to show that precedences, no matter whether they 
-  are newly introduced or modified, are always lower than the one held 
-  by @{term "th"}, which by @{thm th_kept} is preserved along the way.
-*}
-lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s"
-proof(induct rule:ind)
-  case Nil
-  from highest_preced_thread
-  show ?case by simp
-next
-  case (Cons e t)
-    interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
-    interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
-  show ?case
-  proof(cases e)
-    case (Create thread prio')
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      -- {* The following is the common pattern of each branch of the case analysis. *}
-      -- {* The major part is to show that @{text "th"} holds the highest precedence: *}
-      have "Max (?f ` ?A) = ?f th"
-      proof(rule image_Max_eqI)
-        show "finite ?A" using h_e.finite_threads by auto 
-      next
-        show "th \<in> ?A" using h_e.th_kept by auto 
-      next
-        show "\<forall>x\<in>?A. ?f x \<le> ?f th"
-        proof 
-          fix x
-          assume "x \<in> ?A"
-          hence "x = thread \<or> x \<in> threads (t@s)" by (auto simp:Create)
-          thus "?f x \<le> ?f th"
-          proof
-            assume "x = thread"
-            thus ?thesis 
-              apply (simp add:Create the_preced_def preced_def, fold preced_def)
-              using Create h_e.create_low h_t.th_kept lt_tm preced_leI2 
-              preced_th by force
-          next
-            assume h: "x \<in> threads (t @ s)"
-            from Cons(2)[unfolded Create] 
-            have "x \<noteq> thread" using h by (cases, auto)
-            hence "?f x = the_preced (t@s) x" 
-              by (simp add:Create the_preced_def preced_def)
-            hence "?f x \<le> Max (the_preced (t@s) ` threads (t@s))"
-              by (simp add: h_t.finite_threads h)
-            also have "... = ?f th"
-              by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
-            finally show ?thesis .
-          qed
-        qed
-      qed
-     -- {* The minor part is to show that the precedence of @{text "th"} 
-           equals to preserved one, given by the foregoing lemma @{thm th_kept} *}
-      also have "... = ?t" using h_e.th_kept the_preced_def by auto
-      -- {* Then it follows trivially that the precedence preserved
-            for @{term "th"} remains the maximum of all living threads along the way. *}
-      finally show ?thesis .
-    qed 
-  next 
-    case (Exit thread)
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "Max (?f ` ?A) = ?f th"
-      proof(rule image_Max_eqI)
-        show "finite ?A" using h_e.finite_threads by auto 
-      next
-        show "th \<in> ?A" using h_e.th_kept by auto 
-      next
-        show "\<forall>x\<in>?A. ?f x \<le> ?f th"
-        proof 
-          fix x
-          assume "x \<in> ?A"
-          hence "x \<in> threads (t@s)" by (simp add: Exit) 
-          hence "?f x \<le> Max (?f ` threads (t@s))" 
-            by (simp add: h_t.finite_threads) 
-          also have "... \<le> ?f th" 
-            apply (simp add:Exit the_preced_def preced_def, fold preced_def)
-            using Cons.hyps(5) h_t.th_kept the_preced_def by auto
-          finally show "?f x \<le> ?f th" .
-        qed
-      qed
-      also have "... = ?t" using h_e.th_kept the_preced_def by auto
-      finally show ?thesis .
-    qed 
-  next
-    case (P thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def the_preced_def)
-  next
-    case (V thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def the_preced_def)
-  next 
-    case (Set thread prio')
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "Max (?f ` ?A) = ?f th"
-      proof(rule image_Max_eqI)
-        show "finite ?A" using h_e.finite_threads by auto 
-      next
-        show "th \<in> ?A" using h_e.th_kept by auto 
-      next
-        show "\<forall>x\<in>?A. ?f x \<le> ?f th"
-        proof 
-          fix x
-          assume h: "x \<in> ?A"
-          show "?f x \<le> ?f th"
-          proof(cases "x = thread")
-            case True
-            moreover have "the_preced (Set thread prio' # t @ s) thread \<le> the_preced (t @ s) th"
-            proof -
-              have "the_preced (t @ s) th = Prc prio tm"  
-                using h_t.th_kept preced_th by (simp add:the_preced_def)
-              moreover have "prio' \<le> prio" using Set h_e.set_diff_low by auto
-              ultimately show ?thesis by (insert lt_tm, auto simp:the_preced_def preced_def)
-            qed
-            ultimately show ?thesis
-              by (unfold Set, simp add:the_preced_def preced_def)
-          next
-            case False
-            then have "?f x  = the_preced (t@s) x"
-              by (simp add:the_preced_def preced_def Set)
-            also have "... \<le> Max (the_preced (t@s) ` threads (t@s))"
-              using Set h h_t.finite_threads by auto 
-            also have "... = ?f th" by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
-            finally show ?thesis .
-          qed
-        qed
-      qed
-      also have "... = ?t" using h_e.th_kept the_preced_def by auto
-      finally show ?thesis .
-    qed 
-  qed
-qed
-
-lemma max_preced: "preced th (t@s) = Max (the_preced (t@s) ` (threads (t@s)))"
-  by (insert th_kept max_kept, auto)
-
-text {*
-  The reason behind the following lemma is that:
-  Since @{term "cp"} is defined as the maximum precedence 
-  of those threads contained in the sub-tree of node @{term "Th th"} 
-  in @{term "RAG (t@s)"}, and all these threads are living threads, and 
-  @{term "th"} is also among them, the maximum precedence of 
-  them all must be the one for @{text "th"}.
-*}
-lemma th_cp_max_preced: 
-  "cp (t@s) th = Max (the_preced (t@s) ` (threads (t@s)))" (is "?L = ?R") 
-proof -
-  let ?f = "the_preced (t@s)"
-  have "?L = ?f th"
-  proof(unfold cp_alt_def, rule image_Max_eqI)
-    show "finite {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
-    proof -
-      have "{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)} = 
-            the_thread ` {n . n \<in> subtree (RAG (t @ s)) (Th th) \<and>
-                            (\<exists> th'. n = Th th')}"
-      by (smt Collect_cong Setcompr_eq_image mem_Collect_eq the_thread.simps)
-      moreover have "finite ..." by (simp add: vat_t.fsbtRAGs.finite_subtree) 
-      ultimately show ?thesis by simp
-    qed
-  next
-    show "th \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
-      by (auto simp:subtree_def)
-  next
-    show "\<forall>x\<in>{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}.
-               the_preced (t @ s) x \<le> the_preced (t @ s) th"
-    proof
-      fix th'
-      assume "th' \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
-      hence "Th th' \<in> subtree (RAG (t @ s)) (Th th)" by auto
-      moreover have "... \<subseteq> Field (RAG (t @ s)) \<union> {Th th}"
-        by (meson subtree_Field)
-      ultimately have "Th th' \<in> ..." by auto
-      hence "th' \<in> threads (t@s)" 
-      proof
-        assume "Th th' \<in> {Th th}"
-        thus ?thesis using th_kept by auto 
-      next
-        assume "Th th' \<in> Field (RAG (t @ s))"
-        thus ?thesis using vat_t.not_in_thread_isolated by blast 
-      qed
-      thus "the_preced (t @ s) th' \<le> the_preced (t @ s) th"
-        by (metis Max_ge finite_imageI finite_threads image_eqI 
-               max_kept th_kept the_preced_def)
-    qed
-  qed
-  also have "... = ?R" by (simp add: max_preced the_preced_def) 
-  finally show ?thesis .
-qed
-
-lemma th_cp_max[simp]: "Max (cp (t@s) ` threads (t@s)) = cp (t@s) th"
-  using max_cp_eq th_cp_max_preced the_preced_def vt_t by presburger
-
-lemma [simp]: "Max (cp (t@s) ` threads (t@s)) = Max (the_preced (t@s) ` threads (t@s))"
-  by (simp add: th_cp_max_preced)
-  
-lemma [simp]: "Max (the_preced (t@s) ` threads (t@s)) = the_preced (t@s) th"
-  using max_kept th_kept the_preced_def by auto
-
-lemma [simp]: "the_preced (t@s) th = preced th (t@s)"
-  using the_preced_def by auto
-
-lemma [simp]: "preced th (t@s) = preced th s"
-  by (simp add: th_kept)
-
-lemma [simp]: "cp s th = preced th s"
-  by (simp add: eq_cp_s_th)
-
-lemma th_cp_preced [simp]: "cp (t@s) th = preced th s"
-  by (fold max_kept, unfold th_cp_max_preced, simp)
-
-lemma preced_less:
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  shows "preced th' s < preced th s"
-  using assms
-by (metis Max.coboundedI finite_imageI highest not_le order.trans 
-    preced_linorder rev_image_eqI threads_s vat_s.finite_threads 
-    vat_s.le_cp)
-
-section {* The `blocking thread` *}
-
-text {* 
-  The purpose of PIP is to ensure that the most 
-  urgent thread @{term th} is not blocked unreasonably. 
-  Therefore, a clear picture of the blocking thread is essential 
-  to assure people that the purpose is fulfilled. 
-  
-  In this section, we are going to derive a series of lemmas 
-  with finally give rise to a picture of the blocking thread. 
-
-  By `blocking thread`, we mean a thread in running state but 
-  different from thread @{term th}.
-*}
-
-text {*
-  The following lemmas shows that the @{term cp}-value 
-  of the blocking thread @{text th'} equals to the highest
-  precedence in the whole system.
-*}
-lemma runing_preced_inversion:
-  assumes runing': "th' \<in> runing (t@s)"
-  shows "cp (t@s) th' = preced th s" (is "?L = ?R")
-proof -
-  have "?L = Max (cp (t @ s) ` readys (t @ s))" using assms
-      by (unfold runing_def, auto)
-  also have "\<dots> = ?R"
-      by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads) 
-  finally show ?thesis .
-qed
-
-text {*
-
-  The following lemma shows how the counters for @{term "P"} and
-  @{term "V"} operations relate to the running threads in the states
-  @{term s} and @{term "t @ s"}.  The lemma shows that if a thread's
-  @{term "P"}-count equals its @{term "V"}-count (which means it no
-  longer has any resource in its possession), it cannot be a running
-  thread.
-
-  The proof is by contraction with the assumption @{text "th' \<noteq> th"}.
-  The key is the use of @{thm count_eq_dependants} to derive the
-  emptiness of @{text th'}s @{term dependants}-set from the balance of
-  its @{term P} and @{term V} counts.  From this, it can be shown
-  @{text th'}s @{term cp}-value equals to its own precedence.
-
-  On the other hand, since @{text th'} is running, by @{thm
-  runing_preced_inversion}, its @{term cp}-value equals to the
-  precedence of @{term th}.
-
-  Combining the above two resukts we have that @{text th'} and @{term
-  th} have the same precedence. By uniqueness of precedences, we have
-  @{text "th' = th"}, which is in contradiction with the assumption
-  @{text "th' \<noteq> th"}.
-
-*} 
-                      
-lemma eq_pv_blocked: (* ddd *)
-  assumes neq_th': "th' \<noteq> th"
-  and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
-  shows "th' \<notin> runing (t@s)"
-proof
-  assume otherwise: "th' \<in> runing (t@s)"
-  show False
-  proof -
-    have th'_in: "th' \<in> threads (t@s)"
-        using otherwise readys_threads runing_def by auto 
-    have "th' = th"
-    proof(rule preced_unique)
-      -- {* The proof goes like this: 
-            it is first shown that the @{term preced}-value of @{term th'} 
-            equals to that of @{term th}, then by uniqueness 
-            of @{term preced}-values (given by lemma @{thm preced_unique}), 
-            @{term th'} equals to @{term th}: *}
-      show "preced th' (t @ s) = preced th (t @ s)" (is "?L = ?R")
-      proof -
-        -- {* Since the counts of @{term th'} are balanced, the subtree
-              of it contains only itself, so, its @{term cp}-value
-              equals its @{term preced}-value: *}
-        have "?L = cp (t@s) th'"
-          by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp)
-        -- {* Since @{term "th'"} is running, by @{thm runing_preced_inversion},
-              its @{term cp}-value equals @{term "preced th s"}, 
-              which equals to @{term "?R"} by simplification: *}
-        also have "... = ?R" 
-        thm runing_preced_inversion
-            using runing_preced_inversion[OF otherwise] by simp
-        finally show ?thesis .
-      qed
-    qed (auto simp: th'_in th_kept)
-    with `th' \<noteq> th` show ?thesis by simp
- qed
-qed
-
-text {*
-  The following lemma is the extrapolation of @{thm eq_pv_blocked}.
-  It says if a thread, different from @{term th}, 
-  does not hold any resource at the very beginning,
-  it will keep hand-emptied in the future @{term "t@s"}.
-*}
-lemma eq_pv_persist: (* ddd *)
-  assumes neq_th': "th' \<noteq> th"
-  and eq_pv: "cntP s th' = cntV s th'"
-  shows "cntP (t@s) th' = cntV (t@s) th'"
-proof(induction rule:ind) -- {* The proof goes by induction. *}
-  -- {* The nontrivial case is for the @{term Cons}: *}
-  case (Cons e t)
-  -- {* All results derived so far hold for both @{term s} and @{term "t@s"}: *}
-  interpret vat_t: extend_highest_gen s th prio tm t using Cons by simp
-  interpret vat_e: extend_highest_gen s th prio tm "(e # t)" using Cons by simp
-  show ?case
-  proof -
-    -- {* It can be proved that @{term cntP}-value of @{term th'} does not change
-          by the happening of event @{term e}: *}
-    have "cntP ((e#t)@s) th' = cntP (t@s) th'"
-    proof(rule ccontr) -- {* Proof by contradiction. *}
-      -- {* Suppose @{term cntP}-value of @{term th'} is changed by @{term e}: *}
-      assume otherwise: "cntP ((e # t) @ s) th' \<noteq> cntP (t @ s) th'"
-      -- {* Then the actor of @{term e} must be @{term th'} and @{term e}
-            must be a @{term P}-event: *}
-      hence "isP e" "actor e = th'" by (auto simp:cntP_diff_inv) 
-      with vat_t.actor_inv[OF Cons(2)]
-      -- {* According to @{thm actor_inv}, @{term th'} must be running at 
-            the moment @{term "t@s"}: *}
-      have "th' \<in> runing (t@s)" by (cases e, auto)
-      -- {* However, an application of @{thm eq_pv_blocked} to induction hypothesis
-            shows @{term th'} can not be running at moment  @{term "t@s"}: *}
-      moreover have "th' \<notin> runing (t@s)" 
-               using vat_t.eq_pv_blocked[OF neq_th' Cons(5)] .
-      -- {* Contradiction is finally derived: *}
-      ultimately show False by simp
-    qed
-    -- {* It can also be proved that @{term cntV}-value of @{term th'} does not change
-          by the happening of event @{term e}: *}
-    -- {* The proof follows exactly the same pattern as the case for @{term cntP}-value: *}
-    moreover have "cntV ((e#t)@s) th' = cntV (t@s) th'"
-    proof(rule ccontr) -- {* Proof by contradiction. *}
-      assume otherwise: "cntV ((e # t) @ s) th' \<noteq> cntV (t @ s) th'"
-      hence "isV e" "actor e = th'" by (auto simp:cntV_diff_inv) 
-      with vat_t.actor_inv[OF Cons(2)]
-      have "th' \<in> runing (t@s)" by (cases e, auto)
-      moreover have "th' \<notin> runing (t@s)"
-          using vat_t.eq_pv_blocked[OF neq_th' Cons(5)] .
-      ultimately show False by simp
-    qed
-    -- {* Finally, it can be shown that the @{term cntP} and @{term cntV} 
-          value for @{term th'} are still in balance, so @{term th'} 
-          is still hand-emptied after the execution of event @{term e}: *}
-    ultimately show ?thesis using Cons(5) by metis
-  qed
-qed (auto simp:eq_pv)
-
-text {*
-  By combining @{thm  eq_pv_blocked} and @{thm eq_pv_persist},
-  it can be derived easily that @{term th'} can not be running in the future:
-*}
-lemma eq_pv_blocked_persist:
-  assumes neq_th': "th' \<noteq> th"
-  and eq_pv: "cntP s th' = cntV s th'"
-  shows "th' \<notin> runing (t@s)"
-  using assms
-  by (simp add: eq_pv_blocked eq_pv_persist) 
-
-text {*
-  The following lemma shows the blocking thread @{term th'}
-  must hold some resource in the very beginning. 
-*}
-lemma runing_cntP_cntV_inv: (* ddd *)
-  assumes is_runing: "th' \<in> runing (t@s)"
-  and neq_th': "th' \<noteq> th"
-  shows "cntP s th' > cntV s th'"
-  using assms
-proof -
-  -- {* First, it can be shown that the number of @{term P} and
-        @{term V} operations can not be equal for thred @{term th'} *}
-  have "cntP s th' \<noteq> cntV s th'"
-  proof
-     -- {* The proof goes by contradiction, suppose otherwise: *}
-    assume otherwise: "cntP s th' = cntV s th'"
-    -- {* By applying @{thm  eq_pv_blocked_persist} to this: *}
-    from eq_pv_blocked_persist[OF neq_th' otherwise] 
-    -- {* we have that @{term th'} can not be running at moment @{term "t@s"}: *}
-    have "th' \<notin> runing (t@s)" .
-    -- {* This is obvious in contradiction with assumption @{thm is_runing}  *}
-    thus False using is_runing by simp
-  qed
-  -- {* However, the number of @{term V} is always less or equal to @{term P}: *}
-  moreover have "cntV s th' \<le> cntP s th'" using vat_s.cnp_cnv_cncs by auto
-  -- {* Thesis is finally derived by combining the these two results: *}
-  ultimately show ?thesis by auto
-qed
-
-
-text {*
-  The following lemmas shows the blocking thread @{text th'} must be live 
-  at the very beginning, i.e. the moment (or state) @{term s}. 
-
-  The proof is a  simple combination of the results above:
-*}
-lemma runing_threads_inv: 
-  assumes runing': "th' \<in> runing (t@s)"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<in> threads s"
-proof(rule ccontr) -- {* Proof by contradiction: *}
-  assume otherwise: "th' \<notin> threads s" 
-  have "th' \<notin> runing (t @ s)"
-  proof -
-    from vat_s.cnp_cnv_eq[OF otherwise]
-    have "cntP s th' = cntV s th'" .
-    from eq_pv_blocked_persist[OF neq_th' this]
-    show ?thesis .
-  qed
-  with runing' show False by simp
-qed
-
-text {*
-  The following lemma summarizes several foregoing 
-  lemmas to give an overall picture of the blocking thread @{text "th'"}:
-*}
-lemma runing_inversion: (* ddd, one of the main lemmas to present *)
-  assumes runing': "th' \<in> runing (t@s)"
-  and neq_th: "th' \<noteq> th"
-  shows "th' \<in> threads s"
-  and    "\<not>detached s th'"
-  and    "cp (t@s) th' = preced th s"
-proof -
-  from runing_threads_inv[OF assms]
-  show "th' \<in> threads s" .
-next
-  from runing_cntP_cntV_inv[OF runing' neq_th]
-  show "\<not>detached s th'" using vat_s.detached_eq by simp
-next
-  from runing_preced_inversion[OF runing']
-  show "cp (t@s) th' = preced th s" .
-qed
-
-section {* The existence of `blocking thread` *}
-
-text {* 
-  Suppose @{term th} is not running, it is first shown that
-  there is a path in RAG leading from node @{term th} to another thread @{text "th'"} 
-  in the @{term readys}-set (So @{text "th'"} is an ancestor of @{term th}}).
-
-  Now, since @{term readys}-set is non-empty, there must be
-  one in it which holds the highest @{term cp}-value, which, by definition, 
-  is the @{term runing}-thread. However, we are going to show more: this running thread
-  is exactly @{term "th'"}.
-     *}
-lemma th_blockedE: (* ddd, the other main lemma to be presented: *)
-  assumes "th \<notin> runing (t@s)"
-  obtains th' where "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
-                    "th' \<in> runing (t@s)"
-proof -
-  -- {* According to @{thm vat_t.th_chain_to_ready}, either 
-        @{term "th"} is in @{term "readys"} or there is path leading from it to 
-        one thread in @{term "readys"}. *}
-  have "th \<in> readys (t @ s) \<or> (\<exists>th'. th' \<in> readys (t @ s) \<and> (Th th, Th th') \<in> (RAG (t @ s))\<^sup>+)" 
-    using th_kept vat_t.th_chain_to_ready by auto
-  -- {* However, @{term th} can not be in @{term readys}, because otherwise, since 
-       @{term th} holds the highest @{term cp}-value, it must be @{term "runing"}. *}
-  moreover have "th \<notin> readys (t@s)" 
-    using assms runing_def th_cp_max vat_t.max_cp_readys_threads by auto 
-  -- {* So, there must be a path from @{term th} to another thread @{text "th'"} in 
-        term @{term readys}: *}
-  ultimately obtain th' where th'_in: "th' \<in> readys (t@s)"
-                          and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto
-  -- {* We are going to show that this @{term th'} is running. *}
-  have "th' \<in> runing (t@s)"
-  proof -
-    -- {* We only need to show that this @{term th'} holds the highest @{term cp}-value: *}
-    have "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" (is "?L = ?R")
-    proof -
-      have "?L =  Max ((the_preced (t @ s) \<circ> the_thread) ` subtree (tRAG (t @ s)) (Th th'))"
-        by (unfold cp_alt_def1, simp)
-      also have "... = (the_preced (t @ s) \<circ> the_thread) (Th th)"
-      proof(rule image_Max_subset)
-        show "finite (Th ` (threads (t@s)))" by (simp add: vat_t.finite_threads)
-      next
-        show "subtree (tRAG (t @ s)) (Th th') \<subseteq> Th ` threads (t @ s)"
-          by (metis Range.intros dp trancl_range vat_t.range_in vat_t.subtree_tRAG_thread) 
-      next
-        show "Th th \<in> subtree (tRAG (t @ s)) (Th th')" using dp
-                    by (unfold tRAG_subtree_eq, auto simp:subtree_def)
-      next
-        show "Max ((the_preced (t @ s) \<circ> the_thread) ` Th ` threads (t @ s)) =
-                      (the_preced (t @ s) \<circ> the_thread) (Th th)" (is "Max ?L = _")
-        proof -
-          have "?L = the_preced (t @ s) `  threads (t @ s)" 
-                     by (unfold image_comp, rule image_cong, auto)
-          thus ?thesis using max_preced the_preced_def by auto
-        qed
-      qed
-      also have "... = ?R"
-        using th_cp_max th_cp_preced th_kept 
-              the_preced_def vat_t.max_cp_readys_threads by auto
-      finally show ?thesis .
-    qed 
-    -- {* Now, since @{term th'} holds the highest @{term cp} 
-          and we have already show it is in @{term readys},
-          it is @{term runing} by definition. *}
-    with `th' \<in> readys (t@s)` show ?thesis by (simp add: runing_def) 
-  qed
-  -- {* It is easy to show @{term th'} is an ancestor of @{term th}: *}
-  moreover have "Th th' \<in> ancestors (RAG (t @ s)) (Th th)" 
-    using `(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+` by (auto simp:ancestors_def)
-  ultimately show ?thesis using that by metis
-qed
-
-text {*
-  Now it is easy to see there is always a thread to run by case analysis
-  on whether thread @{term th} is running: if the answer is Yes, the 
-  the running thread is obviously @{term th} itself; otherwise, the running
-  thread is the @{text th'} given by lemma @{thm th_blockedE}.
-*}
-lemma live: "runing (t@s) \<noteq> {}"
-proof(cases "th \<in> runing (t@s)") 
-  case True thus ?thesis by auto
-next
-  case False
-  thus ?thesis using th_blockedE by auto
-qed
-
-
-end
-end
-=======
-theory Correctness
-imports PIPBasics
-begin
-
-
-text {* 
-  The following two auxiliary lemmas are used to reason about @{term Max}.
-*}
-lemma image_Max_eqI: 
-  assumes "finite B"
-  and "b \<in> B"
-  and "\<forall> x \<in> B. f x \<le> f b"
-  shows "Max (f ` B) = f b"
-  using assms
-  using Max_eqI by blast 
-
-lemma image_Max_subset:
-  assumes "finite A"
-  and "B \<subseteq> A"
-  and "a \<in> B"
-  and "Max (f ` A) = f a"
-  shows "Max (f ` B) = f a"
-proof(rule image_Max_eqI)
-  show "finite B"
-    using assms(1) assms(2) finite_subset by auto 
-next
-  show "a \<in> B" using assms by simp
-next
-  show "\<forall>x\<in>B. f x \<le> f a"
-    by (metis Max_ge assms(1) assms(2) assms(4) 
-            finite_imageI image_eqI subsetCE) 
-qed
-
-text {*
-  The following locale @{text "highest_gen"} sets the basic context for our
-  investigation: supposing thread @{text th} holds the highest @{term cp}-value
-  in state @{text s}, which means the task for @{text th} is the 
-  most urgent. We want to show that  
-  @{text th} is treated correctly by PIP, which means
-  @{text th} will not be blocked unreasonably by other less urgent
-  threads. 
-*}
-locale highest_gen =
-  fixes s th prio tm
-  assumes vt_s: "vt s"
-  and threads_s: "th \<in> threads s"
-  and highest: "preced th s = Max ((cp s)`threads s)"
-  -- {* The internal structure of @{term th}'s precedence is exposed:*}
-  and preced_th: "preced th s = Prc prio tm" 
-
--- {* @{term s} is a valid trace, so it will inherit all results derived for
-      a valid trace: *}
-sublocale highest_gen < vat_s: valid_trace "s"
-  by (unfold_locales, insert vt_s, simp)
-
-context highest_gen
-begin
-
-text {*
-  @{term tm} is the time when the precedence of @{term th} is set, so 
-  @{term tm} must be a valid moment index into @{term s}.
-*}
-lemma lt_tm: "tm < length s"
-  by (insert preced_tm_lt[OF threads_s preced_th], simp)
-
-text {*
-  Since @{term th} holds the highest precedence and @{text "cp"}
-  is the highest precedence of all threads in the sub-tree of 
-  @{text "th"} and @{text th} is among these threads, 
-  its @{term cp} must equal to its precedence:
-*}
-lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R")
-proof -
-  have "?L \<le> ?R"
-  by (unfold highest, rule Max_ge, 
-        auto simp:threads_s finite_threads)
-  moreover have "?R \<le> ?L"
-    by (unfold vat_s.cp_rec, rule Max_ge, 
-        auto simp:the_preced_def vat_s.fsbttRAGs.finite_children)
-  ultimately show ?thesis by auto
-qed
-
-lemma highest_cp_preced: "cp s th = Max (the_preced s ` threads s)"
-  using eq_cp_s_th highest max_cp_eq the_preced_def by presburger
-  
-
-lemma highest_preced_thread: "preced th s = Max (the_preced s ` threads s)"
-  by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
-
-lemma highest': "cp s th = Max (cp s ` threads s)"
-  by (simp add: eq_cp_s_th highest)
-
-end
-
-locale extend_highest_gen = highest_gen + 
-  fixes t 
-  assumes vt_t: "vt (t@s)"
-  and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
-  and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
-  and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"
-
-sublocale extend_highest_gen < vat_t: valid_trace "t@s"
-  by (unfold_locales, insert vt_t, simp)
-
-lemma step_back_vt_app: 
-  assumes vt_ts: "vt (t@s)" 
-  shows "vt s"
-proof -
-  from vt_ts show ?thesis
-  proof(induct t)
-    case Nil
-    from Nil show ?case by auto
-  next
-    case (Cons e t)
-    assume ih: " vt (t @ s) \<Longrightarrow> vt s"
-      and vt_et: "vt ((e # t) @ s)"
-    show ?case
-    proof(rule ih)
-      show "vt (t @ s)"
-      proof(rule step_back_vt)
-        from vt_et show "vt (e # t @ s)" by simp
-      qed
-    qed
-  qed
-qed
-
-(* locale red_extend_highest_gen = extend_highest_gen +
-   fixes i::nat
-*)
-
-(*
-sublocale red_extend_highest_gen <   red_moment: extend_highest_gen "s" "th" "prio" "tm" "(moment i t)"
-  apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
-  apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
-  by (unfold highest_gen_def, auto dest:step_back_vt_app)
-*)
-
-context extend_highest_gen
-begin
-
- lemma ind [consumes 0, case_names Nil Cons, induct type]:
-  assumes 
-    h0: "R []"
-  and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e; 
-                    extend_highest_gen s th prio tm t; 
-                    extend_highest_gen s th prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
-  shows "R t"
-proof -
-  from vt_t extend_highest_gen_axioms show ?thesis
-  proof(induct t)
-    from h0 show "R []" .
-  next
-    case (Cons e t')
-    assume ih: "\<lbrakk>vt (t' @ s); extend_highest_gen s th prio tm t'\<rbrakk> \<Longrightarrow> R t'"
-      and vt_e: "vt ((e # t') @ s)"
-      and et: "extend_highest_gen s th prio tm (e # t')"
-    from vt_e and step_back_step have stp: "step (t'@s) e" by auto
-    from vt_e and step_back_vt have vt_ts: "vt (t'@s)" by auto
-    show ?case
-    proof(rule h2 [OF vt_ts stp _ _ _ ])
-      show "R t'"
-      proof(rule ih)
-        from et show ext': "extend_highest_gen s th prio tm t'"
-          by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
-      next
-        from vt_ts show "vt (t' @ s)" .
-      qed
-    next
-      from et show "extend_highest_gen s th prio tm (e # t')" .
-    next
-      from et show ext': "extend_highest_gen s th prio tm t'"
-          by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
-    qed
-  qed
-qed
-
-
-lemma th_kept: "th \<in> threads (t @ s) \<and> 
-                 preced th (t@s) = preced th s" (is "?Q t") 
-proof -
-  show ?thesis
-  proof(induct rule:ind)
-    case Nil
-    from threads_s
-    show ?case
-      by auto
-  next
-    case (Cons e t)
-    interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
-    interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
-    show ?case
-    proof(cases e)
-      case (Create thread prio)
-      show ?thesis
-      proof -
-        from Cons and Create have "step (t@s) (Create thread prio)" by auto
-        hence "th \<noteq> thread"
-        proof(cases)
-          case thread_create
-          with Cons show ?thesis by auto
-        qed
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold Create, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:Create)
-      qed
-    next
-      case (Exit thread)
-      from h_e.exit_diff and Exit
-      have neq_th: "thread \<noteq> th" by auto
-      with Cons
-      show ?thesis
-        by (unfold Exit, auto simp:preced_def)
-    next
-      case (P thread cs)
-      with Cons
-      show ?thesis 
-        by (auto simp:P preced_def)
-    next
-      case (V thread cs)
-      with Cons
-      show ?thesis 
-        by (auto simp:V preced_def)
-    next
-      case (Set thread prio')
-      show ?thesis
-      proof -
-        from h_e.set_diff_low and Set
-        have "th \<noteq> thread" by auto
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold Set, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:Set)
-      qed
-    qed
-  qed
-qed
-
-text {*
-  According to @{thm th_kept}, thread @{text "th"} has its living status
-  and precedence kept along the way of @{text "t"}. The following lemma
-  shows that this preserved precedence of @{text "th"} remains as the highest
-  along the way of @{text "t"}.
-
-  The proof goes by induction over @{text "t"} using the specialized
-  induction rule @{thm ind}, followed by case analysis of each possible 
-  operations of PIP. All cases follow the same pattern rendered by the 
-  generalized introduction rule @{thm "image_Max_eqI"}. 
-
-  The very essence is to show that precedences, no matter whether they 
-  are newly introduced or modified, are always lower than the one held 
-  by @{term "th"}, which by @{thm th_kept} is preserved along the way.
-*}
-lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s"
-proof(induct rule:ind)
-  case Nil
-  from highest_preced_thread
-  show ?case by simp
-next
-  case (Cons e t)
-    interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
-    interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
-  show ?case
-  proof(cases e)
-    case (Create thread prio')
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      -- {* The following is the common pattern of each branch of the case analysis. *}
-      -- {* The major part is to show that @{text "th"} holds the highest precedence: *}
-      have "Max (?f ` ?A) = ?f th"
-      proof(rule image_Max_eqI)
-        show "finite ?A" using h_e.finite_threads by auto 
-      next
-        show "th \<in> ?A" using h_e.th_kept by auto 
-      next
-        show "\<forall>x\<in>?A. ?f x \<le> ?f th"
-        proof 
-          fix x
-          assume "x \<in> ?A"
-          hence "x = thread \<or> x \<in> threads (t@s)" by (auto simp:Create)
-          thus "?f x \<le> ?f th"
-          proof
-            assume "x = thread"
-            thus ?thesis 
-              apply (simp add:Create the_preced_def preced_def, fold preced_def)
-              using Create h_e.create_low h_t.th_kept lt_tm preced_leI2 
-              preced_th by force
-          next
-            assume h: "x \<in> threads (t @ s)"
-            from Cons(2)[unfolded Create] 
-            have "x \<noteq> thread" using h by (cases, auto)
-            hence "?f x = the_preced (t@s) x" 
-              by (simp add:Create the_preced_def preced_def)
-            hence "?f x \<le> Max (the_preced (t@s) ` threads (t@s))"
-              by (simp add: h_t.finite_threads h)
-            also have "... = ?f th"
-              by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
-            finally show ?thesis .
-          qed
-        qed
-      qed
-     -- {* The minor part is to show that the precedence of @{text "th"} 
-           equals to preserved one, given by the foregoing lemma @{thm th_kept} *}
-      also have "... = ?t" using h_e.th_kept the_preced_def by auto
-      -- {* Then it follows trivially that the precedence preserved
-            for @{term "th"} remains the maximum of all living threads along the way. *}
-      finally show ?thesis .
-    qed 
-  next 
-    case (Exit thread)
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "Max (?f ` ?A) = ?f th"
-      proof(rule image_Max_eqI)
-        show "finite ?A" using h_e.finite_threads by auto 
-      next
-        show "th \<in> ?A" using h_e.th_kept by auto 
-      next
-        show "\<forall>x\<in>?A. ?f x \<le> ?f th"
-        proof 
-          fix x
-          assume "x \<in> ?A"
-          hence "x \<in> threads (t@s)" by (simp add: Exit) 
-          hence "?f x \<le> Max (?f ` threads (t@s))" 
-            by (simp add: h_t.finite_threads) 
-          also have "... \<le> ?f th" 
-            apply (simp add:Exit the_preced_def preced_def, fold preced_def)
-            using Cons.hyps(5) h_t.th_kept the_preced_def by auto
-          finally show "?f x \<le> ?f th" .
-        qed
-      qed
-      also have "... = ?t" using h_e.th_kept the_preced_def by auto
-      finally show ?thesis .
-    qed 
-  next
-    case (P thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def the_preced_def)
-  next
-    case (V thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def the_preced_def)
-  next 
-    case (Set thread prio')
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "Max (?f ` ?A) = ?f th"
-      proof(rule image_Max_eqI)
-        show "finite ?A" using h_e.finite_threads by auto 
-      next
-        show "th \<in> ?A" using h_e.th_kept by auto 
-      next
-        show "\<forall>x\<in>?A. ?f x \<le> ?f th"
-        proof 
-          fix x
-          assume h: "x \<in> ?A"
-          show "?f x \<le> ?f th"
-          proof(cases "x = thread")
-            case True
-            moreover have "the_preced (Set thread prio' # t @ s) thread \<le> the_preced (t @ s) th"
-            proof -
-              have "the_preced (t @ s) th = Prc prio tm"  
-                using h_t.th_kept preced_th by (simp add:the_preced_def)
-              moreover have "prio' \<le> prio" using Set h_e.set_diff_low by auto
-              ultimately show ?thesis by (insert lt_tm, auto simp:the_preced_def preced_def)
-            qed
-            ultimately show ?thesis
-              by (unfold Set, simp add:the_preced_def preced_def)
-          next
-            case False
-            then have "?f x  = the_preced (t@s) x"
-              by (simp add:the_preced_def preced_def Set)
-            also have "... \<le> Max (the_preced (t@s) ` threads (t@s))"
-              using Set h h_t.finite_threads by auto 
-            also have "... = ?f th" by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
-            finally show ?thesis .
-          qed
-        qed
-      qed
-      also have "... = ?t" using h_e.th_kept the_preced_def by auto
-      finally show ?thesis .
-    qed 
-  qed
-qed
-
-lemma max_preced: "preced th (t@s) = Max (the_preced (t@s) ` (threads (t@s)))"
-  by (insert th_kept max_kept, auto)
-
-text {*
-  The reason behind the following lemma is that:
-  Since @{term "cp"} is defined as the maximum precedence 
-  of those threads contained in the sub-tree of node @{term "Th th"} 
-  in @{term "RAG (t@s)"}, and all these threads are living threads, and 
-  @{term "th"} is also among them, the maximum precedence of 
-  them all must be the one for @{text "th"}.
-*}
-lemma th_cp_max_preced: 
-  "cp (t@s) th = Max (the_preced (t@s) ` (threads (t@s)))" (is "?L = ?R") 
-proof -
-  let ?f = "the_preced (t@s)"
-  have "?L = ?f th"
-  proof(unfold cp_alt_def, rule image_Max_eqI)
-    show "finite {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
-    proof -
-      have "{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)} = 
-            the_thread ` {n . n \<in> subtree (RAG (t @ s)) (Th th) \<and>
-                            (\<exists> th'. n = Th th')}"
-      by (smt Collect_cong Setcompr_eq_image mem_Collect_eq the_thread.simps)
-      moreover have "finite ..." by (simp add: vat_t.fsbtRAGs.finite_subtree) 
-      ultimately show ?thesis by simp
-    qed
-  next
-    show "th \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
-      by (auto simp:subtree_def)
-  next
-    show "\<forall>x\<in>{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}.
-               the_preced (t @ s) x \<le> the_preced (t @ s) th"
-    proof
-      fix th'
-      assume "th' \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
-      hence "Th th' \<in> subtree (RAG (t @ s)) (Th th)" by auto
-      moreover have "... \<subseteq> Field (RAG (t @ s)) \<union> {Th th}"
-        by (meson subtree_Field)
-      ultimately have "Th th' \<in> ..." by auto
-      hence "th' \<in> threads (t@s)" 
-      proof
-        assume "Th th' \<in> {Th th}"
-        thus ?thesis using th_kept by auto 
-      next
-        assume "Th th' \<in> Field (RAG (t @ s))"
-        thus ?thesis using vat_t.not_in_thread_isolated by blast 
-      qed
-      thus "the_preced (t @ s) th' \<le> the_preced (t @ s) th"
-        by (metis Max_ge finite_imageI finite_threads image_eqI 
-               max_kept th_kept the_preced_def)
-    qed
-  qed
-  also have "... = ?R" by (simp add: max_preced the_preced_def) 
-  finally show ?thesis .
-qed
-
-lemma th_cp_max[simp]: "Max (cp (t@s) ` threads (t@s)) = cp (t@s) th"
-  using max_cp_eq th_cp_max_preced the_preced_def vt_t by presburger
-
-lemma [simp]: "Max (cp (t@s) ` threads (t@s)) = Max (the_preced (t@s) ` threads (t@s))"
-  by (simp add: th_cp_max_preced)
-  
-lemma [simp]: "Max (the_preced (t@s) ` threads (t@s)) = the_preced (t@s) th"
-  using max_kept th_kept the_preced_def by auto
-
-lemma [simp]: "the_preced (t@s) th = preced th (t@s)"
-  using the_preced_def by auto
-
-lemma [simp]: "preced th (t@s) = preced th s"
-  by (simp add: th_kept)
-
-lemma [simp]: "cp s th = preced th s"
-  by (simp add: eq_cp_s_th)
-
-lemma th_cp_preced [simp]: "cp (t@s) th = preced th s"
-  by (fold max_kept, unfold th_cp_max_preced, simp)
-
-lemma preced_less:
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  shows "preced th' s < preced th s"
-  using assms
-by (metis Max.coboundedI finite_imageI highest not_le order.trans 
-    preced_linorder rev_image_eqI threads_s vat_s.finite_threads 
-    vat_s.le_cp)
-
-section {* The `blocking thread` *}
-
-text {* 
-
-  The purpose of PIP is to ensure that the most urgent thread @{term
-  th} is not blocked unreasonably. Therefore, below, we will derive
-  properties of the blocking thread. By blocking thread, we mean a
-  thread in running state t @ s, but is different from thread @{term
-  th}.
-
-  The first lemmas shows that the @{term cp}-value of the blocking
-  thread @{text th'} equals to the highest precedence in the whole
-  system.
-
-*}
-
-lemma runing_preced_inversion:
-  assumes runing': "th' \<in> runing (t @ s)"
-  shows "cp (t @ s) th' = preced th s" 
-proof -
-  have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))" 
-    using assms by (unfold runing_def, auto)
-  also have "\<dots> = preced th s"
-    by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads) 
-  finally show ?thesis .
-qed
-
-text {*
-
-  The next lemma shows how the counters for @{term "P"} and @{term
-  "V"} operations relate to the running threads in the states @{term
-  s} and @{term "t @ s"}: if a thread's @{term "P"}-count equals its
-  @{term "V"}-count (which means it no longer has any resource in its
-  possession), it cannot be a running thread.
-
-  The proof is by contraction with the assumption @{text "th' \<noteq> th"}.
-  The key is the use of @{thm count_eq_dependants} to derive the
-  emptiness of @{text th'}s @{term dependants}-set from the balance of
-  its @{term P} and @{term V} counts.  From this, it can be shown
-  @{text th'}s @{term cp}-value equals to its own precedence.
-
-  On the other hand, since @{text th'} is running, by @{thm
-  runing_preced_inversion}, its @{term cp}-value equals to the
-  precedence of @{term th}.
-
-  Combining the above two results we have that @{text th'} and @{term
-  th} have the same precedence. By uniqueness of precedences, we have
-  @{text "th' = th"}, which is in contradiction with the assumption
-  @{text "th' \<noteq> th"}.
-
-*} 
-                      
-lemma eq_pv_blocked: (* ddd *)
-  assumes neq_th': "th' \<noteq> th"
-  and eq_pv: "cntP (t @ s) th' = cntV (t @ s) th'"
-  shows "th' \<notin> runing (t @ s)"
-proof
-  assume otherwise: "th' \<in> runing (t @ s)"
-  show False
-  proof -
-    have th'_in: "th' \<in> threads (t @ s)"
-        using otherwise readys_threads runing_def by auto 
-    have "th' = th"
-    proof(rule preced_unique)
-      -- {* The proof goes like this: 
-            it is first shown that the @{term preced}-value of @{term th'} 
-            equals to that of @{term th}, then by uniqueness 
-            of @{term preced}-values (given by lemma @{thm preced_unique}), 
-            @{term th'} equals to @{term th}: *}
-      show "preced th' (t @ s) = preced th (t @ s)" (is "?L = ?R")
-      proof -
-        -- {* Since the counts of @{term th'} are balanced, the subtree
-              of it contains only itself, so, its @{term cp}-value
-              equals its @{term preced}-value: *}
-        have "?L = cp (t @ s) th'"
-          by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp)
-        -- {* Since @{term "th'"} is running, by @{thm runing_preced_inversion},
-              its @{term cp}-value equals @{term "preced th s"}, 
-              which equals to @{term "?R"} by simplification: *}
-        also have "... = ?R" 
-            using runing_preced_inversion[OF otherwise] by simp
-        finally show ?thesis .
-      qed
-    qed (auto simp: th'_in th_kept)
-    with `th' \<noteq> th` show ?thesis by simp
- qed
-qed
-
-text {*
-  The following lemma is the extrapolation of @{thm eq_pv_blocked}.
-  It says if a thread, different from @{term th}, 
-  does not hold any resource at the very beginning,
-  it will keep hand-emptied in the future @{term "t@s"}.
-*}
-lemma eq_pv_persist: (* ddd *)
-  assumes neq_th': "th' \<noteq> th"
-  and eq_pv: "cntP s th' = cntV s th'"
-  shows "cntP (t @ s) th' = cntV (t @ s) th'"
-proof(induction rule: ind) 
-  -- {* The nontrivial case is for the @{term Cons}: *}
-  case (Cons e t)
-  -- {* All results derived so far hold for both @{term s} and @{term "t@s"}: *}
-  interpret vat_t: extend_highest_gen s th prio tm t using Cons by simp
-  interpret vat_e: extend_highest_gen s th prio tm "(e # t)" using Cons by simp
-  show ?case
-  proof -
-    -- {* It can be proved that @{term cntP}-value of @{term th'} does not change
-          by the happening of event @{term e}: *}
-    have "cntP ((e#t)@s) th' = cntP (t@s) th'"
-    proof(rule ccontr) -- {* Proof by contradiction. *}
-      -- {* Suppose @{term cntP}-value of @{term th'} is changed by @{term e}: *}
-      assume otherwise: "cntP ((e # t) @ s) th' \<noteq> cntP (t @ s) th'"
-      -- {* Then the actor of @{term e} must be @{term th'} and @{term e}
-            must be a @{term P}-event: *}
-      hence "isP e" "actor e = th'" by (auto simp:cntP_diff_inv) 
-      with vat_t.actor_inv[OF Cons(2)]
-      -- {* According to @{thm actor_inv}, @{term th'} must be running at 
-            the moment @{term "t@s"}: *}
-      have "th' \<in> runing (t@s)" by (cases e, auto)
-      -- {* However, an application of @{thm eq_pv_blocked} to induction hypothesis
-            shows @{term th'} can not be running at moment  @{term "t@s"}: *}
-      moreover have "th' \<notin> runing (t@s)" 
-               using vat_t.eq_pv_blocked[OF neq_th' Cons(5)] .
-      -- {* Contradiction is finally derived: *}
-      ultimately show False by simp
-    qed
-    -- {* It can also be proved that @{term cntV}-value of @{term th'} does not change
-          by the happening of event @{term e}: *}
-    -- {* The proof follows exactly the same pattern as the case for @{term cntP}-value: *}
-    moreover have "cntV ((e#t)@s) th' = cntV (t@s) th'"
-    proof(rule ccontr) -- {* Proof by contradiction. *}
-      assume otherwise: "cntV ((e # t) @ s) th' \<noteq> cntV (t @ s) th'"
-      hence "isV e" "actor e = th'" by (auto simp:cntV_diff_inv) 
-      with vat_t.actor_inv[OF Cons(2)]
-      have "th' \<in> runing (t@s)" by (cases e, auto)
-      moreover have "th' \<notin> runing (t@s)"
-          using vat_t.eq_pv_blocked[OF neq_th' Cons(5)] .
-      ultimately show False by simp
-    qed
-    -- {* Finally, it can be shown that the @{term cntP} and @{term cntV} 
-          value for @{term th'} are still in balance, so @{term th'} 
-          is still hand-emptied after the execution of event @{term e}: *}
-    ultimately show ?thesis using Cons(5) by metis
-  qed
-qed (auto simp:eq_pv)
-
-text {*
-
-  By combining @{thm eq_pv_blocked} and @{thm eq_pv_persist}, it can
-  be derived easily that @{term th'} can not be running in the future:
-
-*}
-
-lemma eq_pv_blocked_persist:
-  assumes neq_th': "th' \<noteq> th"
-  and eq_pv: "cntP s th' = cntV s th'"
-  shows "th' \<notin> runing (t @ s)"
-  using assms
-  by (simp add: eq_pv_blocked eq_pv_persist) 
-
-text {*
-
-  The following lemma shows the blocking thread @{term th'} must hold
-  some resource in the very beginning.
-
-*}
-
-lemma runing_cntP_cntV_inv: (* ddd *)
-  assumes is_runing: "th' \<in> runing (t @ s)"
-  and neq_th': "th' \<noteq> th"
-  shows "cntP s th' > cntV s th'"
-  using assms
-proof -
-  -- {* First, it can be shown that the number of @{term P} and
-        @{term V} operations can not be equal for thred @{term th'} *}
-  have "cntP s th' \<noteq> cntV s th'"
-  proof
-     -- {* The proof goes by contradiction, suppose otherwise: *}
-    assume otherwise: "cntP s th' = cntV s th'"
-    -- {* By applying @{thm  eq_pv_blocked_persist} to this: *}
-    from eq_pv_blocked_persist[OF neq_th' otherwise] 
-    -- {* we have that @{term th'} can not be running at moment @{term "t@s"}: *}
-    have "th' \<notin> runing (t@s)" .
-    -- {* This is obvious in contradiction with assumption @{thm is_runing}  *}
-    thus False using is_runing by simp
-  qed
-  -- {* However, the number of @{term V} is always less or equal to @{term P}: *}
-  moreover have "cntV s th' \<le> cntP s th'" using vat_s.cnp_cnv_cncs by auto
-  -- {* Thesis is finally derived by combining the these two results: *}
-  ultimately show ?thesis by auto
-qed
-
-
-text {*
-
-  The following lemmas shows the blocking thread @{text th'} must be
-  live at the very beginning, i.e. the moment (or state) @{term s}.
-  The proof is a  simple combination of the results above:
-
-*}
-
-lemma runing_threads_inv: 
-  assumes runing': "th' \<in> runing (t@s)"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<in> threads s"
-proof(rule ccontr) -- {* Proof by contradiction: *}
-  assume otherwise: "th' \<notin> threads s" 
-  have "th' \<notin> runing (t @ s)"
-  proof -
-    from vat_s.cnp_cnv_eq[OF otherwise]
-    have "cntP s th' = cntV s th'" .
-    from eq_pv_blocked_persist[OF neq_th' this]
-    show ?thesis .
-  qed
-  with runing' show False by simp
-qed
-
-text {*
-
-  The following lemma summarises the above lemmas to give an overall
-  characterisationof the blocking thread @{text "th'"}:
-
-*}
-
-lemma runing_inversion: (* ddd, one of the main lemmas to present *)
-  assumes runing': "th' \<in> runing (t@s)"
-  and neq_th: "th' \<noteq> th"
-  shows "th' \<in> threads s"
-  and    "\<not>detached s th'"
-  and    "cp (t@s) th' = preced th s"
-proof -
-  from runing_threads_inv[OF assms]
-  show "th' \<in> threads s" .
-next
-  from runing_cntP_cntV_inv[OF runing' neq_th]
-  show "\<not>detached s th'" using vat_s.detached_eq by simp
-next
-  from runing_preced_inversion[OF runing']
-  show "cp (t@s) th' = preced th s" .
-qed
-
-
-section {* The existence of `blocking thread` *}
-
-text {* 
-
-  Suppose @{term th} is not running, it is first shown that there is a
-  path in RAG leading from node @{term th} to another thread @{text
-  "th'"} in the @{term readys}-set (So @{text "th'"} is an ancestor of
-  @{term th}}).
-
-  Now, since @{term readys}-set is non-empty, there must be one in it
-  which holds the highest @{term cp}-value, which, by definition, is
-  the @{term runing}-thread. However, we are going to show more: this
-  running thread is exactly @{term "th'"}.
-
-*}
-
-lemma th_blockedE: (* ddd, the other main lemma to be presented: *)
-  assumes "th \<notin> runing (t@s)"
-  obtains th' where "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
-                    "th' \<in> runing (t@s)"
-proof -
-  -- {* According to @{thm vat_t.th_chain_to_ready}, either 
-        @{term "th"} is in @{term "readys"} or there is path leading from it to 
-        one thread in @{term "readys"}. *}
-  have "th \<in> readys (t @ s) \<or> (\<exists>th'. th' \<in> readys (t @ s) \<and> (Th th, Th th') \<in> (RAG (t @ s))\<^sup>+)" 
-    using th_kept vat_t.th_chain_to_ready by auto
-  -- {* However, @{term th} can not be in @{term readys}, because otherwise, since 
-       @{term th} holds the highest @{term cp}-value, it must be @{term "runing"}. *}
-  moreover have "th \<notin> readys (t@s)" 
-    using assms runing_def th_cp_max vat_t.max_cp_readys_threads by auto 
-  -- {* So, there must be a path from @{term th} to another thread @{text "th'"} in 
-        term @{term readys}: *}
-  ultimately obtain th' where th'_in: "th' \<in> readys (t@s)"
-                          and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto
-  -- {* We are going to show that this @{term th'} is running. *}
-  have "th' \<in> runing (t@s)"
-  proof -
-    -- {* We only need to show that this @{term th'} holds the highest @{term cp}-value: *}
-    have "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" (is "?L = ?R")
-    proof -
-      have "?L =  Max ((the_preced (t @ s) \<circ> the_thread) ` subtree (tRAG (t @ s)) (Th th'))"
-        by (unfold cp_alt_def1, simp)
-      also have "... = (the_preced (t @ s) \<circ> the_thread) (Th th)"
-      proof(rule image_Max_subset)
-        show "finite (Th ` (threads (t@s)))" by (simp add: vat_t.finite_threads)
-      next
-        show "subtree (tRAG (t @ s)) (Th th') \<subseteq> Th ` threads (t @ s)"
-          by (metis Range.intros dp trancl_range vat_t.range_in vat_t.subtree_tRAG_thread) 
-      next
-        show "Th th \<in> subtree (tRAG (t @ s)) (Th th')" using dp
-                    by (unfold tRAG_subtree_eq, auto simp:subtree_def)
-      next
-        show "Max ((the_preced (t @ s) \<circ> the_thread) ` Th ` threads (t @ s)) =
-                      (the_preced (t @ s) \<circ> the_thread) (Th th)" (is "Max ?L = _")
-        proof -
-          have "?L = the_preced (t @ s) `  threads (t @ s)" 
-                     by (unfold image_comp, rule image_cong, auto)
-          thus ?thesis using max_preced the_preced_def by auto
-        qed
-      qed
-      also have "... = ?R"
-        using th_cp_max th_cp_preced th_kept 
-              the_preced_def vat_t.max_cp_readys_threads by auto
-      finally show ?thesis .
-    qed 
-    -- {* Now, since @{term th'} holds the highest @{term cp} 
-          and we have already show it is in @{term readys},
-          it is @{term runing} by definition. *}
-    with `th' \<in> readys (t@s)` show ?thesis by (simp add: runing_def) 
-  qed
-  -- {* It is easy to show @{term th'} is an ancestor of @{term th}: *}
-  moreover have "Th th' \<in> ancestors (RAG (t @ s)) (Th th)" 
-    using `(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+` by (auto simp:ancestors_def)
-  ultimately show ?thesis using that by metis
-qed
-
-text {*
-
-  Now it is easy to see there is always a thread to run by case
-  analysis on whether thread @{term th} is running: if the answer is
-  yes, the the running thread is obviously @{term th} itself;
-  otherwise, the running thread is the @{text th'} given by lemma
-  @{thm th_blockedE}.
-
-*}
-
-lemma live: "runing (t@s) \<noteq> {}"
-proof(cases "th \<in> runing (t@s)") 
-  case True thus ?thesis by auto
-next
-  case False
-  thus ?thesis using th_blockedE by auto
-qed
-
-
-end
-end
--- a/PrioGDef.thy~	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,616 +0,0 @@
-chapter {* Definitions *}
-(*<*)
-theory PrioGDef
-imports Precedence_ord Moment
-begin
-(*>*)
-
-text {*
-  In this section, the formal model of  Priority Inheritance Protocol (PIP) is presented. 
-  The model is based on Paulson's inductive protocol verification method, where 
-  the state of the system is modelled as a list of events happened so far with the latest 
-  event put at the head. 
-*}
-
-text {*
-  To define events, the identifiers of {\em threads},
-  {\em priority} and {\em critical resources } (abbreviated as @{text "cs"}) 
-  need to be represented. All three are represetned using standard 
-  Isabelle/HOL type @{typ "nat"}:
-*}
-
-type_synonym thread = nat -- {* Type for thread identifiers. *}
-type_synonym priority = nat  -- {* Type for priorities. *}
-type_synonym cs = nat -- {* Type for critical sections (or critical resources). *}
-
-text {*
-  \noindent
-  The abstraction of Priority Inheritance Protocol (PIP) is set at the system call level.
-  Every system call is represented as an event. The format of events is defined 
-  defined as follows:
-  *}
-
-datatype event = 
-  Create thread priority | -- {* Thread @{text "thread"} is created with priority @{text "priority"}. *}
-  Exit thread | -- {* Thread @{text "thread"} finishing its execution. *}
-  P thread cs | -- {* Thread @{text "thread"} requesting critical resource @{text "cs"}. *}
-  V thread cs | -- {* Thread @{text "thread"}  releasing critical resource @{text "cs"}. *}
-  Set thread priority -- {* Thread @{text "thread"} resets its priority to @{text "priority"}. *}
-
-
-text {* 
-  As mentioned earlier, in Paulson's inductive method, the states of system are represented as lists of events,
-  which is defined by the following type @{text "state"}:
-  *}
-type_synonym state = "event list"
-
-
-text {* 
-\noindent
-  Resource Allocation Graph (RAG for short) is used extensively in our formal analysis. 
-  The following type @{text "node"} is used to represent nodes in RAG.
-  *}
-datatype node = 
-   Th "thread" | -- {* Node for thread. *}
-   Cs "cs" -- {* Node for critical resource. *}
-
-text {*
-  \noindent
-  The following function
-  @{text "threads"} is used to calculate the set of live threads (@{text "threads s"})
-  in state @{text "s"}.
-  *}
-fun threads :: "state \<Rightarrow> thread set"
-  where 
-  -- {* At the start of the system, the set of threads is empty: *}
-  "threads [] = {}" | 
-  -- {* New thread is added to the @{text "threads"}: *}
-  "threads (Create thread prio#s) = {thread} \<union> threads s" | 
-  -- {* Finished thread is removed: *}
-  "threads (Exit thread # s) = (threads s) - {thread}" | 
-  -- {* Other kind of events does not affect the value of @{text "threads"}: *}
-  "threads (e#s) = threads s" 
-
-text {* 
-  \noindent
-  The function @{text "threads"} defined above is one of 
-  the so called {\em observation function}s which forms 
-  the very basis of Paulson's inductive protocol verification method.
-  Each observation function {\em observes} one particular aspect (or attribute)
-  of the system. For example, the attribute observed by  @{text "threads s"}
-  is the set of threads living in state @{text "s"}. 
-  The protocol being modelled 
-  The decision made the protocol being modelled is based on the {\em observation}s
-  returned by {\em observation function}s. Since {\observation function}s forms 
-  the very basis on which Paulson's inductive method is based, there will be 
-  a lot of such observation functions introduced in the following. In fact, any function 
-  which takes event list as argument is a {\em observation function}.
-  *}
-
-text {* \noindent
-  Observation @{text "priority th s"} is
-  the {\em original priority} of thread @{text "th"} in state @{text "s"}. 
-  The {\em original priority} is the priority 
-  assigned to a thread when it is created or when it is reset by system call 
-  (represented by event @{text "Set thread priority"}).
-*}
-
-fun priority :: "thread \<Rightarrow> state \<Rightarrow> priority"
-  where
-  -- {* @{text "0"} is assigned to threads which have never been created: *}
-  "priority thread [] = 0" |
-  "priority thread (Create thread' prio#s) = 
-     (if thread' = thread then prio else priority thread s)" |
-  "priority thread (Set thread' prio#s) = 
-     (if thread' = thread then prio else priority thread s)" |
-  "priority thread (e#s) = priority thread s"
-
-text {*
-  \noindent
-  Observation @{text "last_set th s"} is the last time when the priority of thread @{text "th"} is set, 
-  observed from state @{text "s"}.
-  The time in the system is measured by the number of events happened so far since the very beginning.
-*}
-fun last_set :: "thread \<Rightarrow> state \<Rightarrow> nat"
-  where
-  "last_set thread [] = 0" |
-  "last_set thread ((Create thread' prio)#s) = 
-       (if (thread = thread') then length s else last_set thread s)" |
-  "last_set thread ((Set thread' prio)#s) = 
-       (if (thread = thread') then length s else last_set thread s)" |
-  "last_set thread (_#s) = last_set thread s"
-
-text {*
-  \noindent 
-  The {\em precedence} is a notion derived from {\em priority}, where the {\em precedence} of 
-  a thread is the combination of its {\em original priority} and {\em time} the priority is set. 
-  The intention is to discriminate threads with the same priority by giving threads whose priority
-  is assigned earlier higher precedences, becasue such threads are more urgent to finish. 
-  This explains the following definition:
-  *}
-definition preced :: "thread \<Rightarrow> state \<Rightarrow> precedence"
-  where "preced thread s \<equiv> Prc (priority thread s) (last_set thread s)"
-
-
-text {*
-  \noindent
-  A number of important notions in PIP are represented as the following functions, 
-  defined in terms of the waiting queues of the system, where the waiting queues 
-  , as a whole, is represented by the @{text "wq"} argument of every notion function.
-  The @{text "wq"} argument is itself a functions which maps every critical resource 
-  @{text "cs"} to the list of threads which are holding or waiting for it. 
-  The thread at the head of this list is designated as the thread which is current 
-  holding the resrouce, which is slightly different from tradition where
-  all threads in the waiting queue are considered as waiting for the resource.
-  *}
-
-consts 
-  holding :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool" 
-  waiting :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool"
-  RAG :: "'b \<Rightarrow> (node \<times> node) set"
-  dependants :: "'b \<Rightarrow> thread \<Rightarrow> thread set"
-
-defs (overloaded) 
-  -- {* 
-  \begin{minipage}{0.9\textwidth}
-  This meaning of @{text "wq"} is reflected in the following definition of @{text "holding wq th cs"},
-  where @{text "holding wq th cs"} means thread @{text "th"} is holding the critical 
-  resource @{text "cs"}. This decision is based on @{text "wq"}.
-  \end{minipage}
-  *}
-
-  cs_holding_def: 
-  "holding wq thread cs \<equiv> (thread \<in> set (wq cs) \<and> thread = hd (wq cs))"
-  -- {* 
-  \begin{minipage}{0.9\textwidth}
-  In accordance with the definition of @{text "holding wq th cs"}, 
-  a thread @{text "th"} is considered waiting for @{text "cs"} if 
-  it is in the {\em waiting queue} of critical resource @{text "cs"}, but not at the head.
-  This is reflected in the definition of @{text "waiting wq th cs"} as follows:
-  \end{minipage}
-  *}
-  cs_waiting_def: 
-  "waiting wq thread cs \<equiv> (thread \<in> set (wq cs) \<and> thread \<noteq> hd (wq cs))"
-  -- {* 
-  \begin{minipage}{0.9\textwidth}
-  @{text "RAG wq"} generates RAG (a binary relations on @{text "node"})
-  out of waiting queues of the system (represented by the @{text "wq"} argument):
-  \end{minipage}
-  *}
-  cs_RAG_def: 
-  "RAG (wq::cs \<Rightarrow> thread list) \<equiv>
-      {(Th th, Cs cs) | th cs. waiting wq th cs} \<union> {(Cs cs, Th th) | cs th. holding wq th cs}"
-  -- {* 
-  \begin{minipage}{0.9\textwidth}
-  The following @{text "dependants wq th"} represents the set of threads which are RAGing on
-  thread @{text "th"} in Resource Allocation Graph @{text "RAG wq"}. 
-  Here, "RAGing" means waiting directly or indirectly on the critical resource. 
-  \end{minipage}
-  *}
-  cs_dependants_def: 
-  "dependants (wq::cs \<Rightarrow> thread list) th \<equiv> {th' . (Th th', Th th) \<in> (RAG wq)^+}"
-
-
-text {* \noindent 
-  The following
-  @{text "cpreced s th"} gives the {\em current precedence} of thread @{text "th"} under
-  state @{text "s"}. The definition of @{text "cpreced"} reflects the basic idea of 
-  Priority Inheritance that the {\em current precedence} of a thread is the precedence 
-  inherited from the maximum of all its dependants, i.e. the threads which are waiting 
-  directly or indirectly waiting for some resources from it. If no such thread exits, 
-  @{text "th"}'s {\em current precedence} equals its original precedence, i.e. 
-  @{text "preced th s"}.
-  *}
-
-definition cpreced :: "(cs \<Rightarrow> thread list) \<Rightarrow> state \<Rightarrow> thread \<Rightarrow> precedence"
-  where "cpreced wq s = (\<lambda>th. Max ((\<lambda>th'. preced th' s) ` ({th} \<union> dependants wq th)))"
-
-text {*
-  Notice that the current precedence (@{text "cpreced"}) of one thread @{text "th"} can be boosted 
-  (becoming larger than its own precedence) by those threads in 
-  the @{text "dependants wq th"}-set. If one thread get boosted, we say 
-  it inherits the priority (or, more precisely, the precedence) of 
-  its dependants. This is how the word "Inheritance" in 
-  Priority Inheritance Protocol comes.
-*}
-
-(*<*)
-lemma 
-  cpreced_def2:
-  "cpreced wq s th \<equiv> Max ({preced th s} \<union> {preced th' s | th'. th' \<in> dependants wq th})"
-  unfolding cpreced_def image_def
-  apply(rule eq_reflection)
-  apply(rule_tac f="Max" in arg_cong)
-  by (auto)
-(*>*)
-
-
-text {* \noindent
-  Assuming @{text "qs"} be the waiting queue of a critical resource, 
-  the following abbreviation "release qs" is the waiting queue after the thread 
-  holding the resource (which is thread at the head of @{text "qs"}) released
-  the resource:
-*}
-abbreviation
-  "release qs \<equiv> case qs of
-             [] => [] 
-          |  (_#qs') => (SOME q. distinct q \<and> set q = set qs')"
-text {* \noindent
-  It can be seen from the definition that the thread at the head of @{text "qs"} is removed
-  from the return value, and the value @{term "q"} is an reordering of @{text "qs'"}, the 
-  tail of @{text "qs"}. Through this reordering, one of the waiting threads (those in @{text "qs'"} }
-  is chosen nondeterministically to be the head of the new queue @{text "q"}. 
-  Therefore, this thread is the one who takes over the resource. This is a little better different 
-  from common sense that the thread who comes the earliest should take over.  
-  The intention of this definition is to show that the choice of which thread to take over the 
-  release resource does not affect the correctness of the PIP protocol. 
-*}
-
-text {*
-  The data structure used by the operating system for scheduling is referred to as 
-  {\em schedule state}. It is represented as a record consisting of 
-  a function assigning waiting queue to resources 
-  (to be used as the @{text "wq"} argument in @{text "holding"}, @{text "waiting"} 
-  and  @{text "RAG"}, etc) and a function assigning precedence to threads:
-  *}
-
-record schedule_state = 
-    wq_fun :: "cs \<Rightarrow> thread list" -- {* The function assigning waiting queue. *}
-    cprec_fun :: "thread \<Rightarrow> precedence" -- {* The function assigning precedence. *}
-
-text {* \noindent
-  The following two abbreviations (@{text "all_unlocked"} and @{text "initial_cprec"}) 
-  are used to set the initial values of the @{text "wq_fun"} @{text "cprec_fun"} fields 
-  respectively of the @{text "schedule_state"} record by the following function @{text "sch"},
-  which is used to calculate the system's {\em schedule state}.
-
-  Since there is no thread at the very beginning to make request, all critical resources 
-  are free (or unlocked). This status is represented by the abbreviation
-  @{text "all_unlocked"}. 
-  *}
-abbreviation
-  "all_unlocked \<equiv> \<lambda>_::cs. ([]::thread list)"
-
-
-text {* \noindent
-  The initial current precedence for a thread can be anything, because there is no thread then. 
-  We simply assume every thread has precedence @{text "Prc 0 0"}.
-  *}
-
-abbreviation 
-  "initial_cprec \<equiv> \<lambda>_::thread. Prc 0 0"
-
-
-text {* \noindent
-  The following function @{text "schs"} is used to calculate the system's schedule state @{text "schs s"}
-  out of the current system state @{text "s"}. It is the central function to model Priority Inheritance:
-  *}
-fun schs :: "state \<Rightarrow> schedule_state"
-  where 
-  -- {*
-  \begin{minipage}{0.9\textwidth}
-    Setting the initial value of the @{text "schedule_state"} record (see the explanations above).
-  \end{minipage}
-  *}
-  "schs [] = (| wq_fun = all_unlocked,  cprec_fun = initial_cprec |)" |
-
-  -- {*
-  \begin{minipage}{0.9\textwidth}
-  \begin{enumerate}
-  \item @{text "ps"} is the schedule state of last moment.
-  \item @{text "pwq"} is the waiting queue function of last moment.
-  \item @{text "pcp"} is the precedence function of last moment (NOT USED). 
-  \item @{text "nwq"} is the new waiting queue function. It is calculated using a @{text "case"} statement:
-  \begin{enumerate}
-      \item If the happening event is @{text "P thread cs"}, @{text "thread"} is added to 
-            the end of @{text "cs"}'s waiting queue.
-      \item If the happening event is @{text "V thread cs"} and @{text "s"} is a legal state,
-            @{text "th'"} must equal to @{text "thread"}, 
-            because @{text "thread"} is the one currently holding @{text "cs"}. 
-            The case @{text "[] \<Longrightarrow> []"} may never be executed in a legal state.
-            the @{text "(SOME q. distinct q \<and> set q = set qs)"} is used to choose arbitrarily one 
-            thread in waiting to take over the released resource @{text "cs"}. In our representation,
-            this amounts to rearrange elements in waiting queue, so that one of them is put at the head.
-      \item For other happening event, the schedule state just does not change.
-  \end{enumerate}
-  \item @{text "ncp"} is new precedence function, it is calculated from the newly updated waiting queue 
-        function. The RAGency of precedence function on waiting queue function is the reason to 
-        put them in the same record so that they can evolve together.
-  \end{enumerate}
-  
-
-  The calculation of @{text "cprec_fun"} depends on the value of @{text "wq_fun"}. 
-  Therefore, in the following cases, @{text "wq_fun"} is always calculated first, in 
-  the name of @{text "wq"} (if  @{text "wq_fun"} is not changed
-  by the happening event) or @{text "new_wq"} (if the value of @{text "wq_fun"} is changed).
-  \end{minipage}
-     *}
-   "schs (Create th prio # s) = 
-       (let wq = wq_fun (schs s) in
-          (|wq_fun = wq, cprec_fun = cpreced wq (Create th prio # s)|))"
-|  "schs (Exit th # s) = 
-       (let wq = wq_fun (schs s) in
-          (|wq_fun = wq, cprec_fun = cpreced wq (Exit th # s)|))"
-|  "schs (Set th prio # s) = 
-       (let wq = wq_fun (schs s) in
-          (|wq_fun = wq, cprec_fun = cpreced wq (Set th prio # s)|))"
-   -- {*
-   \begin{minipage}{0.9\textwidth}
-      Different from the forth coming cases, the @{text "wq_fun"} field of the schedule state 
-      is changed. So, the new value is calculated first, in the name of @{text "new_wq"}.
-   \end{minipage}   
-   *}
-|  "schs (P th cs # s) = 
-       (let wq = wq_fun (schs s) in
-        let new_wq = wq(cs := (wq cs @ [th])) in
-          (|wq_fun = new_wq, cprec_fun = cpreced new_wq (P th cs # s)|))"
-|  "schs (V th cs # s) = 
-       (let wq = wq_fun (schs s) in
-        let new_wq = wq(cs := release (wq cs)) in
-          (|wq_fun = new_wq, cprec_fun = cpreced new_wq (V th cs # s)|))"
-
-lemma cpreced_initial: 
-  "cpreced (\<lambda> cs. []) [] = (\<lambda>_. (Prc 0 0))"
-apply(simp add: cpreced_def)
-apply(simp add: cs_dependants_def cs_RAG_def cs_waiting_def cs_holding_def)
-apply(simp add: preced_def)
-done
-
-lemma sch_old_def:
-  "schs (e#s) = (let ps = schs s in 
-                  let pwq = wq_fun ps in 
-                  let nwq = case e of
-                             P th cs \<Rightarrow>  pwq(cs:=(pwq cs @ [th])) |
-                             V th cs \<Rightarrow> let nq = case (pwq cs) of
-                                                      [] \<Rightarrow> [] | 
-                                                      (_#qs) \<Rightarrow> (SOME q. distinct q \<and> set q = set qs)
-                                            in pwq(cs:=nq)                 |
-                              _ \<Rightarrow> pwq
-                  in let ncp = cpreced nwq (e#s) in 
-                     \<lparr>wq_fun = nwq, cprec_fun = ncp\<rparr>
-                 )"
-apply(cases e)
-apply(simp_all)
-done
-
-
-text {* 
-  \noindent
-  The following @{text "wq"} is a shorthand for @{text "wq_fun"}. 
-  *}
-definition wq :: "state \<Rightarrow> cs \<Rightarrow> thread list" 
-  where "wq s = wq_fun (schs s)"
-
-text {* \noindent 
-  The following @{text "cp"} is a shorthand for @{text "cprec_fun"}. 
-  *}
-definition cp :: "state \<Rightarrow> thread \<Rightarrow> precedence"
-  where "cp s \<equiv> cprec_fun (schs s)"
-
-definition "cp_gen s x = Max ((the_preced s \<circ> the_thread) ` subtree (tRAG s) x)"
-
-text {* \noindent
-  Functions @{text "holding"}, @{text "waiting"}, @{text "RAG"} and 
-  @{text "dependants"} still have the 
-  same meaning, but redefined so that they no longer RAG on the 
-  fictitious {\em waiting queue function}
-  @{text "wq"}, but on system state @{text "s"}.
-  *}
-defs (overloaded) 
-  s_holding_abv: 
-  "holding (s::state) \<equiv> holding (wq_fun (schs s))"
-  s_waiting_abv: 
-  "waiting (s::state) \<equiv> waiting (wq_fun (schs s))"
-  s_RAG_abv: 
-  "RAG (s::state) \<equiv> RAG (wq_fun (schs s))"
-  s_dependants_abv: 
-  "dependants (s::state) \<equiv> dependants (wq_fun (schs s))"
-
-
-text {* 
-  The following lemma can be proved easily, and the meaning is obvious.
-  *}
-lemma
-  s_holding_def: 
-  "holding (s::state) th cs \<equiv> (th \<in> set (wq_fun (schs s) cs) \<and> th = hd (wq_fun (schs s) cs))"
-  by (auto simp:s_holding_abv wq_def cs_holding_def)
-
-lemma s_waiting_def: 
-  "waiting (s::state) th cs \<equiv> (th \<in> set (wq_fun (schs s) cs) \<and> th \<noteq> hd (wq_fun (schs s) cs))"
-  by (auto simp:s_waiting_abv wq_def cs_waiting_def)
-
-lemma s_RAG_def: 
-  "RAG (s::state) =
-    {(Th th, Cs cs) | th cs. waiting (wq s) th cs} \<union> {(Cs cs, Th th) | cs th. holding (wq s) th cs}"
-  by (auto simp:s_RAG_abv wq_def cs_RAG_def)
-
-lemma
-  s_dependants_def: 
-  "dependants (s::state) th \<equiv> {th' . (Th th', Th th) \<in> (RAG (wq s))^+}"
-  by (auto simp:s_dependants_abv wq_def cs_dependants_def)
-
-text {*
-  The following function @{text "readys"} calculates the set of ready threads. A thread is {\em ready} 
-  for running if it is a live thread and it is not waiting for any critical resource.
-  *}
-definition readys :: "state \<Rightarrow> thread set"
-  where "readys s \<equiv> {th . th \<in> threads s \<and> (\<forall> cs. \<not> waiting s th cs)}"
-
-text {* \noindent
-  The following function @{text "runing"} calculates the set of running thread, which is the ready 
-  thread with the highest precedence.  
-  *}
-definition runing :: "state \<Rightarrow> thread set"
-  where "runing s \<equiv> {th . th \<in> readys s \<and> cp s th = Max ((cp s) ` (readys s))}"
-
-text {* \noindent
-  Notice that the definition of @{text "running"} reflects the preemptive scheduling strategy,  
-  because, if the @{text "running"}-thread (the one in @{text "runing"} set) 
-  lowered its precedence by resetting its own priority to a lower
-  one, it will lose its status of being the max in @{text "ready"}-set and be superseded.
-*}
-
-text {* \noindent
-  The following function @{text "holdents s th"} returns the set of resources held by thread 
-  @{text "th"} in state @{text "s"}.
-  *}
-definition holdents :: "state \<Rightarrow> thread \<Rightarrow> cs set"
-  where "holdents s th \<equiv> {cs . holding s th cs}"
-
-lemma holdents_test: 
-  "holdents s th = {cs . (Cs cs, Th th) \<in> RAG s}"
-unfolding holdents_def
-unfolding s_RAG_def
-unfolding s_holding_abv
-unfolding wq_def
-by (simp)
-
-text {* \noindent
-  Observation @{text "cntCS s th"} returns the number of resources held by thread @{text "th"} in
-  state @{text "s"}:
-  *}
-definition cntCS :: "state \<Rightarrow> thread \<Rightarrow> nat"
-  where "cntCS s th = card (holdents s th)"
-
-text {* \noindent
-  According to the convention of Paulson's inductive method,
-  the decision made by a protocol that event @{text "e"} is eligible to happen next under state @{text "s"} 
-  is expressed as @{text "step s e"}. The predicate @{text "step"} is inductively defined as 
-  follows (notice how the decision is based on the {\em observation function}s 
-  defined above, and also notice how a complicated protocol is modeled by a few simple 
-  observations, and how such a kind of simplicity gives rise to improved trust on
-  faithfulness):
-  *}
-inductive step :: "state \<Rightarrow> event \<Rightarrow> bool"
-  where
-  -- {* 
-  A thread can be created if it is not a live thread:
-  *}
-  thread_create: "\<lbrakk>thread \<notin> threads s\<rbrakk> \<Longrightarrow> step s (Create thread prio)" |
-  -- {*
-  A thread can exit if it no longer hold any resource:
-  *}
-  thread_exit: "\<lbrakk>thread \<in> runing s; holdents s thread = {}\<rbrakk> \<Longrightarrow> step s (Exit thread)" |
-  -- {*
-  \begin{minipage}{0.9\textwidth}
-  A thread can request for an critical resource @{text "cs"}, if it is running and 
-  the request does not form a loop in the current RAG. The latter condition 
-  is set up to avoid deadlock. The condition also reflects our assumption all threads are 
-  carefully programmed so that deadlock can not happen:
-  \end{minipage}
-  *}
-  thread_P: "\<lbrakk>thread \<in> runing s;  (Cs cs, Th thread)  \<notin> (RAG s)^+\<rbrakk> \<Longrightarrow> 
-                                                                step s (P thread cs)" |
-  -- {*
-  \begin{minipage}{0.9\textwidth}
-  A thread can release a critical resource @{text "cs"} 
-  if it is running and holding that resource:
-  \end{minipage}
-  *}
-  thread_V: "\<lbrakk>thread \<in> runing s;  holding s thread cs\<rbrakk> \<Longrightarrow> step s (V thread cs)" |
-  -- {*
-  \begin{minipage}{0.9\textwidth}
-  A thread can adjust its own priority as long as it is current running. 
-  With the resetting of one thread's priority, its precedence may change. 
-  If this change lowered the precedence, according to the definition of @{text "running"}
-  function, 
-  \end{minipage}
-  *}  
-  thread_set: "\<lbrakk>thread \<in> runing s\<rbrakk> \<Longrightarrow> step s (Set thread prio)"
-
-text {*
-  In Paulson's inductive method, every protocol is defined by such a @{text "step"}
-  predicate. For instance, the predicate @{text "step"} given above 
-  defines the PIP protocol. So, it can also be called "PIP".
-*}
-
-abbreviation
-  "PIP \<equiv> step"
-
-
-text {* \noindent
-  For any protocol defined by a @{text "step"} predicate, 
-  the fact that @{text "s"} is a legal state in 
-  the protocol is expressed as: @{text "vt step s"}, where
-  the predicate @{text "vt"} can be defined as the following:
-  *}
-inductive vt :: "state \<Rightarrow> bool"
-  where
-  -- {* Empty list @{text "[]"} is a legal state in any protocol:*}
-  vt_nil[intro]: "vt []" |
-  -- {* 
-  \begin{minipage}{0.9\textwidth}
-  If @{text "s"} a legal state of the protocol defined by predicate @{text "step"}, 
-  and event @{text "e"} is allowed to happen under state @{text "s"} by the protocol 
-  predicate @{text "step"}, then @{text "e#s"} is a new legal state rendered by the 
-  happening of @{text "e"}:
-  \end{minipage}
-  *}
-  vt_cons[intro]: "\<lbrakk>vt s; step s e\<rbrakk> \<Longrightarrow> vt (e#s)"
-
-text {*  \noindent
-  It is easy to see that the definition of @{text "vt"} is generic. It can be applied to 
-  any specific protocol specified by a @{text "step"}-predicate to get the set of
-  legal states of that particular protocol.
-  *}
-
-text {* 
-  The following are two very basic properties of @{text "vt"}.
-*}
-
-lemma step_back_vt: "vt (e#s) \<Longrightarrow> vt s"
-  by(ind_cases "vt (e#s)", simp)
-
-lemma step_back_step: "vt (e#s) \<Longrightarrow> step s e"
-  by(ind_cases "vt (e#s)", simp)
-
-text {* \noindent
-  The following two auxiliary functions @{text "the_cs"} and @{text "the_th"} are used to extract
-  critical resource and thread respectively out of RAG nodes.
-  *}
-fun the_cs :: "node \<Rightarrow> cs"
-  where "the_cs (Cs cs) = cs"
-
-fun the_th :: "node \<Rightarrow> thread"
-  where "the_th (Th th) = th"
-
-text {* \noindent
-  The following predicate @{text "next_th"} describe the next thread to 
-  take over when a critical resource is released. In @{text "next_th s th cs t"}, 
-  @{text "th"} is the thread to release, @{text "t"} is the one to take over.
-  Notice how this definition is backed up by the @{text "release"} function and its use 
-  in the @{text "V"}-branch of @{text "schs"} function. This @{text "next_th"} function
-  is not needed for the execution of PIP. It is introduced as an auxiliary function 
-  to state lemmas. The correctness of this definition will be confirmed by 
-  lemmas @{text "step_v_hold_inv"}, @{text " step_v_wait_inv"}, 
-  @{text "step_v_get_hold"} and @{text "step_v_not_wait"}.
-  *}
-definition next_th:: "state \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> thread \<Rightarrow> bool"
-  where "next_th s th cs t = (\<exists> rest. wq s cs = th#rest \<and> rest \<noteq> [] \<and> 
-                                     t = hd (SOME q. distinct q \<and> set q = set rest))"
-
-text {* \noindent
-  The aux function @{text "count Q l"} is used to count the occurrence of situation @{text "Q"}
-  in list @{text "l"}:
-  *}
-definition count :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> nat"
-  where "count Q l = length (filter Q l)"
-
-text {* \noindent
-  The following observation @{text "cntP s"} returns the number of operation @{text "P"} happened 
-  before reaching state @{text "s"}.
-  *}
-definition cntP :: "state \<Rightarrow> thread \<Rightarrow> nat"
-  where "cntP s th = count (\<lambda> e. \<exists> cs. e = P th cs) s"
-
-text {* \noindent
-  The following observation @{text "cntV s"} returns the number of operation @{text "V"} happened 
-  before reaching state @{text "s"}.
-  *}
-definition cntV :: "state \<Rightarrow> thread \<Rightarrow> nat"
-  where "cntV s th = count (\<lambda> e. \<exists> cs. e = V th cs) s"
-(*<*)
-
-end
-(*>*)
-
--- a/RTree.thy~	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1880 +0,0 @@
-theory RTree
-imports "~~/src/HOL/Library/Transitive_Closure_Table" Max
-begin
-
-section {* A theory of relational trees *}
-
-inductive_cases path_nilE [elim!]: "rtrancl_path r x [] y"
-inductive_cases path_consE [elim!]: "rtrancl_path r x (z#zs) y"
-
-subsection {* Definitions *}
-
-text {*
-  In this theory, we are going to give a notion of of `Relational Graph` and
-  its derived notion `Relational Tree`. Given a binary relation @{text "r"},
-  the `Relational Graph of @{text "r"}` is the graph, the edges of which
-  are those in @{text "r"}. In this way, any binary relation can be viewed
-  as a `Relational Graph`. Note, this notion of graph includes infinite graphs. 
-
-  A `Relation Graph` @{text "r"} is said to be a `Relational Tree` if it is both
-  {\em single valued} and {\em acyclic}. 
-*}
-
-text {*
-  The following @{text "sgv"} specifies that relation @{text "r"} is {\em single valued}.
-*}
-locale sgv = 
-  fixes r
-  assumes sgv: "single_valued r"
-
-text {*
-  The following @{text "rtree"} specifies that @{text "r"} is a 
-  {\em Relational Tree}.
-*}
-locale rtree = sgv + 
-  assumes acl: "acyclic r"
-
-text {* 
-  The following two auxiliary functions @{text "rel_of"} and @{text "pred_of"} 
-  transfer between the predicate and set representation of binary relations.
-*}
-
-definition "rel_of r = {(x, y) | x y. r x y}"
-
-definition "pred_of r = (\<lambda> x y. (x, y) \<in> r)"
-
-text {*
-  To reason about {\em Relational Graph}, a notion of path is 
-  needed, which is given by the following @{text "rpath"} (short 
-  for `relational path`). 
-  The path @{text "xs"} in proposition @{text "rpath r x xs y"} is 
-  a path leading from @{text "x"} to @{text "y"}, which serves as a 
-  witness of the fact @{text "(x, y) \<in> r^*"}. 
-
-  @{text "rpath"}
-  is simply a wrapper of the @{text "rtrancl_path"} defined in the imported 
-  theory @{text "Transitive_Closure_Table"}, which defines 
-  a notion of path for the predicate form of binary relations. 
-*}
-definition "rpath r x xs y = rtrancl_path (pred_of r) x xs y"
-
-text {*
-  Given a path @{text "ps"}, @{text "edges_on ps"} is the 
-  set of edges along the path, which is defined as follows:
-*}
-
-definition "edges_on ps = {(a,b) | a b. \<exists> xs ys. ps = xs@[a,b]@ys}"
-
-text {*
-   The following @{text "indep"} defines a notion of independence. 
-   Two nodes @{text "x"} and @{text "y"} are said to be independent
-   (expressed as @{text "indep x y"}),  if neither one is reachable 
-   from the other in relational graph @{text "r"}.
-*}
-definition "indep r x y = (((x, y) \<notin> r^*) \<and> ((y, x) \<notin> r^*))"
-
-text {*
-  In relational tree @{text "r"}, the sub tree of node @{text "x"} is written
-  @{text "subtree r x"}, which is defined to be the set of nodes (including itself) 
-  which can reach @{text "x"} by following some path in @{text "r"}:
-*}
-
-definition "subtree r x = {y . (y, x) \<in> r^*}"
-
-definition "ancestors r x = {y. (x, y) \<in> r^+}"
-
-definition "root r x = (ancestors r x = {})"
-
-text {*
-  The following @{text "edge_in r x"} is the set of edges
-  contained in the sub-tree of @{text "x"}, with @{text "r"} as the underlying graph.
-*}
-
-definition "edges_in r x = {(a, b) | a b. (a, b) \<in> r \<and> b \<in> subtree r x}"
-
-text {*
-  The following lemma @{text "edges_in_meaning"} shows the intuitive meaning 
-  of `an edge @{text "(a, b)"} is in the sub-tree of @{text "x"}`, 
-  i.e., both @{text "a"} and @{text "b"} are in the sub-tree.
-*}
-lemma edges_in_meaning: 
-  "edges_in r x = {(a, b) | a b. (a, b) \<in> r \<and> a \<in> subtree r x \<and> b \<in> subtree r x}"
-proof -
-  { fix a b
-    assume h: "(a, b) \<in> r" "b \<in> subtree r x"
-    moreover have "a \<in> subtree r x"
-    proof -
-      from h(2)[unfolded subtree_def] have "(b, x) \<in> r^*" by simp
-      with h(1) have "(a, x) \<in> r^*" by auto
-      thus ?thesis by (auto simp:subtree_def)
-    qed
-    ultimately have "((a, b) \<in> r \<and> a \<in> subtree r x \<and> b \<in> subtree r x)" 
-      by (auto)
-  } thus ?thesis by (auto simp:edges_in_def)
-qed
-
-text {*
-  The following lemma shows the meaning of @{term "edges_in"} from the other side, 
-  which says: for the edge @{text "(a,b)"} to be outside of the sub-tree of @{text "x"}, 
-  it is sufficient to show that @{text "b"} is.
-*}
-lemma edges_in_refutation:
-  assumes "b \<notin> subtree r x"
-  shows "(a, b) \<notin> edges_in r x"
-  using assms by (unfold edges_in_def subtree_def, auto)
-
-definition "children r x = {y. (y, x) \<in> r}"
-
-locale fbranch =
-  fixes r
-  assumes fb: "\<forall> x \<in> Range r . finite (children r x)"
-begin
-
-lemma finite_children: "finite (children r x)"
-proof(cases "children r x = {}")
-  case True
-  thus ?thesis by auto
-next
-  case False
-  then obtain y where "(y, x) \<in> r" by (auto simp:children_def)
-  hence "x \<in> Range r" by auto
-  from fb[rule_format, OF this]
-  show ?thesis .
-qed
-
-end
-
-locale fsubtree = fbranch + 
-   assumes wf: "wf r"
-
-(* ccc *)
-
-subsection {* Auxiliary lemmas *}
-
-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
-  proof(induct i rule:less_induct)
-    case (less t)
-    show ?case
-    proof(cases "\<forall> j < t. \<not> P j")
-      case True
-      with less (2) show ?thesis by blast
-    next
-      case False
-      then obtain j where "j < t" "P j" by auto
-      from less(1)[OF this]
-      show ?thesis .
-    qed
-  qed 
-  with that show ?thesis by metis
-qed
-
-subsection {* Properties of Relational Graphs and Relational Trees *}
-
-subsubsection {* Properties of @{text "rel_of"} and @{text "pred_of"} *}
-
-text {* The following lemmas establish bijectivity of the two functions *}
-
-lemma pred_rel_eq: "pred_of (rel_of r) = r" by (auto simp:rel_of_def pred_of_def)
-
-lemma rel_pred_eq: "rel_of (pred_of r) = r" by (auto simp:rel_of_def pred_of_def)
-
-lemma rel_of_star: "rel_of (r^**) = (rel_of r)^*"
-  by (unfold rel_of_def rtranclp_rtrancl_eq, auto)
-
-lemma pred_of_star: "pred_of (r^*) = (pred_of r)^**"
-proof -
-  { fix x y
-    have "pred_of (r^*) x y = (pred_of r)^** x y"
-    by (unfold pred_of_def rtranclp_rtrancl_eq, auto)
-  } thus ?thesis by auto
-qed
-
-lemma star_2_pstar: "(x, y) \<in> r^* = (pred_of (r^*)) x y"
-  by (simp add: pred_of_def)
-
-subsubsection {* Properties of @{text "rpath"} *}
-
-text {* Induction rule for @{text "rpath"}: *}
-
-lemma rpath_induct [consumes 1, case_names rbase rstep, induct pred: rpath]:
-  assumes "rpath r x1 x2 x3"
-    and "\<And>x. P x [] x"
-    and "\<And>x y ys z. (x, y) \<in> r \<Longrightarrow> rpath r y ys z \<Longrightarrow> P y ys z \<Longrightarrow> P x (y # ys) z"
-  shows "P x1 x2 x3"
-  using assms[unfolded rpath_def]
-  by (induct, auto simp:pred_of_def rpath_def)
-
-lemma rpathE: 
-  assumes "rpath r x xs y"
-  obtains (base) "y = x" "xs = []"
-     | (step) z zs where "(x, z) \<in> r" "rpath r z zs y" "xs = z#zs"
-  using assms
-  by (induct, auto)
-
-text {* Introduction rule for empty path *}
-lemma rbaseI [intro!]: 
-  assumes "x = y"
-  shows "rpath r x [] y"
-  by  (unfold rpath_def assms, 
-         rule Transitive_Closure_Table.rtrancl_path.base)
-
-text {* Introduction rule for non-empty path *}
-lemma rstepI [intro!]: 
-  assumes "(x, y) \<in> r"
-    and "rpath r y ys z"
-  shows "rpath r x (y#ys) z" 
-proof(unfold rpath_def, rule Transitive_Closure_Table.rtrancl_path.step)
-  from assms(1) show "pred_of r x y" by (auto simp:pred_of_def)
-next
-  from assms(2) show "rtrancl_path (pred_of r) y ys z"  
-  by (auto simp:pred_of_def rpath_def)
-qed
-
-text {* Introduction rule for @{text "@"}-path *}
-lemma rpath_appendI [intro]: 
-  assumes "rpath r x xs a" and "rpath r a ys y"
-  shows "rpath r x (xs @ ys) y"
-  using assms 
-  by (unfold rpath_def, auto intro:rtrancl_path_trans)
-
-text {* Elimination rule for empty path *}
-
-lemma rpath_cases [cases pred:rpath]:
-  assumes "rpath r a1 a2 a3"
-  obtains (rbase)  "a1 = a3" and "a2 = []"
-    | (rstep)  y :: "'a" and ys :: "'a list"  
-         where "(a1, y) \<in> r" and "a2 = y # ys" and "rpath r y ys a3"
-  using assms [unfolded rpath_def]
-  by (cases, auto simp:rpath_def pred_of_def)
-
-lemma rpath_nilE [elim!, cases pred:rpath]: 
-  assumes "rpath r x [] y"
-  obtains "y = x"
-  using assms[unfolded rpath_def] by auto
-
--- {* This is a auxiliary lemmas used only in the proof of @{text "rpath_nnl_lastE"} *}
-lemma rpath_nnl_last:
-  assumes "rtrancl_path r x xs y"
-  and "xs \<noteq> []"
-  obtains xs' where "xs = xs'@[y]"
-proof -
-  from append_butlast_last_id[OF `xs \<noteq> []`, symmetric] 
-  obtain xs' y' where eq_xs: "xs = (xs' @ y' # [])" by simp
-  with assms(1)
-  have "rtrancl_path r x ... y" by simp
-  hence "y = y'" by (rule rtrancl_path_appendE, auto)
-  with eq_xs have "xs = xs'@[y]" by simp
-  from that[OF this] show ?thesis .
-qed
-
-text {*
-  Elimination rule for non-empty paths constructed with @{text "#"}.
-*}
-
-lemma rpath_ConsE [elim!, cases pred:rpath]:
-  assumes "rpath r x (y # ys) x2"
-  obtains (rstep) "(x, y) \<in> r" and "rpath r y ys x2"
-  using assms[unfolded rpath_def]
-  by (cases, auto simp:rpath_def pred_of_def)
-
-text {*
-  Elimination rule for non-empty path, where the destination node 
-  @{text "y"} is shown to be at the end of the path.
-*}
-lemma rpath_nnl_lastE: 
-  assumes "rpath r x xs y"
-  and "xs \<noteq> []"
-  obtains xs' where "xs = xs'@[y]"
-  using assms[unfolded rpath_def]
-  by (rule rpath_nnl_last, auto)
-
-text {* Other elimination rules of @{text "rpath"} *}
-
-lemma rpath_appendE:
-  assumes "rpath r x (xs @ [a] @ ys) y"
-  obtains "rpath r x (xs @ [a]) a" and "rpath r a ys y"
-  using rtrancl_path_appendE[OF assms[unfolded rpath_def, simplified], folded rpath_def]
-  by auto
-
-lemma rpath_subE: 
-  assumes "rpath r x (xs @ [a] @ ys @ [b] @ zs) y"
-  obtains "rpath r x (xs @ [a]) a" and "rpath r a (ys @ [b]) b" and "rpath r b zs y" 
-  using assms
- by (elim rpath_appendE, auto)
-
-text {* Every path has a unique end point. *}
-lemma rpath_dest_eq:
-  assumes "rpath r x xs x1"
-  and "rpath r x xs x2"
-  shows "x1 = x2"
-  using assms
-  by (induct, auto)
-
-subsubsection {* Properites of @{text "edges_on"} *}
-
-lemma edges_on_unfold:
-  "edges_on (a # b # xs) = {(a, b)} \<union> edges_on (b # xs)" (is "?L = ?R")
-proof -
-  { fix c d
-    assume "(c, d) \<in> ?L"
-    then obtain l1 l2 where h: "(a # b # xs) = l1 @ [c, d] @ l2" 
-        by (auto simp:edges_on_def)
-    have "(c, d) \<in> ?R"
-    proof(cases "l1")
-      case Nil
-      with h have "(c, d) = (a, b)" by auto
-      thus ?thesis by auto
-    next
-      case (Cons e es)
-      from h[unfolded this] have "b#xs = es@[c, d]@l2" by auto
-      thus ?thesis by (auto simp:edges_on_def)
-    qed
-  } moreover
-  { fix c d
-    assume "(c, d) \<in> ?R"
-    moreover have "(a, b) \<in> ?L" 
-    proof -
-      have "(a # b # xs) = []@[a,b]@xs" by simp
-      hence "\<exists> l1 l2. (a # b # xs) = l1@[a,b]@l2" by auto
-      thus ?thesis by (unfold edges_on_def, simp)
-    qed
-    moreover {
-        assume "(c, d) \<in> edges_on (b#xs)"
-        then obtain l1 l2 where "b#xs = l1@[c, d]@l2" by (unfold edges_on_def, auto)
-        hence "a#b#xs = (a#l1)@[c,d]@l2" by simp
-        hence "\<exists> l1 l2. (a # b # xs) = l1@[c,d]@l2" by metis
-        hence "(c,d) \<in> ?L" by (unfold edges_on_def, simp)
-    }
-    ultimately have "(c, d) \<in> ?L" by auto
-  } ultimately show ?thesis by auto
-qed
-
-lemma edges_on_len:
-  assumes "(a,b) \<in> edges_on l"
-  shows "length l \<ge> 2"
-  using assms
-  by (unfold edges_on_def, auto)
-
-text {* Elimination of @{text "edges_on"} for non-empty path *}
-
-lemma edges_on_consE [elim, cases set:edges_on]:
-  assumes "(a,b) \<in> edges_on (x#xs)"
-  obtains (head)  xs' where "x = a" and "xs = b#xs'"
-      |  (tail)  "(a,b) \<in> edges_on xs"
-proof -
-  from assms obtain l1 l2 
-  where h: "(x#xs) = l1 @ [a,b] @ l2" by (unfold edges_on_def, blast)
-  have "(\<exists> xs'. x = a \<and> xs = b#xs') \<or> ((a,b) \<in> edges_on xs)"
-  proof(cases "l1")
-    case Nil with h 
-    show ?thesis by auto
-  next
-    case (Cons e el)
-    from h[unfolded this] 
-    have "xs = el @ [a,b] @ l2" by auto
-    thus ?thesis 
-      by (unfold edges_on_def, auto)
-  qed
-  thus ?thesis 
-  proof
-    assume "(\<exists>xs'. x = a \<and> xs = b # xs')"
-    then obtain xs' where "x = a" "xs = b#xs'" by blast
-    from that(1)[OF this] show ?thesis .
-  next
-    assume "(a, b) \<in> edges_on xs"
-    from that(2)[OF this] show ?thesis .
-  qed
-qed
-
-text {*
-  Every edges on the path is a graph edges:
-*}
-lemma rpath_edges_on:
-  assumes "rpath r x xs y"
-  shows "(edges_on (x#xs)) \<subseteq> r"
-  using assms
-proof(induct arbitrary:y)
-  case (rbase x)
-  thus ?case by (unfold edges_on_def, auto)
-next
-  case (rstep x y ys z)
-  show ?case
-  proof -
-    { fix a b
-      assume "(a, b) \<in> edges_on (x # y # ys)"
-      hence "(a, b) \<in> r" by (cases, insert rstep, auto)
-    } thus ?thesis by auto
-  qed
-qed
-
-text {* @{text "edges_on"} is mono with respect to @{text "#"}-operation: *}
-lemma edges_on_Cons_mono:
-   shows "edges_on xs \<subseteq> edges_on (x#xs)"
-proof -
-  { fix a b
-    assume "(a, b) \<in> edges_on xs"
-    then obtain l1 l2 where "xs = l1 @ [a,b] @ l2" 
-      by (auto simp:edges_on_def)
-    hence "x # xs = (x#l1) @ [a, b] @ l2" by auto
-    hence "(a, b) \<in> edges_on (x#xs)" 
-      by (unfold edges_on_def, blast)
-  } thus ?thesis by auto
-qed
-
-text {*
-  The following rule @{text "rpath_transfer"} is used to show 
-  that one path is intact as long as all the edges on it are intact
-  with the change of graph.
-
-  If @{text "x#xs"} is path in graph @{text "r1"} and 
-  every edges along the path is also in @{text "r2"}, 
-  then @{text "x#xs"} is also a edge in graph @{text "r2"}:
-*}
-
-lemma rpath_transfer:
-  assumes "rpath r1 x xs y"
-  and "edges_on (x#xs) \<subseteq> r2"
-  shows "rpath r2 x xs y"
-  using assms
-proof(induct)
-  case (rstep x y ys z)
-  show ?case 
-  proof(rule rstepI)
-    show "(x, y) \<in> r2"
-    proof -
-      have "(x, y) \<in> edges_on  (x # y # ys)"
-          by (unfold edges_on_def, auto)
-     with rstep(4) show ?thesis by auto
-    qed
-  next
-    show "rpath r2 y ys z" 
-     using rstep edges_on_Cons_mono[of "y#ys" "x"] by (auto)
-  qed
-qed (unfold rpath_def, auto intro!:Transitive_Closure_Table.rtrancl_path.base)
-
-lemma edges_on_rpathI:
-  assumes "edges_on (a#xs@[b]) \<subseteq> r"
-  shows "rpath r a (xs@[b]) b"
-  using assms
-proof(induct xs arbitrary: a b)
-  case Nil
-  moreover have "(a, b) \<in> edges_on (a # [] @ [b])"
-      by (unfold edges_on_def, auto)
-  ultimately have "(a, b) \<in> r" by auto
-  thus ?case by auto
-next
-  case (Cons x xs a b)
-  from this(2) have "edges_on (x # xs @ [b]) \<subseteq> r" by (simp add:edges_on_unfold)
-  from Cons(1)[OF this] have " rpath r x (xs @ [b]) b" .
-  moreover from Cons(2) have "(a, x) \<in> r" by (auto simp:edges_on_unfold)
-  ultimately show ?case by (auto)
-qed
-
-text {*
-  The following lemma extracts the path from @{text "x"} to @{text "y"}
-  from proposition @{text "(x, y) \<in> r^*"}
-*}
-lemma star_rpath:
-  assumes "(x, y) \<in> r^*"
-  obtains xs where "rpath r x xs y"
-proof -
-  have "\<exists> xs. rpath r x xs y"
-  proof(unfold rpath_def, rule iffD1[OF rtranclp_eq_rtrancl_path])
-    from assms
-    show "(pred_of r)\<^sup>*\<^sup>* x y"
-      apply (fold pred_of_star)
-      by (auto simp:pred_of_def)
-  qed
-  from that and this show ?thesis by blast
-qed
-
-text {*
-  The following lemma uses the path @{text "xs"} from @{text "x"} to @{text "y"}
-  as a witness to show @{text "(x, y) \<in> r^*"}.
-*}
-lemma rpath_star: 
-  assumes "rpath r x xs y"
-  shows "(x, y) \<in> r^*"
-proof -
-  from iffD2[OF rtranclp_eq_rtrancl_path] and assms[unfolded rpath_def]
-  have "(pred_of r)\<^sup>*\<^sup>* x y" by metis
-  thus ?thesis by (simp add: pred_of_star star_2_pstar)
-qed  
-
-lemma subtree_transfer:
-  assumes "a \<in> subtree r1 a'"
-  and "r1 \<subseteq> r2"
-  shows "a \<in> subtree r2 a'"
-proof -
-  from assms(1)[unfolded subtree_def] 
-  have "(a, a') \<in> r1^*" by auto
-  from star_rpath[OF this]
-  obtain xs where rp: "rpath r1 a xs a'" by blast
-  hence "rpath r2 a xs a'"
-  proof(rule rpath_transfer)
-    from rpath_edges_on[OF rp] and assms(2)
-    show "edges_on (a # xs) \<subseteq> r2" by simp
-  qed
-  from rpath_star[OF this]
-  show ?thesis by (auto simp:subtree_def)
-qed
-
-lemma subtree_rev_transfer:
-  assumes "a \<notin> subtree r2 a'"
-  and "r1 \<subseteq> r2"
-  shows "a \<notin> subtree r1 a'"
-  using assms and subtree_transfer by metis
-
-text {*
-  The following lemmas establishes a relation from paths in @{text "r"}
-  to @{text "r^+"} relation.
-*}
-lemma rpath_plus: 
-  assumes "rpath r x xs y"
-  and "xs \<noteq> []"
-  shows "(x, y) \<in> r^+"
-proof -
-  from assms(2) obtain e es where "xs = e#es" by (cases xs, auto)
-  from assms(1)[unfolded this]
-  show ?thesis
-  proof(cases)
-    case rstep
-    show ?thesis
-    proof -
-      from rpath_star[OF rstep(2)] have "(e, y) \<in> r\<^sup>*" .
-      with rstep(1) show "(x, y) \<in> r^+" by auto
-    qed
-  qed
-qed
-
-lemma plus_rpath: 
-  assumes "(x, y) \<in> r^+"
-  obtains xs where "rpath r x xs y" and "xs \<noteq> []"
-proof -
-  from assms
-  show ?thesis
-  proof(cases rule:converse_tranclE[consumes 1])
-    case 1
-    hence "rpath r x [y] y" by auto
-    from that[OF this] show ?thesis by auto
-  next
-    case (2 z)
-    from 2(2) have "(z, y) \<in> r^*" by auto
-    from star_rpath[OF this] obtain xs where "rpath r z xs y" by auto
-    from rstepI[OF 2(1) this]
-    have "rpath r x (z # xs) y" .
-    from that[OF this] show ?thesis by auto
-  qed
-qed
-
-subsubsection {* Properties of @{text "subtree"} and @{term "ancestors"}*}
-
-lemma ancestors_subtreeI:
-  assumes "b \<in> ancestors r a"
-  shows "a \<in> subtree r b"
-  using assms by (auto simp:subtree_def ancestors_def)
-
-lemma ancestors_Field:
-  assumes "b \<in> ancestors r a"
-  obtains "a \<in> Domain r" "b \<in> Range r"
-  using assms 
-  apply (unfold ancestors_def, simp)
-  by (metis Domain.DomainI Range.intros trancl_domain trancl_range)
-
-lemma subtreeE:
-  assumes "a \<in> subtree r b"
-  obtains "a = b"
-      | "a \<noteq> b" and "b \<in> ancestors r a"
-proof -
-  from assms have "(a, b) \<in> r^*" by (auto simp:subtree_def)
-  from rtranclD[OF this]
-  have " a = b \<or> a \<noteq> b \<and> (a, b) \<in> r\<^sup>+" .
-  with that[unfolded ancestors_def] show ?thesis by auto
-qed
-
-
-lemma subtree_Field:
-  "subtree r x \<subseteq> Field r \<union> {x}"
-proof
-  fix y
-  assume "y \<in> subtree r x"
-  thus "y \<in> Field r \<union> {x}"
-  proof(cases rule:subtreeE)
-    case 1
-    thus ?thesis by auto
-  next
-    case 2
-    thus ?thesis apply (auto simp:ancestors_def)
-    using Field_def tranclD by fastforce 
-  qed
-qed
-
-lemma subtree_ancestorsI:
-  assumes "a \<in> subtree r b"
-  and "a \<noteq> b"
-  shows "b \<in> ancestors r a"
-  using assms
-  by (auto elim!:subtreeE)
-
-text {*
-  @{text "subtree"} is mono with respect to the underlying graph.
-*}
-lemma subtree_mono:
-  assumes "r1 \<subseteq> r2"
-  shows "subtree r1 x \<subseteq> subtree r2 x"
-proof
-  fix c
-  assume "c \<in> subtree r1 x"
-  hence "(c, x) \<in> r1^*" by (auto simp:subtree_def)
-  from star_rpath[OF this] obtain xs 
-  where rp:"rpath r1 c xs x" by metis
-  hence "rpath r2 c xs x"
-  proof(rule rpath_transfer)
-    from rpath_edges_on[OF rp] have "edges_on (c # xs) \<subseteq> r1" .
-    with assms show "edges_on (c # xs) \<subseteq> r2" by auto
-  qed
-  thus "c \<in> subtree r2 x"
-    by (rule rpath_star[elim_format], auto simp:subtree_def)
-qed
-
-text {*
-  The following lemma characterizes the change of sub-tree of @{text "x"}
-  with the removal of an outside edge @{text "(a,b)"}. 
-
-  Note that, according to lemma @{thm edges_in_refutation}, the assumption
-  @{term "b \<notin> subtree r x"} amounts to saying @{text "(a, b)"} 
-  is outside the sub-tree of @{text "x"}.
-*}
-lemma subtree_del_outside: (* ddd *)
-    assumes "b \<notin> subtree r x" 
-    shows "subtree (r - {(a, b)}) x = (subtree r x)" 
-proof -
-  { fix c
-    assume "c \<in> (subtree r x)"
-    hence "(c, x) \<in> r^*" by (auto simp:subtree_def)
-    hence "c \<in> subtree (r - {(a, b)}) x"
-    proof(rule star_rpath)
-      fix xs
-      assume rp: "rpath r c xs x"
-      show ?thesis
-      proof -
-        from rp
-        have "rpath  (r - {(a, b)}) c xs x"
-        proof(rule rpath_transfer)
-          from rpath_edges_on[OF rp] have "edges_on (c # xs) \<subseteq> r" .
-          moreover have "(a, b) \<notin> edges_on (c#xs)"
-          proof
-            assume "(a, b) \<in> edges_on (c # xs)"
-            then obtain l1 l2 where h: "c#xs = l1@[a,b]@l2" by (auto simp:edges_on_def)
-            hence "tl (c#xs) = tl (l1@[a,b]@l2)" by simp
-            then obtain l1' where eq_xs_b: "xs = l1'@[b]@l2" by (cases l1, auto)
-            from rp[unfolded this]
-            show False
-            proof(rule rpath_appendE)
-              assume "rpath r b l2 x"
-              thus ?thesis
-              by(rule rpath_star[elim_format], insert assms(1), auto simp:subtree_def)
-            qed
-          qed
-          ultimately show "edges_on (c # xs) \<subseteq> r - {(a,b)}" by auto
-        qed
-        thus ?thesis by (rule rpath_star[elim_format], auto simp:subtree_def)
-      qed
-    qed
-  } moreover {
-    fix c
-    assume "c \<in> subtree (r - {(a, b)}) x"
-    moreover have "... \<subseteq> (subtree r x)" by (rule subtree_mono, auto)
-    ultimately have "c \<in> (subtree r x)" by auto
-  } ultimately show ?thesis by auto
-qed
-
-(* ddd *)
-lemma subset_del_subtree_outside: (* ddd *)
-    assumes "Range r' \<inter> subtree r x = {}" 
-    shows "subtree (r - r') x = (subtree r x)" 
-proof -
-  { fix c
-    assume "c \<in> (subtree r x)"
-    hence "(c, x) \<in> r^*" by (auto simp:subtree_def)
-    hence "c \<in> subtree (r - r') x"
-    proof(rule star_rpath)
-      fix xs
-      assume rp: "rpath r c xs x"
-      show ?thesis
-      proof -
-        from rp
-        have "rpath  (r - r') c xs x"
-        proof(rule rpath_transfer)
-          from rpath_edges_on[OF rp] have "edges_on (c # xs) \<subseteq> r" .
-          moreover {
-              fix a b
-              assume h: "(a, b) \<in> r'"
-              have "(a, b) \<notin> edges_on (c#xs)"
-              proof
-                assume "(a, b) \<in> edges_on (c # xs)"
-                then obtain l1 l2 where "c#xs = (l1@[a])@[b]@l2" by (auto simp:edges_on_def)
-                hence "tl (c#xs) = tl (l1@[a,b]@l2)" by simp
-                then obtain l1' where eq_xs_b: "xs = l1'@[b]@l2" by (cases l1, auto)
-                from rp[unfolded this]
-                show False
-                proof(rule rpath_appendE)
-                  assume "rpath r b l2 x"
-                  from rpath_star[OF this]
-                  have "b \<in> subtree r x" by (auto simp:subtree_def)
-                  with assms (1) and h show ?thesis by (auto)
-                qed
-             qed
-         } ultimately show "edges_on (c # xs) \<subseteq> r - r'" by auto
-        qed
-        thus ?thesis by (rule rpath_star[elim_format], auto simp:subtree_def)
-      qed
-    qed
-  } moreover {
-    fix c
-    assume "c \<in> subtree (r - r') x"
-    moreover have "... \<subseteq> (subtree r x)" by (rule subtree_mono, auto)
-    ultimately have "c \<in> (subtree r x)" by auto
-  } ultimately show ?thesis by auto
-qed
-
-lemma subtree_insert_ext:
-    assumes "b \<in> subtree r x"
-    shows "subtree (r \<union> {(a, b)}) x = (subtree r x) \<union> (subtree r a)" 
-    using assms by (auto simp:subtree_def rtrancl_insert)
-
-lemma subtree_insert_next:
-    assumes "b \<notin> subtree r x"
-    shows "subtree (r \<union> {(a, b)}) x = (subtree r x)" 
-    using assms
-    by (auto simp:subtree_def rtrancl_insert)
-
-lemma set_add_rootI:
-  assumes "root r a"
-  and "a \<notin> Domain r1"
-  shows "root (r \<union> r1) a"
-proof -
-  let ?r = "r \<union> r1"
-  { fix a'
-    assume "a' \<in> ancestors ?r a"
-    hence "(a, a') \<in> ?r^+" by (auto simp:ancestors_def)
-    from tranclD[OF this] obtain z where "(a, z) \<in> ?r" by auto
-    moreover have "(a, z) \<notin> r"
-    proof
-      assume "(a, z) \<in> r"
-      with assms(1) show False 
-        by (auto simp:root_def ancestors_def)
-    qed
-    ultimately have "(a, z) \<in> r1" by auto
-    with assms(2) 
-    have False by (auto)
-  } thus ?thesis by (auto simp:root_def)
-qed
-
-lemma ancestors_mono:
-  assumes "r1 \<subseteq> r2"
-  shows "ancestors r1 x \<subseteq> ancestors r2 x"
-proof
- fix a
- assume "a \<in> ancestors r1 x"
- hence "(x, a) \<in> r1^+" by (auto simp:ancestors_def)
- from plus_rpath[OF this] obtain xs where 
-    h: "rpath r1 x xs a" "xs \<noteq> []" .
- have "rpath r2 x xs a"
- proof(rule rpath_transfer[OF h(1)])
-  from rpath_edges_on[OF h(1)] and assms
-  show "edges_on (x # xs) \<subseteq> r2" by auto
- qed
- from rpath_plus[OF this h(2)]
- show "a \<in> ancestors r2 x" by (auto simp:ancestors_def)
-qed
-
-lemma subtree_refute:
-  assumes "x \<notin> ancestors r y"
-  and "x \<noteq> y"
-  shows "y \<notin> subtree r x"
-proof
-   assume "y \<in> subtree r x"
-   thus False
-     by(elim subtreeE, insert assms, auto)
-qed
-
-subsubsection {* Properties about relational trees *}
-
-context rtree 
-begin
-
-lemma ancestors_headE:
-  assumes "c \<in> ancestors r a"
-  assumes "(a, b) \<in> r"
-  obtains "b = c"
-     |   "c \<in> ancestors r b"
-proof -
-  from assms(1) 
-  have "(a, c) \<in> r^+" by (auto simp:ancestors_def)
-  hence "b = c \<or> c \<in> ancestors r b"
-  proof(cases rule:converse_tranclE[consumes 1])
-    case 1
-    with assms(2) and sgv have "b = c" by (auto simp:single_valued_def)
-    thus ?thesis by auto
-  next
-    case (2 y)
-    from 2(1) and assms(2) and sgv have "y = b" by (auto simp:single_valued_def)
-    from 2(2)[unfolded this] have "c \<in> ancestors r b" by (auto simp:ancestors_def)
-    thus ?thesis by auto
-  qed
-  with that show ?thesis by metis
-qed
-
-lemma ancestors_accum:
-  assumes "(a, b) \<in> r"
-  shows "ancestors r a = ancestors r b \<union> {b}"
-proof -
-  { fix c
-    assume "c \<in> ancestors r a"
-    hence "(a, c) \<in> r^+" by (auto simp:ancestors_def)
-    hence "c \<in> ancestors r b \<union> {b}"
-    proof(cases rule:converse_tranclE[consumes 1])
-      case 1
-      with sgv assms have "c = b" by (unfold single_valued_def, auto)
-      thus ?thesis by auto
-    next
-      case (2 c')
-      with sgv assms have "c' = b" by (unfold single_valued_def, auto)
-      from 2(2)[unfolded this]
-      show ?thesis by (auto simp:ancestors_def)
-    qed
-  } moreover {
-    fix c
-    assume "c \<in> ancestors r b \<union> {b}"
-    hence "c = b \<or> c \<in> ancestors r b" by auto
-    hence "c \<in> ancestors r a"
-    proof
-      assume "c = b"
-      from assms[folded this] 
-      show ?thesis by (auto simp:ancestors_def)
-    next
-      assume "c \<in> ancestors r b"
-      with assms show ?thesis by (auto simp:ancestors_def)
-    qed
-  } ultimately show ?thesis by auto
-qed
-
-lemma rootI:
-  assumes h: "\<And> x'. x' \<noteq> x \<Longrightarrow> x \<notin> subtree r' x'"
-  and "r' \<subseteq> r"
-  shows "root r' x"
-proof -
-  from acyclic_subset[OF acl assms(2)]
-  have acl': "acyclic r'" .
-  { fix x'
-    assume "x' \<in> ancestors r' x"
-    hence h1: "(x, x') \<in> r'^+" by (auto simp:ancestors_def)
-    have "x' \<noteq> x"
-    proof
-      assume eq_x: "x' = x"
-      from h1[unfolded this] and acl'
-      show False by (auto simp:acyclic_def)
-    qed
-    moreover from h1 have "x \<in> subtree r' x'" by (auto simp:subtree_def)
-    ultimately have False using h by auto
-  } thus ?thesis by (auto simp:root_def)
-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"
-  obtains xs3 where "xs2 = xs1 @ xs3"
-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 [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"
-     |    (less_2) xs3 where "xs1 = xs2 @ xs3"
-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, 
-  i.e. there is at most one path between any two nodes.
-  Similar to the proof of @{thm rpath_overlap}, we starts with
-  the one side version first.
-*}
-
-lemma rpath_unique_oneside:
-  assumes "rpath r x xs1 y"
-    and "rpath r x xs2 y"
-    and "length xs1 \<le> length xs2"
-  shows "xs1 = xs2"
-proof -
-  from rpath_overlap_oneside[OF assms] 
-  obtain xs3 where less_1: "xs2 = xs1 @ xs3" by blast
-  show ?thesis
-  proof(cases "xs3 = []") 
-    case True
-    from less_1[unfolded this] show ?thesis by simp
-  next
-    case False
-    note FalseH = this
-    show ?thesis
-    proof(cases "xs1 = []")
-      case True
-      have "(x, x) \<in> r^+"
-      proof(rule rpath_plus)
-        from assms(1)[unfolded True] 
-        have "y = x" by (cases rule:rpath_nilE, simp)
-        from assms(2)[unfolded this] show "rpath r x xs2 x" .
-      next
-        from less_1 and False show "xs2 \<noteq> []" by simp
-      qed
-      with acl show ?thesis by (unfold acyclic_def, auto)
-    next 
-      case False
-      then obtain e es where eq_xs1: "xs1 = es@[e]" using rev_exhaust by auto
-      from assms(2)[unfolded less_1 this]
-      have "rpath r x (es @ [e] @ xs3) y" by simp
-      thus ?thesis
-      proof(cases rule:rpath_appendE)
-        case 1
-        from rpath_dest_eq [OF 1(1)[folded eq_xs1] assms(1)]
-        have "e = y" .
-        from rpath_plus [OF 1(2)[unfolded this] FalseH]
-        have "(y, y) \<in> r^+" .
-        with acl show ?thesis by (unfold acyclic_def, auto)
-      qed
-    qed
-  qed
-qed
-
-text {*
-  The following is the full version of path uniqueness.
-*}
-lemma rpath_unique:
-  assumes "rpath r x xs1 y"
-    and "rpath r x xs2 y"
-  shows "xs1 = xs2"
-proof(cases "length xs1 \<le> length xs2")
-   case True
-   from rpath_unique_oneside[OF assms this] show ?thesis .
-next
-  case False
-  hence "length xs2 \<le> length xs1" by simp
-  from rpath_unique_oneside[OF assms(2,1) this]
-  show ?thesis by simp
-qed
-
-text {*
-  The following lemma shows that the `independence` relation is symmetric.
-  It is an obvious auxiliary lemma which will be used later. 
-*}
-lemma sym_indep: "indep r x y \<Longrightarrow> indep r y x"
-  by (unfold indep_def, auto)
-
-text {*
-  This is another `obvious` lemma about trees, which says trees rooted at 
-  independent nodes are disjoint.
-*}
-lemma subtree_disjoint:
-  assumes "indep r x y"
-  shows "subtree r x \<inter> subtree r y = {}"
-proof -
-  { fix z x y xs1 xs2 xs3
-      assume ind: "indep r x y"
-      and rp1: "rpath r z xs1 x"
-      and rp2: "rpath r z xs2 y"
-      and h: "xs2 = xs1 @ xs3"
-      have False
-      proof(cases "xs1 = []")
-        case True
-        from rp1[unfolded this] have "x = z" by auto
-        from rp2[folded this] rpath_star ind[unfolded indep_def]
-        show ?thesis by metis
-      next
-        case False
-        then obtain e es where eq_xs1: "xs1 = es@[e]" using rev_exhaust by blast
-        from rp2[unfolded h this]
-        have "rpath r z (es @ [e] @ xs3) y" by simp
-        thus ?thesis
-        proof(cases rule:rpath_appendE)
-          case 1
-          have "e = x" using 1(1)[folded eq_xs1] rp1 rpath_dest_eq by metis
-          from rpath_star[OF 1(2)[unfolded this]] ind[unfolded indep_def]
-          show ?thesis by auto
-        qed
-      qed
-  } note my_rule = this
-  { fix z
-    assume h: "z \<in> subtree r x" "z \<in> subtree r y"
-    from h(1) have "(z, x) \<in> r^*" by (unfold subtree_def, auto)
-    then obtain xs1 where rp1: "rpath r z xs1 x" using star_rpath by metis
-    from h(2) have "(z, y) \<in> r^*" by (unfold subtree_def, auto)
-    then obtain xs2 where rp2: "rpath r z xs2 y" using star_rpath by metis
-    from rp1 rp2
-    have False
-    by (cases, insert my_rule[OF sym_indep[OF assms(1)] rp2 rp1] 
-                  my_rule[OF assms(1) rp1 rp2], auto)
-  } thus ?thesis by auto
-qed
-
-text {*
-  The following lemma @{text "subtree_del"} characterizes the change of sub-tree of 
-  @{text "x"} with the removal of an inside edge @{text "(a, b)"}. 
-  Note that, the case for the removal of an outside edge has already been dealt with
-  in lemma @{text "subtree_del_outside"}). 
-
-  This lemma is underpinned by the following two `obvious` facts:
-  \begin{enumearte}
-  \item
-  In graph @{text "r"}, for an inside edge @{text "(a,b) \<in> edges_in r x"},  
-  every node @{text "c"} in the sub-tree of @{text "a"} has a path
-  which goes first from @{text "c"} to @{text "a"}, then through edge @{text "(a, b)"}, and 
-  finally reaches @{text "x"}. By the uniqueness of path in a tree,
-  all paths from sub-tree of @{text "a"} to @{text "x"} are such constructed, therefore 
-  must go through @{text "(a, b)"}. The consequence is: with the removal of @{text "(a,b)"},
-  all such paths will be broken. 
-
-  \item
-  On the other hand, all paths not originate from within the sub-tree of @{text "a"}
-  will not be affected by the removal of edge @{text "(a, b)"}. 
-  The reason is simple: if the path is affected by the removal, it must 
-  contain @{text "(a, b)"}, then it must originate from within the sub-tree of @{text "a"}.
-  \end{enumearte}
-*}
-
-lemma subtree_del_inside: (* ddd *)
-    assumes "(a,b) \<in> edges_in r x"
-    shows "subtree (r - {(a, b)}) x = (subtree r x) - subtree r a"
-proof -
-  from assms have asm: "b \<in> subtree r x" "(a, b) \<in> r" by (auto simp:edges_in_def)
-  -- {* The proof follows a common pattern to prove the equality of sets. *}
-  { -- {* The `left to right` direction.
-       *}
-    fix c
-    -- {* Assuming @{text "c"} is inside the sub-tree of @{text "x"} in the reduced graph *}
-    assume h: "c \<in> subtree (r - {(a, b)}) x" 
-    -- {* We are going to show that @{text "c"} can not be in the sub-tree of @{text "a"} in 
-          the original graph. *}
-    -- {* In other words, all nodes inside the sub-tree of @{text "a"} in the original 
-          graph will be removed from the sub-tree of @{text "x"} in the reduced graph. *}
-    -- {* The reason, as analyzed before, is that all paths from within the 
-          sub-tree of @{text "a"} are broken with the removal of edge @{text "(a,b)"}.
-       *}
-    have "c \<in> (subtree r x) - subtree r a" 
-    proof -
-      let ?r' = "r - {(a, b)}" -- {* The reduced graph is abbreviated as @{text "?r'"} *}
-      from h have "(c, x) \<in> ?r'^*" by (auto simp:subtree_def)
-      -- {* Extract from the reduced graph the path @{text "xs"} from @{text "c"} to @{text "x"}. *}
-      then obtain xs where rp0: "rpath ?r' c xs x" by (rule star_rpath, auto)
-      -- {* It is easy to show @{text "xs"} is also a path in the original graph *}
-      hence rp1: "rpath r c xs x"
-      proof(rule rpath_transfer)
-          from rpath_edges_on[OF rp0] 
-          show "edges_on (c # xs) \<subseteq> r" by auto
-      qed
-      -- {* @{text "xs"} is used as the witness to show that @{text "c"} 
-                   in the sub-tree of @{text "x"} in the original graph. *}
-      hence "c \<in> subtree r x"
-         by (rule rpath_star[elim_format], auto simp:subtree_def)
-      -- {* The next step is to show that @{text "c"} can not be in the sub-tree of @{text "a"}
-            in the original graph. *}
-      -- {* We need to use the fact that all paths originate from within sub-tree of @{text "a"}
-             are broken. *}
-      moreover have "c \<notin> subtree r a"
-      proof
-        -- {* Proof by contradiction, suppose otherwise *}
-        assume otherwise: "c \<in> subtree r a"
-        -- {* Then there is a path in original graph leading from @{text "c"} to @{text "a"} *}
-        obtain xs1 where rp_c: "rpath r c xs1 a" 
-        proof -
-          from otherwise have "(c, a) \<in> r^*" by (auto simp:subtree_def)
-          thus ?thesis by (rule star_rpath, auto intro!:that)
-        qed
-        -- {* Starting from this path, we are going to construct a fictional 
-                  path from @{text "c"} to @{text "x"}, which, as explained before,
-              is broken, so that contradiction can be derived. *}
-        -- {* First, there is a path from @{text "b"} to @{text "x"} *}
-        obtain ys where rp_b: "rpath r b ys x" 
-        proof -
-          from asm have "(b, x) \<in> r^*" by (auto simp:subtree_def)
-          thus ?thesis by (rule star_rpath, auto intro!:that)
-        qed
-        -- {* The paths @{text "xs1"} and @{text "ys"} can be 
-                 tied together using @{text "(a,b)"} to form a path 
-               from @{text "c"} to @{text "x"}: *}
-        have "rpath r c (xs1 @ b # ys) x"
-        proof -
-          from rstepI[OF asm(2) rp_b] have "rpath r a (b # ys) x" .
-          from rpath_appendI[OF rp_c this]
-          show ?thesis .
-        qed
-        -- {* By the uniqueness of path between two nodes of a tree, we have: *}
-        from rpath_unique[OF rp1 this] have eq_xs: "xs = xs1 @ b # ys" .
-        -- {* Contradiction can be derived from from this fictional path . *}
-        show False
-        proof -
-          -- {* It can be shown that @{term "(a,b)"} is on this fictional path. *}
-          have "(a, b) \<in> edges_on (c#xs)"
-          proof(cases "xs1 = []")
-            case True
-            from rp_c[unfolded this] have "rpath r c [] a" .
-            hence eq_c: "c = a" by (rule rpath_nilE, simp)
-            hence "c#xs = a#xs" by simp
-            from this and eq_xs have "c#xs = a # xs1 @ b # ys" by simp
-            from this[unfolded True] have "c#xs = []@[a,b]@ys" by simp
-            thus ?thesis by (auto simp:edges_on_def)
-          next
-            case False
-            from rpath_nnl_lastE[OF rp_c this]
-            obtain xs' where "xs1 = xs'@[a]" by auto
-            from eq_xs[unfolded this] have "c#xs = (c#xs')@[a,b]@ys" by simp
-            thus ?thesis by (unfold edges_on_def, blast)
-          qed
-          -- {* It can also be shown that @{term "(a,b)"} is not on this fictional path. *}
-          moreover have "(a, b) \<notin> edges_on (c#xs)"
-              using rpath_edges_on[OF rp0] by auto
-          -- {* Contradiction is thus derived. *}
-          ultimately show False by auto
-        qed
-      qed
-      ultimately show ?thesis by auto
-    qed
-  } moreover {
-    -- {* The `right to left` direction.
-       *} 
-     fix c
-   -- {* Assuming that @{text "c"} is in the sub-tree of @{text "x"}, but
-         outside of the sub-tree of @{text "a"} in the original graph, *}
-   assume h: "c \<in> (subtree r x) - subtree r a"
-   -- {* we need to show that in the reduced graph, @{text "c"} is still in 
-         the sub-tree of @{text "x"}. *}
-   have "c \<in> subtree (r - {(a, b)}) x"
-   proof -
-      -- {* The proof goes by showing that the path from @{text "c"} to @{text "x"}
-            in the original graph is not affected by the removal of @{text "(a,b)"}.
-         *}
-      from h have "(c, x) \<in> r^*" by (unfold subtree_def, auto)
-      -- {* Extract the path @{text "xs"} from @{text "c"} to @{text "x"} in the original graph. *}
-      from star_rpath[OF this] obtain xs where rp: "rpath r c xs x" by auto
-      -- {* Show that it is also a path in the reduced graph. *}
-      hence "rpath (r - {(a, b)}) c xs x"
-      -- {* The proof goes by using rule @{thm rpath_transfer} *} 
-      proof(rule rpath_transfer)
-        -- {* We need to show all edges on the path are still in the reduced graph. *}
-        show "edges_on (c # xs) \<subseteq> r - {(a, b)}"
-        proof -
-          -- {* It is easy to show that all the edges are in the original graph. *}
-          from rpath_edges_on [OF rp] have " edges_on (c # xs) \<subseteq> r" .
-          -- {* The essential part is to show that @{text "(a, b)"} is not on the path. *}
-          moreover have "(a,b) \<notin> edges_on (c#xs)"
-          proof
-            -- {* Proof by contradiction, suppose otherwise: *}
-            assume otherwise: "(a, b) \<in> edges_on (c#xs)"
-            -- {* Then @{text "(a, b)"} is in the middle of the path. 
-                  with @{text "l1"} and @{text "l2"} be the nodes in 
-                  the front and rear respectively. *}
-              then obtain l1 l2 where eq_xs: 
-                "c#xs = l1 @ [a, b] @ l2" by (unfold edges_on_def, blast)
-            -- {* From this, it can be shown that @{text "c"} is 
-                      in the sub-tree of @{text "a"} *}
-            have "c \<in> subtree r a" 
-            proof(cases "l1 = []")
-              case True
-              -- {* If @{text "l1"} is null, it can be derived that @{text "c = a"}. *}
-              with eq_xs have "c = a" by auto
-              -- {* So, @{text "c"} is obviously in the sub-tree of @{text "a"}. *}
-              thus ?thesis by (unfold subtree_def, auto)
-            next
-              case False
-              -- {* When @{text "l1"} is not null, it must have a tail @{text "es"}: *}
-              then obtain e es where "l1 = e#es" by (cases l1, auto)
-              -- {* The relation of this tail with @{text "xs"} is derived: *}
-              with eq_xs have "xs = es@[a,b]@l2" by auto
-              -- {* From this, a path from @{text "c"} to @{text "a"} is made visible: *}
-              from rp[unfolded this] have "rpath r c (es @ [a] @ (b#l2)) x" by simp
-              thus ?thesis
-              proof(cases rule:rpath_appendE)
-                -- {* The path from @{text "c"} to @{text "a"} is extraced 
-                             using @{thm "rpath_appendE"}: *}
-                case 1
-                from rpath_star[OF this(1)] 
-                -- {* The extracted path servers as a witness that @{text "c"} is 
-                          in the sub-tree of @{text "a"}: *}
-                show ?thesis by (simp add:subtree_def)
-            qed
-          qed with h show False by auto         
-         qed ultimately show ?thesis by auto
-       qed
-     qed
-     -- {* From , it is shown that @{text "c"} is in the sub-tree of @{text "x"}
-           inthe reduced graph. *}
-     from rpath_star[OF this] show ?thesis by (auto simp:subtree_def)
-    qed
-  } 
-  -- {* The equality of sets is derived from the two directions just proved. *}
-  ultimately show ?thesis by auto
-qed 
-
-lemma  set_del_rootI:
-  assumes "r1 \<subseteq> r"
-  and "a \<in> Domain r1"
-  shows "root (r - r1) a"
-proof -
-   let ?r = "r - r1"
-  { fix a' 
-    assume neq: "a' \<noteq> a"
-    have "a \<notin> subtree ?r a'"
-    proof
-      assume "a \<in> subtree ?r a'"
-      hence "(a, a') \<in> ?r^*" by (auto simp:subtree_def)
-      from star_rpath[OF this] obtain xs
-      where rp: "rpath ?r a xs a'" by auto
-      from rpathE[OF this] and neq
-      obtain z zs where h: "(a, z) \<in> ?r" "rpath ?r z zs a'" "xs = z#zs" by auto
-      from assms(2) obtain z' where z'_in: "(a, z') \<in> r1" by (auto simp:DomainE)
-      with assms(1) have "(a, z') \<in> r" by auto
-      moreover from h(1) have "(a, z) \<in> r" by simp 
-      ultimately have "z' = z" using sgv by (auto simp:single_valued_def)
-      from z'_in[unfolded this] and h(1) show False by auto
-   qed
-  } thus ?thesis by (intro rootI, auto)
-qed
-
-lemma edge_del_no_rootI:
-  assumes "(a, b) \<in> r"
-  shows "root (r - {(a, b)}) a"
-  by (rule set_del_rootI, insert assms, auto)
-
-lemma ancestors_children_unique:
-  assumes "z1 \<in> ancestors r x \<inter> children r y"
-  and "z2 \<in> ancestors r x \<inter> children r y"
-  shows "z1 = z2"
-proof -
-  from assms have h:
-     "(x, z1) \<in> r^+" "(z1, y) \<in> r" 
-     "(x, z2) \<in> r^+" "(z2, y) \<in> r" 
-  by (auto simp:ancestors_def children_def)
-
-  -- {* From this, a path containing @{text "z1"} is obtained. *}
-  from plus_rpath[OF h(1)] obtain xs1 
-     where h1: "rpath r x xs1 z1" "xs1 \<noteq> []" by auto
-  from rpath_nnl_lastE[OF this] obtain xs1' where eq_xs1: "xs1 = xs1' @ [z1]"
-    by auto
-  from h(2) have h2: "rpath r z1 [y] y" by auto
-  from rpath_appendI[OF h1(1) h2, unfolded eq_xs1]
-  have rp1: "rpath r x (xs1' @ [z1, y]) y" by simp
-
-  -- {* Then, another path containing @{text "z2"} is obtained. *}
-  from plus_rpath[OF h(3)] obtain xs2
-     where h3: "rpath r x xs2 z2" "xs2 \<noteq> []" by auto
-  from rpath_nnl_lastE[OF this] obtain xs2' where eq_xs2: "xs2 = xs2' @ [z2]"
-    by auto
-  from h(4) have h4: "rpath r z2 [y] y" by auto
-  from rpath_appendI[OF h3(1) h4, unfolded eq_xs2]
-     have "rpath r x (xs2' @ [z2, y]) y" by simp
-
-  -- {* Finally @{text "z1 = z2"} is proved by uniqueness of path. *}
-  from rpath_unique[OF rp1 this]
-  have "xs1' @ [z1, y] = xs2' @ [z2, y]" .
-  thus ?thesis by auto
-qed
-
-lemma ancestors_childrenE:
-  assumes "y \<in> ancestors r x"
-  obtains "x \<in> children r y"
-      | z where "z \<in> ancestors r x \<inter> children r y"
-proof -
-  from assms(1) have "(x, y) \<in> r^+" by (auto simp:ancestors_def)
-  from tranclD2[OF this] obtain z where 
-     h: "(x, z) \<in> r\<^sup>*" "(z, y) \<in> r" by auto
-  from h(1)
-  show ?thesis
-  proof(cases rule:rtranclE)
-    case base
-    from h(2)[folded this] have "x \<in> children r y" 
-              by (auto simp:children_def)
-    thus ?thesis by (intro that, auto)
-  next
-    case (step u)
-    hence "z \<in> ancestors r x" by (auto simp:ancestors_def)
-    moreover from h(2) have "z \<in> children r y" 
-              by (auto simp:children_def)
-    ultimately show ?thesis by (intro that, auto)
-  qed
-qed
-
-
-end (* of rtree *)
-
-lemma subtree_children:
-  "subtree r x = {x} \<union> (\<Union> (subtree r ` (children r x)))" (is "?L = ?R")
-proof -
-  { fix z
-    assume "z \<in> ?L"
-    hence "z \<in> ?R"
-    proof(cases rule:subtreeE[consumes 1])
-      case 2
-      hence "(z, x) \<in> r^+" by (auto simp:ancestors_def)
-      thus ?thesis
-      proof(rule tranclE)
-        assume "(z, x) \<in> r"
-        hence "z \<in> children r x" by (unfold children_def, auto)
-        moreover have "z \<in> subtree r z" by (auto simp:subtree_def)
-        ultimately show ?thesis by auto
-      next
-        fix c
-        assume h: "(z, c) \<in> r\<^sup>+" "(c, x) \<in> r"
-        hence "c \<in> children r x" by (auto simp:children_def)
-        moreover from h have "z \<in> subtree r c" by (auto simp:subtree_def)
-        ultimately show ?thesis by auto
-      qed
-    qed auto
-  } moreover {
-    fix z
-    assume h: "z \<in> ?R"
-    have "x \<in> subtree r x" by (auto simp:subtree_def)
-    moreover {
-       assume "z \<in> \<Union>(subtree r ` children r x)"
-       then obtain y where "(y, x) \<in> r" "(z, y) \<in> r^*" 
-        by (auto simp:subtree_def children_def)
-       hence "(z, x) \<in> r^*" by auto
-       hence "z \<in> ?L" by (auto simp:subtree_def)
-    } ultimately have "z \<in> ?L" using h by auto
-  } ultimately show ?thesis by auto
-qed
-
-context fsubtree 
-begin
-  
-lemma finite_subtree:
-  shows "finite (subtree r x)"
-proof(induct rule:wf_induct[OF wf])
-  case (1 x)
-  have "finite (\<Union>(subtree r ` children r x))"
-  proof(rule finite_Union)
-    show "finite (subtree r ` children r x)"
-    proof(cases "children r x = {}")
-      case True
-      thus ?thesis by auto
-    next
-      case False
-      hence "x \<in> Range r" by (auto simp:children_def)
-      from fb[rule_format, OF this] 
-      have "finite (children r x)" .
-      thus ?thesis by (rule finite_imageI)
-    qed
-  next
-    fix M 
-    assume "M \<in> subtree r ` children r x"
-    then obtain y where h: "y \<in> children r x" "M = subtree r y" by auto
-    hence "(y, x) \<in> r" by (auto simp:children_def)
-    from 1[rule_format, OF this, folded h(2)]
-    show "finite M" .
-  qed
-  thus ?case
-    by (unfold subtree_children finite_Un, auto)
-qed
-
-end
-
-definition "pairself f = (\<lambda>(a, b). (f a, f b))"
-
-definition "rel_map f r = (pairself f ` r)"
-
-lemma rel_mapE: 
-  assumes "(a, b) \<in> rel_map f r"
-  obtains c d 
-  where "(c, d) \<in> r" "(a, b) = (f c, f d)"
-  using assms
-  by (unfold rel_map_def pairself_def, auto)
-
-lemma rel_mapI: 
-  assumes "(a, b) \<in> r"
-    and "c = f a"
-    and "d = f b"
-  shows "(c, d) \<in> rel_map f r"
-  using assms
-  by (unfold rel_map_def pairself_def, auto)
-
-lemma map_appendE:
-  assumes "map f zs = xs @ ys"
-  obtains xs' ys' 
-  where "zs = xs' @ ys'" "xs = map f xs'" "ys = map f ys'"
-proof -
-  have "\<exists> xs' ys'. zs = xs' @ ys' \<and> xs = map f xs' \<and> ys = map f ys'"
-  using assms
-  proof(induct xs arbitrary:zs ys)
-    case (Nil zs ys)
-    thus ?case by auto
-  next
-    case (Cons x xs zs ys)
-    note h = this
-    show ?case
-    proof(cases zs)
-      case (Cons e es)
-      with h have eq_x: "map f es = xs @ ys" "x = f e" by auto
-      from h(1)[OF this(1)]
-      obtain xs' ys' where "es = xs' @ ys'" "xs = map f xs'" "ys = map f ys'"
-        by blast
-      with Cons eq_x
-      have "zs = (e#xs') @ ys' \<and> x # xs = map f (e#xs') \<and> ys = map f ys'" by auto
-      thus ?thesis by metis
-    qed (insert h, auto)
-  qed
-  thus ?thesis by (auto intro!:that)
-qed
-
-lemma rel_map_mono:
-  assumes "r1 \<subseteq> r2"
-  shows "rel_map f r1 \<subseteq> rel_map f r2"
-  using assms
-  by (auto simp:rel_map_def pairself_def)
-
-lemma rel_map_compose [simp]:
-    shows "rel_map f1 (rel_map f2 r) = rel_map (f1 o f2) r"
-    by (auto simp:rel_map_def pairself_def)
-
-lemma edges_on_map: "edges_on (map f xs) = rel_map f (edges_on xs)"
-proof -
-  { fix a b
-    assume "(a, b) \<in> edges_on (map f xs)"
-    then obtain l1 l2 where eq_map: "map f xs = l1 @ [a, b] @ l2" 
-      by (unfold edges_on_def, auto)
-    hence "(a, b) \<in> rel_map f (edges_on xs)"
-      by (auto elim!:map_appendE intro!:rel_mapI simp:edges_on_def)
-  } moreover { 
-    fix a b
-    assume "(a, b) \<in> rel_map f (edges_on xs)"
-    then obtain c d where 
-        h: "(c, d) \<in> edges_on xs" "(a, b) = (f c, f d)" 
-             by (elim rel_mapE, auto)
-    then obtain l1 l2 where
-        eq_xs: "xs = l1 @ [c, d] @ l2" 
-             by (auto simp:edges_on_def)
-    hence eq_map: "map f xs = map f l1 @ [f c, f d] @ map f l2" by auto
-    have "(a, b) \<in> edges_on (map f xs)"
-    proof -
-      from h(2) have "[f c, f d] = [a, b]" by simp
-      from eq_map[unfolded this] show ?thesis by (auto simp:edges_on_def)
-    qed
-  } ultimately show ?thesis by auto
-qed
-
-lemma image_id:
-  assumes "\<And> x. x \<in> A \<Longrightarrow> f x = x"
-  shows "f ` A = A"
-  using assms by (auto simp:image_def)
-
-lemma rel_map_inv_id:
-  assumes "inj_on f ((Domain r) \<union> (Range r))"
-  shows "(rel_map (inv_into ((Domain r) \<union> (Range r)) f \<circ> f) r) = r"
-proof -
- let ?f = "(inv_into (Domain r \<union> Range r) f \<circ> f)"
- {
-  fix a b
-  assume h0: "(a, b) \<in> r"
-  have "pairself ?f (a, b) = (a, b)"
-  proof -
-    from assms h0 have "?f a = a" by (auto intro:inv_into_f_f)
-    moreover have "?f b = b"
-      by (insert h0, simp, intro inv_into_f_f[OF assms], auto intro!:RangeI)
-    ultimately show ?thesis by (auto simp:pairself_def)
-  qed
- } thus ?thesis by (unfold rel_map_def, intro image_id, case_tac x, auto)
-qed 
-
-lemma rel_map_acyclic:
-  assumes "acyclic r"
-  and "inj_on f ((Domain r) \<union> (Range r))"
-  shows "acyclic (rel_map f r)"
-proof -
-  let ?D = "Domain r \<union> Range r"
-  { fix a 
-    assume "(a, a) \<in> (rel_map f r)^+" 
-    from plus_rpath[OF this]
-    obtain xs where rp: "rpath (rel_map f r) a xs a" "xs \<noteq> []" by auto
-    from rpath_nnl_lastE[OF this] obtain xs' where eq_xs: "xs = xs'@[a]" by auto
-    from rpath_edges_on[OF rp(1)]
-    have h: "edges_on (a # xs) \<subseteq> rel_map f r" .
-    from edges_on_map[of "inv_into ?D f" "a#xs"]
-    have "edges_on (map (inv_into ?D f) (a # xs)) = rel_map (inv_into ?D f) (edges_on (a # xs))" .
-    with rel_map_mono[OF h, of "inv_into ?D f"]
-    have "edges_on (map (inv_into ?D f) (a # xs)) \<subseteq> rel_map ((inv_into ?D f) o f) r" by simp
-    from this[unfolded eq_xs]
-    have subr: "edges_on (map (inv_into ?D f) (a # xs' @ [a])) \<subseteq> rel_map (inv_into ?D f \<circ> f) r" .
-    have "(map (inv_into ?D f) (a # xs' @ [a])) = (inv_into ?D f a) # map (inv_into ?D f) xs' @ [inv_into ?D f a]"
-      by simp
-    from edges_on_rpathI[OF subr[unfolded this]]
-    have "rpath (rel_map (inv_into ?D f \<circ> f) r) 
-                      (inv_into ?D f a) (map (inv_into ?D f) xs' @ [inv_into ?D f a]) (inv_into ?D f a)" .
-    hence "(inv_into ?D f a, inv_into ?D f a) \<in> (rel_map (inv_into ?D f \<circ> f) r)^+"
-        by (rule rpath_plus, simp)
-    moreover have "(rel_map (inv_into ?D f \<circ> f) r) = r" by (rule rel_map_inv_id[OF assms(2)])
-    moreover note assms(1) 
-    ultimately have False by (unfold acyclic_def, auto)
-  } thus ?thesis by (auto simp:acyclic_def)
-qed
-
-lemma relpow_mult: 
-  "((r::'a rel) ^^ m) ^^ n = r ^^ (m*n)"
-proof(induct n arbitrary:m)
-  case (Suc k m)
-  thus ?case
-  proof -
-    have h: "(m * k + m) = (m + m * k)" by auto
-    show ?thesis 
-      apply (simp add:Suc relpow_add[symmetric])
-      by (unfold h, simp)
-  qed
-qed simp
-
-lemma compose_relpow_2:
-  assumes "r1 \<subseteq> r"
-  and "r2 \<subseteq> r"
-  shows "r1 O r2 \<subseteq> r ^^ (2::nat)"
-proof -
-  { fix a b
-    assume "(a, b) \<in> r1 O r2"
-    then obtain e where "(a, e) \<in> r1" "(e, b) \<in> r2"
-      by auto
-    with assms have "(a, e) \<in> r" "(e, b) \<in> r" by auto
-    hence "(a, b) \<in> r ^^ (Suc (Suc 0))" by auto
-  } thus ?thesis by (auto simp:numeral_2_eq_2)
-qed
-
-lemma acyclic_compose:
-  assumes "acyclic r"
-  and "r1 \<subseteq> r"
-  and "r2 \<subseteq> r"
-  shows "acyclic (r1 O r2)"
-proof -
-  { fix a
-    assume "(a, a) \<in> (r1 O r2)^+"
-    from trancl_mono[OF this compose_relpow_2[OF assms(2, 3)]]
-    have "(a, a) \<in> (r ^^ 2) ^+" .
-    from trancl_power[THEN iffD1, OF this]
-    obtain n where h: "(a, a) \<in> (r ^^ 2) ^^ n" "n > 0" by blast
-    from this(1)[unfolded relpow_mult] have h2: "(a, a) \<in> r ^^ (2 * n)" .
-    have "(a, a) \<in> r^+" 
-    proof(cases rule:trancl_power[THEN iffD2])
-      from h(2) h2 show "\<exists>n>0. (a, a) \<in> r ^^ n" 
-        by (rule_tac x = "2*n" in exI, auto)
-    qed
-    with assms have "False" by (auto simp:acyclic_def)
-  } thus ?thesis by (auto simp:acyclic_def)
-qed
-
-lemma children_compose_unfold: 
-  "children (r1 O r2) x = \<Union> (children r1 ` (children r2 x))"
-  by (auto simp:children_def)
-
-lemma fbranch_compose:
-  assumes "fbranch r1"
-  and "fbranch r2"
-  shows "fbranch (r1 O r2)"
-proof -
-  {  fix x
-     assume "x\<in>Range (r1 O r2)"
-     then obtain y z where h: "(y, z) \<in> r1" "(z, x) \<in> r2" by auto
-     have "finite (children (r1 O r2) x)"
-     proof(unfold children_compose_unfold, rule finite_Union)
-      show "finite (children r1 ` children r2 x)"
-      proof(rule finite_imageI)
-        from h(2) have "x \<in> Range r2" by auto
-        from assms(2)[unfolded fbranch_def, rule_format, OF this]
-        show "finite (children r2 x)" .
-      qed
-     next
-       fix M
-       assume "M \<in> children r1 ` children r2 x"
-       then obtain y where h1: "y \<in> children r2 x" "M = children r1 y" by auto
-       show "finite M"
-       proof(cases "children r1 y = {}")
-          case True
-          with h1(2) show ?thesis by auto
-       next
-          case False
-          hence "y \<in> Range r1" by (unfold children_def, auto)
-          from assms(1)[unfolded fbranch_def, rule_format, OF this, folded h1(2)]
-          show ?thesis .
-       qed
-     qed
-  } thus ?thesis by (unfold fbranch_def, auto)
-qed
-
-lemma finite_fbranchI:
-  assumes "finite r"
-  shows "fbranch r"
-proof -
-  { fix x 
-    assume "x \<in>Range r"
-    have "finite (children r x)"
-    proof -
-      have "{y. (y, x) \<in> r} \<subseteq> Domain r" by (auto)
-      from rev_finite_subset[OF finite_Domain[OF assms] this]
-      have "finite {y. (y, x) \<in> r}" .
-      thus ?thesis by (unfold children_def, simp)
-    qed
-  } thus ?thesis by (auto simp:fbranch_def)
-qed
-
-lemma subset_fbranchI:
-  assumes "fbranch r1"
-  and "r2 \<subseteq> r1"
-  shows "fbranch r2"
-proof -
-  { fix x
-    assume "x \<in>Range r2"
-    with assms(2) have "x \<in> Range r1" by auto
-    from assms(1)[unfolded fbranch_def, rule_format, OF this]
-    have "finite (children r1 x)" .
-    hence "finite (children r2 x)"
-    proof(rule rev_finite_subset)
-      from assms(2)
-      show "children r2 x \<subseteq> children r1 x" by (auto simp:children_def)
-    qed
-  } thus ?thesis by (auto simp:fbranch_def)
-qed
-
-lemma children_subtree: 
-  shows "children r x \<subseteq> subtree r x"
-  by (auto simp:children_def subtree_def)
-
-lemma children_union_kept:
-  assumes "x \<notin> Range r'"
-  shows "children (r \<union> r') x = children r x"
-  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/log	Thu Jan 28 21:14:17 2016 +0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,464 +0,0 @@
-修改集:      89:2056d9f481e2
-标签:        tip
-用户:        zhangx
-日期:        Thu Jan 28 16:36:46 2016 +0800
-摘要:        Slightly modified ExtGG.thy and PrioG.thy.
-
-修改集:      88:83dd5345d5d0
-父亲:        83:d239aa953315
-父亲:        87:33cb65e00ac0
-用户:        zhangx
-日期:        Thu Jan 28 16:33:49 2016 +0800
-摘要:        Merged back ExtGG.thy and PrioG.thy.
-
-修改集:      87:33cb65e00ac0
-用户:        zhangx
-日期:        Thu Jan 28 15:36:48 2016 +0800
-摘要:        Tracking ExtGG.thy etc., so that a update to 83 is possible.
-
-修改集:      86:2106021bae53
-用户:        zhangx
-日期:        Thu Jan 28 07:46:05 2016 +0800
-摘要:        Added PrioG.thy again
-
-修改集:      85:61a4429e7d4d
-父亲:        84:c0a4e840aefe
-父亲:        33:7f87232d9424
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed Jan 27 13:50:02 2016 +0000
-摘要:        merged
-
-修改集:      84:c0a4e840aefe
-父亲:        77:b6ea51cd2e88
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed Jan 27 13:47:08 2016 +0000
-摘要:        some small changes to Correctness and Paper
-
-修改集:      83:d239aa953315
-用户:        zhangx
-日期:        Thu Jan 28 07:43:05 2016 +0800
-摘要:        Added PrioG.thy as a parallel copy of Correctness.thy
-
-修改集:      82:cfd644dfc3b4
-用户:        zhangx
-日期:        Wed Jan 27 23:34:23 2016 +0800
-摘要:        The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
-
-修改集:      81:c495eb16beb6
-用户:        zhangx
-日期:        Wed Jan 27 19:28:42 2016 +0800
-摘要:        CpsG.thy restored. It was deleted. But now restored as a temporary holder of PIPBasics.thy.
-
-修改集:      80:17305a85493d
-用户:        zhangx
-日期:        Wed Jan 27 19:26:56 2016 +0800
-摘要:        CpsG.thy retrofiting almost completed. An important mile stone.
-
-修改集:      79:8067efcb43da
-用户:        zhangx
-日期:        Sun Jan 17 22:18:35 2016 +0800
-摘要:        Still improving CpsG.thy
-
-修改集:      78:df0334468335
-父亲:        75:d37703e0c5c4
-父亲:        77:b6ea51cd2e88
-用户:        zhangx
-日期:        Sat Jan 16 11:02:17 2016 +0800
-摘要:        Merged with 77
-
-修改集:      77:b6ea51cd2e88
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Fri Jan 15 02:05:29 2016 +0000
-摘要:        some small changes to the paper
-
-修改集:      76:2aa37de77f31
-父亲:        74:83ba2d8c859a
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Jan 14 03:29:22 2016 +0000
-摘要:        updated paper
-
-修改集:      75:d37703e0c5c4
-用户:        zhangx
-日期:        Sat Jan 16 10:59:03 2016 +0800
-摘要:        CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
-
-修改集:      74:83ba2d8c859a
-用户:        zhangx
-日期:        Thu Jan 14 00:55:54 2016 +0800
-摘要:        Moment.thy further simplified.
-
-修改集:      73:b0054fb0d1ce
-用户:        zhangx
-日期:        Wed Jan 13 23:39:59 2016 +0800
-摘要:        Moment.thy further improved.
-
-修改集:      72:3fa70b12c117
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed Jan 13 15:22:14 2016 +0000
-摘要:        another simplification
-
-修改集:      71:04caf0ccb3ae
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed Jan 13 15:16:59 2016 +0000
-摘要:        some small change
-
-修改集:      70:92ca2410b3d9
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed Jan 13 14:20:58 2016 +0000
-摘要:        further simplificaton of Moment.thy
-
-修改集:      69:1dc801552dfd
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed Jan 13 13:20:45 2016 +0000
-摘要:        simplified Moment.thy
-
-修改集:      68:db196b066b97
-用户:        zhangx
-日期:        Tue Jan 12 08:35:36 2016 +0800
-摘要:        Before retrofiting PIPBasics.thy
-
-修改集:      67:25fd656667a7
-用户:        zhangx
-日期:        Sat Jan 09 22:19:27 2016 +0800
-摘要:        Correctness simplified a great deal.
-
-修改集:      66:2af87bb52fca
-用户:        zhangx
-日期:        Thu Jan 07 22:10:06 2016 +0800
-摘要:        Some small improvements in Correctness.thy.
-
-修改集:      65:633b1fc8631b
-用户:        zhangx
-日期:        Thu Jan 07 08:33:13 2016 +0800
-摘要:        Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
-
-修改集:      64:b4bcd1edbb6d
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed Jan 06 16:34:26 2016 +0000
-摘要:        renamed files
-
-修改集:      63:b620a2a0806a
-用户:        zhangx
-日期:        Wed Jan 06 20:46:14 2016 +0800
-摘要:        ExtGG.thy finished, but more comments are needed.
-
-修改集:      62:031d2ae9c9b8
-用户:        zhangx
-日期:        Tue Dec 22 23:13:31 2015 +0800
-摘要:        In the middle of retrofiting ExtGG.thy.
-
-修改集:      61:f8194fd6214f
-用户:        zhangx
-日期:        Fri Dec 18 22:47:32 2015 +0800
-摘要:        CpsG.thy has been cleaned up.
-
-修改集:      60:f98a95f3deae
-用户:        zhangx
-日期:        Fri Dec 18 19:13:19 2015 +0800
-摘要:        Main proofs in CpsG.thy completed.
-
-修改集:      59:0a069a667301
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Dec 15 15:10:40 2015 +0000
-摘要:        removed some fixes about which Isabelle complains
-
-修改集:      58:ad57323fd4d6
-用户:        zhangx
-日期:        Tue Dec 15 21:45:46 2015 +0800
-摘要:        Extended RTree.thy
-
-修改集:      57:f1b39d77db00
-用户:        xingyuan zhang <xingyuanzhang@126.com>
-日期:        Thu Dec 03 14:34:29 2015 +0800
-摘要:        Added generic theory "RTree.thy"
-
-修改集:      56:0fd478e14e87
-用户:        xingyuan zhang <xingyuanzhang@126.com>
-日期:        Thu Dec 03 14:34:00 2015 +0800
-摘要:        Before switching to generic theory of relational trees.
-
-修改集:      55:b85cfbd58f59
-用户:        xingyuan zhang <xingyuanzhang@126.com>
-日期:        Fri Oct 30 20:40:11 2015 +0800
-摘要:        Comments for Set-operation finished
-
-修改集:      54:fee01b2858a2
-父亲:        50:8142e80f5d58
-父亲:        53:78adeef368c1
-用户:        xingyuan zhang <xingyuanzhang@126.com>
-日期:        Sat Oct 17 16:14:30 2015 +0800
-摘要:        Merge with tip
-
-修改集:      53:78adeef368c1
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Oct 06 14:22:34 2015 +0100
-摘要:        test
-
-修改集:      52:d462d449505f
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Oct 06 14:13:52 2015 +0100
-摘要:        another test
-
-修改集:      51:38ad30559775
-父亲:        49:8679d75b1d76
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Oct 06 14:11:28 2015 +0100
-摘要:        test
-
-修改集:      50:8142e80f5d58
-用户:        xingyuan zhang <xingyuanzhang@126.com>
-日期:        Sat Oct 17 16:10:33 2015 +0800
-摘要:        Finished comments on PrioGDef.thy
-
-修改集:      49:8679d75b1d76
-用户:        xingyuan zhang <xingyuanzhang@126.com>
-日期:        Tue Oct 06 18:52:04 2015 +0800
-摘要:        A little more change.
-
-修改集:      48:c0f14399c12f
-用户:        xingyuan zhang <xingyuanzhang@126.com>
-日期:        Tue Oct 06 13:08:00 2015 +0800
-摘要:        Some changes in the PrioGDef.thy.
-
-修改集:      47:2e6c8d530216
-用户:        xingyuan zhang <xingyuanzhang@126.com>
-日期:        Tue Oct 06 11:26:18 2015 +0800
-摘要:        Just a test change
-
-修改集:      46:331137d43625
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Sun Oct 04 23:02:57 2015 +0100
-摘要:        added one more reference to an incorrect specification
-
-修改集:      45:fc83f79009bd
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed Sep 09 11:24:19 2015 +0100
-摘要:        updated for Isabelle 2015
-
-修改集:      44:f676a68935a0
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Jul 15 17:25:53 2014 +0200
-摘要:        updated teh theories to newer Isabelle version
-
-修改集:      43:45e1d324c493
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Jun 12 10:14:50 2014 +0100
-摘要:        a few additions
-
-修改集:      42:0069bca6dd51
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Jun 10 10:44:48 2014 +0100
-摘要:        added scheduling book
-
-修改集:      41:66ed924aaa5c
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Mon Jun 09 16:01:28 2014 +0100
-摘要:        added another book that makes the error, some more proofs
-
-修改集:      40:0781a2fc93f1
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Jun 03 15:00:12 2014 +0100
-摘要:        added a library about graphs
-
-修改集:      39:7ea6b019ce24
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Mon Jun 02 14:58:42 2014 +0100
-摘要:        updated
-
-修改集:      38:c89013dca1aa
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Fri May 30 07:56:39 2014 +0100
-摘要:        finished proof of acyclity
-
-修改集:      37:c820ac0f3088
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Sat May 24 12:39:12 2014 +0100
-摘要:        simplified RAG_acyclic proof
-
-修改集:      36:af38526275f8
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Fri May 23 15:19:32 2014 +0100
-摘要:        added a test theory for polishing teh proofs
-
-修改集:      35:92f61f6a0fe7
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu May 22 17:40:39 2014 +0100
-摘要:        added a bit more text to the paper and separated a theory about Max
-
-修改集:      34:313acffe63b6
-父亲:        32:9b9f2117561f
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue May 20 12:49:21 2014 +0100
-摘要:        updated ROOT file
-
-修改集:      33:7f87232d9424
-父亲:        29:408ff78ce28f
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed May 14 11:52:53 2014 +0100
-摘要:        test
-
-修改集:      32:9b9f2117561f
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu May 15 16:02:44 2014 +0100
-摘要:        simplified the cp_rec proof
-
-修改集:      31:e861aff29655
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue May 06 14:36:40 2014 +0100
-摘要:        made some modifications.
-
-修改集:      30:8f026b608378
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed Mar 12 10:08:20 2014 +0000
-摘要:        added paper
-
-修改集:      29:408ff78ce28f
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Mar 04 16:47:54 2014 +0000
-摘要:        updated readme
-
-修改集:      28:7fa738a9615a
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Mar 04 16:38:38 2014 +0000
-摘要:        updated
-
-修改集:      27:6b1141c5e24c
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Mar 04 15:49:36 2014 +0000
-摘要:        cleaned up
-
-修改集:      26:da7a6ccfa7a9
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Mar 04 15:30:24 2014 +0000
-摘要:        updated
-
-修改集:      25:a9c0eeb00cc3
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Mar 04 15:27:59 2014 +0000
-摘要:        added two more references
-
-修改集:      24:6f50e6a8c6e0
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Mar 04 09:40:40 2014 +0000
-摘要:        some additions
-
-修改集:      23:24e6884d9258
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Mar 04 08:45:11 2014 +0000
-摘要:        made some small chages
-
-修改集:      22:9f0b78fcc894
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Mon Mar 03 16:22:48 2014 +0000
-摘要:        updated
-
-修改集:      21:55d1591b17f0
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Fri Feb 28 12:49:58 2014 +0000
-摘要:        added llncs to journal
-
-修改集:      20:b56616fd88dd
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Tue Feb 25 20:01:47 2014 +0000
-摘要:        added
-
-修改集:      19:3cc70bd49588
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Jun 20 23:28:26 2013 -0400
-摘要:        added paper
-
-修改集:      18:598409a21f4c
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Jun 20 13:50:01 2013 -0400
-摘要:        added nasa talk
-
-修改集:      17:105715a0a807
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Sat Dec 22 14:50:29 2012 +0000
-摘要:        updated
-
-修改集:      16:9764023f719e
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Sat Dec 22 01:58:45 2012 +0000
-摘要:        added
-
-修改集:      15:9e664c268e25
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Fri Dec 21 23:32:58 2012 +0000
-摘要:        added
-
-修改集:      14:1bf194825a4e
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Fri Dec 21 18:06:00 2012 +0000
-摘要:        more one the implementation
-
-修改集:      13:735e36c64a71
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Fri Dec 21 13:30:14 2012 +0000
-摘要:        added explanation of the code
-
-修改集:      12:85116bc854c0
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Fri Dec 21 00:24:30 2012 +0000
-摘要:        updated
-
-修改集:      11:8e02fb168350
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Dec 20 14:54:06 2012 +0000
-摘要:        added
-
-修改集:      10:242a781135ba
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Dec 20 12:23:44 2012 +0000
-摘要:        added two papers about PIP on multiprocs
-
-修改集:      9:a8e8ec87a933
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Dec 20 10:53:31 2012 +0000
-摘要:        added original Sha paper
-
-修改集:      8:5ba3d79622da
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed Dec 19 23:46:36 2012 +0000
-摘要:        added a paragraph about RAGS
-
-修改集:      7:0514be2ad83e
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Wed Dec 19 12:51:06 2012 +0000
-摘要:        started code explanation
-
-修改集:      6:7f2493296c39
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Mon Dec 17 12:34:24 2012 +0000
-摘要:        updated
-
-修改集:      5:0f2d4b78f839
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Mon Dec 10 21:27:22 2012 +0000
-摘要:        updated
-
-修改集:      4:9d667d545e32
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Dec 06 16:30:57 2012 +0000
-摘要:        added
-
-修改集:      3:51019d035a79
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Dec 06 15:49:20 2012 +0000
-摘要:        made everything working
-
-修改集:      2:a04084de4946
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Dec 06 15:12:49 2012 +0000
-摘要:        added
-
-修改集:      1:c4783e4ef43f
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Dec 06 15:11:51 2012 +0000
-摘要:        added
-
-修改集:      0:110247f9d47e
-用户:        Christian Urban <christian dot urban at kcl dot ac dot uk>
-日期:        Thu Dec 06 15:11:21 2012 +0000
-摘要:        added
-